diff --git a/src/HOL/Analysis/Abstract_Euclidean_Space.thy b/src/HOL/Analysis/Abstract_Euclidean_Space.thy --- a/src/HOL/Analysis/Abstract_Euclidean_Space.thy +++ b/src/HOL/Analysis/Abstract_Euclidean_Space.thy @@ -1,405 +1,405 @@ (* Author: LCP, ported from HOL Light *) section\Euclidean space and n-spheres, as subtopologies of n-dimensional space\ theory Abstract_Euclidean_Space imports Homotopy Locally begin subsection \Euclidean spaces as abstract topologies\ definition Euclidean_space :: "nat \ (nat \ real) topology" where "Euclidean_space n \ subtopology (powertop_real UNIV) {x. \i\n. x i = 0}" lemma topspace_Euclidean_space: "topspace(Euclidean_space n) = {x. \i\n. x i = 0}" by (simp add: Euclidean_space_def) lemma nonempty_Euclidean_space: "topspace(Euclidean_space n) \ {}" by (force simp: topspace_Euclidean_space) lemma subset_Euclidean_space [simp]: "topspace(Euclidean_space m) \ topspace(Euclidean_space n) \ m \ n" apply (simp add: topspace_Euclidean_space subset_iff, safe) apply (drule_tac x="(\i. if i < m then 1 else 0)" in spec) apply auto using not_less by fastforce lemma topspace_Euclidean_space_alt: "topspace(Euclidean_space n) = (\i \ {n..}. {x. x \ topspace(powertop_real UNIV) \ x i \ {0}})" by (auto simp: topspace_Euclidean_space) lemma closedin_Euclidean_space: "closedin (powertop_real UNIV) (topspace(Euclidean_space n))" proof - have "closedin (powertop_real UNIV) {x. x i = 0}" if "n \ i" for i proof - have "closedin (powertop_real UNIV) {x \ topspace (powertop_real UNIV). x i \ {0}}" proof (rule closedin_continuous_map_preimage) show "continuous_map (powertop_real UNIV) euclideanreal (\x. x i)" by (metis UNIV_I continuous_map_product_coordinates) show "closedin euclideanreal {0}" by simp qed then show ?thesis by auto qed then show ?thesis unfolding topspace_Euclidean_space_alt by force qed lemma closedin_Euclidean_imp_closed: "closedin (Euclidean_space m) S \ closed S" by (metis Euclidean_space_def closed_closedin closedin_Euclidean_space closedin_closed_subtopology euclidean_product_topology topspace_Euclidean_space) lemma closedin_Euclidean_space_iff: "closedin (Euclidean_space m) S \ closed S \ S \ topspace (Euclidean_space m)" (is "?lhs \ ?rhs") proof show "?lhs \ ?rhs" using closedin_closed_subtopology topspace_Euclidean_space by (fastforce simp: topspace_Euclidean_space_alt closedin_Euclidean_imp_closed) show "?rhs \ ?lhs" apply (simp add: closedin_subtopology Euclidean_space_def) by (metis (no_types) closed_closedin euclidean_product_topology inf.orderE) qed lemma continuous_map_componentwise_Euclidean_space: "continuous_map X (Euclidean_space n) (\x i. if i < n then f x i else 0) \ (\i < n. continuous_map X euclideanreal (\x. f x i))" proof - have *: "continuous_map X euclideanreal (\x. if k < n then f x k else 0)" if "\i. i continuous_map X euclideanreal (\x. f x i)" for k by (intro continuous_intros that) show ?thesis unfolding Euclidean_space_def continuous_map_in_subtopology by (fastforce simp: continuous_map_componentwise_UNIV * elim: continuous_map_eq) qed lemma continuous_map_Euclidean_space_add [continuous_intros]: "\continuous_map X (Euclidean_space n) f; continuous_map X (Euclidean_space n) g\ \ continuous_map X (Euclidean_space n) (\x i. f x i + g x i)" unfolding Euclidean_space_def continuous_map_in_subtopology by (fastforce simp add: continuous_map_componentwise_UNIV continuous_map_add) lemma continuous_map_Euclidean_space_diff [continuous_intros]: "\continuous_map X (Euclidean_space n) f; continuous_map X (Euclidean_space n) g\ \ continuous_map X (Euclidean_space n) (\x i. f x i - g x i)" unfolding Euclidean_space_def continuous_map_in_subtopology by (fastforce simp add: continuous_map_componentwise_UNIV continuous_map_diff) lemma continuous_map_Euclidean_space_iff: "continuous_map (Euclidean_space m) euclidean g = continuous_on (topspace (Euclidean_space m)) g" proof assume "continuous_map (Euclidean_space m) euclidean g" then have "continuous_map (top_of_set {f. \n\m. f n = 0}) euclidean g" by (simp add: Euclidean_space_def euclidean_product_topology) then show "continuous_on (topspace (Euclidean_space m)) g" by (metis continuous_map_subtopology_eu subtopology_topspace topspace_Euclidean_space) next assume "continuous_on (topspace (Euclidean_space m)) g" then have "continuous_map (top_of_set {f. \n\m. f n = 0}) euclidean g" by (metis (lifting) continuous_map_into_fulltopology continuous_map_subtopology_eu order_refl topspace_Euclidean_space) then show "continuous_map (Euclidean_space m) euclidean g" by (simp add: Euclidean_space_def euclidean_product_topology) qed lemma cm_Euclidean_space_iff_continuous_on: "continuous_map (subtopology (Euclidean_space m) S) (Euclidean_space n) f \ continuous_on (topspace (subtopology (Euclidean_space m) S)) f \ f ` (topspace (subtopology (Euclidean_space m) S)) \ topspace (Euclidean_space n)" (is "?P \ ?Q \ ?R") proof - have ?Q if ?P proof - have "\n. Euclidean_space n = top_of_set {f. \m\n. f m = 0}" by (simp add: Euclidean_space_def euclidean_product_topology) with that show ?thesis by (simp add: subtopology_subtopology) qed moreover have ?R if ?P using that by (simp add: image_subset_iff continuous_map_def) moreover have ?P if ?Q ?R proof - have "continuous_map (top_of_set (topspace (subtopology (subtopology (powertop_real UNIV) {f. \n\m. f n = 0}) S))) (top_of_set (topspace (subtopology (powertop_real UNIV) {f. \na\n. f na = 0}))) f" using Euclidean_space_def that by auto then show ?thesis by (simp add: Euclidean_space_def euclidean_product_topology subtopology_subtopology) qed ultimately show ?thesis by auto qed lemma homeomorphic_Euclidean_space_product_topology: "Euclidean_space n homeomorphic_space product_topology (\i. euclideanreal) {..i. euclideanreal) {..x. if k < n then x k else 0)" for k by (auto intro: continuous_map_if continuous_map_product_projection) show ?thesis unfolding homeomorphic_space_def homeomorphic_maps_def apply (rule_tac x="\f. restrict f {.. n = 0" by (auto simp: homeomorphic_compact_space [OF homeomorphic_Euclidean_space_product_topology] compact_space_product_topology) subsection\n-dimensional spheres\ definition nsphere where "nsphere n \ subtopology (Euclidean_space (Suc n)) { x. (\i\n. x i ^ 2) = 1 }" lemma nsphere: "nsphere n = subtopology (powertop_real UNIV) {x. (\i\n. x i ^ 2) = 1 \ (\i>n. x i = 0)}" by (simp add: nsphere_def Euclidean_space_def subtopology_subtopology Suc_le_eq Collect_conj_eq Int_commute) lemma continuous_map_nsphere_projection: "continuous_map (nsphere n) euclideanreal (\x. x k)" unfolding nsphere by (blast intro: continuous_map_from_subtopology [OF continuous_map_product_projection]) lemma in_topspace_nsphere: "(\n. if n = 0 then 1 else 0) \ topspace (nsphere n)" by (simp add: nsphere_def topspace_Euclidean_space power2_eq_square if_distrib [where f = "\x. x * _"] cong: if_cong) lemma nonempty_nsphere [simp]: "~ (topspace(nsphere n) = {})" using in_topspace_nsphere by auto lemma subtopology_nsphere_equator: "subtopology (nsphere (Suc n)) {x. x(Suc n) = 0} = nsphere n" proof - have "({x. (\i\n. (x i)\<^sup>2) + (x (Suc n))\<^sup>2 = 1 \ (\i>Suc n. x i = 0)} \ {x. x (Suc n) = 0}) = {x. (\i\n. (x i)\<^sup>2) = 1 \ (\i>n. x i = (0::real))}" using Suc_lessI [of n] by (fastforce simp: set_eq_iff) then show ?thesis by (simp add: nsphere subtopology_subtopology) qed lemma topspace_nsphere_minus1: assumes x: "x \ topspace (nsphere n)" and "x n = 0" shows "x \ topspace (nsphere (n - Suc 0))" proof (cases "n = 0") case True then show ?thesis using x by auto next case False have subt_eq: "nsphere (n - Suc 0) = subtopology (nsphere n) {x. x n = 0}" by (metis False Suc_pred le_zero_eq not_le subtopology_nsphere_equator) with x show ?thesis by (simp add: assms) qed lemma continuous_map_nsphere_reflection: "continuous_map (nsphere n) (nsphere n) (\x i. if i = k then -x i else x i)" proof - have cm: "continuous_map (powertop_real UNIV) euclideanreal (\x. if j = k then - x j else x j)" for j proof (cases "j=k") case True then show ?thesis by simp (metis UNIV_I continuous_map_product_projection) next case False then show ?thesis by (auto intro: continuous_map_product_projection) qed have eq: "(if i = k then x k * x k else x i * x i) = x i * x i" for i and x :: "nat \ real" by simp show ?thesis apply (simp add: nsphere continuous_map_in_subtopology continuous_map_componentwise_UNIV continuous_map_from_subtopology cm) apply (intro conjI allI impI continuous_intros continuous_map_from_subtopology continuous_map_product_projection) apply (auto simp: power2_eq_square if_distrib [where f = "\x. x * _"] eq cong: if_cong) done qed proposition contractible_space_upper_hemisphere: assumes "k \ n" shows "contractible_space(subtopology (nsphere n) {x. x k \ 0})" proof - define p:: "nat \ real" where "p \ \i. if i = k then 1 else 0" have "p \ topspace(nsphere n)" using assms - by (simp add: nsphere topspace_subtopology p_def power2_eq_square if_distrib [where f = "\x. x * _"] cong: if_cong) + by (simp add: nsphere p_def power2_eq_square if_distrib [where f = "\x. x * _"] cong: if_cong) let ?g = "\x i. x i / sqrt(\j\n. x j ^ 2)" let ?h = "\(t,q) i. (1 - t) * q i + t * p i" let ?Y = "subtopology (Euclidean_space (Suc n)) {x. 0 \ x k \ (\i\n. x i \ 0)}" have "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (nsphere n) {x. 0 \ x k})) (subtopology (nsphere n) {x. 0 \ x k}) (?g \ ?h)" proof (rule continuous_map_compose) have *: "\0 \ b k; (\i\n. (b i)\<^sup>2) = 1; \i>n. b i = 0; 0 \ a; a \ 1\ \ \i. (i = k \ (1 - a) * b k + a \ 0) \ (i \ k \ i \ n \ a \ 1 \ b i \ 0)" for a::real and b apply (cases "a \ 1 \ b k = 0"; simp) apply (metis (no_types, lifting) atMost_iff sum.neutral zero_power2) by (metis add.commute add_le_same_cancel2 diff_ge_0_iff_ge diff_zero less_eq_real_def mult_eq_0_iff mult_nonneg_nonneg not_le numeral_One zero_neq_numeral) show "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (nsphere n) {x. 0 \ x k})) ?Y ?h" using assms - apply (auto simp: * nsphere topspace_subtopology continuous_map_componentwise_UNIV + apply (auto simp: * nsphere continuous_map_componentwise_UNIV prod_topology_subtopology subtopology_subtopology case_prod_unfold continuous_map_in_subtopology Euclidean_space_def p_def if_distrib [where f = "\x. _ * x"] cong: if_cong) apply (intro continuous_map_prod_snd continuous_intros continuous_map_from_subtopology) apply auto done next have 1: "\x i. \ i \ n; x i \ 0\ \ (\i\n. (x i / sqrt (\j\n. (x j)\<^sup>2))\<^sup>2) = 1" by (force simp: sum_nonneg sum_nonneg_eq_0_iff field_split_simps simp flip: sum_divide_distrib) have cm: "continuous_map ?Y (nsphere n) (\x i. x i / sqrt (\j\n. (x j)\<^sup>2))" unfolding Euclidean_space_def nsphere subtopology_subtopology continuous_map_in_subtopology proof (intro continuous_intros conjI) show "continuous_map (subtopology (powertop_real UNIV) ({x. \i\Suc n. x i = 0} \ {x. 0 \ x k \ (\i\n. x i \ 0)})) (powertop_real UNIV) (\x i. x i / sqrt (\j\n. (x j)\<^sup>2))" unfolding continuous_map_componentwise by (intro continuous_intros conjI ballI) (auto simp: sum_nonneg_eq_0_iff) qed (auto simp: 1) show "continuous_map ?Y (subtopology (nsphere n) {x. 0 \ x k}) (\x i. x i / sqrt (\j\n. (x j)\<^sup>2))" by (force simp: cm sum_nonneg continuous_map_in_subtopology if_distrib [where f = "\x. _ * x"] cong: if_cong) qed moreover have "(?g \ ?h) (0, x) = x" if "x \ topspace (subtopology (nsphere n) {x. 0 \ x k})" for x using that - by (simp add: assms topspace_subtopology nsphere) + by (simp add: assms nsphere) moreover have "(?g \ ?h) (1, x) = p" if "x \ topspace (subtopology (nsphere n) {x. 0 \ x k})" for x by (force simp: assms p_def power2_eq_square if_distrib [where f = "\x. x * _"] cong: if_cong) ultimately show ?thesis apply (simp add: contractible_space_def homotopic_with) apply (rule_tac x=p in exI) apply (rule_tac x="?g \ ?h" in exI, force) done qed corollary contractible_space_lower_hemisphere: assumes "k \ n" shows "contractible_space(subtopology (nsphere n) {x. x k \ 0})" proof - have "contractible_space (subtopology (nsphere n) {x. 0 \ x k}) = ?thesis" proof (rule homeomorphic_space_contractibility) show "subtopology (nsphere n) {x. 0 \ x k} homeomorphic_space subtopology (nsphere n) {x. x k \ 0}" unfolding homeomorphic_space_def homeomorphic_maps_def apply (rule_tac x="\x i. if i = k then -(x i) else x i" in exI)+ apply (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology continuous_map_nsphere_reflection) done qed then show ?thesis using contractible_space_upper_hemisphere [OF assms] by metis qed proposition nullhomotopic_nonsurjective_sphere_map: assumes f: "continuous_map (nsphere p) (nsphere p) f" and fim: "f ` (topspace(nsphere p)) \ topspace(nsphere p)" obtains a where "homotopic_with (\x. True) (nsphere p) (nsphere p) f (\x. a)" proof - obtain a where a: "a \ topspace(nsphere p)" "a \ f ` (topspace(nsphere p))" using fim continuous_map_image_subset_topspace f by blast then have a1: "(\i\p. (a i)\<^sup>2) = 1" and a0: "\i. i > p \ a i = 0" by (simp_all add: nsphere) have f1: "(\j\p. (f x j)\<^sup>2) = 1" if "x \ topspace (nsphere p)" for x proof - have "f x \ topspace (nsphere p)" using continuous_map_image_subset_topspace f that by blast then show ?thesis by (simp add: nsphere) qed show thesis proof let ?g = "\x i. x i / sqrt(\j\p. x j ^ 2)" let ?h = "\(t,x) i. (1 - t) * f x i - t * a i" let ?Y = "subtopology (Euclidean_space(Suc p)) (- {\i. 0})" let ?T01 = "top_of_set {0..1::real}" have 1: "continuous_map (prod_topology ?T01 (nsphere p)) (nsphere p) (?g \ ?h)" proof (rule continuous_map_compose) have "continuous_map (prod_topology ?T01 (nsphere p)) euclideanreal ((\x. f x k) \ snd)" for k unfolding nsphere apply (simp add: continuous_map_of_snd) apply (rule continuous_map_compose [of _ "nsphere p" f, unfolded o_def]) using f apply (simp add: nsphere) by (simp add: continuous_map_nsphere_projection) then have "continuous_map (prod_topology ?T01 (nsphere p)) euclideanreal (\r. ?h r k)" for k unfolding case_prod_unfold o_def by (intro continuous_map_into_fulltopology [OF continuous_map_fst] continuous_intros) auto moreover have "?h ` ({0..1} \ topspace (nsphere p)) \ {x. \i\Suc p. x i = 0}" using continuous_map_image_subset_topspace [OF f] by (auto simp: nsphere image_subset_iff a0) moreover have "(\i. 0) \ ?h ` ({0..1} \ topspace (nsphere p))" proof clarify fix t b assume eq: "(\i. 0) = (\i. (1 - t) * f b i - t * a i)" and "t \ {0..1}" and b: "b \ topspace (nsphere p)" have "(1 - t)\<^sup>2 = (\i\p. ((1 - t) * f b i)^2)" using f1 [OF b] by (simp add: power_mult_distrib flip: sum_distrib_left) also have "\ = (\i\p. (t * a i)^2)" using eq by (simp add: fun_eq_iff) also have "\ = t\<^sup>2" using a1 by (simp add: power_mult_distrib flip: sum_distrib_left) finally have "1 - t = t" by (simp add: power2_eq_iff) then have *: "t = 1/2" by simp have fba: "f b \ a" using a(2) b by blast then show False using eq unfolding * by (simp add: fun_eq_iff) qed ultimately show "continuous_map (prod_topology ?T01 (nsphere p)) ?Y ?h" by (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV) next have *: "\\i\Suc p. x i = 0; x \ (\i. 0)\ \ (\j\p. (x j)\<^sup>2) \ 0" for x :: "nat \ real" by (force simp: fun_eq_iff not_less_eq_eq sum_nonneg_eq_0_iff) show "continuous_map ?Y (nsphere p) ?g" apply (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV nsphere continuous_map_componentwise subtopology_subtopology) apply (intro conjI allI continuous_intros continuous_map_from_subtopology [OF continuous_map_product_projection]) apply (simp_all add: *) apply (force simp: sum_nonneg fun_eq_iff not_less_eq_eq sum_nonneg_eq_0_iff power_divide simp flip: sum_divide_distrib) done qed have 2: "(?g \ ?h) (0, x) = f x" if "x \ topspace (nsphere p)" for x using that f1 by simp have 3: "(?g \ ?h) (1, x) = (\i. - a i)" for x using a by (force simp: field_split_simps nsphere) then show "homotopic_with (\x. True) (nsphere p) (nsphere p) f (\x. (\i. - a i))" by (force simp: homotopic_with intro: 1 2 3) qed qed lemma Hausdorff_Euclidean_space: "Hausdorff_space (Euclidean_space n)" unfolding Euclidean_space_def by (rule Hausdorff_space_subtopology) (metis Hausdorff_space_euclidean euclidean_product_topology) end diff --git a/src/HOL/Analysis/Abstract_Limits.thy b/src/HOL/Analysis/Abstract_Limits.thy --- a/src/HOL/Analysis/Abstract_Limits.thy +++ b/src/HOL/Analysis/Abstract_Limits.thy @@ -1,309 +1,309 @@ theory Abstract_Limits imports Abstract_Topology begin subsection\nhdsin and atin\ definition nhdsin :: "'a topology \ 'a \ 'a filter" where "nhdsin X a = (if a \ topspace X then (INF S\{S. openin X S \ a \ S}. principal S) else bot)" definition atin :: "'a topology \ 'a \ 'a filter" where "atin X a \ inf (nhdsin X a) (principal (topspace X - {a}))" lemma nhdsin_degenerate [simp]: "a \ topspace X \ nhdsin X a = bot" and atin_degenerate [simp]: "a \ topspace X \ atin X a = bot" by (simp_all add: nhdsin_def atin_def) lemma eventually_nhdsin: "eventually P (nhdsin X a) \ a \ topspace X \ (\S. openin X S \ a \ S \ (\x\S. P x))" proof (cases "a \ topspace X") case True hence "nhdsin X a = (INF S\{S. openin X S \ a \ S}. principal S)" by (simp add: nhdsin_def) also have "eventually P \ \ (\S. openin X S \ a \ S \ (\x\S. P x))" using True by (subst eventually_INF_base) (auto simp: eventually_principal) finally show ?thesis using True by simp qed auto lemma eventually_atin: "eventually P (atin X a) \ a \ topspace X \ (\U. openin X U \ a \ U \ (\x \ U - {a}. P x))" proof (cases "a \ topspace X") case True hence "eventually P (atin X a) \ (\S. openin X S \ a \ S \ (\x\S. x \ topspace X \ x \ a \ P x))" by (simp add: atin_def eventually_inf_principal eventually_nhdsin) also have "\ \ (\U. openin X U \ a \ U \ (\x \ U - {a}. P x))" using openin_subset by (intro ex_cong) auto finally show ?thesis by (simp add: True) qed auto subsection\Limits in a topological space\ definition limitin :: "'a topology \ ('b \ 'a) \ 'a \ 'b filter \ bool" where "limitin X f l F \ l \ topspace X \ (\U. openin X U \ l \ U \ eventually (\x. f x \ U) F)" lemma limitin_canonical_iff [simp]: "limitin euclidean f l F \ (f \ l) F" by (auto simp: limitin_def tendsto_def) lemma limitin_topspace: "limitin X f l F \ l \ topspace X" by (simp add: limitin_def) lemma limitin_const_iff [simp]: "limitin X (\a. l) l F \ l \ topspace X" by (simp add: limitin_def) lemma limitin_const: "limitin euclidean (\a. l) l F" by simp lemma limitin_eventually: "\l \ topspace X; eventually (\x. f x = l) F\ \ limitin X f l F" by (auto simp: limitin_def eventually_mono) lemma limitin_subsequence: "\strict_mono r; limitin X f l sequentially\ \ limitin X (f \ r) l sequentially" unfolding limitin_def using eventually_subseq by fastforce lemma limitin_subtopology: "limitin (subtopology X S) f l F \ l \ S \ eventually (\a. f a \ S) F \ limitin X f l F" (is "?lhs = ?rhs") proof (cases "l \ S \ topspace X") case True show ?thesis proof assume L: ?lhs with True have "\\<^sub>F b in F. f b \ topspace X \ S" by (metis (no_types) limitin_def openin_topspace topspace_subtopology) with L show ?rhs - apply (clarsimp simp add: limitin_def eventually_mono topspace_subtopology openin_subtopology_alt) + apply (clarsimp simp add: limitin_def eventually_mono openin_subtopology_alt) apply (drule_tac x="S \ U" in spec, force simp: elim: eventually_mono) done next assume ?rhs then show ?lhs using eventually_elim2 - by (fastforce simp add: limitin_def topspace_subtopology openin_subtopology_alt) + by (fastforce simp add: limitin_def openin_subtopology_alt) qed -qed (auto simp: limitin_def topspace_subtopology) +qed (auto simp: limitin_def) lemma limitin_canonical_iff_gen [simp]: assumes "open S" shows "limitin (top_of_set S) f l F \ (f \ l) F \ l \ S" using assms by (auto simp: limitin_subtopology tendsto_def) lemma limitin_sequentially: "limitin X S l sequentially \ l \ topspace X \ (\U. openin X U \ l \ U \ (\N. \n. N \ n \ S n \ U))" by (simp add: limitin_def eventually_sequentially) lemma limitin_sequentially_offset: "limitin X f l sequentially \ limitin X (\i. f (i + k)) l sequentially" unfolding limitin_sequentially by (metis add.commute le_add2 order_trans) lemma limitin_sequentially_offset_rev: assumes "limitin X (\i. f (i + k)) l sequentially" shows "limitin X f l sequentially" proof - have "\N. \n\N. f n \ U" if U: "openin X U" "l \ U" for U proof - obtain N where "\n. n\N \ f (n + k) \ U" using assms U unfolding limitin_sequentially by blast then have "\n\N+k. f n \ U" by (metis add_leD2 le_add_diff_inverse ordered_cancel_comm_monoid_diff_class.le_diff_conv2 add.commute) then show ?thesis .. qed with assms show ?thesis unfolding limitin_sequentially by simp qed lemma limitin_atin: "limitin Y f y (atin X x) \ y \ topspace Y \ (x \ topspace X \ (\V. openin Y V \ y \ V \ (\U. openin X U \ x \ U \ f ` (U - {x}) \ V)))" by (auto simp: limitin_def eventually_atin image_subset_iff) lemma limitin_atin_self: "limitin Y f (f a) (atin X a) \ f a \ topspace Y \ (a \ topspace X \ (\V. openin Y V \ f a \ V \ (\U. openin X U \ a \ U \ f ` U \ V)))" unfolding limitin_atin by fastforce lemma limitin_trivial: "\trivial_limit F; y \ topspace X\ \ limitin X f y F" by (simp add: limitin_def) lemma limitin_transform_eventually: "\eventually (\x. f x = g x) F; limitin X f l F\ \ limitin X g l F" unfolding limitin_def using eventually_elim2 by fastforce lemma continuous_map_limit: assumes "continuous_map X Y g" and f: "limitin X f l F" shows "limitin Y (g \ f) (g l) F" proof - have "g l \ topspace Y" by (meson assms continuous_map_def limitin_topspace) moreover have "\U. \\V. openin X V \ l \ V \ (\\<^sub>F x in F. f x \ V); openin Y U; g l \ U\ \ \\<^sub>F x in F. g (f x) \ U" using assms eventually_mono by (fastforce simp: limitin_def dest!: openin_continuous_map_preimage) ultimately show ?thesis using f by (fastforce simp add: limitin_def) qed subsection\Pointwise continuity in topological spaces\ definition topcontinuous_at where "topcontinuous_at X Y f x \ x \ topspace X \ (\x \ topspace X. f x \ topspace Y) \ (\V. openin Y V \ f x \ V \ (\U. openin X U \ x \ U \ (\y \ U. f y \ V)))" lemma topcontinuous_at_atin: "topcontinuous_at X Y f x \ x \ topspace X \ (\x \ topspace X. f x \ topspace Y) \ limitin Y f (f x) (atin X x)" unfolding topcontinuous_at_def by (fastforce simp add: limitin_atin)+ lemma continuous_map_eq_topcontinuous_at: "continuous_map X Y f \ (\x \ topspace X. topcontinuous_at X Y f x)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: continuous_map_def topcontinuous_at_def) next assume R: ?rhs then show ?lhs apply (auto simp: continuous_map_def topcontinuous_at_def) apply (subst openin_subopen, safe) apply (drule bspec, assumption) using openin_subset[of X] apply (auto simp: subset_iff dest!: spec) done qed lemma continuous_map_atin: "continuous_map X Y f \ (\x \ topspace X. limitin Y f (f x) (atin X x))" by (auto simp: limitin_def topcontinuous_at_atin continuous_map_eq_topcontinuous_at) lemma limitin_continuous_map: "\continuous_map X Y f; a \ topspace X; f a = b\ \ limitin Y f b (atin X a)" by (auto simp: continuous_map_atin) subsection\Combining theorems for continuous functions into the reals\ lemma continuous_map_canonical_const [continuous_intros]: "continuous_map X euclidean (\x. c)" by simp lemma continuous_map_real_mult [continuous_intros]: "\continuous_map X euclideanreal f; continuous_map X euclideanreal g\ \ continuous_map X euclideanreal (\x. f x * g x)" by (simp add: continuous_map_atin tendsto_mult) lemma continuous_map_real_pow [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. f x ^ n)" by (induction n) (auto simp: continuous_map_real_mult) lemma continuous_map_real_mult_left: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. c * f x)" by (simp add: continuous_map_atin tendsto_mult) lemma continuous_map_real_mult_left_eq: "continuous_map X euclideanreal (\x. c * f x) \ c = 0 \ continuous_map X euclideanreal f" proof (cases "c = 0") case False have "continuous_map X euclideanreal (\x. c * f x) \ continuous_map X euclideanreal f" apply (frule continuous_map_real_mult_left [where c="inverse c"]) apply (simp add: field_simps False) done with False show ?thesis using continuous_map_real_mult_left by blast qed simp lemma continuous_map_real_mult_right: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. f x * c)" by (simp add: continuous_map_atin tendsto_mult) lemma continuous_map_real_mult_right_eq: "continuous_map X euclideanreal (\x. f x * c) \ c = 0 \ continuous_map X euclideanreal f" by (simp add: mult.commute flip: continuous_map_real_mult_left_eq) lemma continuous_map_minus [continuous_intros]: fixes f :: "'a\'b::real_normed_vector" shows "continuous_map X euclidean f \ continuous_map X euclidean (\x. - f x)" by (simp add: continuous_map_atin tendsto_minus) lemma continuous_map_minus_eq [simp]: fixes f :: "'a\'b::real_normed_vector" shows "continuous_map X euclidean (\x. - f x) \ continuous_map X euclidean f" using continuous_map_minus add.inverse_inverse continuous_map_eq by fastforce lemma continuous_map_add [continuous_intros]: fixes f :: "'a\'b::real_normed_vector" shows "\continuous_map X euclidean f; continuous_map X euclidean g\ \ continuous_map X euclidean (\x. f x + g x)" by (simp add: continuous_map_atin tendsto_add) lemma continuous_map_diff [continuous_intros]: fixes f :: "'a\'b::real_normed_vector" shows "\continuous_map X euclidean f; continuous_map X euclidean g\ \ continuous_map X euclidean (\x. f x - g x)" by (simp add: continuous_map_atin tendsto_diff) lemma continuous_map_norm [continuous_intros]: fixes f :: "'a\'b::real_normed_vector" shows "continuous_map X euclidean f \ continuous_map X euclidean (\x. norm(f x))" by (simp add: continuous_map_atin tendsto_norm) lemma continuous_map_real_abs [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. abs(f x))" by (simp add: continuous_map_atin tendsto_rabs) lemma continuous_map_real_max [continuous_intros]: "\continuous_map X euclideanreal f; continuous_map X euclideanreal g\ \ continuous_map X euclideanreal (\x. max (f x) (g x))" by (simp add: continuous_map_atin tendsto_max) lemma continuous_map_real_min [continuous_intros]: "\continuous_map X euclideanreal f; continuous_map X euclideanreal g\ \ continuous_map X euclideanreal (\x. min (f x) (g x))" by (simp add: continuous_map_atin tendsto_min) lemma continuous_map_sum [continuous_intros]: fixes f :: "'a\'b\'c::real_normed_vector" shows "\finite I; \i. i \ I \ continuous_map X euclidean (\x. f x i)\ \ continuous_map X euclidean (\x. sum (f x) I)" by (simp add: continuous_map_atin tendsto_sum) lemma continuous_map_prod [continuous_intros]: "\finite I; \i. i \ I \ continuous_map X euclideanreal (\x. f x i)\ \ continuous_map X euclideanreal (\x. prod (f x) I)" by (simp add: continuous_map_atin tendsto_prod) lemma continuous_map_real_inverse [continuous_intros]: "\continuous_map X euclideanreal f; \x. x \ topspace X \ f x \ 0\ \ continuous_map X euclideanreal (\x. inverse(f x))" by (simp add: continuous_map_atin tendsto_inverse) lemma continuous_map_real_divide [continuous_intros]: "\continuous_map X euclideanreal f; continuous_map X euclideanreal g; \x. x \ topspace X \ g x \ 0\ \ continuous_map X euclideanreal (\x. f x / g x)" by (simp add: continuous_map_atin tendsto_divide) end diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy +++ b/src/HOL/Analysis/Abstract_Topology.thy @@ -1,4607 +1,4607 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ theory Abstract_Topology imports Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" begin subsection \General notion of a topology as a value\ definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ (\S T. L S \ L T \ L (S \ T)) \ (\\. (\K\\. L K) \ L (\\))" typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast lemma istopology_openin[intro]: "istopology(openin U)" using openin[of U] by blast lemma istopology_open: "istopology open" by (auto simp: istopology_def) lemma topology_inverse': "istopology U \ openin (topology U) = U" using topology_inverse[unfolded mem_Collect_eq] . lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" using topology_inverse[of U] istopology_openin[of "topology U"] by auto lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" proof assume "T1 = T2" then show "\S. openin T1 S \ openin T2 S" by simp next assume H: "\S. openin T1 S \ openin T2 S" then have "openin T1 = openin T2" by (simp add: fun_eq_iff) then have "topology (openin T1) = topology (openin T2)" by simp then show "T1 = T2" unfolding openin_inverse . qed text\The "universe": the union of all sets in the topology.\ definition "topspace T = \{S. openin T S}" subsubsection \Main properties of open sets\ proposition openin_clauses: fixes U :: "'a topology" shows "openin U {}" "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def by auto lemma openin_subset[intro]: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (rule openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" by (rule openin_clauses) lemma openin_Union[intro]: "(\S. S \ K \ openin U S) \ openin U (\K)" using openin_clauses by blast lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (force simp: openin_Union topspace_def) lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by auto next assume H: ?rhs let ?t = "\{T. openin U T \ T \ S}" have "openin U ?t" by (force simp: openin_Union) also have "?t = S" using H by auto finally show "openin U S" . qed lemma openin_INT [intro]: assumes "finite I" "\i. i \ I \ openin T (U i)" shows "openin T ((\i \ I. U i) \ topspace T)" using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) lemma openin_INT2 [intro]: assumes "finite I" "I \ {}" "\i. i \ I \ openin T (U i)" shows "openin T (\i \ I. U i)" proof - have "(\i \ I. U i) \ topspace T" using \I \ {}\ openin_subset[OF assms(3)] by auto then show ?thesis using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) qed lemma openin_Inter [intro]: assumes "finite \" "\ \ {}" "\X. X \ \ \ openin T X" shows "openin T (\\)" by (metis (full_types) assms openin_INT2 image_ident) lemma openin_Int_Inter: assumes "finite \" "openin T U" "\X. X \ \ \ openin T X" shows "openin T (U \ \\)" using openin_Inter [of "insert U \"] assms by auto subsubsection \Closed sets\ definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" by (simp add: closedin_def) lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" by (auto simp: Diff_Un closedin_def) lemma Diff_Inter[intro]: "A - \S = \{A - s|s. s\S}" by auto lemma closedin_Union: assumes "finite S" "\T. T \ S \ closedin U T" shows "closedin U (\S)" using assms by induction auto lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S. S \K \ closedin U S" shows "closedin U (\K)" using Ke Kc unfolding closedin_def Diff_Inter by auto lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" apply (rule closedin_Inter) using assms apply auto done lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" using closedin_Inter[of "{S,T}" U] by auto lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" apply (auto simp: closedin_def Diff_Diff_Int inf_absorb2) apply (metis openin_subset subset_eq) done lemma topology_finer_closedin: "topspace X = topspace Y \ (\S. openin Y S \ openin X S) \ (\S. closedin Y S \ closedin X S)" apply safe apply (simp add: closedin_def) by (simp add: openin_closedin_eq) lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" by (simp add: openin_closedin_eq) lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" proof - have "S - T = S \ (topspace U - T)" using openin_subset[of U S] oS cT by (auto simp: topspace_def openin_subset) then show ?thesis using oS cT by (auto simp: closedin_def) qed lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" proof - have "S - T = S \ (topspace U - T)" using closedin_subset[of U S] oS cT by (auto simp: topspace_def) then show ?thesis using oS cT by (auto simp: openin_closedin_eq) qed subsection\The discrete topology\ definition discrete_topology where "discrete_topology U \ topology (\S. S \ U)" lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \ S \ U" proof - have "istopology (\S. S \ U)" by (auto simp: istopology_def) then show ?thesis by (simp add: discrete_topology_def topology_inverse') qed lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U" by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym) lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \ S \ U" by (simp add: closedin_def) lemma discrete_topology_unique: "discrete_topology U = X \ topspace X = U \ (\x \ U. openin X {x})" (is "?lhs = ?rhs") proof assume R: ?rhs then have "openin X S" if "S \ U" for S using openin_subopen subsetD that by fastforce moreover have "x \ topspace X" if "openin X S" and "x \ S" for x S using openin_subset that by blast ultimately show ?lhs using R by (auto simp: topology_eq) qed auto lemma discrete_topology_unique_alt: "discrete_topology U = X \ topspace X \ U \ (\x \ U. openin X {x})" using openin_subset by (auto simp: discrete_topology_unique) lemma subtopology_eq_discrete_topology_empty: "X = discrete_topology {} \ topspace X = {}" using discrete_topology_unique [of "{}" X] by auto lemma subtopology_eq_discrete_topology_sing: "X = discrete_topology {a} \ topspace X = {a}" by (metis discrete_topology_unique openin_topspace singletonD) subsection \Subspace topology\ definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" (is "istopology ?L") proof - have "?L {}" by blast { fix A B assume A: "?L A" and B: "?L B" from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast have "A \ B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ then have "?L (A \ B)" by blast } moreover { fix K assume K: "K \ Collect ?L" have th0: "Collect ?L = (\S. S \ V) ` Collect (openin U)" by blast from K[unfolded th0 subset_image_iff] obtain Sk where Sk: "Sk \ Collect (openin U)" "K = (\S. S \ V) ` Sk" by blast have "\K = (\Sk) \ V" using Sk by auto moreover have "openin U (\Sk)" using Sk by (auto simp: subset_eq) ultimately have "?L (\K)" by blast } ultimately show ?thesis unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" unfolding subtopology_def topology_inverse'[OF istopology_subtopology] by auto lemma openin_subtopology_Int: "openin X S \ openin (subtopology X T) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_Int2: "openin X T \ openin (subtopology X S) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_diff_closed: "\S \ topspace X; closedin X T\ \ openin (subtopology X S) (S - T)" unfolding closedin_def openin_subtopology by (rule_tac x="topspace X - T" in exI) auto lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)" by (force simp: relative_to_def openin_subtopology) lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \ V" by (auto simp: topspace_def openin_subtopology) lemma topspace_subtopology_subset: "S \ topspace X \ topspace(subtopology X S) = S" - by (simp add: inf.absorb_iff2 topspace_subtopology) + by (simp add: inf.absorb_iff2) lemma closedin_subtopology: "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) lemma subtopology_subtopology: "subtopology (subtopology X S) T = subtopology X (S \ T)" proof - have eq: "\T'. (\S'. T' = S' \ T \ (\T. openin X T \ S' = T \ S)) = (\Sa. T' = Sa \ (S \ T) \ openin X Sa)" by (metis inf_assoc) have "subtopology (subtopology X S) T = topology (\Ta. \Sa. Ta = Sa \ T \ openin (subtopology X S) Sa)" by (simp add: subtopology_def) also have "\ = subtopology X (S \ T)" by (simp add: openin_subtopology eq) (simp add: subtopology_def) finally show ?thesis . qed lemma openin_subtopology_alt: "openin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (openin X)" by (simp add: image_iff inf_commute openin_subtopology) lemma closedin_subtopology_alt: "closedin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (closedin X)" by (simp add: image_iff inf_commute closedin_subtopology) lemma subtopology_superset: assumes UV: "topspace U \ V" shows "subtopology U V = U" proof - { fix S { fix T assume T: "openin U T" "S = T \ V" from T openin_subset[OF T(1)] UV have eq: "S = T" by blast have "openin U S" unfolding eq using T by blast } moreover { assume S: "openin U S" then have "\T. openin U T \ S = T \ V" using openin_subset[OF S] UV by auto } ultimately have "(\T. openin U T \ S = T \ V) \ openin U S" by blast } then show ?thesis unfolding topology_eq openin_subtopology by blast qed lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" by (simp add: subtopology_superset) lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" by (simp add: subtopology_superset) lemma subtopology_restrict: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) lemma openin_subtopology_empty: "openin (subtopology U {}) S \ S = {}" by (metis Int_empty_right openin_empty openin_subtopology) lemma closedin_subtopology_empty: "closedin (subtopology U {}) S \ S = {}" by (metis Int_empty_right closedin_empty closedin_subtopology) lemma closedin_subtopology_refl [simp]: "closedin (subtopology U X) X \ X \ topspace U" by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) lemma closedin_topspace_empty: "topspace T = {} \ (closedin T S \ S = {})" by (simp add: closedin_def) lemma open_in_topspace_empty: "topspace X = {} \ openin X S \ S = {}" by (simp add: openin_closedin_eq) lemma openin_imp_subset: "openin (subtopology U S) T \ T \ S" by (metis Int_iff openin_subtopology subsetI) lemma closedin_imp_subset: "closedin (subtopology U S) T \ T \ S" -by (simp add: closedin_def topspace_subtopology) +by (simp add: closedin_def) lemma openin_open_subtopology: "openin X S \ openin (subtopology X S) T \ openin X T \ T \ S" by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology) lemma closedin_closed_subtopology: "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" by (simp add: openin_subtopology) blast lemma closedin_subtopology_Un: "\closedin (subtopology X T) S; closedin (subtopology X U) S\ \ closedin (subtopology X (T \ U)) S" by (simp add: closedin_subtopology) blast lemma openin_trans_full: "\openin (subtopology X U) S; openin X U\ \ openin X S" by (simp add: openin_open_subtopology) subsection \The canonical topology from the underlying type class\ abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" where "top_of_set \ subtopology (topology open)" lemma open_openin: "open S \ openin euclidean S" apply (rule cong[where x=S and y=S]) apply (rule topology_inverse[symmetric]) apply (auto simp: istopology_def) done declare open_openin [symmetric, simp] lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" - by (simp add: topspace_subtopology) + by (simp) lemma closed_closedin: "closed S \ closedin euclidean S" by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) declare closed_closedin [symmetric, simp] lemma openin_subtopology_self [simp]: "openin (top_of_set S) S" by (metis openin_topspace topspace_euclidean_subtopology) subsubsection\The most basic facts about the usual topology and metric on R\ abbreviation euclideanreal :: "real topology" where "euclideanreal \ topology open" subsection \Basic "localization" results are handy for connectedness.\ lemma openin_open: "openin (top_of_set U) S \ (\T. open T \ (S = U \ T))" by (auto simp: openin_subtopology) lemma openin_Int_open: "\openin (top_of_set U) S; open T\ \ openin (top_of_set U) (S \ T)" by (metis open_Int Int_assoc openin_open) lemma openin_open_Int[intro]: "open S \ openin (top_of_set U) (U \ S)" by (auto simp: openin_open) lemma open_openin_trans[trans]: "open S \ open T \ T \ S \ openin (top_of_set S) T" by (metis Int_absorb1 openin_open_Int) lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_closed: "closedin (top_of_set U) S \ (\T. closed T \ S = U \ T)" by (simp add: closedin_subtopology Int_ac) lemma closedin_closed_Int: "closed S \ closedin (top_of_set U) (U \ S)" by (metis closedin_closed) lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: "\closedin (top_of_set U) V; T \ U; S = V \ T\ \ closedin (top_of_set T) S" by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (top_of_set T) S" by (simp add: finite_imp_closed closed_subset) lemma closedin_singleton [simp]: fixes a :: "'a::t1_space" shows "closedin (top_of_set U) {a} \ a \ U" using closedin_subset by (force intro: closed_subset) lemma openin_euclidean_subtopology_iff: fixes S U :: "'a::metric_space set" shows "openin (top_of_set U) S \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs unfolding openin_open open_dist by blast next define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp apply (rule_tac x="d - dist x a" in exI) apply (clarsimp simp add: less_diff_eq) by (metis dist_commute dist_triangle_lt) assume ?rhs then have 2: "S = U \ T" unfolding T_def by auto (metis dist_self) from 1 2 show ?lhs unfolding openin_open open_dist by fast qed lemma connected_openin: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe) by (simp_all, blast+) (* SLOW *) lemma connected_openin_eq: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_openin, safe, blast) by (metis Int_lower1 Un_subset_iff openin_open subset_antisym) lemma connected_closedin: "connected S \ (\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp add: connected_closed closedin_closed) next assume R: ?rhs then show ?lhs proof (clarsimp simp add: connected_closed closedin_closed) fix A B assume s_sub: "S \ A \ B" "B \ S \ {}" and disj: "A \ B \ S = {}" and cl: "closed A" "closed B" have "S \ (A \ B) = S" using s_sub(1) by auto have "S - A = B \ S" using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto then have "S \ A = {}" by (metis Diff_Diff_Int Diff_disjoint Un_Diff_Int R cl closedin_closed_Int inf_commute order_refl s_sub(2)) then show "A \ S = {}" by blast qed qed lemma connected_closedin_eq: "connected S \ \(\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_closedin, safe, blast) by (metis Int_lower1 Un_subset_iff closedin_closed subset_antisym) text \These "transitivity" results are handy too\ lemma openin_trans[trans]: "openin (top_of_set T) S \ openin (top_of_set U) T \ openin (top_of_set U) S" by (metis openin_Int_open openin_open) lemma openin_open_trans: "openin (top_of_set T) S \ open T \ open S" by (auto simp: openin_open intro: openin_trans) lemma closedin_trans[trans]: "closedin (top_of_set T) S \ closedin (top_of_set U) T \ closedin (top_of_set U) S" by (auto simp: closedin_closed closed_Inter Int_assoc) lemma closedin_closed_trans: "closedin (top_of_set T) S \ closed T \ closed S" by (auto simp: closedin_closed intro: closedin_trans) lemma openin_subtopology_Int_subset: "\openin (top_of_set u) (u \ S); v \ u\ \ openin (top_of_set v) (v \ S)" by (auto simp: openin_subtopology) lemma openin_open_eq: "open s \ (openin (top_of_set s) t \ open t \ t \ s)" using open_subset openin_open_trans openin_subset by fastforce subsection\Derived set (set of limit points)\ definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" lemma derived_set_of_restrict [simp]: "X derived_set_of (topspace X \ S) = X derived_set_of S" by (simp add: derived_set_of_def) (metis openin_subset subset_iff) lemma in_derived_set_of: "x \ X derived_set_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))" by (simp add: derived_set_of_def) lemma derived_set_of_subset_topspace: "X derived_set_of S \ topspace X" by (auto simp add: derived_set_of_def) lemma derived_set_of_subtopology: "(subtopology X U) derived_set_of S = U \ (X derived_set_of (U \ S))" - by (simp add: derived_set_of_def openin_subtopology topspace_subtopology) blast + by (simp add: derived_set_of_def openin_subtopology) blast lemma derived_set_of_subset_subtopology: "(subtopology X S) derived_set_of T \ S" by (simp add: derived_set_of_subtopology) lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}" by (auto simp: derived_set_of_def) lemma derived_set_of_mono: "S \ T \ X derived_set_of S \ X derived_set_of T" unfolding derived_set_of_def by blast lemma derived_set_of_Un: "X derived_set_of (S \ T) = X derived_set_of S \ X derived_set_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" apply (clarsimp simp: in_derived_set_of) by (metis IntE IntI openin_Int) show "?rhs \ ?lhs" by (simp add: derived_set_of_mono) qed lemma derived_set_of_Union: "finite \ \ X derived_set_of (\\) = (\S \ \. X derived_set_of S)" proof (induction \ rule: finite_induct) case (insert S \) then show ?case by (simp add: derived_set_of_Un) qed auto lemma derived_set_of_topspace: "X derived_set_of (topspace X) = {x \ topspace X. \ openin X {x}}" apply (auto simp: in_derived_set_of) by (metis Set.set_insert all_not_in_conv insertCI openin_subset subsetCE) lemma discrete_topology_unique_derived_set: "discrete_topology U = X \ topspace X = U \ X derived_set_of U = {}" by (auto simp: discrete_topology_unique derived_set_of_topspace) lemma subtopology_eq_discrete_topology_eq: "subtopology X U = discrete_topology U \ U \ topspace X \ U \ X derived_set_of U = {}" using discrete_topology_unique_derived_set [of U "subtopology X U"] - by (auto simp: eq_commute topspace_subtopology derived_set_of_subtopology) + by (auto simp: eq_commute derived_set_of_subtopology) lemma subtopology_eq_discrete_topology: "S \ topspace X \ S \ X derived_set_of S = {} \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq) lemma subtopology_eq_discrete_topology_gen: "S \ X derived_set_of S = {} \ subtopology X S = discrete_topology(topspace X \ S)" by (metis Int_lower1 derived_set_of_restrict inf_assoc inf_bot_right subtopology_eq_discrete_topology_eq subtopology_subtopology subtopology_topspace) lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \ S)" proof - have "(\T. \Sa. T = Sa \ S \ Sa \ U) = (\Sa. Sa \ U \ Sa \ S)" by force then show ?thesis by (simp add: subtopology_def) (simp add: discrete_topology_def) qed lemma openin_Int_derived_set_of_subset: "openin X S \ S \ X derived_set_of T \ X derived_set_of (S \ T)" by (auto simp: derived_set_of_def) lemma openin_Int_derived_set_of_eq: "openin X S \ S \ X derived_set_of T = S \ X derived_set_of (S \ T)" apply auto apply (meson IntI openin_Int_derived_set_of_subset subsetCE) by (meson derived_set_of_mono inf_sup_ord(2) subset_eq) subsection\ Closure with respect to a topological space\ definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" unfolding closure_of_def apply safe apply (meson IntI openin_subset subset_iff) by auto lemma in_closure_of: "x \ X closure_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y. y \ S \ y \ T))" by (auto simp: closure_of_def) lemma closure_of: "X closure_of S = topspace X \ (S \ X derived_set_of S)" by (fastforce simp: in_closure_of in_derived_set_of) lemma closure_of_alt: "X closure_of S = topspace X \ S \ X derived_set_of S" using derived_set_of_subset_topspace [of X S] unfolding closure_of_def in_derived_set_of by safe (auto simp: in_derived_set_of) lemma derived_set_of_subset_closure_of: "X derived_set_of S \ X closure_of S" by (fastforce simp: closure_of_def in_derived_set_of) lemma closure_of_subtopology: "(subtopology X U) closure_of S = U \ (X closure_of (U \ S))" unfolding closure_of_def topspace_subtopology openin_subtopology by safe (metis (full_types) IntI Int_iff inf.commute)+ lemma closure_of_empty [simp]: "X closure_of {} = {}" by (simp add: closure_of_alt) lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X" by (simp add: closure_of) lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X" by (simp add: closure_of) lemma closure_of_subset_topspace: "X closure_of S \ topspace X" by (simp add: closure_of) lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \ S" by (simp add: closure_of_subtopology) lemma closure_of_mono: "S \ T \ X closure_of S \ X closure_of T" by (fastforce simp add: closure_of_def) lemma closure_of_subtopology_subset: "(subtopology X U) closure_of S \ (X closure_of S)" unfolding closure_of_subtopology by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2) lemma closure_of_subtopology_mono: "T \ U \ (subtopology X T) closure_of S \ (subtopology X U) closure_of S" unfolding closure_of_subtopology by auto (meson closure_of_mono inf_mono subset_iff) lemma closure_of_Un [simp]: "X closure_of (S \ T) = X closure_of S \ X closure_of T" by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1) lemma closure_of_Union: "finite \ \ X closure_of (\\) = (\S \ \. X closure_of S)" by (induction \ rule: finite_induct) auto lemma closure_of_subset: "S \ topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_Int: "topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_eq: "S \ topspace X \ X closure_of S \ S \ closedin X S" proof (cases "S \ topspace X") case True then have "\x. x \ topspace X \ (\T. x \ T \ openin X T \ (\y\S. y \ T)) \ x \ S \ openin X (topspace X - S)" apply (subst openin_subopen, safe) by (metis DiffI subset_eq openin_subset [of X]) then show ?thesis by (auto simp: closedin_def closure_of_def) next case False then show ?thesis by (simp add: closedin_def) qed lemma closure_of_eq: "X closure_of S = S \ closedin X S" proof (cases "S \ topspace X") case True then show ?thesis by (metis closure_of_subset closure_of_subset_eq set_eq_subset) next case False then show ?thesis using closure_of closure_of_subset_eq by fastforce qed lemma closedin_contains_derived_set: "closedin X S \ X derived_set_of S \ S \ S \ topspace X" proof (intro iffI conjI) show "closedin X S \ X derived_set_of S \ S" using closure_of_eq derived_set_of_subset_closure_of by fastforce show "closedin X S \ S \ topspace X" using closedin_subset by blast show "X derived_set_of S \ S \ S \ topspace X \ closedin X S" by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE) qed lemma derived_set_subset_gen: "X derived_set_of S \ S \ closedin X (topspace X \ S)" by (simp add: closedin_contains_derived_set derived_set_of_restrict derived_set_of_subset_topspace) lemma derived_set_subset: "S \ topspace X \ (X derived_set_of S \ S \ closedin X S)" by (simp add: closedin_contains_derived_set) lemma closedin_derived_set: "closedin (subtopology X T) S \ S \ topspace X \ S \ T \ (\x. x \ X derived_set_of S \ x \ T \ x \ S)" - by (auto simp: closedin_contains_derived_set topspace_subtopology derived_set_of_subtopology Int_absorb1) + by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1) lemma closedin_Int_closure_of: "closedin (subtopology X S) T \ S \ X closure_of T = T" by (metis Int_left_absorb closure_of_eq closure_of_subtopology) lemma closure_of_closedin: "closedin X S \ X closure_of S = S" by (simp add: closure_of_eq) lemma closure_of_eq_diff: "X closure_of S = topspace X - \{T. openin X T \ disjnt S T}" by (auto simp: closure_of_def disjnt_iff) lemma closedin_closure_of [simp]: "closedin X (X closure_of S)" unfolding closure_of_eq_diff by blast lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S" by (simp add: closure_of_eq) lemma closure_of_hull: assumes "S \ topspace X" shows "X closure_of S = (closedin X) hull S" proof (rule hull_unique [THEN sym]) show "S \ X closure_of S" by (simp add: closure_of_subset assms) next show "closedin X (X closure_of S)" by simp show "\T. \S \ T; closedin X T\ \ X closure_of S \ T" by (metis closure_of_eq closure_of_mono) qed lemma closure_of_minimal: "\S \ T; closedin X T\ \ (X closure_of S) \ T" by (metis closure_of_eq closure_of_mono) lemma closure_of_minimal_eq: "\S \ topspace X; closedin X T\ \ (X closure_of S) \ T \ S \ T" by (meson closure_of_minimal closure_of_subset subset_trans) lemma closure_of_unique: "\S \ T; closedin X T; \T'. \S \ T'; closedin X T'\ \ T \ T'\ \ X closure_of S = T" by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans) lemma closure_of_eq_empty_gen: "X closure_of S = {} \ disjnt (topspace X) S" unfolding disjnt_def closure_of_restrict [where S=S] using closure_of by fastforce lemma closure_of_eq_empty: "S \ topspace X \ X closure_of S = {} \ S = {}" using closure_of_subset by fastforce lemma openin_Int_closure_of_subset: assumes "openin X S" shows "S \ X closure_of T \ X closure_of (S \ T)" proof - have "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" by (meson assms openin_Int_derived_set_of_eq) moreover have "S \ (S \ T) = S \ T" by fastforce ultimately show ?thesis by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1) qed lemma closure_of_openin_Int_closure_of: assumes "openin X S" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" proof show "X closure_of (S \ X closure_of T) \ X closure_of (S \ T)" by (simp add: assms closure_of_minimal openin_Int_closure_of_subset) next show "X closure_of (S \ T) \ X closure_of (S \ X closure_of T)" by (metis Int_lower1 Int_subset_iff assms closedin_closure_of closure_of_minimal_eq closure_of_mono inf_le2 le_infI1 openin_subset) qed lemma openin_Int_closure_of_eq: "openin X S \ S \ X closure_of T = S \ X closure_of (S \ T)" apply (rule equalityI) apply (simp add: openin_Int_closure_of_subset) by (meson closure_of_mono inf.cobounded2 inf_mono subset_refl) lemma openin_Int_closure_of_eq_empty: "openin X S \ S \ X closure_of T = {} \ S \ T = {}" apply (subst openin_Int_closure_of_eq, auto) by (meson IntI closure_of_subset_Int disjoint_iff_not_equal openin_subset subset_eq) lemma closure_of_openin_Int_superset: "openin X S \ S \ X closure_of T \ X closure_of (S \ T) = X closure_of S" by (metis closure_of_openin_Int_closure_of inf.orderE) lemma closure_of_openin_subtopology_Int_closure_of: assumes S: "openin (subtopology X U) S" and "T \ U" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" (is "?lhs = ?rhs") proof obtain S0 where S0: "openin X S0" "S = S0 \ U" using assms by (auto simp: openin_subtopology) show "?lhs \ ?rhs" proof - have "S0 \ X closure_of T = S0 \ X closure_of (S0 \ T)" by (meson S0(1) openin_Int_closure_of_eq) moreover have "S0 \ T = S0 \ U \ T" using \T \ U\ by fastforce ultimately have "S \ X closure_of T \ X closure_of (S \ T)" using S0(2) by auto then show ?thesis by (meson closedin_closure_of closure_of_minimal) qed next show "?rhs \ ?lhs" proof - have "T \ S \ T \ X derived_set_of T" by force then show ?thesis by (metis Int_subset_iff S closure_of closure_of_mono inf.cobounded2 inf.coboundedI2 inf_commute openin_closedin_eq topspace_subtopology) qed qed lemma closure_of_subtopology_open: "openin X U \ S \ U \ (subtopology X U) closure_of S = U \ X closure_of S" by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq) lemma discrete_topology_closure_of: "(discrete_topology U) closure_of S = U \ S" by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl) text\ Interior with respect to a topological space. \ definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: "X interior_of S = X interior_of (topspace X \ S)" using openin_subset by (auto simp: interior_of_def) lemma interior_of_eq: "(X interior_of S = S) \ openin X S" unfolding interior_of_def using openin_subopen by blast lemma interior_of_openin: "openin X S \ X interior_of S = S" by (simp add: interior_of_eq) lemma interior_of_empty [simp]: "X interior_of {} = {}" by (simp add: interior_of_eq) lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X" by (simp add: interior_of_eq) lemma openin_interior_of [simp]: "openin X (X interior_of S)" unfolding interior_of_def using openin_subopen by fastforce lemma interior_of_interior_of [simp]: "X interior_of X interior_of S = X interior_of S" by (simp add: interior_of_eq) lemma interior_of_subset: "X interior_of S \ S" by (auto simp: interior_of_def) lemma interior_of_subset_closure_of: "X interior_of S \ X closure_of S" by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset) lemma subset_interior_of_eq: "S \ X interior_of S \ openin X S" by (metis interior_of_eq interior_of_subset subset_antisym) lemma interior_of_mono: "S \ T \ X interior_of S \ X interior_of T" by (auto simp: interior_of_def) lemma interior_of_maximal: "\T \ S; openin X T\ \ T \ X interior_of S" by (auto simp: interior_of_def) lemma interior_of_maximal_eq: "openin X T \ T \ X interior_of S \ T \ S" by (meson interior_of_maximal interior_of_subset order_trans) lemma interior_of_unique: "\T \ S; openin X T; \T'. \T' \ S; openin X T'\ \ T' \ T\ \ X interior_of S = T" by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym) lemma interior_of_subset_topspace: "X interior_of S \ topspace X" by (simp add: openin_subset) lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \ S" by (meson openin_imp_subset openin_interior_of) lemma interior_of_Int: "X interior_of (S \ T) = X interior_of S \ X interior_of T" apply (rule equalityI) apply (simp add: interior_of_mono) apply (auto simp: interior_of_maximal_eq openin_Int interior_of_subset le_infI1 le_infI2) done lemma interior_of_Inter_subset: "X interior_of (\\) \ (\S \ \. X interior_of S)" by (simp add: INT_greatest Inf_lower interior_of_mono) lemma union_interior_of_subset: "X interior_of S \ X interior_of T \ X interior_of (S \ T)" by (simp add: interior_of_mono) lemma interior_of_eq_empty: "X interior_of S = {} \ (\T. openin X T \ T \ S \ T = {})" by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of) lemma interior_of_eq_empty_alt: "X interior_of S = {} \ (\T. openin X T \ T \ {} \ T - S \ {})" by (auto simp: interior_of_eq_empty) lemma interior_of_Union_openin_subsets: "\{T. openin X T \ T \ S} = X interior_of S" by (rule interior_of_unique [symmetric]) auto lemma interior_of_complement: "X interior_of (topspace X - S) = topspace X - X closure_of S" by (auto simp: interior_of_def closure_of_def) lemma interior_of_closure_of: "X interior_of S = topspace X - X closure_of (topspace X - S)" unfolding interior_of_complement [symmetric] by (metis Diff_Diff_Int interior_of_restrict) lemma closure_of_interior_of: "X closure_of S = topspace X - X interior_of (topspace X - S)" by (simp add: interior_of_complement Diff_Diff_Int closure_of) lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S" unfolding interior_of_def closure_of_def by (blast dest: openin_subset) lemma interior_of_eq_empty_complement: "X interior_of S = {} \ X closure_of (topspace X - S) = topspace X" using interior_of_subset_topspace [of X S] closure_of_complement by fastforce lemma closure_of_eq_topspace: "X closure_of S = topspace X \ X interior_of (topspace X - S) = {}" using closure_of_subset_topspace [of X S] interior_of_complement by fastforce lemma interior_of_subtopology_subset: "U \ X interior_of S \ (subtopology X U) interior_of S" by (auto simp: interior_of_def openin_subtopology) lemma interior_of_subtopology_subsets: "T \ U \ T \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology) lemma interior_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets) lemma interior_of_subtopology_open: assumes "openin X U" shows "(subtopology X U) interior_of S = U \ X interior_of S" proof - have "\A. U \ X closure_of (U \ A) = U \ X closure_of A" using assms openin_Int_closure_of_eq by blast then have "topspace X \ U - U \ X closure_of (topspace X \ U - S) = U \ (topspace X - X closure_of (topspace X - S))" by (metis (no_types) Diff_Int_distrib Int_Diff inf_commute) then show ?thesis unfolding interior_of_closure_of closure_of_subtopology_open topspace_subtopology using openin_Int_closure_of_eq [OF assms] by (metis assms closure_of_subtopology_open) qed lemma dense_intersects_open: "X closure_of S = topspace X \ (\T. openin X T \ T \ {} \ S \ T \ {})" proof - have "X closure_of S = topspace X \ (topspace X - X interior_of (topspace X - S) = topspace X)" by (simp add: closure_of_interior_of) also have "\ \ X interior_of (topspace X - S) = {}" by (simp add: closure_of_complement interior_of_eq_empty_complement) also have "\ \ (\T. openin X T \ T \ {} \ S \ T \ {})" unfolding interior_of_eq_empty_alt using openin_subset by fastforce finally show ?thesis . qed lemma interior_of_closedin_union_empty_interior_of: assumes "closedin X S" and disj: "X interior_of T = {}" shows "X interior_of (S \ T) = X interior_of S" proof - have "X closure_of (topspace X - T) = topspace X" by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of) then show ?thesis unfolding interior_of_closure_of by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset) qed lemma interior_of_union_eq_empty: "closedin X S \ (X interior_of (S \ T) = {} \ X interior_of S = {} \ X interior_of T = {})" by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset) lemma discrete_topology_interior_of [simp]: "(discrete_topology U) interior_of S = U \ S" by (simp add: interior_of_restrict [of _ S] interior_of_eq) subsection \Frontier with respect to topological space \ definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: "X frontier_of S = X closure_of S \ X closure_of (topspace X - S)" by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of) lemma interior_of_union_frontier_of [simp]: "X interior_of S \ X frontier_of S = X closure_of S" by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym) lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \ S)" by (metis closure_of_restrict frontier_of_def interior_of_restrict) lemma closedin_frontier_of: "closedin X (X frontier_of S)" by (simp add: closedin_Int frontier_of_closures) lemma frontier_of_subset_topspace: "X frontier_of S \ topspace X" by (simp add: closedin_frontier_of closedin_subset) lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \ S" by (metis (no_types) closedin_derived_set closedin_frontier_of) lemma frontier_of_subtopology_subset: "U \ (subtopology X U) frontier_of S \ (X frontier_of S)" proof - have "U \ X interior_of S - subtopology X U interior_of S = {}" by (simp add: interior_of_subtopology_subset) moreover have "X closure_of S \ subtopology X U closure_of S = subtopology X U closure_of S" by (meson closure_of_subtopology_subset inf.absorb_iff2) ultimately show ?thesis unfolding frontier_of_def by blast qed lemma frontier_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X T) frontier_of S \ (subtopology X U) frontier_of S" by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono) lemma clopenin_eq_frontier_of: "closedin X S \ openin X S \ S \ topspace X \ X frontier_of S = {}" proof (cases "S \ topspace X") case True then show ?thesis by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right) next case False then show ?thesis by (simp add: frontier_of_closures openin_closedin_eq) qed lemma frontier_of_eq_empty: "S \ topspace X \ (X frontier_of S = {} \ closedin X S \ openin X S)" by (simp add: clopenin_eq_frontier_of) lemma frontier_of_openin: "openin X S \ X frontier_of S = X closure_of S - S" by (metis (no_types) frontier_of_def interior_of_eq) lemma frontier_of_openin_straddle_Int: assumes "openin X U" "U \ X frontier_of S \ {}" shows "U \ S \ {}" "U - S \ {}" proof - have "U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}" using assms by (simp add: frontier_of_closures) then show "U \ S \ {}" using assms openin_Int_closure_of_eq_empty by fastforce show "U - S \ {}" proof - have "\A. X closure_of (A - S) \ U \ {}" using \U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}\ by blast then have "\ U \ S" by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty) then show ?thesis by blast qed qed lemma frontier_of_subset_closedin: "closedin X S \ (X frontier_of S) \ S" using closure_of_eq frontier_of_def by fastforce lemma frontier_of_empty [simp]: "X frontier_of {} = {}" by (simp add: frontier_of_def) lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}" by (simp add: frontier_of_def) lemma frontier_of_subset_eq: assumes "S \ topspace X" shows "(X frontier_of S) \ S \ closedin X S" proof show "X frontier_of S \ S \ closedin X S" by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff) show "closedin X S \ X frontier_of S \ S" by (simp add: frontier_of_subset_closedin) qed lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S" by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute) lemma frontier_of_disjoint_eq: assumes "S \ topspace X" shows "((X frontier_of S) \ S = {} \ openin X S)" proof assume "X frontier_of S \ S = {}" then have "closedin X (topspace X - S)" using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce then show "openin X S" using assms by (simp add: openin_closedin) next show "openin X S \ X frontier_of S \ S = {}" by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute) qed lemma frontier_of_disjoint_eq_alt: "S \ (topspace X - X frontier_of S) \ openin X S" proof (cases "S \ topspace X") case True show ?thesis using True frontier_of_disjoint_eq by auto next case False then show ?thesis by (meson Diff_subset openin_subset subset_trans) qed lemma frontier_of_Int: "X frontier_of (S \ T) = X closure_of (S \ T) \ (X frontier_of S \ X frontier_of T)" proof - have *: "U \ S \ U \ T \ U \ (S \ A \ T \ B) = U \ (A \ B)" for U S T A B :: "'a set" by blast show ?thesis by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un) qed lemma frontier_of_Int_subset: "X frontier_of (S \ T) \ X frontier_of S \ X frontier_of T" by (simp add: frontier_of_Int) lemma frontier_of_Int_closedin: "\closedin X S; closedin X T\ \ X frontier_of(S \ T) = X frontier_of S \ T \ S \ X frontier_of T" apply (simp add: frontier_of_Int closedin_Int closure_of_closedin) using frontier_of_subset_closedin by blast lemma frontier_of_Un_subset: "X frontier_of(S \ T) \ X frontier_of S \ X frontier_of T" by (metis Diff_Un frontier_of_Int_subset frontier_of_complement) lemma frontier_of_Union_subset: "finite \ \ X frontier_of (\\) \ (\T \ \. X frontier_of T)" proof (induction \ rule: finite_induct) case (insert A \) then show ?case using frontier_of_Un_subset by fastforce qed simp lemma frontier_of_frontier_of_subset: "X frontier_of (X frontier_of S) \ X frontier_of S" by (simp add: closedin_frontier_of frontier_of_subset_closedin) lemma frontier_of_subtopology_open: "openin X U \ (subtopology X U) frontier_of S = U \ X frontier_of S" by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open) lemma discrete_topology_frontier_of [simp]: "(discrete_topology U) frontier_of S = {}" by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures) subsection\Locally finite collections\ definition locally_finite_in where "locally_finite_in X \ \ (\\ \ topspace X) \ (\x \ topspace X. \V. openin X V \ x \ V \ finite {U \ \. U \ V \ {}})" lemma finite_imp_locally_finite_in: "\finite \; \\ \ topspace X\ \ locally_finite_in X \" by (auto simp: locally_finite_in_def) lemma locally_finite_in_subset: assumes "locally_finite_in X \" "\ \ \" shows "locally_finite_in X \" proof - have "finite {U \ \. U \ V \ {}} \ finite {U \ \. U \ V \ {}}" for V apply (erule rev_finite_subset) using \\ \ \\ by blast then show ?thesis using assms unfolding locally_finite_in_def by (fastforce simp add:) qed lemma locally_finite_in_refinement: assumes \: "locally_finite_in X \" and f: "\S. S \ \ \ f S \ S" shows "locally_finite_in X (f ` \)" proof - show ?thesis unfolding locally_finite_in_def proof safe fix x assume "x \ topspace X" then obtain V where "openin X V" "x \ V" "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast moreover have "{U \ \. f U \ V \ {}} \ {U \ \. U \ V \ {}}" for V using f by blast ultimately have "finite {U \ \. f U \ V \ {}}" using finite_subset by blast moreover have "f ` {U \ \. f U \ V \ {}} = {U \ f ` \. U \ V \ {}}" by blast ultimately have "finite {U \ f ` \. U \ V \ {}}" by (metis (no_types, lifting) finite_imageI) then show "\V. openin X V \ x \ V \ finite {U \ f ` \. U \ V \ {}}" using \openin X V\ \x \ V\ by blast next show "\x xa. \xa \ \; x \ f xa\ \ x \ topspace X" by (meson Sup_upper \ f locally_finite_in_def subset_iff) qed qed lemma locally_finite_in_subtopology: assumes \: "locally_finite_in X \" "\\ \ S" shows "locally_finite_in (subtopology X S) \" unfolding locally_finite_in_def proof safe fix x assume x: "x \ topspace (subtopology X S)" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def topspace_subtopology by blast show "\V. openin (subtopology X S) V \ x \ V \ finite {U \ \. U \ V \ {}}" proof (intro exI conjI) show "openin (subtopology X S) (S \ V)" by (simp add: \openin X V\ openin_subtopology_Int2) have "{U \ \. U \ (S \ V) \ {}} \ {U \ \. U \ V \ {}}" by auto with fin show "finite {U \ \. U \ (S \ V) \ {}}" using finite_subset by auto show "x \ S \ V" - using x \x \ V\ by (simp add: topspace_subtopology) + using x \x \ V\ by (simp) qed next show "\x A. \x \ A; A \ \\ \ x \ topspace (subtopology X S)" using assms unfolding locally_finite_in_def topspace_subtopology by blast qed lemma closedin_locally_finite_Union: assumes clo: "\S. S \ \ \ closedin X S" and \: "locally_finite_in X \" shows "closedin X (\\)" using \ unfolding locally_finite_in_def closedin_def proof clarify show "openin X (topspace X - \\)" proof (subst openin_subopen, clarify) fix x assume "x \ topspace X" and "x \ \\" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast let ?T = "V - \{S \ \. S \ V \ {}}" show "\T. openin X T \ x \ T \ T \ topspace X - \\" proof (intro exI conjI) show "openin X ?T" by (metis (no_types, lifting) fin \openin X V\ clo closedin_Union mem_Collect_eq openin_diff) show "x \ ?T" using \x \ \\\ \x \ V\ by auto show "?T \ topspace X - \\" using \openin X V\ openin_subset by auto qed qed qed lemma locally_finite_in_closure: assumes \: "locally_finite_in X \" shows "locally_finite_in X ((\S. X closure_of S) ` \)" using \ unfolding locally_finite_in_def proof (intro conjI; clarsimp) fix x A assume "x \ X closure_of A" then show "x \ topspace X" by (meson in_closure_of) next fix x assume "x \ topspace X" and "\\ \ topspace X" then obtain V where V: "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast have eq: "{y \ f ` \. Q y} = f ` {x. x \ \ \ Q(f x)}" for f Q by blast have eq2: "{A \ \. X closure_of A \ V \ {}} = {A \ \. A \ V \ {}}" using openin_Int_closure_of_eq_empty V by blast have "finite {U \ (closure_of) X ` \. U \ V \ {}}" by (simp add: eq eq2 fin) with V show "\V. openin X V \ x \ V \ finite {U \ (closure_of) X ` \. U \ V \ {}}" by blast qed lemma closedin_Union_locally_finite_closure: "locally_finite_in X \ \ closedin X (\((\S. X closure_of S) ` \))" by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" by clarify (meson Union_upper closure_of_mono subsetD) lemma closure_of_locally_finite_Union: "locally_finite_in X \ \ X closure_of (\\) = \((\S. X closure_of S) ` \)" apply (rule closure_of_unique) apply (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def) apply (simp add: closedin_Union_locally_finite_closure) by (simp add: Sup_le_iff closure_of_minimal) subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ definition continuous_map where "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" lemma continuous_map: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" by (auto simp: continuous_map_def) lemma continuous_map_image_subset_topspace: "continuous_map X Y f \ f ` (topspace X) \ topspace Y" by (auto simp: continuous_map_def) lemma continuous_map_on_empty: "topspace X = {} \ continuous_map X Y f" by (auto simp: continuous_map_def) lemma continuous_map_closedin: "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" proof - have "(\U. openin Y U \ openin X {x \ topspace X. f x \ U}) = (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" if "\x. x \ topspace X \ f x \ topspace Y" proof - have eq: "{x \ topspace X. f x \ topspace Y \ f x \ C} = (topspace X - {x \ topspace X. f x \ C})" for C using that by blast show ?thesis proof (intro iffI allI impI) fix C assume "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and "closedin Y C" then have "openin X {x \ topspace X. f x \ topspace Y - C}" by blast then show "closedin X {x \ topspace X. f x \ C}" by (auto simp add: closedin_def eq) next fix U assume "\C. closedin Y C \ closedin X {x \ topspace X. f x \ C}" and "openin Y U" then have "closedin X {x \ topspace X. f x \ topspace Y - U}" by blast then show "openin X {x \ topspace X. f x \ U}" by (auto simp add: openin_closedin_eq eq) qed qed then show ?thesis by (auto simp: continuous_map_def) qed lemma openin_continuous_map_preimage: "\continuous_map X Y f; openin Y U\ \ openin X {x \ topspace X. f x \ U}" by (simp add: continuous_map_def) lemma closedin_continuous_map_preimage: "\continuous_map X Y f; closedin Y C\ \ closedin X {x \ topspace X. f x \ C}" by (simp add: continuous_map_closedin) lemma openin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "openin X U" "openin Y V" shows "openin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) openin_closedin_eq by fastforce show ?thesis unfolding eq using assms openin_continuous_map_preimage by fastforce qed lemma closedin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "closedin X U" "closedin Y V" shows "closedin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) closedin_def by fastforce show ?thesis unfolding eq using assms closedin_continuous_map_preimage by fastforce qed lemma continuous_map_image_closure_subset: assumes "continuous_map X Y f" shows "f ` (X closure_of S) \ Y closure_of f ` S" proof - have *: "f ` (topspace X) \ topspace Y" by (meson assms continuous_map) have "X closure_of T \ {x \ X closure_of T. f x \ Y closure_of (f ` T)}" if "T \ topspace X" for T proof (rule closure_of_minimal) show "T \ {x \ X closure_of T. f x \ Y closure_of f ` T}" using closure_of_subset * that by (fastforce simp: in_closure_of) next show "closedin X {x \ X closure_of T. f x \ Y closure_of f ` T}" using assms closedin_continuous_map_preimage_gen by fastforce qed then have "f ` (X closure_of (topspace X \ S)) \ Y closure_of (f ` (topspace X \ S))" by blast also have "\ \ Y closure_of (topspace Y \ f ` S)" using * by (blast intro!: closure_of_mono) finally have "f ` (X closure_of (topspace X \ S)) \ Y closure_of (topspace Y \ f ` S)" . then show ?thesis by (metis closure_of_restrict) qed lemma continuous_map_subset_aux1: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_image_closure_subset by blast lemma continuous_map_subset_aux2: assumes "\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S" shows "continuous_map X Y f" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "f x \ topspace Y" using assms closure_of_subset_topspace by fastforce next fix C assume "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" proof (clarsimp simp flip: closure_of_subset_eq, intro conjI) fix x assume x: "x \ X closure_of {x \ topspace X. f x \ C}" and "C \ topspace Y" and "Y closure_of C \ C" show "x \ topspace X" by (meson x in_closure_of) have "{a \ topspace X. f a \ C} \ topspace X" by simp moreover have "Y closure_of f ` {a \ topspace X. f a \ C} \ C" by (simp add: \closedin Y C\ closure_of_minimal image_subset_iff) ultimately have "f ` (X closure_of {a \ topspace X. f a \ C}) \ C" using assms by blast then show "f x \ C" using x by auto qed qed lemma continuous_map_eq_image_closure_subset: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_alt: "continuous_map X Y f \ (\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_gen: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis lemma continuous_map_closure_preimage_subset: "continuous_map X Y f \ X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" unfolding continuous_map_closedin by (rule closure_of_minimal) (use in_closure_of in \fastforce+\) lemma continuous_map_frontier_frontier_preimage_subset: assumes "continuous_map X Y f" shows "X frontier_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y frontier_of T}" proof - have eq: "topspace X - {x \ topspace X. f x \ T} = {x \ topspace X. f x \ topspace Y - T}" using assms unfolding continuous_map_def by blast have "X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" by (simp add: assms continuous_map_closure_preimage_subset) moreover have "X closure_of (topspace X - {x \ topspace X. f x \ T}) \ {x \ topspace X. f x \ Y closure_of (topspace Y - T)}" using continuous_map_closure_preimage_subset [OF assms] eq by presburger ultimately show ?thesis by (auto simp: frontier_of_closures) qed lemma topology_finer_continuous_id: "topspace X = topspace Y \ ((\S. openin X S \ openin Y S) \ continuous_map Y X id)" unfolding continuous_map_def apply auto using openin_subopen openin_subset apply fastforce using openin_subopen topspace_def by fastforce lemma continuous_map_const [simp]: "continuous_map X Y (\x. C) \ topspace X = {} \ C \ topspace Y" proof (cases "topspace X = {}") case False show ?thesis proof (cases "C \ topspace Y") case True with openin_subopen show ?thesis by (auto simp: continuous_map_def) next case False then show ?thesis unfolding continuous_map_def by fastforce qed qed (auto simp: continuous_map_on_empty) declare continuous_map_const [THEN iffD2, continuous_intros] lemma continuous_map_compose [continuous_intros]: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" shows "continuous_map X X'' (g \ f)" unfolding continuous_map_def proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "(g \ f) x \ topspace X''" using assms unfolding continuous_map_def by force next fix U assume "openin X'' U" have eq: "{x \ topspace X. (g \ f) x \ U} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U}}" by auto (meson f continuous_map_def) show "openin X {x \ topspace X. (g \ f) x \ U}" unfolding eq using assms unfolding continuous_map_def using \openin X'' U\ by blast qed lemma continuous_map_eq: assumes "continuous_map X X' f" and "\x. x \ topspace X \ f x = g x" shows "continuous_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms by (simp add: continuous_map_def eq) qed lemma restrict_continuous_map [simp]: "topspace X \ S \ continuous_map X X' (restrict f S) \ continuous_map X X' f" by (auto simp: elim!: continuous_map_eq) lemma continuous_map_in_subtopology: "continuous_map X (subtopology X' S) f \ continuous_map X X' f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof - have "\A. f ` (X closure_of A) \ subtopology X' S closure_of f ` A" by (meson L continuous_map_image_closure_subset) then show ?thesis by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset dual_order.trans) qed next assume R: ?rhs then have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. f x \ U \ f x \ S}" for U by auto show ?lhs using R unfolding continuous_map - by (auto simp: topspace_subtopology openin_subtopology eq) + by (auto simp: openin_subtopology eq) qed lemma continuous_map_from_subtopology: "continuous_map X X' f \ continuous_map (subtopology X S) X' f" - by (auto simp: continuous_map topspace_subtopology openin_subtopology) + by (auto simp: continuous_map openin_subtopology) lemma continuous_map_into_fulltopology: "continuous_map X (subtopology X' T) f \ continuous_map X X' f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_into_subtopology: "\continuous_map X X' f; f ` topspace X \ T\ \ continuous_map X (subtopology X' T) f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_from_subtopology_mono: "\continuous_map (subtopology X T) X' f; S \ T\ \ continuous_map (subtopology X S) X' f" by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology) lemma continuous_map_from_discrete_topology [simp]: "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g" by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant) lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g" by (metis continuous_map_iff_continuous subtopology_UNIV) lemma continuous_map_openin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X (topspace X \ f -` U))" by (auto simp: continuous_map_def vimage_def Int_def) lemma continuous_map_closedin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. closedin Y U \ closedin X (topspace X \ f -` U))" by (auto simp: continuous_map_closedin vimage_def Int_def) lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt" by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt) lemma continuous_map_sqrt [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. sqrt(f x))" by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply) lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id" unfolding continuous_map_def using openin_subopen topspace_def by fastforce declare continuous_map_id [unfolded id_def, simp, continuous_intros] lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id" by (simp add: continuous_map_from_subtopology) declare continuous_map_id_subt [unfolded id_def, simp] lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ topspace T2)" by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute) lemma continuous_map_open [intro]: "continuous_map T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" unfolding continuous_map_alt by auto lemma continuous_map_preimage_topspace [intro]: assumes "continuous_map T1 T2 f" shows "f-`(topspace T2) \ topspace T1 = topspace T1" using assms unfolding continuous_map_def by auto subsection\Open and closed maps (not a priori assumed continuous)\ definition open_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "open_map X1 X2 f \ \U. openin X1 U \ openin X2 (f ` U)" definition closed_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "closed_map X1 X2 f \ \U. closedin X1 U \ closedin X2 (f ` U)" lemma open_map_imp_subset_topspace: "open_map X1 X2 f \ f ` (topspace X1) \ topspace X2" unfolding open_map_def by (simp add: openin_subset) lemma open_map_on_empty: "topspace X = {} \ open_map X Y f" by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset) lemma closed_map_on_empty: "topspace X = {} \ closed_map X Y f" by (simp add: closed_map_def closedin_topspace_empty) lemma closed_map_const: "closed_map X Y (\x. c) \ topspace X = {} \ closedin Y {c}" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: closed_map_on_empty) next case False then show ?thesis by (auto simp: closed_map_def image_constant_conv) qed lemma open_map_imp_subset: "\open_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" by (meson order_trans open_map_imp_subset_topspace subset_image_iff) lemma topology_finer_open_id: "(\S. openin X S \ openin X' S) \ open_map X X' id" unfolding open_map_def by auto lemma open_map_id: "open_map X X id" unfolding open_map_def by auto lemma open_map_eq: "\open_map X X' f; \x. x \ topspace X \ f x = g x\ \ open_map X X' g" unfolding open_map_def by (metis image_cong openin_subset subset_iff) lemma open_map_inclusion_eq: "open_map (subtopology X S) X id \ openin X (topspace X \ S)" proof - have *: "openin X (T \ S)" if "openin X (S \ topspace X)" "openin X T" for T proof - have "T \ topspace X" using that by (simp add: openin_subset) with that show "openin X (T \ S)" by (metis inf.absorb1 inf.left_commute inf_commute openin_Int) qed show ?thesis by (fastforce simp add: open_map_def Int_commute openin_subtopology_alt intro: *) qed lemma open_map_inclusion: "openin X S \ open_map (subtopology X S) X id" by (simp add: open_map_inclusion_eq openin_Int) lemma open_map_compose: "\open_map X X' f; open_map X' X'' g\ \ open_map X X'' (g \ f)" by (metis (no_types, lifting) image_comp open_map_def) lemma closed_map_imp_subset_topspace: "closed_map X1 X2 f \ f ` (topspace X1) \ topspace X2" by (simp add: closed_map_def closedin_subset) lemma closed_map_imp_subset: "\closed_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" using closed_map_imp_subset_topspace by blast lemma topology_finer_closed_id: "(\S. closedin X S \ closedin X' S) \ closed_map X X' id" by (simp add: closed_map_def) lemma closed_map_id: "closed_map X X id" by (simp add: closed_map_def) lemma closed_map_eq: "\closed_map X X' f; \x. x \ topspace X \ f x = g x\ \ closed_map X X' g" unfolding closed_map_def by (metis image_cong closedin_subset subset_iff) lemma closed_map_compose: "\closed_map X X' f; closed_map X' X'' g\ \ closed_map X X'' (g \ f)" by (metis (no_types, lifting) closed_map_def image_comp) lemma closed_map_inclusion_eq: "closed_map (subtopology X S) X id \ closedin X (topspace X \ S)" proof - have *: "closedin X (T \ S)" if "closedin X (S \ topspace X)" "closedin X T" for T proof - have "T \ topspace X" using that by (simp add: closedin_subset) with that show "closedin X (T \ S)" by (metis inf.absorb1 inf.left_commute inf_commute closedin_Int) qed show ?thesis by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *) qed lemma closed_map_inclusion: "closedin X S \ closed_map (subtopology X S) X id" by (simp add: closed_map_inclusion_eq closedin_Int) lemma open_map_into_subtopology: "\open_map X X' f; f ` topspace X \ S\ \ open_map X (subtopology X' S) f" unfolding open_map_def openin_subtopology using openin_subset by fastforce lemma closed_map_into_subtopology: "\closed_map X X' f; f ` topspace X \ S\ \ closed_map X (subtopology X' S) f" unfolding closed_map_def closedin_subtopology using closedin_subset by fastforce lemma open_map_into_discrete_topology: "open_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding open_map_def openin_discrete_topology using openin_subset by blast lemma closed_map_into_discrete_topology: "closed_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast lemma bijective_open_imp_closed_map: "\open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ closed_map X X' f" unfolding open_map_def closed_map_def closedin_def by auto (metis Diff_subset inj_on_image_set_diff) lemma bijective_closed_imp_open_map: "\closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ open_map X X' f" unfolding closed_map_def open_map_def openin_closedin_eq by auto (metis Diff_subset inj_on_image_set_diff) lemma open_map_from_subtopology: "\open_map X X' f; openin X U\ \ open_map (subtopology X U) X' f" unfolding open_map_def openin_subtopology_alt by blast lemma closed_map_from_subtopology: "\closed_map X X' f; closedin X U\ \ closed_map (subtopology X U) X' f" unfolding closed_map_def closedin_subtopology_alt by blast lemma open_map_restriction: "\open_map X X' f; {x. x \ topspace X \ f x \ V} = U\ \ open_map (subtopology X U) (subtopology X' V) f" unfolding open_map_def openin_subtopology_alt apply clarify apply (rename_tac T) apply (rule_tac x="f ` T" in image_eqI) using openin_closedin_eq by fastforce+ lemma closed_map_restriction: "\closed_map X X' f; {x. x \ topspace X \ f x \ V} = U\ \ closed_map (subtopology X U) (subtopology X' V) f" unfolding closed_map_def closedin_subtopology_alt apply clarify apply (rename_tac T) apply (rule_tac x="f ` T" in image_eqI) using closedin_def by fastforce+ subsection\Quotient maps\ definition quotient_map where "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (openin X {x. x \ topspace X \ f x \ U} \ openin X' U))" lemma quotient_map_eq: assumes "quotient_map X X' f" "\x. x \ topspace X \ f x = g x" shows "quotient_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms unfolding quotient_map_def by (metis (mono_tags, lifting) eq image_cong) qed lemma quotient_map_compose: assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g" shows "quotient_map X X'' (g \ f)" unfolding quotient_map_def proof (intro conjI allI impI) show "(g \ f) ` topspace X = topspace X''" using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def) next fix U'' assume "U'' \ topspace X''" define U' where "U' \ {y \ topspace X'. g y \ U''}" have "U' \ topspace X'" by (auto simp add: U'_def) then have U': "openin X {x \ topspace X. f x \ U'} = openin X' U'" using assms unfolding quotient_map_def by simp have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U''} = {x \ topspace X. (g \ f) x \ U''}" using f quotient_map_def by fastforce have "openin X {x \ topspace X. (g \ f) x \ U''} = openin X {x \ topspace X. f x \ U'}" using assms by (simp add: quotient_map_def U'_def eq) also have "\ = openin X'' U''" using U'_def \U'' \ topspace X''\ U' g quotient_map_def by fastforce finally show "openin X {x \ topspace X. (g \ f) x \ U''} = openin X'' U''" . qed lemma quotient_map_from_composition: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \ f)" shows "quotient_map X' X'' g" unfolding quotient_map_def proof (intro conjI allI impI) show "g ` topspace X' = topspace X''" using assms unfolding continuous_map_def quotient_map_def by fastforce next fix U'' :: "'c set" assume U'': "U'' \ topspace X''" have eq: "{x \ topspace X. g (f x) \ U''} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U''}}" using continuous_map_def f by fastforce show "openin X' {x \ topspace X'. g x \ U''} = openin X'' U''" using assms unfolding continuous_map_def quotient_map_def by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq) qed lemma quotient_imp_continuous_map: "quotient_map X X' f \ continuous_map X X' f" by (simp add: continuous_map openin_subset quotient_map_def) lemma quotient_imp_surjective_map: "quotient_map X X' f \ f ` (topspace X) = topspace X'" by (simp add: quotient_map_def) lemma quotient_map_closedin: "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (closedin X {x. x \ topspace X \ f x \ U} \ closedin X' U))" proof - have eq: "(topspace X - {x \ topspace X. f x \ U'}) = {x \ topspace X. f x \ topspace X' \ f x \ U'}" if "f ` topspace X = topspace X'" "U' \ topspace X'" for U' using that by auto have "(\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U) = (\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U)" if "f ` topspace X = topspace X'" proof (rule iffI; intro allI impI subsetI) fix U' assume *[rule_format]: "\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U" and U': "U' \ topspace X'" show "closedin X {x \ topspace X. f x \ U'} = closedin X' U'" using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that]) next fix U' :: "'b set" assume *[rule_format]: "\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U" and U': "U' \ topspace X'" show "openin X {x \ topspace X. f x \ U'} = openin X' U'" using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that]) qed then show ?thesis unfolding quotient_map_def by force qed lemma continuous_open_imp_quotient_map: assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - { fix U assume U: "U \ topspace X'" and "openin X {x \ topspace X. f x \ U}" then have ope: "openin X' (f ` {x \ topspace X. f x \ U})" using om unfolding open_map_def by blast then have "openin X' U" using U feq by (subst openin_subopen) force } moreover have "openin X {x \ topspace X. f x \ U}" if "U \ topspace X'" and "openin X' U" for U using that assms unfolding continuous_map_def by blast ultimately show ?thesis unfolding quotient_map_def using assms by blast qed lemma continuous_closed_imp_quotient_map: assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - have "f ` {x \ topspace X. f x \ U} = U" if "U \ topspace X'" for U using that feq by auto with assms show ?thesis unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto qed lemma continuous_open_quotient_map: "\continuous_map X X' f; open_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_open_imp_quotient_map quotient_map_def) lemma continuous_closed_quotient_map: "\continuous_map X X' f; closed_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_closed_imp_quotient_map quotient_map_def) lemma injective_quotient_map: assumes "inj_on f (topspace X)" shows "quotient_map X X' f \ continuous_map X X' f \ open_map X X' f \ closed_map X X' f \ f ` (topspace X) = topspace X'" (is "?lhs = ?rhs") proof assume L: ?lhs have "open_map X X' f" proof (clarsimp simp add: open_map_def) fix U assume "openin X U" then have "U \ topspace X" by (simp add: openin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "openin X' (f ` U)" using L unfolding quotient_map_def by (metis (no_types, lifting) Collect_cong \openin X U\ image_mono) qed moreover have "closed_map X X' f" proof (clarsimp simp add: closed_map_def) fix U assume "closedin X U" then have "U \ topspace X" by (simp add: closedin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "closedin X' (f ` U)" using L unfolding quotient_map_closedin by (metis (no_types, lifting) Collect_cong \closedin X U\ image_mono) qed ultimately show ?rhs using L by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map) next assume ?rhs then show ?lhs by (simp add: continuous_closed_imp_quotient_map) qed lemma continuous_compose_quotient_map: assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \ f)" shows "continuous_map X' X'' g" unfolding quotient_map_def continuous_map_def proof (intro conjI ballI allI impI) show "\x'. x' \ topspace X' \ g x' \ topspace X''" using assms unfolding quotient_map_def by (metis (no_types, hide_lams) continuous_map_image_subset_topspace image_comp image_subset_iff) next fix U'' :: "'c set" assume U'': "openin X'' U''" have "f ` topspace X = topspace X'" by (simp add: f quotient_imp_surjective_map) then have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U} = {x \ topspace X. g (f x) \ U}" for U by auto have "openin X {x \ topspace X. f x \ topspace X' \ g (f x) \ U''}" unfolding eq using U'' g openin_continuous_map_preimage by fastforce then have *: "openin X {x \ topspace X. f x \ {x \ topspace X'. g x \ U''}}" by auto show "openin X' {x \ topspace X'. g x \ U''}" using f unfolding quotient_map_def by (metis (no_types) Collect_subset *) qed lemma continuous_compose_quotient_map_eq: "quotient_map X X' f \ continuous_map X X'' (g \ f) \ continuous_map X' X'' g" using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast lemma quotient_map_compose_eq: "quotient_map X X' f \ quotient_map X X'' (g \ f) \ quotient_map X' X'' g" apply safe apply (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_from_composition) by (simp add: quotient_map_compose) lemma quotient_map_restriction: assumes quo: "quotient_map X Y f" and U: "{x \ topspace X. f x \ V} = U" and disj: "openin Y V \ closedin Y V" shows "quotient_map (subtopology X U) (subtopology Y V) f" using disj proof assume V: "openin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: openin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" using quo unfolding quotient_map_def by auto have "openin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_def proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" - using sub U fim by (auto simp: topspace_subtopology) + using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" - by (simp_all add: topspace_subtopology) + by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "openin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = openin (subtopology Y V) Y'" using U V Y \openin X U\ \Y' \ topspace Y\ \Y' \ V\ - by (simp add: topspace_subtopology openin_open_subtopology eq) (auto simp: openin_closedin_eq) + by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq) qed next assume V: "closedin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: closedin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin X {x \ topspace X. f x \ U} = closedin Y U" using quo unfolding quotient_map_closedin by auto have "closedin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_closedin proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" - using sub U fim by (auto simp: topspace_subtopology) + using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" - by (simp_all add: topspace_subtopology) + by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "closedin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = closedin (subtopology Y V) Y'" using U V Y \closedin X U\ \Y' \ topspace Y\ \Y' \ V\ - by (simp add: topspace_subtopology closedin_closed_subtopology eq) (auto simp: closedin_def) + by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def) qed qed lemma quotient_map_saturated_open: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. openin X U \ {x \ topspace X. f x \ f ` U} \ U \ openin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin Y U = openin X {x \ topspace X. f x \ U}" unfolding quotient_map_def by auto show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "openin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim openin_subset by fastforce+ show "openin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then have YX: "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \openin X U; {x \ topspace X. f x \ f ` U} \ U\ \ openin Y (f ` U)" by (auto simp: quotient_map_def continuous_map_def) show ?lhs proof (simp add: quotient_map_def fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "openin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "openin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "openin Y U" show "openin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" where "separatedin X S T \ S \ topspace X \ T \ topspace X \ S \ X closure_of T = {} \ T \ X closure_of S = {}" lemma separatedin_empty [simp]: "separatedin X S {} \ S \ topspace X" "separatedin X {} S \ S \ topspace X" by (simp_all add: separatedin_def) lemma separatedin_refl [simp]: "separatedin X S S \ S = {}" proof - have "\x. \separatedin X S S; x \ S\ \ False" by (metis all_not_in_conv closure_of_subset inf.orderE separatedin_def) then show ?thesis by auto qed lemma separatedin_sym: "separatedin X S T \ separatedin X T S" by (auto simp: separatedin_def) lemma separatedin_imp_disjoint: "separatedin X S T \ disjnt S T" by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def) lemma separatedin_mono: "\separatedin X S T; S' \ S; T' \ T\ \ separatedin X S' T'" unfolding separatedin_def using closure_of_mono by blast lemma separatedin_open_sets: "\openin X S; openin X T\ \ separatedin X S T \ disjnt S T" unfolding disjnt_def separatedin_def by (auto simp: openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closed_sets: "\closedin X S; closedin X T\ \ separatedin X S T \ disjnt S T" by (metis closedin_def closure_of_eq disjnt_def inf_commute separatedin_def) lemma separatedin_subtopology: "separatedin (subtopology X U) S T \ S \ U \ T \ U \ separatedin X S T" - apply (simp add: separatedin_def closure_of_subtopology topspace_subtopology) + apply (simp add: separatedin_def closure_of_subtopology) apply (safe; metis Int_absorb1 inf.assoc inf.orderE insert_disjoint(2) mk_disjoint_insert) done lemma separatedin_discrete_topology: "separatedin (discrete_topology U) S T \ S \ U \ T \ U \ disjnt S T" by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology) lemma separated_eq_distinguishable: "separatedin X {x} {y} \ x \ topspace X \ y \ topspace X \ (\U. openin X U \ x \ U \ (y \ U)) \ (\v. openin X v \ y \ v \ (x \ v))" by (force simp: separatedin_def closure_of_def) lemma separatedin_Un [simp]: "separatedin X S (T \ U) \ separatedin X S T \ separatedin X S U" "separatedin X (S \ T) U \ separatedin X S U \ separatedin X T U" by (auto simp: separatedin_def) lemma separatedin_Union: "finite \ \ separatedin X S (\\) \ S \ topspace X \ (\T \ \. separatedin X S T)" "finite \ \ separatedin X (\\) S \ (\T \ \. separatedin X S T) \ S \ topspace X" by (auto simp: separatedin_def closure_of_Union) lemma separatedin_openin_diff: "\openin X S; openin X T\ \ separatedin X (S - T) (T - S)" unfolding separatedin_def apply (intro conjI) apply (meson Diff_subset openin_subset subset_trans)+ using openin_Int_closure_of_eq_empty by fastforce+ lemma separatedin_closedin_diff: "\closedin X S; closedin X T\ \ separatedin X (S - T) (T - S)" apply (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2) apply (meson Diff_subset closedin_subset subset_trans) done lemma separation_closedin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ closedin (subtopology X (S \ T)) S \ closedin (subtopology X (S \ T)) T" apply (simp add: separatedin_def closedin_Int_closure_of disjnt_iff) using closure_of_subset apply blast done lemma separation_openin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ openin (subtopology X (S \ T)) S \ openin (subtopology X (S \ T)) T" unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ definition homeomorphic_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "homeomorphic_map X Y f \ quotient_map X Y f \ inj_on f (topspace X)" definition homeomorphic_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "homeomorphic_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" lemma homeomorphic_map_eq: "\homeomorphic_map X Y f; \x. x \ topspace X \ f x = g x\ \ homeomorphic_map X Y g" by (meson homeomorphic_map_def inj_on_cong quotient_map_eq) lemma homeomorphic_maps_eq: "\homeomorphic_maps X Y f g; \x. x \ topspace X \ f x = f' x; \y. y \ topspace Y \ g y = g' y\ \ homeomorphic_maps X Y f' g'" apply (simp add: homeomorphic_maps_def) by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff) lemma homeomorphic_maps_sym: "homeomorphic_maps X Y f g \ homeomorphic_maps Y X g f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_maps_id: "homeomorphic_maps X Y id id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have "topspace X = topspace Y" by (auto simp: homeomorphic_maps_def continuous_map_def) with L show ?rhs unfolding homeomorphic_maps_def by (metis topology_finer_continuous_id topology_eq) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by auto qed lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have eq: "topspace X = topspace Y" by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def) then have "\S. openin X S \ openin Y S" by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id) then show ?rhs using L unfolding homeomorphic_map_def by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id) next assume ?rhs then show ?lhs unfolding homeomorphic_map_def by (simp add: closed_map_id continuous_closed_imp_quotient_map) qed lemma homeomorphic_maps_i [simp]:"homeomorphic_maps X Y id id \ Y = X" by (metis (full_types) eq_id_iff homeomorphic_maps_id) lemma homeomorphic_map_i [simp]: "homeomorphic_map X Y id \ Y = X" by (metis (no_types) eq_id_iff homeomorphic_map_id) lemma homeomorphic_map_compose: assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g" shows "homeomorphic_map X X'' (g \ f)" proof - have "inj_on g (f ` topspace X)" by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map) then show ?thesis using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq) qed lemma homeomorphic_maps_compose: "homeomorphic_maps X Y f h \ homeomorphic_maps Y X'' g k \ homeomorphic_maps X X'' (g \ f) (h \ k)" unfolding homeomorphic_maps_def by (auto simp: continuous_map_compose; simp add: continuous_map_def) lemma homeomorphic_eq_everything_map: "homeomorphic_map X Y f \ continuous_map X Y f \ open_map X Y f \ closed_map X Y f \ f ` (topspace X) = topspace Y \ inj_on f (topspace X)" unfolding homeomorphic_map_def by (force simp: injective_quotient_map intro: injective_quotient_map) lemma homeomorphic_imp_continuous_map: "homeomorphic_map X Y f \ continuous_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_open_map: "homeomorphic_map X Y f \ open_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_closed_map: "homeomorphic_map X Y f \ closed_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_surjective_map: "homeomorphic_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_injective_map: "homeomorphic_map X Y f \ inj_on f (topspace X)" by (simp add: homeomorphic_eq_everything_map) lemma bijective_open_imp_homeomorphic_map: "\continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map) lemma bijective_closed_imp_homeomorphic_map: "\continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_closed_quotient_map homeomorphic_map_def) lemma open_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "open_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y open_map_def continuous_map_def eq) qed lemma closed_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "closed_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "closedin X U" for U using closedin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y closed_map_def continuous_map_closedin eq) qed lemma homeomorphic_maps_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f \ homeomorphic_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "continuous_map Y X g" "\x\topspace X. g (f x) = x" "\x'\topspace Y. f (g x') = x'" by (auto simp: homeomorphic_maps_def) show ?rhs proof (intro conjI bijective_open_imp_homeomorphic_map L) show "open_map X Y f" using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def) show "open_map Y X g" using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def) show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X" using L by (force simp: continuous_map_closedin)+ show "inj_on f (topspace X)" "inj_on g (topspace Y)" using L unfolding inj_on_def by metis+ qed next assume ?rhs then show ?lhs by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map) qed lemma homeomorphic_maps_imp_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f" using homeomorphic_maps_map by blast lemma homeomorphic_map_maps: "homeomorphic_map X Y f \ (\g. homeomorphic_maps X Y f g)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f" "f ` (topspace X) = topspace Y" "inj_on f (topspace X)" by (auto simp: homeomorphic_eq_everything_map) have X: "\x. x \ topspace X \ f x \ topspace Y \ inv_into (topspace X) f (f x) = x" using L by auto have Y: "\y. y \ topspace Y \ inv_into (topspace X) f y \ topspace X \ f (inv_into (topspace X) f y) = y" by (simp add: L f_inv_into_f inv_into_into) have "homeomorphic_maps X Y f (inv_into (topspace X) f)" unfolding homeomorphic_maps_def proof (intro conjI L) show "continuous_map Y X (inv_into (topspace X) f)" by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f]) next show "\x\topspace X. inv_into (topspace X) f (f x) = x" "\y\topspace Y. f (inv_into (topspace X) f y) = y" using X Y by auto qed then show ?rhs by metis next assume ?rhs then show ?lhs using homeomorphic_maps_map by blast qed lemma homeomorphic_maps_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_maps X X f f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_map_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_map X X f" using homeomorphic_maps_involution homeomorphic_maps_map by blast lemma homeomorphic_map_openness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "openin Y (f ` U) \ openin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "openin X U \ openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast show "openin Y (f ` U) = openin X U" proof assume L: "openin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "openin X U" by (metis L homeomorphic_imp_open_map open_map_def g) next assume "openin X U" then show "openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast qed qed lemma homeomorphic_map_closedness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "closedin Y (f ` U) \ closedin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "closedin X U \ closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast show "closedin Y (f ` U) = closedin X U" proof assume L: "closedin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "closedin X U" by (metis L homeomorphic_imp_closed_map closed_map_def g) next assume "closedin X U" then show "closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast qed qed lemma homeomorphic_map_openness_eq: "homeomorphic_map X Y f \ openin X U \ U \ topspace X \ openin Y (f ` U)" by (meson homeomorphic_map_openness openin_closedin_eq) lemma homeomorphic_map_closedness_eq: "homeomorphic_map X Y f \ closedin X U \ U \ topspace X \ closedin Y (f ` U)" by (meson closedin_subset homeomorphic_map_closedness) lemma all_openin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. openin Y V \ P V) \ (\U. openin X U \ P(f ` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson assms homeomorphic_map_openness_eq) next assume ?rhs then show ?lhs by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff) qed lemma all_closedin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. closedin Y V \ P V) \ (\U. closedin X U \ P(f ` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson assms homeomorphic_map_closedness_eq) next assume ?rhs then show ?lhs by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff) qed lemma homeomorphic_map_derived_set_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)" proof - have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)" using hom by (auto simp: homeomorphic_eq_everything_map) have iff: "(\T. x \ T \ openin X T \ (\y. y \ x \ y \ S \ y \ T)) = (\T. T \ topspace Y \ f x \ T \ openin Y T \ (\y. y \ f x \ y \ f ` S \ y \ T))" if "x \ topspace X" for x proof - have 1: "(x \ T \ openin X T) = (T \ topspace X \ f x \ f ` T \ openin Y (f ` T))" for T by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that) have 2: "(\y. y \ x \ y \ S \ y \ T) = (\y. y \ f x \ y \ f ` S \ y \ f ` T)" (is "?lhs = ?rhs") if "T \ topspace X \ f x \ f ` T \ openin Y (f ` T)" for T proof show "?lhs \ ?rhs" by (meson "1" imageI inj inj_on_eq_iff inj_on_subset that) show "?rhs \ ?lhs" using S inj inj_onD that by fastforce qed show ?thesis apply (simp flip: fim add: all_subset_image) apply (simp flip: imp_conjL) by (intro all_cong1 imp_cong 1 2) qed have *: "\T = f ` S; \x. x \ S \ P x \ Q(f x)\ \ {y. y \ T \ Q y} = f ` {x \ S. P x}" for T S P Q by auto show ?thesis unfolding derived_set_of_def apply (rule *) using fim apply blast using iff openin_subset by force qed lemma homeomorphic_map_closure_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y closure_of (f ` S) = f ` (X closure_of S)" unfolding closure_of using homeomorphic_imp_surjective_map [OF hom] S by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms]) lemma homeomorphic_map_interior_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y interior_of (f ` S) = f ` (X interior_of S)" proof - { fix y assume "y \ topspace Y" and "y \ Y closure_of (topspace Y - f ` S)" then have "y \ f ` (topspace X - X closure_of (topspace X - S))" using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom] by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) } moreover { fix x assume "x \ topspace X" then have "f x \ topspace Y" using hom homeomorphic_imp_surjective_map by blast } moreover { fix x assume "x \ topspace X" and "x \ X closure_of (topspace X - S)" and "f x \ Y closure_of (topspace Y - f ` S)" then have "False" using homeomorphic_map_closure_of [OF hom] hom unfolding homeomorphic_eq_everything_map by (metis (no_types, lifting) Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff_alt inj_on_image_set_diff) } ultimately show ?thesis by (auto simp: interior_of_closure_of) qed lemma homeomorphic_map_frontier_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y frontier_of (f ` S) = f ` (X frontier_of S)" unfolding frontier_of_def proof (intro equalityI subsetI DiffI) fix y assume "y \ Y closure_of f ` S - Y interior_of f ` S" then show "y \ f ` (X closure_of S - X interior_of S)" using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce next fix y assume "y \ f ` (X closure_of S - X interior_of S)" then show "y \ Y closure_of f ` S" using S hom homeomorphic_map_closure_of by fastforce next fix x assume "x \ f ` (X closure_of S - X interior_of S)" then obtain y where y: "x = f y" "y \ X closure_of S" "y \ X interior_of S" by blast then have "y \ topspace X" by (simp add: in_closure_of) then have "f y \ f ` (X interior_of S)" by (meson hom homeomorphic_eq_everything_map inj_on_image_mem_iff_alt interior_of_subset_topspace y(3)) then show "x \ Y interior_of f ` S" using S hom homeomorphic_map_interior_of y(1) by blast qed lemma homeomorphic_maps_subtopologies: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def - by (force simp: continuous_map_from_subtopology topspace_subtopology continuous_map_in_subtopology) + by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_maps_subtopologies_alt: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) \ T; g ` (topspace Y \ T) \ S\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def - by (force simp: continuous_map_from_subtopology topspace_subtopology continuous_map_in_subtopology) + by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_map_subtopologies: "\homeomorphic_map X Y f; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies) lemma homeomorphic_map_subtopologies_alt: "\homeomorphic_map X Y f; \x. \x \ topspace X; f x \ topspace Y\ \ f x \ T \ x \ S\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" unfolding homeomorphic_map_maps apply (erule ex_forward) apply (rule homeomorphic_maps_subtopologies) apply (auto simp: homeomorphic_maps_def continuous_map_def) by (metis IntI image_iff) subsection\Relation of homeomorphism between topological spaces\ definition homeomorphic_space (infixr "homeomorphic'_space" 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl: "X homeomorphic_space X" by (meson homeomorphic_maps_id homeomorphic_space_def) lemma homeomorphic_space_sym: "X homeomorphic_space Y \ Y homeomorphic_space X" unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym) lemma homeomorphic_space_trans [trans]: "\X1 homeomorphic_space X2; X2 homeomorphic_space X3\ \ X1 homeomorphic_space X3" unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose) lemma homeomorphic_space: "X homeomorphic_space Y \ (\f. homeomorphic_map X Y f)" by (simp add: homeomorphic_map_maps homeomorphic_space_def) lemma homeomorphic_maps_imp_homeomorphic_space: "homeomorphic_maps X Y f g \ X homeomorphic_space Y" unfolding homeomorphic_space_def by metis lemma homeomorphic_map_imp_homeomorphic_space: "homeomorphic_map X Y f \ X homeomorphic_space Y" unfolding homeomorphic_map_maps using homeomorphic_space_def by blast lemma homeomorphic_empty_space: "X homeomorphic_space Y \ topspace X = {} \ topspace Y = {}" by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty) lemma homeomorphic_empty_space_eq: assumes "topspace X = {}" shows "X homeomorphic_space Y \ topspace Y = {}" proof - have "\f t. continuous_map X (t::'b topology) f" using assms continuous_map_on_empty by blast then show ?thesis by (metis (no_types) assms continuous_map_on_empty empty_iff homeomorphic_empty_space homeomorphic_maps_def homeomorphic_space_def) qed subsection\Connected topological spaces\ definition connected_space :: "'a topology \ bool" where "connected_space X \ \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where "connectedin X S \ S \ topspace X \ connected_space (subtopology X S)" lemma connected_spaceD: "\connected_space X; openin X U; openin X V; topspace X \ U \ V; U \ V = {}; U \ {}; V \ {}\ \ False" by (auto simp: connected_space_def) lemma connectedin_subset_topspace: "connectedin X S \ S \ topspace X" by (simp add: connectedin_def) lemma connectedin_topspace: "connectedin X (topspace X) \ connected_space X" by (simp add: connectedin_def) lemma connected_space_subtopology: "connectedin X S \ connected_space (subtopology X S)" by (simp add: connectedin_def) lemma connectedin_subtopology: "connectedin (subtopology X S) T \ connectedin X T \ T \ S" - by (force simp: connectedin_def subtopology_subtopology topspace_subtopology inf_absorb2) + by (force simp: connectedin_def subtopology_subtopology inf_absorb2) lemma connected_space_eq: "connected_space X \ (\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_space_def by (metis openin_Un openin_subset subset_antisym) lemma connected_space_closedin: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then have L: "\E1 E2. \openin X E1; E1 \ E2 = {}; topspace X \ E1 \ E2; openin X E2\ \ E1 = {} \ E2 = {}" by (simp add: connected_space_def) show ?rhs unfolding connected_space_def proof clarify fix E1 E2 assume "closedin X E1" and "closedin X E2" and "topspace X \ E1 \ E2" and "E1 \ E2 = {}" and "E1 \ {}" and "E2 \ {}" have "E1 \ E2 = topspace X" by (meson Un_subset_iff \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ closedin_def subset_antisym) then have "topspace X - E2 = E1" using \E1 \ E2 = {}\ by fastforce then have "topspace X = E1" using \E1 \ {}\ L \closedin X E1\ \closedin X E2\ by blast then show "False" using \E1 \ E2 = {}\ \E1 \ E2 = topspace X\ \E2 \ {}\ by blast qed next assume R: ?rhs show ?lhs unfolding connected_space_def proof clarify fix E1 E2 assume "openin X E1" and "openin X E2" and "topspace X \ E1 \ E2" and "E1 \ E2 = {}" and "E1 \ {}" and "E2 \ {}" have "E1 \ E2 = topspace X" by (meson Un_subset_iff \openin X E1\ \openin X E2\ \topspace X \ E1 \ E2\ openin_closedin_eq subset_antisym) then have "topspace X - E2 = E1" using \E1 \ E2 = {}\ by fastforce then have "topspace X = E1" using \E1 \ {}\ R \openin X E1\ \openin X E2\ by blast then show "False" using \E1 \ E2 = {}\ \E1 \ E2 = topspace X\ \E2 \ {}\ by blast qed qed lemma connected_space_closedin_eq: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" apply (simp add: connected_space_closedin) apply (intro all_cong) using closedin_subset apply blast done lemma connected_space_clopen_in: "connected_space X \ (\T. openin X T \ closedin X T \ T = {} \ T = topspace X)" proof - have eq: "openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ P \ E2 = topspace X - E1 \ openin X E1 \ openin X E2 \ P" for E1 E2 P using openin_subset by blast show ?thesis unfolding connected_space_eq eq closedin_def by (auto simp: openin_closedin_eq) qed lemma connectedin: "connectedin X S \ S \ topspace X \ (\E1 E2. openin X E1 \ openin X E2 \ S \ E1 \ E2 \ E1 \ E2 \ S = {} \ E1 \ S \ {} \ E2 \ S \ {})" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology Not_eq_iff * apply (intro conj_cong arg_cong [where f=Not] ex_cong1 refl) apply (blast elim: dest!: openin_subset)+ done qed lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology Not_eq_iff * apply (intro conj_cong arg_cong [where f=Not] ex_cong1 refl) apply (blast elim: dest!: openin_subset)+ done qed lemma connectedin_empty [simp]: "connectedin X {}" by (simp add: connectedin) lemma connected_space_topspace_empty: "topspace X = {} \ connected_space X" using connectedin_topspace by fastforce lemma connectedin_sing [simp]: "connectedin X {a} \ a \ topspace X" by (simp add: connectedin) lemma connectedin_absolute [simp]: "connectedin (subtopology X S) S \ connectedin X S" apply (simp only: connectedin_def topspace_subtopology subtopology_subtopology) apply (intro conj_cong imp_cong arg_cong [where f=Not] all_cong1 ex_cong1 refl) by auto lemma connectedin_Union: assumes \: "\S. S \ \ \ connectedin X S" and ne: "\\ \ {}" shows "connectedin X (\\)" proof - have "\\ \ topspace X" using \ by (simp add: Union_least connectedin_def) moreover have False if "openin X E1" "openin X E2" and cover: "\\ \ E1 \ E2" and disj: "E1 \ E2 \ \\ = {}" and overlap1: "E1 \ \\ \ {}" and overlap2: "E2 \ \\ \ {}" for E1 E2 proof - have disjS: "E1 \ E2 \ S = {}" if "S \ \" for S using Diff_triv that disj by auto have coverS: "S \ E1 \ E2" if "S \ \" for S using that cover by blast have "\ \ {}" using overlap1 by blast obtain a where a: "\U. U \ \ \ a \ U" using ne by force with \\ \ {}\ have "a \ \\" by blast then consider "a \ E1" | "a \ E2" using \\\ \ E1 \ E2\ by auto then show False proof cases case 1 then obtain b S where "b \ E2" "b \ S" "S \ \" using overlap2 by blast then show ?thesis using "1" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) next case 2 then obtain b S where "b \ E1" "b \ S" "S \ \" using overlap1 by blast then show ?thesis using "2" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) qed qed ultimately show ?thesis unfolding connectedin by blast qed lemma connectedin_Un: "\connectedin X S; connectedin X T; S \ T \ {}\ \ connectedin X (S \ T)" using connectedin_Union [of "{S,T}"] by auto lemma connected_space_subconnected: "connected_space X \ (\x \ topspace X. \y \ topspace X. \S. connectedin X S \ x \ S \ y \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connectedin_topspace by blast next assume R [rule_format]: ?rhs have False if "openin X U" "openin X V" and disj: "U \ V = {}" and cover: "topspace X \ U \ V" and "U \ {}" "V \ {}" for U V proof - obtain u v where "u \ U" "v \ V" using \U \ {}\ \V \ {}\ by auto then obtain T where "u \ T" "v \ T" and T: "connectedin X T" using R [of u v] that by (meson \openin X U\ \openin X V\ subsetD openin_subset) then show False using that unfolding connectedin by (metis IntI \u \ U\ \v \ V\ empty_iff inf_bot_left subset_trans) qed then show ?lhs by (auto simp: connected_space_def) qed lemma connectedin_intermediate_closure_of: assumes "connectedin X S" "S \ T" "T \ X closure_of S" shows "connectedin X T" proof - have S: "S \ topspace X"and T: "T \ topspace X" using assms by (meson closure_of_subset_topspace dual_order.trans)+ show ?thesis using assms apply (simp add: connectedin closure_of_subset_topspace S T) apply (elim all_forward imp_forward2 asm_rl) apply (blast dest: openin_Int_closure_of_eq_empty [of X _ S])+ done qed lemma connectedin_closure_of: "connectedin X S \ connectedin X (X closure_of S)" by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl) lemma connectedin_separation: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ C1 \ X closure_of C2 = {} \ C2 \ X closure_of C1 = {})" (is "?lhs = ?rhs") unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology apply (intro conj_cong refl arg_cong [where f=Not]) apply (intro ex_cong1 iffI, blast) using closure_of_subset_Int by force lemma connectedin_eq_not_separated: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" apply (simp add: separatedin_def connectedin_separation) apply (intro conj_cong all_cong1 refl, blast) done lemma connectedin_eq_not_separated_subset: "connectedin X S \ S \ topspace X \ (\C1 C2. S \ C1 \ C2 \ S \ C1 \ {} \ S \ C2 \ {} \ separatedin X C1 C2)" proof - have *: "\C1 C2. S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" if "\C1 C2. C1 \ C2 = S \ C1 = {} \ C2 = {} \ \ separatedin X C1 C2" proof (intro allI) fix C1 C2 show "S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" using that [of "S \ C1" "S \ C2"] by (auto simp: separatedin_mono) qed show ?thesis apply (simp add: connectedin_eq_not_separated) apply (intro conj_cong refl iffI *) apply (blast elim!: all_forward)+ done qed lemma connected_space_eq_not_separated: "connected_space X \ (\C1 C2. C1 \ C2 = topspace X \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (simp add: connectedin_eq_not_separated flip: connectedin_topspace) lemma connected_space_eq_not_separated_subset: "connected_space X \ (\C1 C2. topspace X \ C1 \ C2 \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" apply (simp add: connected_space_eq_not_separated) apply (intro all_cong1) by (metis Un_absorb dual_order.antisym separatedin_def subset_refl sup_mono) lemma connectedin_subset_separated_union: "\connectedin X C; separatedin X S T; C \ S \ T\ \ C \ S \ C \ T" unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: "\connectedin X S; connectedin X T; \separatedin X S T\ \ connectedin X (S \ T)" apply (simp add: connectedin_eq_not_separated_subset, auto) apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono sup_commute) apply (metis (no_types, hide_lams) Diff_subset_conv Diff_triv disjoint_iff_not_equal separatedin_mono separatedin_sym sup_commute) by (meson disjoint_iff_not_equal) lemma connected_space_closures: "connected_space X \ (\e1 e2. e1 \ e2 = topspace X \ X closure_of e1 \ X closure_of e2 = {} \ e1 \ {} \ e2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding connected_space_closedin_eq by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace) next assume ?rhs then show ?lhs unfolding connected_space_closedin_eq by (metis closure_of_eq) qed lemma connectedin_inter_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) have "S - (topspace X \ T) \ {}" using assms(3) by blast moreover have "S \ topspace X \ T \ {}" using assms(1) assms(2) connectedin by fastforce moreover have False if "S \ T \ {}" "S - T \ {}" "T \ topspace X" "S \ X frontier_of T = {}" for T proof - have null: "S \ (X closure_of T - X interior_of T) = {}" using that unfolding frontier_of_def by blast have 1: "X interior_of T \ (topspace X - X closure_of T) \ S = {}" by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty) have 2: "S \ X interior_of T \ (topspace X - X closure_of T)" using that \S \ topspace X\ null by auto have 3: "S \ X interior_of T \ {}" using closure_of_subset that(1) that(3) null by fastforce show ?thesis using null \S \ topspace X\ that * [of "X interior_of T" "topspace X - X closure_of T"] apply (clarsimp simp add: openin_diff 1 2) apply (simp add: Int_commute Diff_Int_distrib 3) by (metis Int_absorb2 contra_subsetD interior_of_subset) qed ultimately show ?thesis by (metis Int_lower1 frontier_of_restrict inf_assoc) qed lemma connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and "connectedin X S" shows "connectedin Y (f ` S)" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) show ?thesis unfolding connectedin connected_space_def proof (intro conjI notI; clarify) show "f x \ topspace Y" if "x \ S" for x using \S \ topspace X\ continuous_map_image_subset_topspace f that by blast next fix U V let ?U = "{x \ topspace X. f x \ U}" let ?V = "{x \ topspace X. f x \ V}" assume UV: "openin Y U" "openin Y V" "f ` S \ U \ V" "U \ V \ f ` S = {}" "U \ f ` S \ {}" "V \ f ` S \ {}" then have 1: "?U \ ?V \ S = {}" by auto have 2: "openin X ?U" "openin X ?V" using \openin Y U\ \openin Y V\ continuous_map f by fastforce+ show "False" using * [of ?U ?V] UV \S \ topspace X\ by (auto simp: 1 2) qed qed lemma homeomorphic_connected_space: "X homeomorphic_space Y \ connected_space X \ connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def apply safe apply (metis connectedin_continuous_map_image connected_space_subconnected continuous_map_image_subset_topspace image_eqI image_subset_iff) by (metis (no_types, hide_lams) connectedin_continuous_map_image connectedin_topspace continuous_map_def continuous_map_image_subset_topspace imageI set_eq_subset subsetI) lemma homeomorphic_map_connectedness: assumes f: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "connectedin Y (f ` U) \ connectedin X U" proof - have 1: "f ` U \ topspace Y \ U \ topspace X" using U f homeomorphic_imp_surjective_map by blast moreover have "connected_space (subtopology Y (f ` U)) \ connected_space (subtopology X U)" proof (rule homeomorphic_connected_space) have "f ` U \ topspace Y" by (simp add: U 1) then have "topspace Y \ f ` U = f ` U" by (simp add: subset_antisym) then show "subtopology Y (f ` U) homeomorphic_space subtopology X U" by (metis (no_types) Int_subset_iff U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym subset_antisym subset_refl) qed ultimately show ?thesis by (auto simp: connectedin_def) qed lemma homeomorphic_map_connectedness_eq: "homeomorphic_map X Y f \ connectedin X U \ U \ topspace X \ connectedin Y (f ` U)" using homeomorphic_map_connectedness connectedin_subset_topspace by metis lemma connectedin_discrete_topology: "connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" proof (cases "S \ U") case True show ?thesis proof (cases "S = {}") case False moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" apply safe using False connectedin_inter_frontier_of insert_Diff apply fastforce using True by auto ultimately show ?thesis by auto qed simp next case False then show ?thesis by (simp add: connectedin_def) qed lemma connected_space_discrete_topology: "connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology) subsection\Compact sets\ definition compactin where "compactin X S \ S \ topspace X \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\))" definition compact_space where "compact_space X \ compactin X (topspace X)" lemma compact_space_alt: "compact_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ topspace X \ \\))" by (simp add: compact_space_def compactin_def) lemma compact_space: "compact_space X \ (\\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. finite \ \ \ \ \ \ \\ = topspace X))" unfolding compact_space_alt using openin_subset by fastforce lemma compactinD: "\compactin X S; \U. U \ \ \ openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" by (auto simp: compactin_def) lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: "compactin (subtopology X S) S \ compactin X S" proof - have eq: "(\U \ \. \Y. openin X Y \ U = Y \ S) \ \ \ (\Y. Y \ S) ` {y. openin X y}" for \ by auto show ?thesis - by (auto simp: compactin_def topspace_subtopology openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) + by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) qed lemma compactin_subspace: "compactin X S \ S \ topspace X \ compact_space (subtopology X S)" unfolding compact_space_def topspace_subtopology by (metis compactin_absolute compactin_def inf.absorb2) lemma compact_space_subtopology: "compactin X S \ compact_space (subtopology X S)" by (simp add: compactin_subspace) lemma compactin_subtopology: "compactin (subtopology X S) T \ compactin X T \ T \ S" -apply (simp add: compactin_subspace topspace_subtopology) +apply (simp add: compactin_subspace) by (metis inf.orderE inf_commute subtopology_subtopology) lemma compactin_subset_topspace: "compactin X S \ S \ topspace X" by (simp add: compactin_subspace) lemma compactin_contractive: "\compactin X' S; topspace X' = topspace X; \U. openin X U \ openin X' U\ \ compactin X S" by (simp add: compactin_def) lemma finite_imp_compactin: "\S \ topspace X; finite S\ \ compactin X S" by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology) lemma compactin_empty [iff]: "compactin X {}" by (simp add: finite_imp_compactin) lemma compact_space_topspace_empty: "topspace X = {} \ compact_space X" by (simp add: compact_space_def) lemma finite_imp_compactin_eq: "finite S \ (compactin X S \ S \ topspace X)" using compactin_subset_topspace finite_imp_compactin by blast lemma compactin_sing [simp]: "compactin X {a} \ a \ topspace X" by (simp add: finite_imp_compactin_eq) lemma closed_compactin: assumes XK: "compactin X K" and "C \ K" and XC: "closedin X C" shows "compactin X C" unfolding compactin_def proof (intro conjI allI impI) show "C \ topspace X" by (simp add: XC closedin_subset) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ C \ \\" have "(\U\insert (topspace X - C) \. openin X U)" using XC \ by blast moreover have "K \ \(insert (topspace X - C) \)" using \ XK compactin_subset_topspace by fastforce ultimately obtain \ where "finite \" "\ \ insert (topspace X - C) \" "K \ \\" using assms unfolding compactin_def by metis moreover have "openin X (topspace X - C)" using XC by auto ultimately show "\\. finite \ \ \ \ \ \ C \ \\" using \C \ K\ by (rule_tac x="\ - {topspace X - C}" in exI) auto qed lemma closedin_compact_space: "\compact_space X; closedin X S\ \ compactin X S" by (simp add: closed_compactin closedin_subset compact_space_def) lemma compact_Int_closedin: assumes "compactin X S" "closedin X T" shows "compactin X (S \ T)" proof - have "compactin (subtopology X S) (S \ T)" by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute) then show ?thesis by (simp add: compactin_subtopology) qed lemma closed_Int_compactin: "\closedin X S; compactin X T\ \ compactin X (S \ T)" by (metis compact_Int_closedin inf_commute) lemma compactin_Un: assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \ T)" unfolding compactin_def proof (intro conjI allI impI) show "S \ T \ topspace X" using assms by (auto simp: compactin_def) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ S \ T \ \\" with S obtain \ where \: "finite \" "\ \ \" "S \ \\" unfolding compactin_def by (meson sup.bounded_iff) obtain \ where "finite \" "\ \ \" "T \ \\" using \ T unfolding compactin_def by (meson sup.bounded_iff) with \ show "\\. finite \ \ \ \ \ \ S \ T \ \\" by (rule_tac x="\ \ \" in exI) auto qed lemma compactin_Union: "\finite \; \S. S \ \ \ compactin X S\ \ compactin X (\\)" by (induction rule: finite_induct) (simp_all add: compactin_Un) lemma compactin_subtopology_imp_compact: assumes "compactin (subtopology X S) K" shows "compactin X K" using assms -proof (clarsimp simp add: compactin_def topspace_subtopology) +proof (clarsimp simp add: compactin_def) fix \ define \ where "\ \ (\U. U \ S) ` \" assume "K \ topspace X" and "K \ S" and "\x\\. openin X x" and "K \ \\" then have "\V \ \. openin (subtopology X S) V" "K \ \\" unfolding \_def by (auto simp: openin_subtopology) moreover assume "\\. (\x\\. openin (subtopology X S) x) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson then have \: "\U. U \ \ \ V = U \ S" if "V \ \" for V unfolding \_def using that by blast let ?\ = "(\F. @U. U \ \ \ F = U \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using someI_ex [OF \] by blast show "K \ \?\" proof clarsimp fix x assume "x \ K" then show "\V \ \. x \ (SOME U. U \ \ \ V = U \ S)" using \K \ \\\ someI_ex [OF \] by (metis (no_types, lifting) IntD1 Union_iff subsetCE) qed qed qed lemma compact_imp_compactin_subtopology: assumes "compactin X K" "K \ S" shows "compactin (subtopology X S) K" using assms -proof (clarsimp simp add: compactin_def topspace_subtopology) +proof (clarsimp simp add: compactin_def) fix \ :: "'a set set" define \ where "\ \ {V. openin X V \ (\U \ \. U = V \ S)}" assume "K \ S" and "K \ topspace X" and "\U\\. openin (subtopology X S) U" and "K \ \\" then have "\V \ \. openin X V" "K \ \\" unfolding \_def by (fastforce simp: subset_eq openin_subtopology)+ moreover assume "\\. (\U\\. openin X U) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson let ?\ = "(\F. F \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using \_def \\ \ \\ by blast show "K \ \?\" using \K \ \\\ assms(2) by auto qed qed proposition compact_space_fip: "compact_space X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" (is "_ = ?rhs") proof (cases "topspace X = {}") case True then show ?thesis apply (clarsimp simp add: compact_space_def closedin_topspace_empty) by (metis finite.emptyI finite_insert infinite_super insertI1 subsetI) next case False show ?thesis proof safe fix \ :: "'a set set" assume * [rule_format]: "\\. finite \ \ \ \ \ \ \\ \ {}" define \ where "\ \ (\S. topspace X - S) ` \" assume clo: "\C\\. closedin X C" and [simp]: "\\ = {}" then have "\V \ \. openin X V" "topspace X \ \\" by (auto simp: \_def) moreover assume [unfolded compact_space_alt, rule_format, of \]: "compact_space X" ultimately obtain \ where \: "finite \" "\ \ \" "topspace X \ topspace X - \\" by (auto simp: ex_finite_subset_image \_def) moreover have "\ \ {}" using \ \topspace X \ {}\ by blast ultimately show "False" using * [of \] by auto (metis Diff_iff Inter_iff clo closedin_def subsetD) next assume R [rule_format]: ?rhs show "compact_space X" unfolding compact_space_alt proof clarify fix \ :: "'a set set" define \ where "\ \ (\S. topspace X - S) ` \" assume "\C\\. openin X C" and "topspace X \ \\" with \topspace X \ {}\ have *: "\V \ \. closedin X V" "\ \ {}" by (auto simp: \_def) show "\\. finite \ \ \ \ \ \ topspace X \ \\" proof (rule ccontr; simp) assume "\\\\. finite \ \ \ topspace X \ \\" then have "\\. finite \ \ \ \ \ \ \\ \ {}" by (simp add: \_def all_finite_subset_image) with \topspace X \ \\\ show False using R [of \] * by (simp add: \_def) qed qed qed qed corollary compactin_fip: "compactin X S \ S \ topspace X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof (cases "S = {}") case False show ?thesis proof (cases "S \ topspace X") case True then have "compactin X S \ (\\. \ \ (\T. S \ T) ` {T. closedin X T} \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL) also have "\ = (\\\Collect (closedin X). (\\. finite \ \ \ \ (\) S ` \ \ \\ \ {}) \ \ ((\) S ` \) \ {})" by (simp add: all_subset_image) also have "\ = (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof - have eq: "((\\. finite \ \ \ \ \ \ \ ((\) S ` \) \ {}) \ \ ((\) S ` \) \ {}) \ ((\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" for \ by simp (use \S \ {}\ in blast) show ?thesis apply (simp only: imp_conjL [symmetric] all_finite_subset_image eq) apply (simp add: subset_eq) done qed finally show ?thesis using True by simp qed (simp add: compactin_subspace) qed force corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" and ne: "\n. C n \ {}" and inc: "\m n. m \ n \ C n \ C m" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" have "closedin X A" if "A \ ?\" for A using that clo by auto moreover have "(\n\K. \m \ n. C m) \ {}" if "finite K" for K proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast with inc have "C n \ (\n\K. \m \ n. C m)" by blast with ne [of n] show ?thesis by blast qed ultimately show ?thesis using \compact_space X\ [unfolded compact_space_fip, rule_format, of ?\] by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps) qed lemma compactin_discrete_topology: "compactin (discrete_topology X) S \ S \ X \ finite S" (is "?lhs = ?rhs") proof (intro iffI conjI) assume L: ?lhs then show "S \ X" by (auto simp: compactin_def) have *: "\\. Ball \ (openin (discrete_topology X)) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\)" using L by (auto simp: compactin_def) show "finite S" using * [of "(\x. {x}) ` X"] \S \ X\ by clarsimp (metis UN_singleton finite_subset_image infinite_super) next assume ?rhs then show ?lhs by (simp add: finite_imp_compactin) qed lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \ finite X" by (simp add: compactin_discrete_topology compact_space_def) lemma compact_space_imp_Bolzano_Weierstrass: assumes "compact_space X" "infinite S" "S \ topspace X" shows "X derived_set_of S \ {}" proof assume X: "X derived_set_of S = {}" then have "closedin X S" by (simp add: closedin_contains_derived_set assms) then have "compactin X S" by (rule closedin_compact_space [OF \compact_space X\]) with X show False by (metis \infinite S\ compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq) qed lemma compactin_imp_Bolzano_Weierstrass: "\compactin X S; infinite T \ T \ S\ \ S \ X derived_set_of T \ {}" using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"] - by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2 topspace_subtopology) + by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2) lemma compact_closure_of_imp_Bolzano_Weierstrass: "\compactin X (X closure_of S); infinite T; T \ S; T \ topspace X\ \ X derived_set_of T \ {}" using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce lemma discrete_compactin_eq_finite: "S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" apply (rule iffI) using compactin_imp_Bolzano_Weierstrass compactin_subset_topspace apply blast by (simp add: finite_imp_compactin_eq) lemma discrete_compact_space_eq_finite: "X derived_set_of (topspace X) = {} \ (compact_space X \ finite(topspace X))" by (metis compact_space_discrete_topology discrete_topology_unique_derived_set) lemma image_compactin: assumes cpt: "compactin X S" and cont: "continuous_map X Y f" shows "compactin Y (f ` S)" unfolding compactin_def proof (intro conjI allI impI) show "f ` S \ topspace Y" using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast next fix \ :: "'b set set" assume \: "Ball \ (openin Y) \ f ` S \ \\" define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "S \ topspace X" and *: "\\. \\U\\. openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using cpt by (auto simp: compactin_def) obtain \ where \: "finite \" "\ \ \" "S \ \\" proof - have 1: "\U\\. openin X U" unfolding \_def using \ cont[unfolded continuous_map] by blast have 2: "S \ \\" unfolding \_def using compactin_subset_topspace cpt \ by fastforce show thesis using * [OF 1 2] that by metis qed have "\v \ \. \U. U \ \ \ v = {x \ topspace X. f x \ U}" using \_def by blast then obtain U where U: "\v \ \. U v \ \ \ v = {x \ topspace X. f x \ U v}" by metis show "\\. finite \ \ \ \ \ \ f ` S \ \\" proof (intro conjI exI) show "finite (U ` \)" by (simp add: \finite \\) next show "U ` \ \ \" using \\ \ \\ U by auto next show "f ` S \ \ (U ` \)" using \(2-3) U UnionE subset_eq U by fastforce qed qed lemma homeomorphic_compact_space: assumes "X homeomorphic_space Y" shows "compact_space X \ compact_space Y" using homeomorphic_space_sym by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin) lemma homeomorphic_map_compactness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "compactin Y (f ` U) \ compactin X U" proof - have "f ` U \ topspace Y" using hom U homeomorphic_imp_surjective_map by blast moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f" using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies) then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)" using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast ultimately show ?thesis by (simp add: compactin_subspace U) qed lemma homeomorphic_map_compactness_eq: "homeomorphic_map X Y f \ compactin X U \ U \ topspace X \ compactin Y (f ` U)" by (meson compactin_subset_topspace homeomorphic_map_compactness) subsection\Embedding maps\ definition embedding_map where "embedding_map X Y f \ homeomorphic_map X (subtopology Y (f ` (topspace X))) f" lemma embedding_map_eq: "\embedding_map X Y f; \x. x \ topspace X \ f x = g x\ \ embedding_map X Y g" unfolding embedding_map_def by (metis homeomorphic_map_eq image_cong) lemma embedding_map_compose: assumes "embedding_map X X' f" "embedding_map X' X'' g" shows "embedding_map X X'' (g \ f)" proof - have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g" using assms by (auto simp: embedding_map_def) then obtain C where "g ` topspace X' \ C = (g \ f) ` topspace X" by (metis (no_types) Int_absorb1 continuous_map_image_subset_topspace continuous_map_in_subtopology homeomorphic_eq_everything_map image_comp image_mono) then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \ f) ` topspace X)) g" by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology) then show ?thesis unfolding embedding_map_def using hm(1) homeomorphic_map_compose by blast qed lemma surjective_embedding_map: "embedding_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (force simp: embedding_map_def homeomorphic_eq_everything_map) lemma embedding_map_in_subtopology: "embedding_map X (subtopology Y S) f \ embedding_map X Y f \ f ` (topspace X) \ S" apply (auto simp: embedding_map_def subtopology_subtopology Int_absorb1) apply (metis (no_types) homeomorphic_imp_surjective_map subtopology_subtopology subtopology_topspace topspace_subtopology) apply (simp add: continuous_map_def homeomorphic_eq_everything_map) done lemma injective_open_imp_embedding_map: "\continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def apply (rule bijective_open_imp_homeomorphic_map) using continuous_map_in_subtopology apply blast apply (auto simp: continuous_map_in_subtopology open_map_into_subtopology continuous_map) done lemma injective_closed_imp_embedding_map: "\continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def apply (rule bijective_closed_imp_homeomorphic_map) apply (simp_all add: continuous_map_into_subtopology closed_map_into_subtopology) apply (simp add: continuous_map inf.absorb_iff2) done lemma embedding_map_imp_homeomorphic_space: "embedding_map X Y f \ X homeomorphic_space (subtopology Y (f ` (topspace X)))" unfolding embedding_map_def using homeomorphic_space by blast lemma embedding_imp_closed_map: "\embedding_map X Y f; closedin Y (f ` topspace X)\ \ closed_map X Y f" unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) subsection\Retraction and section maps\ definition retraction_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "retraction_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace Y. f(g x) = x)" definition section_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "section_map X Y f \ \g. retraction_maps Y X g f" definition retraction_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "retraction_map X Y f \ \g. retraction_maps X Y f g" lemma retraction_maps_eq: "\retraction_maps X Y f g; \x. x \ topspace X \ f x = f' x; \x. x \ topspace Y \ g x = g' x\ \ retraction_maps X Y f' g'" unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq) lemma section_map_eq: "\section_map X Y f; \x. x \ topspace X \ f x = g x\ \ section_map X Y g" unfolding section_map_def using retraction_maps_eq by blast lemma retraction_map_eq: "\retraction_map X Y f; \x. x \ topspace X \ f x = g x\ \ retraction_map X Y g" unfolding retraction_map_def using retraction_maps_eq by blast lemma homeomorphic_imp_retraction_maps: "homeomorphic_maps X Y f g \ retraction_maps X Y f g" by (simp add: homeomorphic_maps_def retraction_maps_def) lemma section_and_retraction_eq_homeomorphic_map: "section_map X Y f \ retraction_map X Y f \ homeomorphic_map X Y f" apply (auto simp: section_map_def retraction_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def) by (metis (full_types) continuous_map_image_subset_topspace image_subset_iff) lemma section_imp_embedding_map: "section_map X Y f \ embedding_map X Y f" unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def - by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology topspace_subtopology) + by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology) lemma retraction_imp_quotient_map: assumes "retraction_map X Y f" shows "quotient_map X Y f" unfolding quotient_map_def proof (intro conjI subsetI allI impI) show "f ` topspace X = topspace Y" using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def) next fix U assume U: "U \ topspace Y" have "openin Y U" if "\x\topspace Y. g x \ topspace X" "\x\topspace Y. f (g x) = x" "openin Y {x \ topspace Y. g x \ {x \ topspace X. f x \ U}}" for g using openin_subopen U that by fastforce then show "openin X {x \ topspace X. f x \ U} = openin Y U" using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def) qed lemma retraction_maps_compose: "\retraction_maps X Y f f'; retraction_maps Y Z g g'\ \ retraction_maps X Z (g \ f) (f' \ g')" by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def) lemma retraction_map_compose: "\retraction_map X Y f; retraction_map Y Z g\ \ retraction_map X Z (g \ f)" by (meson retraction_map_def retraction_maps_compose) lemma section_map_compose: "\section_map X Y f; section_map Y Z g\ \ section_map X Z (g \ f)" by (meson retraction_maps_compose section_map_def) lemma surjective_section_eq_homeomorphic_map: "section_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map) lemma surjective_retraction_or_section_map: "f ` (topspace X) = topspace Y \ retraction_map X Y f \ section_map X Y f \ retraction_map X Y f" using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce lemma retraction_imp_surjective_map: "retraction_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map) lemma section_imp_injective_map: "\section_map X Y f; x \ topspace X; y \ topspace X\ \ f x = f y \ x = y" by (metis (mono_tags, hide_lams) retraction_maps_def section_map_def) lemma retraction_maps_to_retract_maps: "retraction_maps X Y r s \ retraction_maps X (subtopology X (s ` (topspace Y))) (s \ r) id" unfolding retraction_maps_def by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology) subsection \Continuity\ lemma continuous_on_open: "continuous_on S f \ (\T. openin (top_of_set (f ` S)) T \ openin (top_of_set S) (S \ f -` T))" unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_closed: "continuous_on S f \ (\T. closedin (top_of_set (f ` S)) T \ closedin (top_of_set S) (S \ f -` T))" unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_imp_closedin: assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T" shows "closedin (top_of_set S) (S \ f -` T)" using assms continuous_on_closed by blast lemma continuous_map_subtopology_eu [simp]: "continuous_map (top_of_set S) (subtopology euclidean T) h \ continuous_on S h \ h ` S \ T" apply safe apply (metis continuous_map_closedin_preimage_eq continuous_map_in_subtopology continuous_on_closed order_refl topspace_euclidean_subtopology) apply (simp add: continuous_map_closedin_preimage_eq image_subset_iff) by (metis (no_types, hide_lams) continuous_map_closedin_preimage_eq continuous_map_in_subtopology continuous_on_closed order_refl topspace_euclidean_subtopology) lemma continuous_map_euclidean_top_of_set: assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f" shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "openin (top_of_set (f ` S)) (T \ f ` S)" using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto then show ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "closedin (top_of_set (f ` S)) (T \ f ` S)" using closedin_closed_Int[of T "f ` S", OF assms(2)] by (simp add: Int_commute) then show ?thesis using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_openin_preimage_eq: "continuous_on S f \ (\T. open T \ openin (top_of_set S) (S \ f -` T))" apply safe apply (simp add: continuous_openin_preimage_gen) apply (fastforce simp add: continuous_on_open openin_open) done lemma continuous_closedin_preimage_eq: "continuous_on S f \ (\T. closed T \ closedin (top_of_set S) (S \ f -` T))" apply safe apply (simp add: continuous_closedin_preimage) apply (fastforce simp add: continuous_on_closed closedin_closed) done lemma continuous_open_preimage: assumes contf: "continuous_on S f" and "open S" "open T" shows "open (S \ f -` T)" proof- obtain U where "open U" "(S \ f -` T) = S \ U" using continuous_openin_preimage_gen[OF contf \open T\] unfolding openin_open by auto then show ?thesis using open_Int[of S U, OF \open S\] by auto qed lemma continuous_closed_preimage: assumes contf: "continuous_on S f" and "closed S" "closed T" shows "closed (S \ f -` T)" proof- obtain U where "closed U" "(S \ f -` T) = S \ U" using continuous_closedin_preimage[OF contf \closed T\] unfolding closedin_closed by auto then show ?thesis using closed_Int[of S U, OF \closed S\] by auto qed lemma continuous_open_vimage: "open S \ (\x. continuous (at x) f) \ open (f -` S)" by (metis continuous_on_eq_continuous_within open_vimage) lemma continuous_closed_vimage: "closed S \ (\x. continuous (at x) f) \ closed (f -` S)" by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: assumes "(x, y) \ U" "openin (top_of_set (S \ T)) U" obtains V W where "openin (top_of_set S) V" "x \ V" "openin (top_of_set T) W" "y \ W" "(V \ W) \ U" proof - from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" by (auto simp: openin_open) from open_prod_elim[OF \open E\ \(x, y) \ E\] obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" by blast show ?thesis proof show "openin (top_of_set S) (E1 \ S)" "openin (top_of_set T) (E2 \ T)" using \open E1\ \open E2\ by (auto simp: openin_open) show "x \ E1 \ S" "y \ E2 \ T" using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto show "(E1 \ S) \ (E2 \ T) \ U" using \E1 \ E2 \ E\ \U = _\ by (auto simp: ) qed qed lemma closedin_Times: "closedin (top_of_set S) S' \ closedin (top_of_set T) T' \ closedin (top_of_set (S \ T)) (S' \ T')" unfolding closedin_closed using closed_Times by blast lemma openin_Times: "openin (top_of_set S) S' \ openin (top_of_set T) T' \ openin (top_of_set (S \ T)) (S' \ T')" unfolding openin_open using open_Times by blast lemma openin_Times_eq: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (top_of_set (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (top_of_set S) S' \ openin (top_of_set T) T'" (is "?lhs = ?rhs") proof (cases "S' = {} \ T' = {}") case True then show ?thesis by auto next case False then obtain x y where "x \ S'" "y \ T'" by blast show ?thesis proof assume ?lhs have "openin (top_of_set S) S'" apply (subst openin_subopen, clarify) apply (rule Times_in_interior_subtopology [OF _ \?lhs\]) using \y \ T'\ apply auto done moreover have "openin (top_of_set T) T'" apply (subst openin_subopen, clarify) apply (rule Times_in_interior_subtopology [OF _ \?lhs\]) using \x \ S'\ apply auto done ultimately show ?rhs by simp next assume ?rhs with False show ?lhs by (simp add: openin_Times) qed qed lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" and "openin (top_of_set T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - have "\\<^sub>F x in at a within T. x \ T \ x \ a" by (simp add: eventually_at_filter) moreover from \openin _ _\ obtain U where "open U" "S = T \ U" by (auto simp: openin_open) then have "a \ U" using \a \ S\ by auto from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] have "\\<^sub>F x in at a within T. x \ U" by auto ultimately have "\\<^sub>F x in at a within T. f x = g x" by eventually_elim (auto simp: \S = _\ eq) with f show ?thesis by (rule Lim_transform_eventually) qed lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (top_of_set T) U \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (clarsimp simp add: continuous_openin_preimage_eq openin_open) (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) then show "openin (top_of_set S) (S \ f -` U)" by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: "\continuous_on S f; f ` S \ T; openin (top_of_set T) U\ \ openin (top_of_set S) (S \ f -` U)" by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast show ?thesis proof assume L: ?lhs show ?rhs proof clarify fix U assume "closedin (top_of_set T) U" then show "closedin (top_of_set S) (S \ f -` U)" using L unfolding continuous_on_open_gen [OF assms] by (metis * closedin_def inf_le1 topspace_euclidean_subtopology) qed next assume R [rule_format]: ?rhs show ?lhs unfolding continuous_on_open_gen [OF assms] by (metis * R inf_le1 openin_closedin_eq topspace_euclidean_subtopology) qed qed lemma continuous_closedin_preimage_gen: assumes "continuous_on S f" "f ` S \ T" "closedin (top_of_set T) U" shows "closedin (top_of_set S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: assumes "continuous (at a within T) f" and "openin (top_of_set T) S" "a \ S" and eq: "\x. x \ S \ f x = g x" shows "continuous (at a within T) g" using assms by (simp add: Lim_transform_within_openin continuous_within) subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ the empty set. However, it is convenient to have, and is never a problem in proofs, so I prefer to write it down explicitly. We do not require \UNIV\ to be an open set, as this will not be the case in applications. (We are thinking of a topology on a subset of \UNIV\, the remaining part of \UNIV\ being irrelevant.)\ inductive generate_topology_on for S where Empty: "generate_topology_on S {}" | Int: "generate_topology_on S a \ generate_topology_on S b \ generate_topology_on S (a \ b)" | UN: "(\k. k \ K \ generate_topology_on S k) \ generate_topology_on S (\K)" | Basis: "s \ S \ generate_topology_on S s" lemma istopology_generate_topology_on: "istopology (generate_topology_on S)" unfolding istopology_def by (auto intro: generate_topology_on.intros) text \The basic property of the topology generated by a set \S\ is that it is the smallest topology containing all the elements of \S\:\ lemma generate_topology_on_coarsest: assumes "istopology T" "\s. s \ S \ T s" "generate_topology_on S s0" shows "T s0" using assms(3) apply (induct rule: generate_topology_on.induct) using assms(1) assms(2) unfolding istopology_def by auto abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ topology (generate_topology_on S)" lemma openin_topology_generated_by_iff: "openin (topology_generated_by S) s \ generate_topology_on S s" using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp lemma openin_topology_generated_by: "openin (topology_generated_by S) s \ generate_topology_on S s" using openin_topology_generated_by_iff by auto lemma topology_generated_by_topspace [simp]: "topspace (topology_generated_by S) = (\S)" proof { fix s assume "openin (topology_generated_by S) s" then have "generate_topology_on S s" by (rule openin_topology_generated_by) then have "s \ (\S)" by (induct, auto) } then show "topspace (topology_generated_by S) \ (\S)" unfolding topspace_def by auto next have "generate_topology_on S (\S)" using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp then show "(\S) \ topspace (topology_generated_by S)" unfolding topspace_def using openin_topology_generated_by_iff by auto qed lemma topology_generated_by_Basis: "s \ S \ openin (topology_generated_by S) s" by (simp only: openin_topology_generated_by_iff, auto simp: generate_topology_on.Basis) lemma generate_topology_on_Inter: "\finite \; \K. K \ \ \ generate_topology_on \ K; \ \ {}\ \ generate_topology_on \ (\\)" by (induction \ rule: finite_induct; force intro: generate_topology_on.intros) subsection\Topology bases and sub-bases\ lemma istopology_base_alt: "istopology (arbitrary union_of P) \ (\S T. (arbitrary union_of P) S \ (arbitrary union_of P) T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union) lemma istopology_base_eq: "istopology (arbitrary union_of P) \ (\S T. P S \ P T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_base_alt arbitrary_union_of_Int_eq) lemma istopology_base: "(\S T. \P S; P T\ \ P(S \ T)) \ istopology (arbitrary union_of P)" by (simp add: arbitrary_def istopology_base_eq union_of_inc) lemma openin_topology_base_unique: "openin X = arbitrary union_of P \ (\V. P V \ openin X V) \ (\U x. openin X U \ x \ U \ (\V. P V \ x \ V \ V \ U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: union_of_def arbitrary_def) next assume R: ?rhs then have *: "\\\Collect P. \\ = S" if "openin X S" for S using that by (rule_tac x="{V. P V \ V \ S}" in exI) fastforce from R show ?lhs by (fastforce simp add: union_of_def arbitrary_def intro: *) qed lemma topology_base_unique: "\\S. P S \ openin X S; \U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U\ \ topology(arbitrary union_of P) = X" by (metis openin_topology_base_unique openin_inverse [of X]) lemma topology_bases_eq_aux: "\(arbitrary union_of P) S; \U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U\ \ (arbitrary union_of Q) S" by (metis arbitrary_union_of_alt arbitrary_union_of_idempot) lemma topology_bases_eq: "\\U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U; \V x. \Q V; x \ V\ \ \U. P U \ x \ U \ U \ V\ \ topology (arbitrary union_of P) = topology (arbitrary union_of Q)" by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux) lemma istopology_subbase: "istopology (arbitrary union_of (finite intersection_of P relative_to S))" by (simp add: finite_intersection_of_Int istopology_base relative_to_Int) lemma openin_subbase: "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S \ (arbitrary union_of (finite intersection_of B relative_to U)) S" by (simp add: istopology_subbase topology_inverse') lemma topspace_subbase [simp]: "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _") proof show "?lhs \ U" by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset) show "U \ ?lhs" by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase openin_subset relative_to_inc subset_UNIV topology_inverse') qed lemma minimal_topology_subbase: "\\S. P S \ openin X S; openin X U; openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S\ \ openin X S" apply (simp add: istopology_subbase topology_inverse) apply (simp add: union_of_def intersection_of_def relative_to_def) apply (blast intro: openin_Int_Inter) done lemma istopology_subbase_UNIV: "istopology (arbitrary union_of (finite intersection_of P))" by (simp add: istopology_base finite_intersection_of_Int) lemma generate_topology_on_eq: "generate_topology_on S = arbitrary union_of finite' intersection_of (\x. x \ S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix A assume "?lhs A" then show "?rhs A" proof induction case (Int a b) then show ?case by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base) next case (UN K) then show ?case by (simp add: arbitrary_union_of_Union) next case (Basis s) then show ?case by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset) qed auto next fix A assume "?rhs A" then obtain \ where \: "\T. T \ \ \ \\. finite' \ \ \ \ S \ \\ = T" and eq: "A = \\" unfolding union_of_def intersection_of_def by auto show "?lhs A" unfolding eq proof (rule generate_topology_on.UN) fix T assume "T \ \" with \ obtain \ where "finite' \" "\ \ S" "\\ = T" by blast have "generate_topology_on S (\\)" proof (rule generate_topology_on_Inter) show "finite \" "\ \ {}" by (auto simp: \finite' \\) show "\K. K \ \ \ generate_topology_on S K" by (metis \\ \ S\ generate_topology_on.simps subset_iff) qed then show "generate_topology_on S T" using \\\ = T\ by blast qed qed lemma continuous_on_generated_topo_iff: "continuous_map T1 (topology_generated_by S) f \ ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" unfolding continuous_map_alt topology_generated_by_topspace proof (auto simp add: topology_generated_by_Basis) assume H: "\U. U \ S \ openin T1 (f -` U \ topspace T1)" fix U assume "openin (topology_generated_by S) U" then have "generate_topology_on S U" by (rule openin_topology_generated_by) then show "openin T1 (f -` U \ topspace T1)" proof (induct) fix a b assume H: "openin T1 (f -` a \ topspace T1)" "openin T1 (f -` b \ topspace T1)" have "f -` (a \ b) \ topspace T1 = (f-`a \ topspace T1) \ (f-`b \ topspace T1)" by auto then show "openin T1 (f -` (a \ b) \ topspace T1)" using H by auto next fix K assume H: "openin T1 (f -` k \ topspace T1)" if "k\ K" for k define L where "L = {f -` k \ topspace T1|k. k \ K}" have *: "openin T1 l" if "l \L" for l using that H unfolding L_def by auto have "openin T1 (\L)" using openin_Union[OF *] by simp moreover have "(\L) = (f -` \K \ topspace T1)" unfolding L_def by auto ultimately show "openin T1 (f -` \K \ topspace T1)" by simp qed (auto simp add: H) qed lemma continuous_on_generated_topo: assumes "\U. U \S \ openin T1 (f-`U \ topspace(T1))" "f`(topspace T1) \ (\ S)" shows "continuous_map T1 (topology_generated_by S) f" using assms continuous_on_generated_topo_iff by blast subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as we will need it to define the strong operator topology on the space of continuous linear operators, by pulling back the product topology on the space of all functions.\ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: "istopology (\S. \U. openin T U \ S = f-`U \ A)" unfolding istopology_def proof (auto) fix K assume "\S\K. \U. openin T U \ S = f -` U \ A" then have "\U. \S\K. openin T (U S) \ S = f-`(U S) \ A" by (rule bchoice) then obtain U where U: "\S\K. openin T (U S) \ S = f-`(U S) \ A" by blast define V where "V = (\S\K. U S)" have "openin T V" "\K = f -` V \ A" unfolding V_def using U by auto then show "\V. openin T V \ \K = f -` V \ A" by auto qed lemma openin_pullback_topology: "openin (pullback_topology A f T) S \ (\U. openin T U \ S = f-`U \ A)" unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto lemma topspace_pullback_topology: "topspace (pullback_topology A f T) = f-`(topspace T) \ A" by (auto simp add: topspace_def openin_pullback_topology) proposition continuous_map_pullback [intro]: assumes "continuous_map T1 T2 g" shows "continuous_map (pullback_topology A f T1) T2 (g o f)" unfolding continuous_map_alt proof (auto) fix U::"'b set" assume "openin T2 U" then have "openin T1 (g-`U \ topspace T1)" using assms unfolding continuous_map_alt by auto have "(g o f)-`U \ topspace (pullback_topology A f T1) = (g o f)-`U \ A \ f-`(topspace T1)" unfolding topspace_pullback_topology by auto also have "... = f-`(g-`U \ topspace T1) \ A " by auto also have "openin (pullback_topology A f T1) (...)" unfolding openin_pullback_topology using \openin T1 (g-`U \ topspace T1)\ by auto finally show "openin (pullback_topology A f T1) ((g \ f) -` U \ topspace (pullback_topology A f T1))" by auto next fix x assume "x \ topspace (pullback_topology A f T1)" then have "f x \ topspace T1" unfolding topspace_pullback_topology by auto then show "g (f x) \ topspace T2" using assms unfolding continuous_map_def by auto qed proposition continuous_map_pullback' [intro]: assumes "continuous_map T1 T2 (f o g)" "topspace T1 \ g-`A" shows "continuous_map T1 (pullback_topology A f T2) g" unfolding continuous_map_alt proof (auto) fix U assume "openin (pullback_topology A f T2) U" then have "\V. openin T2 V \ U = f-`V \ A" unfolding openin_pullback_topology by auto then obtain V where "openin T2 V" "U = f-`V \ A" by blast then have "g -` U \ topspace T1 = g-`(f-`V \ A) \ topspace T1" by blast also have "... = (f o g)-`V \ (g-`A \ topspace T1)" by auto also have "... = (f o g)-`V \ topspace T1" using assms(2) by auto also have "openin T1 (...)" using assms(1) \openin T2 V\ by auto finally show "openin T1 (g -` U \ topspace T1)" by simp next fix x assume "x \ topspace T1" have "(f o g) x \ topspace T2" using assms(1) \x \ topspace T1\ unfolding continuous_map_def by auto then have "g x \ f-`(topspace T2)" unfolding comp_def by blast moreover have "g x \ A" using assms(2) \x \ topspace T1\ by blast ultimately show "g x \ topspace (pullback_topology A f T2)" unfolding topspace_pullback_topology by blast qed subsection\Proper maps (not a priori assumed continuous) \ definition proper_map where "proper_map X Y f \ closed_map X Y f \ (\y \ topspace Y. compactin X {x \ topspace X. f x = y})" lemma proper_imp_closed_map: "proper_map X Y f \ closed_map X Y f" by (simp add: proper_map_def) lemma proper_map_imp_subset_topspace: "proper_map X Y f \ f ` (topspace X) \ topspace Y" by (simp add: closed_map_imp_subset_topspace proper_map_def) lemma closed_injective_imp_proper_map: assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)" shows "proper_map X Y f" unfolding proper_map_def proof (clarsimp simp: f) show "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using inj_on_eq_iff [OF inj] by auto then show ?thesis using that by (metis (no_types, lifting) compactin_empty compactin_sing) qed qed lemma injective_imp_proper_eq_closed_map: "inj_on f (topspace X) \ (proper_map X Y f \ closed_map X Y f)" using closed_injective_imp_proper_map proper_imp_closed_map by blast lemma homeomorphic_imp_proper_map: "homeomorphic_map X Y f \ proper_map X Y f" by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map) lemma compactin_proper_map_preimage: assumes f: "proper_map X Y f" and "compactin Y K" shows "compactin X {x. x \ topspace X \ f x \ K}" proof - have "f ` (topspace X) \ topspace Y" by (simp add: f proper_map_imp_subset_topspace) have *: "\y. y \ topspace Y \ compactin X {x \ topspace X. f x = y}" using f by (auto simp: proper_map_def) show ?thesis unfolding compactin_def proof clarsimp show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x \ K} \ \\" if \: "\U\\. openin X U" and sub: "{x \ topspace X. f x \ K} \ \\" for \ proof - have "\y \ K. \\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" proof fix y assume "y \ K" then have "compactin X {x \ topspace X. f x = y}" by (metis "*" \compactin Y K\ compactin_subspace subsetD) with \y \ K\ show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" unfolding compactin_def using \ sub by fastforce qed then obtain \ where \: "\y. y \ K \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by (metis (full_types)) define F where "F \ \y. topspace Y - f ` (topspace X - \(\ y))" have "\\. finite \ \ \ \ F ` K \ K \ \\" proof (rule compactinD [OF \compactin Y K\]) have "\x. x \ K \ closedin Y (f ` (topspace X - \(\ x)))" using f unfolding proper_map_def closed_map_def by (meson \ \ openin_Union openin_closedin_eq subsetD) then show "openin Y U" if "U \ F ` K" for U using that by (auto simp: F_def) show "K \ \(F ` K)" using \ \compactin Y K\ unfolding F_def compactin_def by fastforce qed then obtain J where "finite J" "J \ K" and J: "K \ \(F ` J)" by (auto simp: ex_finite_subset_image) show ?thesis unfolding F_def proof (intro exI conjI) show "finite (\(\ ` J))" using \ \J \ K\ \finite J\ by blast show "\(\ ` J) \ \" using \ \J \ K\ by blast show "{x \ topspace X. f x \ K} \ \(\(\ ` J))" using J \J \ K\ unfolding F_def by auto qed qed qed qed lemma compact_space_proper_map_preimage: assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y" shows "compact_space X" proof - have eq: "topspace X = {x \ topspace X. f x \ topspace Y}" using fim by blast moreover have "compactin Y (topspace Y)" using \compact_space Y\ compact_space_def by auto ultimately show ?thesis unfolding compact_space_def using eq f compactin_proper_map_preimage by fastforce qed lemma proper_map_alt: "proper_map X Y f \ closed_map X Y f \ (\K. compactin Y K \ compactin X {x. x \ topspace X \ f x \ K})" proof (intro iffI conjI allI impI) show "compactin X {x \ topspace X. f x \ K}" if "proper_map X Y f" and "compactin Y K" for K using that by (simp add: compactin_proper_map_preimage) show "proper_map X Y f" if f: "closed_map X Y f \ (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" proof - have "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "compactin X {x \ topspace X. f x \ {y}}" using f compactin_sing that by fastforce then show ?thesis by auto qed with f show ?thesis by (auto simp: proper_map_def) qed qed (simp add: proper_imp_closed_map) lemma proper_map_on_empty: "topspace X = {} \ proper_map X Y f" by (auto simp: proper_map_def closed_map_on_empty) lemma proper_map_id [simp]: "proper_map X X id" proof (clarsimp simp: proper_map_alt closed_map_id) fix K assume K: "compactin X K" then have "{a \ topspace X. a \ K} = K" by (simp add: compactin_subspace subset_antisym subset_iff) then show "compactin X {a \ topspace X. a \ K}" using K by auto qed lemma proper_map_compose: assumes "proper_map X Y f" "proper_map Y Z g" shows "proper_map X Z (g \ f)" proof - have "closed_map X Y f" and f: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" and "closed_map Y Z g" and g: "\K. compactin Z K \ compactin Y {x \ topspace Y. g x \ K}" using assms by (auto simp: proper_map_alt) show ?thesis unfolding proper_map_alt proof (intro conjI allI impI) show "closed_map X Z (g \ f)" using \closed_map X Y f\ \closed_map Y Z g\ closed_map_compose by blast have "{x \ topspace X. g (f x) \ K} = {x \ topspace X. f x \ {b \ topspace Y. g b \ K}}" for K using \closed_map X Y f\ closed_map_imp_subset_topspace by blast then show "compactin X {x \ topspace X. (g \ f) x \ K}" if "compactin Z K" for K using f [OF g [OF that]] by auto qed qed lemma proper_map_const: "proper_map X Y (\x. c) \ compact_space X \ (topspace X = {} \ closedin Y {c})" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: compact_space_topspace_empty proper_map_on_empty) next case False have *: "compactin X {x \ topspace X. c = y}" if "compact_space X" for y proof (cases "c = y") case True then show ?thesis using compact_space_def \compact_space X\ by auto qed auto then show ?thesis using closed_compactin closedin_subset by (force simp: False proper_map_def closed_map_const compact_space_def) qed lemma proper_map_inclusion: "s \ topspace X \ proper_map (subtopology X s) X id \ closedin X s \ (\k. compactin X k \ compactin X (s \ k))" by (auto simp: proper_map_alt closed_map_inclusion_eq inf.absorb_iff2 Collect_conj_eq compactin_subtopology intro: closed_Int_compactin) subsection\Perfect maps (proper, continuous and surjective)\ definition perfect_map where "perfect_map X Y f \ continuous_map X Y f \ proper_map X Y f \ f ` (topspace X) = topspace Y" lemma homeomorphic_imp_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f" by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def) lemma perfect_imp_quotient_map: "perfect_map X Y f \ quotient_map X Y f" by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def) lemma homeomorphic_eq_injective_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f \ inj_on f (topspace X)" using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast lemma perfect_injective_eq_homeomorphic_map: "perfect_map X Y f \ inj_on f (topspace X) \ homeomorphic_map X Y f" by (simp add: homeomorphic_eq_injective_perfect_map) lemma perfect_map_id [simp]: "perfect_map X X id" by (simp add: homeomorphic_imp_perfect_map) lemma perfect_map_compose: "\perfect_map X Y f; perfect_map Y Z g\ \ perfect_map X Z (g \ f)" by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def) lemma perfect_imp_continuous_map: "perfect_map X Y f \ continuous_map X Y f" using perfect_map_def by blast lemma perfect_imp_closed_map: "perfect_map X Y f \ closed_map X Y f" by (simp add: perfect_map_def proper_map_def) lemma perfect_imp_proper_map: "perfect_map X Y f \ proper_map X Y f" by (simp add: perfect_map_def) lemma perfect_imp_surjective_map: "perfect_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: perfect_map_def) end diff --git a/src/HOL/Analysis/Abstract_Topology_2.thy b/src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy +++ b/src/HOL/Analysis/Abstract_Topology_2.thy @@ -1,1616 +1,1616 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Abstract Topology 2\ theory Abstract_Topology_2 imports Elementary_Topology Abstract_Topology "HOL-Library.Indicator_Function" begin text \Combination of Elementary and Abstract Topology\ (* FIXME: move elsewhere *) lemma approachable_lt_le: "(\(d::real) > 0. \x. f x < d \ P x) \ (\d>0. \x. f x \ d \ P x)" apply auto apply (rule_tac x="d/2" in exI) apply auto done lemma approachable_lt_le2: \ \like the above, but pushes aside an extra formula\ "(\(d::real) > 0. \x. Q x \ f x < d \ P x) \ (\d>0. \x. f x \ d \ Q x \ P x)" apply auto apply (rule_tac x="d/2" in exI, auto) done lemma triangle_lemma: fixes x y z :: real assumes x: "0 \ x" and y: "0 \ y" and z: "0 \ z" and xy: "x\<^sup>2 \ y\<^sup>2 + z\<^sup>2" shows "x \ y + z" proof - have "y\<^sup>2 + z\<^sup>2 \ y\<^sup>2 + 2 * y * z + z\<^sup>2" using z y by simp with xy have th: "x\<^sup>2 \ (y + z)\<^sup>2" by (simp add: power2_eq_square field_simps) from y z have yz: "y + z \ 0" by arith from power2_le_imp_le[OF th yz] show ?thesis . qed lemma isCont_indicator: fixes x :: "'a::t2_space" shows "isCont (indicator A :: 'a \ real) x = (x \ frontier A)" proof auto fix x assume cts_at: "isCont (indicator A :: 'a \ real) x" and fr: "x \ frontier A" with continuous_at_open have 1: "\V::real set. open V \ indicator A x \ V \ (\U::'a set. open U \ x \ U \ (\y\U. indicator A y \ V))" by auto show False proof (cases "x \ A") assume x: "x \ A" hence "indicator A x \ ({0<..<2} :: real set)" by simp hence "\U. open U \ x \ U \ (\y\U. indicator A y \ ({0<..<2} :: real set))" using 1 open_greaterThanLessThan by blast then guess U .. note U = this hence "\y\U. indicator A y > (0::real)" unfolding greaterThanLessThan_def by auto hence "U \ A" using indicator_eq_0_iff by force hence "x \ interior A" using U interiorI by auto thus ?thesis using fr unfolding frontier_def by simp next assume x: "x \ A" hence "indicator A x \ ({-1<..<1} :: real set)" by simp hence "\U. open U \ x \ U \ (\y\U. indicator A y \ ({-1<..<1} :: real set))" using 1 open_greaterThanLessThan by blast then guess U .. note U = this hence "\y\U. indicator A y < (1::real)" unfolding greaterThanLessThan_def by auto hence "U \ -A" by auto hence "x \ interior (-A)" using U interiorI by auto thus ?thesis using fr interior_complement unfolding frontier_def by auto qed next assume nfr: "x \ frontier A" hence "x \ interior A \ x \ interior (-A)" by (auto simp: frontier_def closure_interior) thus "isCont ((indicator A)::'a \ real) x" proof assume int: "x \ interior A" then obtain U where U: "open U" "x \ U" "U \ A" unfolding interior_def by auto hence "\y\U. indicator A y = (1::real)" unfolding indicator_def by auto - hence "continuous_on U (indicator A)" by (simp add: continuous_on_const indicator_eq_1_iff) + hence "continuous_on U (indicator A)" by (simp add: indicator_eq_1_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto next assume ext: "x \ interior (-A)" then obtain U where U: "open U" "x \ U" "U \ -A" unfolding interior_def by auto then have "continuous_on U (indicator A)" using continuous_on_topological by (auto simp: subset_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto qed qed lemma closedin_limpt: "closedin (top_of_set T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" apply (simp add: closedin_closed, safe) apply (simp add: closed_limpt islimpt_subset) apply (rule_tac x="closure S" in exI, simp) apply (force simp: closure_def) done lemma closedin_closed_eq: "closed S \ closedin (top_of_set S) T \ closed T \ T \ S" by (meson closedin_limpt closed_subset closedin_closed_trans) lemma connected_closed_set: "closed S \ connected S \ (\A B. closed A \ closed B \ A \ {} \ B \ {} \ A \ B = S \ A \ B = {})" unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast text \If a connnected set is written as the union of two nonempty closed sets, then these sets have to intersect.\ lemma connected_as_closed_union: assumes "connected C" "C = A \ B" "closed A" "closed B" "A \ {}" "B \ {}" shows "A \ B \ {}" by (metis assms closed_Un connected_closed_set) lemma closedin_subset_trans: "closedin (top_of_set U) S \ S \ T \ T \ U \ closedin (top_of_set T) S" by (meson closedin_limpt subset_iff) lemma openin_subset_trans: "openin (top_of_set U) S \ S \ T \ T \ U \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_compact: "\compact S; closedin (top_of_set S) T\ \ compact T" by (metis closedin_closed compact_Int_closed) lemma closedin_compact_eq: fixes S :: "'a::t2_space set" shows "compact S \ (closedin (top_of_set S) T \ compact T \ T \ S)" by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) subsection \Closure\ lemma euclidean_closure_of [simp]: "euclidean closure_of S = closure S" by (auto simp: closure_of_def closure_def islimpt_def) lemma closure_openin_Int_closure: assumes ope: "openin (top_of_set U) S" and "T \ U" shows "closure(S \ closure T) = closure(S \ T)" proof obtain V where "open V" and S: "S = U \ V" using ope using openin_open by metis show "closure (S \ closure T) \ closure (S \ T)" proof (clarsimp simp: S) fix x assume "x \ closure (U \ V \ closure T)" then have "V \ closure T \ A \ x \ closure A" for A by (metis closure_mono subsetD inf.coboundedI2 inf_assoc) then have "x \ closure (T \ V)" by (metis \open V\ closure_closure inf_commute open_Int_closure_subset) then show "x \ closure (U \ V \ T)" by (metis \T \ U\ inf.absorb_iff2 inf_assoc inf_commute) qed next show "closure (S \ T) \ closure (S \ closure T)" by (meson Int_mono closure_mono closure_subset order_refl) qed corollary infinite_openin: fixes S :: "'a :: t1_space set" shows "\openin (top_of_set U) S; x \ S; x islimpt U\ \ infinite S" by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute) lemma closure_Int_ballI: assumes "\U. \openin (top_of_set S) U; U \ {}\ \ T \ U \ {}" shows "S \ closure T" proof (clarsimp simp: closure_iff_nhds_not_empty) fix x and A and V assume "x \ S" "V \ A" "open V" "x \ V" "T \ A = {}" then have "openin (top_of_set S) (A \ V \ S)" by (auto simp: openin_open intro!: exI[where x="V"]) moreover have "A \ V \ S \ {}" using \x \ V\ \V \ A\ \x \ S\ by auto ultimately have "T \ (A \ V \ S) \ {}" by (rule assms) with \T \ A = {}\ show False by auto qed subsection \Frontier\ lemma euclidean_interior_of [simp]: "euclidean interior_of S = interior S" by (auto simp: interior_of_def interior_def) lemma euclidean_frontier_of [simp]: "euclidean frontier_of S = frontier S" by (auto simp: frontier_of_def frontier_def) lemma connected_Int_frontier: "\connected s; s \ t \ {}; s - t \ {}\ \ (s \ frontier t \ {})" apply (simp add: frontier_interiors connected_openin, safe) apply (drule_tac x="s \ interior t" in spec, safe) apply (drule_tac [2] x="s \ interior (-t)" in spec) apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) done subsection \Compactness\ lemma openin_delete: fixes a :: "'a :: t1_space" shows "openin (top_of_set u) s \ openin (top_of_set u) (s - {a})" by (metis Int_Diff open_delete openin_open) lemma compact_eq_openin_cover: "compact S \ (\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D))" proof safe fix C assume "compact S" and "\c\C. openin (top_of_set S) c" and "S \ \C" then have "\c\{T. open T \ S \ T \ C}. open c" and "S \ \{T. open T \ S \ T \ C}" unfolding openin_open by force+ with \compact S\ obtain D where "D \ {T. open T \ S \ T \ C}" and "finite D" and "S \ \D" by (meson compactE) then have "image (\T. S \ T) D \ C \ finite (image (\T. S \ T) D) \ S \ \(image (\T. S \ T) D)" by auto then show "\D\C. finite D \ S \ \D" .. next assume 1: "\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D)" show "compact S" proof (rule compactI) fix C let ?C = "image (\T. S \ T) C" assume "\t\C. open t" and "S \ \C" then have "(\c\?C. openin (top_of_set S) c) \ S \ \?C" unfolding openin_open by auto with 1 obtain D where "D \ ?C" and "finite D" and "S \ \D" by metis let ?D = "inv_into C (\T. S \ T) ` D" have "?D \ C \ finite ?D \ S \ \?D" proof (intro conjI) from \D \ ?C\ show "?D \ C" by (fast intro: inv_into_into) from \finite D\ show "finite ?D" by (rule finite_imageI) from \S \ \D\ show "S \ \?D" apply (rule subset_trans, clarsimp) apply (frule subsetD [OF \D \ ?C\, THEN f_inv_into_f]) apply (erule rev_bexI, fast) done qed then show "\D\C. finite D \ S \ \D" .. qed qed subsection \Continuity\ lemma interior_image_subset: assumes "inj f" "\x. continuous (at x) f" shows "interior (f ` S) \ f ` (interior S)" proof fix x assume "x \ interior (f ` S)" then obtain T where as: "open T" "x \ T" "T \ f ` S" .. then have "x \ f ` S" by auto then obtain y where y: "y \ S" "x = f y" by auto have "open (f -` T)" using assms \open T\ by (simp add: continuous_at_imp_continuous_on open_vimage) moreover have "y \ vimage f T" using \x = f y\ \x \ T\ by simp moreover have "vimage f T \ S" using \T \ image f S\ \inj f\ unfolding inj_on_def subset_eq by auto ultimately have "y \ interior S" .. with \x = f y\ show "x \ f ` interior S" .. qed subsection\<^marker>\tag unimportant\ \Equality of continuous functions on closure and related results\ lemma continuous_closedin_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closedin (top_of_set S) {x \ S. f x = a}" using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_closed_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closed S \ closed {x \ S. f x = a}" using continuous_closed_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_constant_on_closure: fixes f :: "_ \ 'b::t1_space" assumes "continuous_on (closure S) f" and "\x. x \ S \ f x = a" and "x \ closure S" shows "f x = a" using continuous_closed_preimage_constant[of "closure S" f a] assms closure_minimal[of S "{x \ closure S. f x = a}"] closure_subset unfolding subset_eq by auto lemma image_closure_subset: assumes contf: "continuous_on (closure S) f" and "closed T" and "(f ` S) \ T" shows "f ` (closure S) \ T" proof - have "S \ {x \ closure S. f x \ T}" using assms(3) closure_subset by auto moreover have "closed (closure S \ f -` T)" using continuous_closed_preimage[OF contf] \closed T\ by auto ultimately have "closure S = (closure S \ f -` T)" using closure_minimal[of S "(closure S \ f -` T)"] by auto then show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \A function constant on a set\ definition constant_on (infixl "(constant'_on)" 50) where "f constant_on A \ \y. \x\A. f x = y" lemma constant_on_subset: "\f constant_on A; B \ A\ \ f constant_on B" unfolding constant_on_def by blast lemma injective_not_constant: fixes S :: "'a::{perfect_space} set" shows "\open S; inj_on f S; f constant_on S\ \ S = {}" unfolding constant_on_def by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def) lemma constant_on_closureI: fixes f :: "_ \ 'b::t1_space" assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f" shows "f constant_on (closure S)" using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def by metis subsection\<^marker>\tag unimportant\ \Continuity relative to a union.\ lemma continuous_on_Un_local: "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; continuous_on s f; continuous_on t f\ \ continuous_on (s \ t) f" unfolding continuous_on closedin_limpt by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) lemma continuous_on_cases_local: "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; continuous_on s f; continuous_on t g; \x. \x \ s \ \P x \ x \ t \ P x\ \ f x = g x\ \ continuous_on (s \ t) (\x. if P x then f x else g x)" by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) lemma continuous_on_cases_le: fixes h :: "'a :: topological_space \ real" assumes "continuous_on {t \ s. h t \ a} f" and "continuous_on {t \ s. a \ h t} g" and h: "continuous_on s h" and "\t. \t \ s; h t = a\ \ f t = g t" shows "continuous_on s (\t. if h t \ a then f(t) else g(t))" proof - have s: "s = (s \ h -` atMost a) \ (s \ h -` atLeast a)" by force have 1: "closedin (top_of_set s) (s \ h -` atMost a)" by (rule continuous_closedin_preimage [OF h closed_atMost]) have 2: "closedin (top_of_set s) (s \ h -` atLeast a)" by (rule continuous_closedin_preimage [OF h closed_atLeast]) have eq: "s \ h -` {..a} = {t \ s. h t \ a}" "s \ h -` {a..} = {t \ s. a \ h t}" by auto show ?thesis apply (rule continuous_on_subset [of s, OF _ order_refl]) apply (subst s) apply (rule continuous_on_cases_local) using 1 2 s assms apply (auto simp: eq) done qed lemma continuous_on_cases_1: fixes s :: "real set" assumes "continuous_on {t \ s. t \ a} f" and "continuous_on {t \ s. a \ t} g" and "a \ s \ f a = g a" shows "continuous_on s (\t. if t \ a then f(t) else g(t))" using assms -by (auto simp: continuous_on_id intro: continuous_on_cases_le [where h = id, simplified]) +by (auto intro: continuous_on_cases_le [where h = id, simplified]) subsection\<^marker>\tag unimportant\\Inverse function property for open/closed maps\ lemma continuous_on_inverse_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g (f x) = x" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) qed lemma continuous_on_inverse_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g(f x) = x" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) qed lemma homeomorphism_injective_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_injective_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_imp_open_map: assumes hom: "homeomorphism S T f g" and oo: "openin (top_of_set S) U" shows "openin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_open oo) qed lemma homeomorphism_imp_closed_map: assumes hom: "homeomorphism S T f g" and oo: "closedin (top_of_set S) U" shows "closedin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_closed oo) qed subsection\<^marker>\tag unimportant\ \Seperability\ lemma subset_second_countable: obtains \ :: "'a:: second_countable_topology set set" where "countable \" "{} \ \" "\C. C \ \ \ openin(top_of_set S) C" "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and opeB: "\C. C \ \ \ openin(top_of_set S) C" and \: "\T. openin(top_of_set 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 \ (\) S ` \ \ openin (top_of_set S) C" using ope by auto show "\T. openin (top_of_set S) T \ \\\(\) 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 (top_of_set S) C" by (simp add: \\C. C \ \ \ openin (top_of_set S) C\) show "\\\\ - {{}}. T = \\" if "openin (top_of_set S) T" for T using \ [OF that] apply clarify apply (rule_tac x="\ - {{}}" in exI, auto) done qed auto qed lemma Lindelof_openin: fixes \ :: "'a::second_countable_topology set set" assumes "\S. S \ \ \ openin (top_of_set U) S" obtains \' where "\' \ \" "countable \'" "\\' = \\" proof - have "\S. S \ \ \ \T. open T \ S = U \ T" using assms by (simp add: openin_open) then obtain tf where tf: "\S. S \ \ \ open (tf S) \ (S = U \ tf S)" by metis have [simp]: "\\'. \' \ \ \ \\' = U \ \(tf ` \')" using tf by fastforce obtain \ where "countable \ \ \ \ tf ` \" "\\ = \(tf ` \)" using tf by (force intro: Lindelof [of "tf ` \"]) then obtain \' where \': "\' \ \" "countable \'" "\\' = \\" by (clarsimp simp add: countable_subset_image) then show ?thesis .. qed subsection\<^marker>\tag unimportant\\Closed Maps\ lemma continuous_imp_closed_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "closedin (top_of_set S) U" "continuous_on S f" "f ` S = T" "compact S" shows "closedin (top_of_set T) (f ` U)" by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff) lemma closed_map_restrict: assumes cloU: "closedin (top_of_set (S \ f -` T')) U" and cc: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" and "T' \ T" shows "closedin (top_of_set T') (f ` U)" proof - obtain V where "closed V" "U = S \ f -` T' \ V" using cloU by (auto simp: closedin_closed) with cc [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: closedin_closed) qed subsection\<^marker>\tag unimportant\\Open Maps\ lemma open_map_restrict: assumes opeU: "openin (top_of_set (S \ f -` T')) U" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" and "T' \ T" shows "openin (top_of_set T') (f ` U)" proof - obtain V where "open V" "U = S \ f -` T' \ V" using opeU by (auto simp: openin_open) with oo [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: openin_open) qed subsection\<^marker>\tag unimportant\\Quotient maps\ lemma quotient_map_imp_continuous_open: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis using ope [OF T] apply (simp add: continuous_on_open) by (meson ope openin_imp_subset openin_trans) qed lemma quotient_map_imp_continuous_closed: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (closedin (top_of_set S) (S \ f -` U) \ closedin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis using ope [OF T] apply (simp add: continuous_on_closed) by (metis (no_types, lifting) ope closedin_imp_subset closedin_trans) qed lemma open_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) = openin (top_of_set (f ` S)) T" proof - have "T = f ` (S \ f -` T)" using T by blast then show ?thesis using "ope" contf continuous_on_open by metis qed lemma closed_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. closedin (top_of_set S) T \ closedin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" (is "?lhs = ?rhs") proof assume ?lhs then have *: "closedin (top_of_set S) (S - (S \ f -` T))" using closedin_diff by fastforce have [simp]: "(f ` S - f ` (S - (S \ f -` T))) = T" using T by blast show ?rhs using ope [OF *, unfolded closedin_def] by auto next assume ?rhs with contf show ?lhs by (auto simp: continuous_on_open) qed lemma continuous_right_inverse_imp_quotient_map: assumes contf: "continuous_on S f" and imf: "f ` S \ T" and contg: "continuous_on T g" and img: "g ` T \ S" and fg [simp]: "\y. y \ T \ f(g y) = y" and U: "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" (is "?lhs = ?rhs") proof - have f: "\Z. openin (top_of_set (f ` S)) Z \ openin (top_of_set S) (S \ f -` Z)" and g: "\Z. openin (top_of_set (g ` T)) Z \ openin (top_of_set T) (T \ g -` Z)" using contf contg by (auto simp: continuous_on_open) show ?thesis proof have "T \ g -` (g ` T \ (S \ f -` U)) = {x \ T. f (g x) \ U}" using imf img by blast also have "... = U" using U by auto finally have eq: "T \ g -` (g ` T \ (S \ f -` U)) = U" . assume ?lhs then have *: "openin (top_of_set (g ` T)) (g ` T \ (S \ f -` U))" by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) show ?rhs using g [OF *] eq by auto next assume rhs: ?rhs show ?lhs by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs) qed qed lemma continuous_left_inverse_imp_quotient_map: assumes "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g(f x) = x" and "U \ f ` S" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set (f ` S)) U" apply (rule continuous_right_inverse_imp_quotient_map) using assms apply force+ done lemma continuous_imp_quotient_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "continuous_on S f" "f ` S = T" "compact S" "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map) subsection\<^marker>\tag unimportant\\Pasting lemmas for functions, for of casewise definitions\ subsubsection\on open sets\ lemma pasting_lemma: assumes ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_openin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace topspace_subtopology by fastforce next fix U assume Y: "openin Y U" have T: "T i \ topspace X" if "i \ I" for i using ope by (simp add: openin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have "\i. i \ I \ openin X (T i \ f i -` U)" using cont unfolding continuous_map_openin_preimage_eq by (metis Y T inf.commute inf_absorb1 ope topspace_subtopology openin_trans_full) then show "openin X (topspace X \ g -` U)" by (auto simp: *) qed lemma pasting_lemma_exists: assumes X: "topspace X \ (\i \ I. T i)" and ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map (subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof let ?h = "\x. f (SOME i. i \ I \ x \ T i) x" show "continuous_map X Y ?h" apply (rule pasting_lemma [OF ope cont]) apply (blast intro: f)+ by (metis (no_types, lifting) UN_E X subsetD someI_ex) show "f (SOME i. i \ I \ x \ T i) x = f i x" if "i \ I" "x \ topspace X \ T i" for i x by (metis (no_types, lifting) IntD2 IntI f someI_ex that) qed lemma pasting_lemma_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_closedin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace topspace_subtopology by fastforce next fix U assume Y: "closedin Y U" have T: "T i \ topspace X" if "i \ I" for i using clo by (simp add: closedin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have cTf: "\i. i \ I \ closedin X (T i \ f i -` U)" using cont unfolding continuous_map_closedin_preimage_eq topspace_subtopology by (simp add: Int_absorb1 T Y clo closedin_closed_subtopology) have sub: "{Z \ (\i. T i \ f i -` U) ` I. Z \ V \ {}} \ (\i. T i \ f i -` U) ` {i \ I. T i \ V \ {}}" for V by auto have 1: "(\i\I. T i \ f i -` U) \ topspace X" using T by blast then have lf: "locally_finite_in X ((\i. T i \ f i -` U) ` I)" unfolding locally_finite_in_def using finite_subset [OF sub] fin by force show "closedin X (topspace X \ g -` U)" apply (subst *) apply (rule closedin_locally_finite_Union) apply (auto intro: cTf lf) done qed subsubsection\Likewise on closed sets, with a finiteness assumption\ lemma pasting_lemma_closed: assumes fin: "finite I" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" using pasting_lemma_locally_finite [OF _ clo cont f g] fin by auto lemma pasting_lemma_exists_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f(@i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_locally_finite [OF fin]) apply (blast intro: assms)+ by (metis (no_types, lifting) UN_E X set_rev_mp someI_ex) next fix x i assume "i \ I" and "x \ topspace X \ T i" show "f (SOME i. i \ I \ x \ T i) x = f i x" apply (rule someI2_ex) using \i \ I\ \x \ topspace X \ T i\ apply blast by (meson Int_iff \i \ I\ \x \ topspace X \ T i\ f) qed lemma pasting_lemma_exists_closed: assumes fin: "finite I" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f (SOME i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_closed [OF \finite I\ clo cont]) apply (blast intro: f)+ by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff) next fix x i assume "i \ I" "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (no_types, lifting) IntD2 IntI f someI_ex) qed lemma continuous_map_cases: assumes f: "continuous_map (subtopology X (X closure_of {x. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x. \ P x})) Y g" and fg: "\x. x \ X frontier_of {x. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" proof (rule pasting_lemma_closed) let ?f = "\b. if b then f else g" let ?g = "\x. if P x then f x else g x" let ?T = "\b. if b then X closure_of {x. P x} else X closure_of {x. ~P x}" show "finite {True,False}" by auto have eq: "topspace X - Collect P = topspace X \ {x. \ P x}" by blast show "?f i x = ?f j x" if "i \ {True,False}" "j \ {True,False}" and x: "x \ topspace X \ ?T i \ ?T j" for i j x proof - have "f x = g x" if "i" "\ j" apply (rule fg) unfolding frontier_of_closures eq using x that closure_of_restrict by fastforce moreover have "g x = f x" if "x \ X closure_of {x. \ P x}" "x \ X closure_of Collect P" "\ i" "j" for x apply (rule fg [symmetric]) unfolding frontier_of_closures eq using x that closure_of_restrict by fastforce ultimately show ?thesis using that by (auto simp flip: closure_of_restrict) qed show "\j. j \ {True,False} \ x \ ?T j \ (if P x then f x else g x) = ?f j x" if "x \ topspace X" for x apply simp apply safe apply (metis Int_iff closure_of inf_sup_absorb mem_Collect_eq that) by (metis DiffI eq closure_of_subset_Int contra_subsetD mem_Collect_eq that) qed (auto simp: f g) lemma continuous_map_cases_alt: assumes f: "continuous_map (subtopology X (X closure_of {x \ topspace X. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x \ topspace X. ~P x})) Y g" and fg: "\x. x \ X frontier_of {x \ topspace X. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" apply (rule continuous_map_cases) using assms apply (simp_all add: Collect_conj_eq closure_of_restrict [symmetric] frontier_of_restrict [symmetric]) done lemma continuous_map_cases_function: assumes contp: "continuous_map X Z p" and contf: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of U}) Y f" and contg: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of (topspace Z - U)}) Y g" and fg: "\x. \x \ topspace X; p x \ Z frontier_of U\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ U then f x else g x)" proof (rule continuous_map_cases_alt) show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y f" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of U}" show "continuous_map (subtopology X ?T) Y f" by (simp add: contf) show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule continuous_map_closure_preimage_subset [OF contp]) qed show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y g" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of (topspace Z - U)}" show "continuous_map (subtopology X ?T) Y g" by (simp add: contg) have "X closure_of {x \ topspace X. p x \ U} \ X closure_of {x \ topspace X. p x \ topspace Z - U}" apply (rule closure_of_mono) using continuous_map_closedin contp by fastforce then show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]]) qed next show "f x = g x" if "x \ X frontier_of {x \ topspace X. p x \ U}" for x using that continuous_map_frontier_frontier_preimage_subset [OF contp, of U] fg by blast qed subsection \Retractions\ definition\<^marker>\tag important\ retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" where "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" definition\<^marker>\tag important\ retract_of (infixl "retract'_of" 50) where "T retract_of S \ (\r. retraction S T r)" lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" unfolding retraction_def by auto text \Preservation of fixpoints under (more general notion of) retraction\ lemma invertible_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes contt: "continuous_on T i" and "i ` T \ S" and contr: "continuous_on S r" and "r ` S \ T" and ri: "\y. y \ T \ r (i y) = y" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - have "\x\S. (i \ g \ r) x = x" proof (rule FP) show "continuous_on S (i \ g \ r)" by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) show "(i \ g \ r) ` S \ S" using assms(2,4,8) by force qed then obtain x where x: "x \ S" "(i \ g \ r) x = x" .. then have *: "g (r x) \ T" using assms(4,8) by auto have "r ((i \ g \ r) x) = r x" using x by auto then show ?thesis using "*" ri that by auto qed lemma homeomorphic_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes "S homeomorphic T" shows "(\f. continuous_on S f \ f ` S \ S \ (\x\S. f x = x)) \ (\g. continuous_on T g \ g ` T \ T \ (\y\T. g y = y))" (is "?lhs = ?rhs") proof - obtain r i where r: "\x\S. i (r x) = x" "r ` S = T" "continuous_on S r" "\y\T. r (i y) = y" "i ` T = S" "continuous_on T i" using assms unfolding homeomorphic_def homeomorphism_def by blast show ?thesis proof assume ?lhs with r show ?rhs by (metis invertible_fixpoint_property[of T i S r] order_refl) next assume ?rhs with r show ?lhs by (metis invertible_fixpoint_property[of S r T i] order_refl) qed qed lemma retract_fixpoint_property: fixes f :: "'a::topological_space \ 'b::topological_space" and S :: "'a set" assumes "T retract_of S" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - obtain h where "retraction S T h" using assms(1) unfolding retract_of_def .. then show ?thesis unfolding retraction_def using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] by (metis assms(4) contg image_ident that) qed lemma retraction: "retraction S T r \ T \ S \ continuous_on S r \ r ` S = T \ (\x \ T. r x = x)" by (force simp: retraction_def) lemma retractionE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "retraction S T r" obtains "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof (rule that) from retraction [of S T r] assms have "T \ S" "continuous_on S r" "r ` S = T" and "\x \ T. r x = x" by simp_all then show "T = r ` S" "r ` S \ S" "continuous_on S r" by simp_all from \\x \ T. r x = x\ have "r x = x" if "x \ T" for x using that by simp with \r ` S = T\ show "r (r x) = r x" if "x \ S" for x using that by auto qed lemma retract_ofE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "T retract_of S" obtains r where "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof - from assms obtain r where "retraction S T r" by (auto simp add: retract_of_def) with that show thesis by (auto elim: retractionE) qed lemma retract_of_imp_extensible: assumes "S retract_of T" and "continuous_on S f" and "f ` S \ U" obtains g where "continuous_on T g" "g ` T \ U" "\x. x \ S \ g x = f x" proof - from \S retract_of T\ obtain r where "retraction T S r" by (auto simp add: retract_of_def) show thesis by (rule that [of "f \ r"]) (use \continuous_on S f\ \f ` S \ U\ \retraction T S r\ in \auto simp: continuous_on_compose2 retraction\) qed lemma idempotent_imp_retraction: assumes "continuous_on S f" and "f ` S \ S" and "\x. x \ S \ f(f x) = f x" shows "retraction S (f ` S) f" by (simp add: assms retraction) lemma retraction_subset: assumes "retraction S T r" and "T \ s'" and "s' \ S" shows "retraction s' T r" unfolding retraction_def by (metis assms continuous_on_subset image_mono retraction) lemma retract_of_subset: assumes "T retract_of S" and "T \ s'" and "s' \ S" shows "T retract_of s'" by (meson assms retract_of_def retraction_subset) lemma retraction_refl [simp]: "retraction S S (\x. x)" by (simp add: retraction) lemma retract_of_refl [iff]: "S retract_of S" unfolding retract_of_def retraction_def using continuous_on_id by blast lemma retract_of_imp_subset: "S retract_of T \ S \ T" by (simp add: retract_of_def retraction_def) lemma retract_of_empty [simp]: "({} retract_of S) \ S = {}" "(S retract_of {}) \ S = {}" by (auto simp: retract_of_def retraction_def) lemma retract_of_singleton [iff]: "({x} retract_of S) \ x \ S" unfolding retract_of_def retraction_def by force lemma retraction_comp: "\retraction S T f; retraction T U g\ \ retraction S U (g \ f)" apply (auto simp: retraction_def intro: continuous_on_compose2) by blast lemma retract_of_trans [trans]: assumes "S retract_of T" and "T retract_of U" shows "S retract_of U" using assms by (auto simp: retract_of_def intro: retraction_comp) lemma closedin_retract: fixes S :: "'a :: t2_space set" assumes "S retract_of T" shows "closedin (top_of_set T) S" proof - obtain r where r: "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" using assms by (auto simp: retract_of_def retraction_def) have "S = {x\T. x = r x}" using r by auto also have "\ = T \ ((\x. (x, r x)) -` ({y. \x. y = (x, x)}))" unfolding vimage_def mem_Times_iff fst_conv snd_conv using r by auto also have "closedin (top_of_set T) \" by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r) finally show ?thesis . qed lemma closedin_self [simp]: "closedin (top_of_set S) S" by simp lemma retract_of_closed: fixes S :: "'a :: t2_space set" shows "\closed T; S retract_of T\ \ closed S" by (metis closedin_retract closedin_closed_eq) lemma retract_of_compact: "\compact T; S retract_of T\ \ compact S" by (metis compact_continuous_image retract_of_def retraction) lemma retract_of_connected: "\connected T; S retract_of T\ \ connected S" by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) lemma retraction_openin_vimage_iff: "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if retraction: "retraction S T r" and "U \ T" using retraction apply (rule retractionE) apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) using \U \ T\ apply (auto elim: continuous_on_subset) done lemma retract_of_Times: "\S retract_of s'; T retract_of t'\ \ (S \ T) retract_of (s' \ t')" apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) apply (rename_tac f g) apply (rule_tac x="\z. ((f \ fst) z, (g \ snd) z)" in exI) apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ done subsection\Retractions on a topological space\ definition retract_of_space :: "'a set \ 'a topology \ bool" (infix "retract'_of'_space" 50) where "S retract_of_space X \ S \ topspace X \ (\r. continuous_map X (subtopology X S) r \ (\x \ S. r x = x))" lemma retract_of_space_retraction_maps: "S retract_of_space X \ S \ topspace X \ (\r. retraction_maps X (subtopology X S) r id)" by (auto simp: retract_of_space_def retraction_maps_def) lemma retract_of_space_section_map: "S retract_of_space X \ S \ topspace X \ section_map (subtopology X S) X id" unfolding retract_of_space_def retraction_maps_def section_map_def by (auto simp: continuous_map_from_subtopology) lemma retract_of_space_imp_subset: "S retract_of_space X \ S \ topspace X" by (simp add: retract_of_space_def) lemma retract_of_space_topspace: "topspace X retract_of_space X" using retract_of_space_def by force lemma retract_of_space_empty [simp]: "{} retract_of_space X \ topspace X = {}" by (auto simp: continuous_map_def retract_of_space_def) lemma retract_of_space_singleton [simp]: "{a} retract_of_space X \ a \ topspace X" proof - have "continuous_map X (subtopology X {a}) (\x. a) \ (\x. a) a = a" if "a \ topspace X" using that by simp then show ?thesis by (force simp: retract_of_space_def) qed lemma retract_of_space_clopen: assumes "openin X S" "closedin X S" "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (cases "S = {}") case False then obtain a where "a \ S" by blast show ?thesis unfolding retract_of_space_def proof (intro exI conjI) show "S \ topspace X" by (simp add: assms closedin_subset) have "continuous_map X X (\x. if x \ S then x else a)" proof (rule continuous_map_cases) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. x)" by (simp add: continuous_map_from_subtopology) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. a)" using \S \ topspace X\ \a \ S\ by force show "x = a" if "x \ X frontier_of {x. x \ S}" for x using assms that clopenin_eq_frontier_of by fastforce qed then show "continuous_map X (subtopology X S) (\x. if x \ S then x else a)" using \S \ topspace X\ \a \ S\ by (auto simp: continuous_map_in_subtopology) qed auto qed (use assms in auto) lemma retract_of_space_disjoint_union: assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \ T = topspace X" and "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (rule retract_of_space_clopen) have "S \ T = {}" by (meson ST disjnt_def) then have "S = topspace X - T" using ST by auto then show "closedin X S" using \openin X T\ by blast qed (auto simp: assms) lemma retraction_maps_section_image1: assumes "retraction_maps X Y r s" shows "s ` (topspace Y) retract_of_space X" unfolding retract_of_space_section_map proof show "s ` topspace Y \ topspace X" using assms continuous_map_image_subset_topspace retraction_maps_def by blast show "section_map (subtopology X (s ` topspace Y)) X id" unfolding section_map_def using assms retraction_maps_to_retract_maps by blast qed lemma retraction_maps_section_image2: "retraction_maps X Y r s \ subtopology X (s ` (topspace Y)) homeomorphic_space Y" using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map section_map_def by blast subsection\Paths and path-connectedness\ definition pathin :: "'a topology \ (real \ 'a) \ bool" where "pathin X g \ continuous_map (subtopology euclideanreal {0..1}) X g" lemma pathin_compose: "\pathin X g; continuous_map X Y f\ \ pathin Y (f \ g)" by (simp add: continuous_map_compose pathin_def) lemma pathin_subtopology: "pathin (subtopology X S) g \ pathin X g \ (\x \ {0..1}. g x \ S)" by (auto simp: pathin_def continuous_map_in_subtopology) lemma pathin_const: "pathin X (\x. a) \ a \ topspace X" by (simp add: pathin_def) lemma path_start_in_topspace: "pathin X g \ g 0 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_finish_in_topspace: "pathin X g \ g 1 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_image_subset_topspace: "pathin X g \ g ` ({0..1}) \ topspace X" by (force simp: pathin_def continuous_map) definition path_connected_space :: "'a topology \ bool" where "path_connected_space X \ \x \ topspace X. \ y \ topspace X. \g. pathin X g \ g 0 = x \ g 1 = y" definition path_connectedin :: "'a topology \ 'a set \ bool" where "path_connectedin X S \ S \ topspace X \ path_connected_space(subtopology X S)" lemma path_connectedin_absolute [simp]: "path_connectedin (subtopology X S) S \ path_connectedin X S" - by (simp add: path_connectedin_def subtopology_subtopology topspace_subtopology) + by (simp add: path_connectedin_def subtopology_subtopology) lemma path_connectedin_subset_topspace: "path_connectedin X S \ S \ topspace X" by (simp add: path_connectedin_def) lemma path_connectedin_subtopology: "path_connectedin (subtopology X S) T \ path_connectedin X T \ T \ S" - by (auto simp: path_connectedin_def subtopology_subtopology topspace_subtopology inf.absorb2) + by (auto simp: path_connectedin_def subtopology_subtopology inf.absorb2) lemma path_connectedin: "path_connectedin X S \ S \ topspace X \ (\x \ S. \y \ S. \g. pathin X g \ g ` {0..1} \ S \ g 0 = x \ g 1 = y)" unfolding path_connectedin_def path_connected_space_def pathin_def continuous_map_in_subtopology - by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2 topspace_subtopology) + by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2) lemma path_connectedin_topspace: "path_connectedin X (topspace X) \ path_connected_space X" by (simp add: path_connectedin_def) lemma path_connected_imp_connected_space: assumes "path_connected_space X" shows "connected_space X" proof - have *: "\S. connectedin X S \ g 0 \ S \ g 1 \ S" if "pathin X g" for g proof (intro exI conjI) have "continuous_map (subtopology euclideanreal {0..1}) X g" using connectedin_absolute that by (simp add: pathin_def) then show "connectedin X (g ` {0..1})" by (rule connectedin_continuous_map_image) auto qed auto show ?thesis using assms by (auto intro: * simp add: path_connected_space_def connected_space_subconnected Ball_def) qed lemma path_connectedin_imp_connectedin: "path_connectedin X S \ connectedin X S" by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def) lemma path_connected_space_topspace_empty: "topspace X = {} \ path_connected_space X" by (simp add: path_connected_space_def) lemma path_connectedin_empty [simp]: "path_connectedin X {}" by (simp add: path_connectedin) lemma path_connectedin_singleton [simp]: "path_connectedin X {a} \ a \ topspace X" proof show "path_connectedin X {a} \ a \ topspace X" by (simp add: path_connectedin) show "a \ topspace X \ path_connectedin X {a}" unfolding path_connectedin using pathin_const by fastforce qed lemma path_connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and S: "path_connectedin X S" shows "path_connectedin Y (f ` S)" proof - have fX: "f ` (topspace X) \ topspace Y" by (metis f continuous_map_image_subset_topspace) show ?thesis unfolding path_connectedin proof (intro conjI ballI; clarify?) fix x assume "x \ S" show "f x \ topspace Y" by (meson S fX \x \ S\ image_subset_iff path_connectedin_subset_topspace set_mp) next fix x y assume "x \ S" and "y \ S" then obtain g where g: "pathin X g" "g ` {0..1} \ S" "g 0 = x" "g 1 = y" using S by (force simp: path_connectedin) show "\g. pathin Y g \ g ` {0..1} \ f ` S \ g 0 = f x \ g 1 = f y" proof (intro exI conjI) show "pathin Y (f \ g)" using \pathin X g\ f pathin_compose by auto qed (use g in auto) qed qed lemma path_connectedin_discrete_topology: "path_connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" apply safe using path_connectedin_subset_topspace apply fastforce apply (meson connectedin_discrete_topology path_connectedin_imp_connectedin) using subset_singletonD by fastforce lemma path_connected_space_discrete_topology: "path_connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis path_connectedin_discrete_topology path_connectedin_topspace path_connected_space_topspace_empty subset_singletonD topspace_discrete_topology) lemma homeomorphic_path_connected_space_imp: "\path_connected_space X; X homeomorphic_space Y\ \ path_connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis (no_types, hide_lams) continuous_map_closedin continuous_map_image_subset_topspace imageI order_class.order.antisym path_connectedin_continuous_map_image path_connectedin_topspace subsetI) lemma homeomorphic_path_connected_space: "X homeomorphic_space Y \ path_connected_space X \ path_connected_space Y" by (meson homeomorphic_path_connected_space_imp homeomorphic_space_sym) lemma homeomorphic_map_path_connectedness: assumes "homeomorphic_map X Y f" "U \ topspace X" shows "path_connectedin Y (f ` U) \ path_connectedin X U" unfolding path_connectedin_def proof (intro conj_cong homeomorphic_path_connected_space) show "(f ` U \ topspace Y) = (U \ topspace X)" using assms homeomorphic_imp_surjective_map by blast next assume "U \ topspace X" show "subtopology Y (f ` U) homeomorphic_space subtopology X U" using assms unfolding homeomorphic_eq_everything_map by (metis (no_types, hide_lams) assms homeomorphic_map_subtopologies homeomorphic_space homeomorphic_space_sym image_mono inf.absorb_iff2) qed lemma homeomorphic_map_path_connectedness_eq: "homeomorphic_map X Y f \ path_connectedin X U \ U \ topspace X \ path_connectedin Y (f ` U)" by (meson homeomorphic_map_path_connectedness path_connectedin_def) subsection\Connected components\ definition connected_component_of :: "'a topology \ 'a \ 'a \ bool" where "connected_component_of X x y \ \T. connectedin X T \ x \ T \ y \ T" abbreviation connected_component_of_set where "connected_component_of_set X x \ Collect (connected_component_of X x)" definition connected_components_of :: "'a topology \ ('a set) set" where "connected_components_of X \ connected_component_of_set X ` topspace X" lemma connected_component_in_topspace: "connected_component_of X x y \ x \ topspace X \ y \ topspace X" by (meson connected_component_of_def connectedin_subset_topspace in_mono) lemma connected_component_of_refl: "connected_component_of X x x \ x \ topspace X" by (meson connected_component_in_topspace connected_component_of_def connectedin_sing insertI1) lemma connected_component_of_sym: "connected_component_of X x y \ connected_component_of X y x" by (meson connected_component_of_def) lemma connected_component_of_trans: "\connected_component_of X x y; connected_component_of X y z\ \ connected_component_of X x z" unfolding connected_component_of_def using connectedin_Un by fastforce lemma connected_component_of_mono: "\connected_component_of (subtopology X S) x y; S \ T\ \ connected_component_of (subtopology X T) x y" by (metis connected_component_of_def connectedin_subtopology inf.absorb_iff2 subtopology_subtopology) lemma connected_component_of_set: "connected_component_of_set X x = {y. \T. connectedin X T \ x \ T \ y \ T}" by (meson connected_component_of_def) lemma connected_component_of_subset_topspace: "connected_component_of_set X x \ topspace X" using connected_component_in_topspace by force lemma connected_component_of_eq_empty: "connected_component_of_set X x = {} \ (x \ topspace X)" using connected_component_in_topspace connected_component_of_refl by fastforce lemma connected_space_iff_connected_component: "connected_space X \ (\x \ topspace X. \y \ topspace X. connected_component_of X x y)" by (simp add: connected_component_of_def connected_space_subconnected) lemma connected_space_imp_connected_component_of: "\connected_space X; a \ topspace X; b \ topspace X\ \ connected_component_of X a b" by (simp add: connected_space_iff_connected_component) lemma connected_space_connected_component_set: "connected_space X \ (\x \ topspace X. connected_component_of_set X x = topspace X)" using connected_component_of_subset_topspace connected_space_iff_connected_component by fastforce lemma connected_component_of_maximal: "\connectedin X S; x \ S\ \ S \ connected_component_of_set X x" by (meson Ball_Collect connected_component_of_def) lemma connected_component_of_equiv: "connected_component_of X x y \ x \ topspace X \ y \ topspace X \ connected_component_of X x = connected_component_of X y" apply (simp add: connected_component_in_topspace fun_eq_iff) by (meson connected_component_of_refl connected_component_of_sym connected_component_of_trans) lemma connected_component_of_disjoint: "disjnt (connected_component_of_set X x) (connected_component_of_set X y) \ ~(connected_component_of X x y)" using connected_component_of_equiv unfolding disjnt_iff by force lemma connected_component_of_eq: "connected_component_of X x = connected_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ connected_component_of X x y" by (metis Collect_empty_eq_bot connected_component_of_eq_empty connected_component_of_equiv) lemma connectedin_connected_component_of: "connectedin X (connected_component_of_set X x)" proof - have "connected_component_of_set X x = \ {T. connectedin X T \ x \ T}" by (auto simp: connected_component_of_def) then show ?thesis apply (rule ssubst) by (blast intro: connectedin_Union) qed lemma Union_connected_components_of: "\(connected_components_of X) = topspace X" unfolding connected_components_of_def apply (rule equalityI) apply (simp add: SUP_least connected_component_of_subset_topspace) using connected_component_of_refl by fastforce lemma connected_components_of_maximal: "\C \ connected_components_of X; connectedin X S; ~disjnt C S\ \ S \ C" unfolding connected_components_of_def disjnt_def apply clarify by (metis Int_emptyI connected_component_of_def connected_component_of_trans mem_Collect_eq) lemma pairwise_disjoint_connected_components_of: "pairwise disjnt (connected_components_of X)" unfolding connected_components_of_def pairwise_def apply clarify by (metis connected_component_of_disjoint connected_component_of_equiv) lemma complement_connected_components_of_Union: "C \ connected_components_of X \ topspace X - C = \ (connected_components_of X - {C})" apply (rule equalityI) using Union_connected_components_of apply fastforce by (metis Diff_cancel Diff_subset Union_connected_components_of cSup_singleton diff_Union_pairwise_disjoint equalityE insert_subsetI pairwise_disjoint_connected_components_of) lemma nonempty_connected_components_of: "C \ connected_components_of X \ C \ {}" unfolding connected_components_of_def by (metis (no_types, lifting) connected_component_of_eq_empty imageE) lemma connected_components_of_subset: "C \ connected_components_of X \ C \ topspace X" using Union_connected_components_of by fastforce lemma connectedin_connected_components_of: assumes "C \ connected_components_of X" shows "connectedin X C" proof - have "C \ connected_component_of_set X ` topspace X" using assms connected_components_of_def by blast then show ?thesis using connectedin_connected_component_of by fastforce qed lemma connected_component_in_connected_components_of: "connected_component_of_set X a \ connected_components_of X \ a \ topspace X" apply (rule iffI) using connected_component_of_eq_empty nonempty_connected_components_of apply fastforce by (simp add: connected_components_of_def) lemma connected_space_iff_components_eq: "connected_space X \ (\C \ connected_components_of X. \C' \ connected_components_of X. C = C')" apply (rule iffI) apply (force simp: connected_components_of_def connected_space_connected_component_set image_iff) by (metis connected_component_in_connected_components_of connected_component_of_refl connected_space_iff_connected_component mem_Collect_eq) lemma connected_components_of_eq_empty: "connected_components_of X = {} \ topspace X = {}" by (simp add: connected_components_of_def) lemma connected_components_of_empty_space: "topspace X = {} \ connected_components_of X = {}" by (simp add: connected_components_of_eq_empty) lemma connected_components_of_subset_sing: "connected_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: connected_components_of_empty_space connected_space_topspace_empty) next case False then show ?thesis by (metis (no_types, hide_lams) Union_connected_components_of ccpo_Sup_singleton connected_components_of_eq_empty connected_space_iff_components_eq insertI1 singletonD subsetI subset_singleton_iff) qed lemma connected_space_iff_components_subset_singleton: "connected_space X \ (\a. connected_components_of X \ {a})" by (simp add: connected_components_of_subset_sing) lemma connected_components_of_eq_singleton: "connected_components_of X = {S} \ connected_space X \ topspace X \ {} \ S = topspace X" by (metis ccpo_Sup_singleton connected_components_of_subset_sing insert_not_empty subset_singleton_iff) lemma connected_components_of_connected_space: "connected_space X \ connected_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: connected_components_of_eq_empty connected_components_of_eq_singleton) lemma exists_connected_component_of_superset: assumes "connectedin X S" and ne: "topspace X \ {}" shows "\C. C \ connected_components_of X \ S \ C" proof (cases "S = {}") case True then show ?thesis using ne connected_components_of_def by blast next case False then show ?thesis by (meson all_not_in_conv assms(1) connected_component_in_connected_components_of connected_component_of_maximal connectedin_subset_topspace in_mono) qed lemma closedin_connected_components_of: assumes "C \ connected_components_of X" shows "closedin X C" proof - obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" using assms by (auto simp: connected_components_of_def) have "connected_component_of_set X x \ topspace X" by (simp add: connected_component_of_subset_topspace) moreover have "X closure_of connected_component_of_set X x \ connected_component_of_set X x" proof (rule connected_component_of_maximal) show "connectedin X (X closure_of connected_component_of_set X x)" by (simp add: connectedin_closure_of connectedin_connected_component_of) show "x \ X closure_of connected_component_of_set X x" by (simp add: \x \ topspace X\ closure_of connected_component_of_refl) qed ultimately show ?thesis using closure_of_subset_eq x by auto qed lemma closedin_connected_component_of: "closedin X (connected_component_of_set X x)" by (metis closedin_connected_components_of closedin_empty connected_component_in_connected_components_of connected_component_of_eq_empty) lemma connected_component_of_eq_overlap: "connected_component_of_set X x = connected_component_of_set X y \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x \ connected_component_of_set X y = {})" using connected_component_of_equiv by fastforce lemma connected_component_of_nonoverlap: "connected_component_of_set X x \ connected_component_of_set X y = {} \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x = connected_component_of_set X y)" by (metis connected_component_of_eq_empty connected_component_of_eq_overlap inf.idem) lemma connected_component_of_overlap: "~(connected_component_of_set X x \ connected_component_of_set X y = {}) \ x \ topspace X \ y \ topspace X \ connected_component_of_set X x = connected_component_of_set X y" by (meson connected_component_of_nonoverlap) end \ No newline at end of file diff --git a/src/HOL/Analysis/Bounded_Continuous_Function.thy b/src/HOL/Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Analysis/Bounded_Continuous_Function.thy +++ b/src/HOL/Analysis/Bounded_Continuous_Function.thy @@ -1,321 +1,321 @@ section \Bounded Continuous Functions\ theory Bounded_Continuous_Function imports Topology_Euclidean_Space Uniform_Limit begin subsection \Definition\ definition\<^marker>\tag important\ "bcontfun = {f. continuous_on UNIV f \ bounded (range f)}" typedef (overloaded) ('a, 'b) bcontfun ("(_ \\<^sub>C /_)" [22, 21] 21) = "bcontfun::('a::topological_space \ 'b::metric_space) set" morphisms apply_bcontfun Bcontfun by (auto intro: continuous_intros simp: bounded_def bcontfun_def) declare [[coercion "apply_bcontfun :: ('a::topological_space \\<^sub>C'b::metric_space) \ 'a \ 'b"]] setup_lifting type_definition_bcontfun lemma continuous_on_apply_bcontfun[intro, simp]: "continuous_on T (apply_bcontfun x)" and bounded_apply_bcontfun[intro, simp]: "bounded (range (apply_bcontfun x))" using apply_bcontfun[of x] by (auto simp: bcontfun_def intro: continuous_on_subset) lemma bcontfun_eqI: "(\x. apply_bcontfun f x = apply_bcontfun g x) \ f = g" by transfer auto lemma bcontfunE: assumes "f \ bcontfun" obtains g where "f = apply_bcontfun g" by (blast intro: apply_bcontfun_cases assms ) lemma const_bcontfun: "(\x. b) \ bcontfun" - by (auto simp: bcontfun_def continuous_on_const image_def) + by (auto simp: bcontfun_def image_def) lift_definition const_bcontfun::"'b::metric_space \ ('a::topological_space \\<^sub>C 'b)" is "\c _. c" by (rule const_bcontfun) (* TODO: Generalize to uniform spaces? *) instantiation bcontfun :: (topological_space, metric_space) metric_space begin lift_definition\<^marker>\tag important\ dist_bcontfun :: "'a \\<^sub>C 'b \ 'a \\<^sub>C 'b \ real" is "\f g. (SUP x. dist (f x) (g x))" . definition uniformity_bcontfun :: "('a \\<^sub>C 'b \ 'a \\<^sub>C 'b) filter" where "uniformity_bcontfun = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_bcontfun :: "('a \\<^sub>C 'b) set \ bool" where "open_bcontfun S = (\x\S. \\<^sub>F (x', y) in uniformity. x' = x \ y \ S)" lemma bounded_dist_le_SUP_dist: "bounded (range f) \ bounded (range g) \ dist (f x) (g x) \ (SUP x. dist (f x) (g x))" by (auto intro!: cSUP_upper bounded_imp_bdd_above bounded_dist_comp) lemma dist_bounded: fixes f g :: "'a \\<^sub>C 'b" shows "dist (f x) (g x) \ dist f g" by transfer (auto intro!: bounded_dist_le_SUP_dist simp: bcontfun_def) lemma dist_bound: fixes f g :: "'a \\<^sub>C 'b" assumes "\x. dist (f x) (g x) \ b" shows "dist f g \ b" using assms by transfer (auto intro!: cSUP_least) lemma dist_fun_lt_imp_dist_val_lt: fixes f g :: "'a \\<^sub>C 'b" assumes "dist f g < e" shows "dist (f x) (g x) < e" using dist_bounded assms by (rule le_less_trans) instance proof fix f g h :: "'a \\<^sub>C 'b" show "dist f g = 0 \ f = g" proof have "\x. dist (f x) (g x) \ dist f g" by (rule dist_bounded) also assume "dist f g = 0" finally show "f = g" by (auto simp: apply_bcontfun_inject[symmetric]) qed (auto simp: dist_bcontfun_def intro!: cSup_eq) show "dist f g \ dist f h + dist g h" proof (rule dist_bound) fix x have "dist (f x) (g x) \ dist (f x) (h x) + dist (g x) (h x)" by (rule dist_triangle2) also have "dist (f x) (h x) \ dist f h" by (rule dist_bounded) also have "dist (g x) (h x) \ dist g h" by (rule dist_bounded) finally show "dist (f x) (g x) \ dist f h + dist g h" by simp qed qed (rule open_bcontfun_def uniformity_bcontfun_def)+ end lift_definition PiC::"'a::topological_space set \ ('a \ 'b set) \ ('a \\<^sub>C 'b::metric_space) set" is "\I X. Pi I X \ bcontfun" by auto lemma mem_PiC_iff: "x \ PiC I X \ apply_bcontfun x \ Pi I X" by transfer simp lemmas mem_PiCD = mem_PiC_iff[THEN iffD1] and mem_PiCI = mem_PiC_iff[THEN iffD2] lemma tendsto_bcontfun_uniform_limit: fixes f::"'i \ 'a::topological_space \\<^sub>C 'b::metric_space" assumes "(f \ l) F" shows "uniform_limit UNIV f l F" proof (rule uniform_limitI) fix e::real assume "e > 0" from tendstoD[OF assms this] have "\\<^sub>F x in F. dist (f x) l < e" . then show "\\<^sub>F n in F. \x\UNIV. dist ((f n) x) (l x) < e" by eventually_elim (auto simp: dist_fun_lt_imp_dist_val_lt) qed lemma uniform_limit_tendsto_bcontfun: fixes f::"'i \ 'a::topological_space \\<^sub>C 'b::metric_space" and l::"'a::topological_space \\<^sub>C 'b::metric_space" assumes "uniform_limit UNIV f l F" shows "(f \ l) F" proof (rule tendstoI) fix e::real assume "e > 0" then have "e / 2 > 0" by simp from uniform_limitD[OF assms this] have "\\<^sub>F i in F. \x. dist (f i x) (l x) < e / 2" by simp then have "\\<^sub>F x in F. dist (f x) l \ e / 2" by eventually_elim (blast intro: dist_bound less_imp_le) then show "\\<^sub>F x in F. dist (f x) l < e" by eventually_elim (use \0 < e\ in auto) qed lemma uniform_limit_bcontfunE: fixes f::"'i \ 'a::topological_space \\<^sub>C 'b::metric_space" and l::"'a::topological_space \ 'b::metric_space" assumes "uniform_limit UNIV f l F" "F \ bot" obtains l'::"'a::topological_space \\<^sub>C 'b::metric_space" where "l = l'" "(f \ l') F" by (metis (mono_tags, lifting) always_eventually apply_bcontfun apply_bcontfun_cases assms bcontfun_def mem_Collect_eq uniform_limit_bounded uniform_limit_tendsto_bcontfun uniform_limit_theorem) lemma closed_PiC: fixes I :: "'a::metric_space set" and X :: "'a \ 'b::complete_space set" assumes "\i. i \ I \ closed (X i)" shows "closed (PiC I X)" unfolding closed_sequential_limits proof safe fix f l assume seq: "\n. f n \ PiC I X" and lim: "f \ l" show "l \ PiC I X" proof (safe intro!: mem_PiCI) fix x assume "x \ I" then have "closed (X x)" using assms by simp moreover have "eventually (\i. f i x \ X x) sequentially" using seq \x \ I\ by (auto intro!: eventuallyI dest!: mem_PiCD simp: Pi_iff) moreover note sequentially_bot moreover have "(\n. (f n) x) \ l x" using tendsto_bcontfun_uniform_limit[OF lim] by (rule tendsto_uniform_limitI) simp ultimately show "l x \ X x" by (rule Lim_in_closed_set) qed qed subsection \Complete Space\ instance\<^marker>\tag important\ bcontfun :: (metric_space, complete_space) complete_space proof fix f :: "nat \ ('a, 'b) bcontfun" assume "Cauchy f" \ \Cauchy equals uniform convergence\ then obtain g where "uniform_limit UNIV f g sequentially" using uniformly_convergent_eq_cauchy[of "\_. True" f] unfolding Cauchy_def uniform_limit_sequentially_iff by (metis dist_fun_lt_imp_dist_val_lt) from uniform_limit_bcontfunE[OF this sequentially_bot] obtain l' where "g = apply_bcontfun l'" "(f \ l')" by metis then show "convergent f" by (intro convergentI) qed subsection\<^marker>\tag unimportant\ \Supremum norm for a normed vector space\ instantiation\<^marker>\tag unimportant\ bcontfun :: (topological_space, real_normed_vector) real_vector begin lemma uminus_cont: "f \ bcontfun \ (\x. - f x) \ bcontfun" for f::"'a \ 'b" by (auto simp: bcontfun_def intro!: continuous_intros) lemma plus_cont: "f \ bcontfun \ g \ bcontfun \ (\x. f x + g x) \ bcontfun" for f g::"'a \ 'b" by (auto simp: bcontfun_def intro!: continuous_intros bounded_plus_comp) lemma minus_cont: "f \ bcontfun \ g \ bcontfun \ (\x. f x - g x) \ bcontfun" for f g::"'a \ 'b" by (auto simp: bcontfun_def intro!: continuous_intros bounded_minus_comp) lemma scaleR_cont: "f \ bcontfun \ (\x. a *\<^sub>R f x) \ bcontfun" for f :: "'a \ 'b" by (auto simp: bcontfun_def intro!: continuous_intros bounded_scaleR_comp) lemma bcontfun_normI: "continuous_on UNIV f \ (\x. norm (f x) \ b) \ f \ bcontfun" by (auto simp: bcontfun_def intro: boundedI) lift_definition uminus_bcontfun::"('a \\<^sub>C 'b) \ 'a \\<^sub>C 'b" is "\f x. - f x" by (rule uminus_cont) lift_definition plus_bcontfun::"('a \\<^sub>C 'b) \ ('a \\<^sub>C 'b) \ 'a \\<^sub>C 'b" is "\f g x. f x + g x" by (rule plus_cont) lift_definition minus_bcontfun::"('a \\<^sub>C 'b) \ ('a \\<^sub>C 'b) \ 'a \\<^sub>C 'b" is "\f g x. f x - g x" by (rule minus_cont) lift_definition zero_bcontfun::"'a \\<^sub>C 'b" is "\_. 0" by (rule const_bcontfun) lemma const_bcontfun_0_eq_0[simp]: "const_bcontfun 0 = 0" by transfer simp lift_definition scaleR_bcontfun::"real \ ('a \\<^sub>C 'b) \ 'a \\<^sub>C 'b" is "\r g x. r *\<^sub>R g x" by (rule scaleR_cont) lemmas [simp] = const_bcontfun.rep_eq uminus_bcontfun.rep_eq plus_bcontfun.rep_eq minus_bcontfun.rep_eq zero_bcontfun.rep_eq scaleR_bcontfun.rep_eq instance by standard (auto intro!: bcontfun_eqI simp: algebra_simps) end lemma bounded_norm_le_SUP_norm: "bounded (range f) \ norm (f x) \ (SUP x. norm (f x))" by (auto intro!: cSUP_upper bounded_imp_bdd_above simp: bounded_norm_comp) instantiation\<^marker>\tag unimportant\ bcontfun :: (topological_space, real_normed_vector) real_normed_vector begin definition norm_bcontfun :: "('a, 'b) bcontfun \ real" where "norm_bcontfun f = dist f 0" definition "sgn (f::('a,'b) bcontfun) = f /\<^sub>R norm f" instance proof fix a :: real fix f g :: "('a, 'b) bcontfun" show "dist f g = norm (f - g)" unfolding norm_bcontfun_def by transfer (simp add: dist_norm) show "norm (f + g) \ norm f + norm g" unfolding norm_bcontfun_def by transfer (auto intro!: cSUP_least norm_triangle_le add_mono bounded_norm_le_SUP_norm simp: dist_norm bcontfun_def) show "norm (a *\<^sub>R f) = \a\ * norm f" unfolding norm_bcontfun_def apply transfer by (rule trans[OF _ continuous_at_Sup_mono[symmetric]]) (auto intro!: monoI mult_left_mono continuous_intros bounded_imp_bdd_above simp: bounded_norm_comp bcontfun_def image_comp) qed (auto simp: norm_bcontfun_def sgn_bcontfun_def) end lemma norm_bounded: fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun" shows "norm (apply_bcontfun f x) \ norm f" using dist_bounded[of f x 0] by (simp add: dist_norm) lemma norm_bound: fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun" assumes "\x. norm (apply_bcontfun f x) \ b" shows "norm f \ b" using dist_bound[of f 0 b] assms by (simp add: dist_norm) subsection\<^marker>\tag unimportant\ \(bounded) continuous extenstion\ lemma continuous_on_cbox_bcontfunE: fixes f::"'a::euclidean_space \ 'b::metric_space" assumes "continuous_on (cbox a b) f" obtains g::"'a \\<^sub>C 'b" where "\x. x \ cbox a b \ g x = f x" "\x. g x = f (clamp a b x)" proof - define g where "g \ ext_cont f a b" have "g \ bcontfun" using assms by (auto intro!: continuous_on_ext_cont simp: g_def bcontfun_def) (auto simp: g_def ext_cont_def intro!: clamp_bounded compact_imp_bounded[OF compact_continuous_image] assms) then obtain h where h: "g = apply_bcontfun h" by (rule bcontfunE) then have "h x = f x" if "x \ cbox a b" for x by (auto simp: h[symmetric] g_def that) moreover have "h x = f (clamp a b x)" for x by (auto simp: h[symmetric] g_def ext_cont_def) ultimately show ?thesis .. qed lifting_update bcontfun.lifting lifting_forget bcontfun.lifting end diff --git a/src/HOL/Analysis/Bounded_Linear_Function.thy b/src/HOL/Analysis/Bounded_Linear_Function.thy --- a/src/HOL/Analysis/Bounded_Linear_Function.thy +++ b/src/HOL/Analysis/Bounded_Linear_Function.thy @@ -1,843 +1,843 @@ (* Title: HOL/Analysis/Bounded_Linear_Function.thy Author: Fabian Immler, TU München *) section \Bounded Linear Function\ theory Bounded_Linear_Function imports Topology_Euclidean_Space Operator_Norm Uniform_Limit Function_Topology begin lemma onorm_componentwise: assumes "bounded_linear f" shows "onorm f \ (\i\Basis. norm (f i))" proof - { fix i::'a assume "i \ Basis" hence "onorm (\x. (x \ i) *\<^sub>R f i) \ onorm (\x. (x \ i)) * norm (f i)" by (auto intro!: onorm_scaleR_left_lemma bounded_linear_inner_left) also have "\ \ norm i * norm (f i)" by (rule mult_right_mono) (auto simp: ac_simps Cauchy_Schwarz_ineq2 intro!: onorm_le) finally have "onorm (\x. (x \ i) *\<^sub>R f i) \ norm (f i)" using \i \ Basis\ by simp } hence "onorm (\x. \i\Basis. (x \ i) *\<^sub>R f i) \ (\i\Basis. norm (f i))" by (auto intro!: order_trans[OF onorm_sum_le] bounded_linear_scaleR_const sum_mono bounded_linear_inner_left) also have "(\x. \i\Basis. (x \ i) *\<^sub>R f i) = (\x. f (\i\Basis. (x \ i) *\<^sub>R i))" by (simp add: linear_sum bounded_linear.linear assms linear_simps) also have "\ = f" by (simp add: euclidean_representation) finally show ?thesis . qed lemmas onorm_componentwise_le = order_trans[OF onorm_componentwise] subsection\<^marker>\tag unimportant\ \Intro rules for \<^term>\bounded_linear\\ named_theorems bounded_linear_intros lemma onorm_inner_left: assumes "bounded_linear r" shows "onorm (\x. r x \ f) \ onorm r * norm f" proof (rule onorm_bound) fix x have "norm (r x \ f) \ norm (r x) * norm f" by (simp add: Cauchy_Schwarz_ineq2) also have "\ \ onorm r * norm x * norm f" by (intro mult_right_mono onorm assms norm_ge_zero) finally show "norm (r x \ f) \ onorm r * norm f * norm x" by (simp add: ac_simps) qed (intro mult_nonneg_nonneg norm_ge_zero onorm_pos_le assms) lemma onorm_inner_right: assumes "bounded_linear r" shows "onorm (\x. f \ r x) \ norm f * onorm r" apply (subst inner_commute) apply (rule onorm_inner_left[OF assms, THEN order_trans]) apply simp done lemmas [bounded_linear_intros] = bounded_linear_zero bounded_linear_add bounded_linear_const_mult bounded_linear_mult_const bounded_linear_scaleR_const bounded_linear_const_scaleR bounded_linear_ident bounded_linear_sum bounded_linear_Pair bounded_linear_sub bounded_linear_fst_comp bounded_linear_snd_comp bounded_linear_inner_left_comp bounded_linear_inner_right_comp subsection\<^marker>\tag unimportant\ \declaration of derivative/continuous/tendsto introduction rules for bounded linear functions\ attribute_setup bounded_linear = \Scan.succeed (Thm.declaration_attribute (fn thm => fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r)) [ (@{thm bounded_linear.has_derivative}, \<^named_theorems>\derivative_intros\), (@{thm bounded_linear.tendsto}, \<^named_theorems>\tendsto_intros\), (@{thm bounded_linear.continuous}, \<^named_theorems>\continuous_intros\), (@{thm bounded_linear.continuous_on}, \<^named_theorems>\continuous_intros\), (@{thm bounded_linear.uniformly_continuous_on}, \<^named_theorems>\continuous_intros\), (@{thm bounded_linear_compose}, \<^named_theorems>\bounded_linear_intros\) ]))\ attribute_setup bounded_bilinear = \Scan.succeed (Thm.declaration_attribute (fn thm => fold (fn (r, s) => Named_Theorems.add_thm s (thm RS r)) [ (@{thm bounded_bilinear.FDERIV}, \<^named_theorems>\derivative_intros\), (@{thm bounded_bilinear.tendsto}, \<^named_theorems>\tendsto_intros\), (@{thm bounded_bilinear.continuous}, \<^named_theorems>\continuous_intros\), (@{thm bounded_bilinear.continuous_on}, \<^named_theorems>\continuous_intros\), (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_left]}, \<^named_theorems>\bounded_linear_intros\), (@{thm bounded_linear_compose[OF bounded_bilinear.bounded_linear_right]}, \<^named_theorems>\bounded_linear_intros\), (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_left]}, \<^named_theorems>\continuous_intros\), (@{thm bounded_linear.uniformly_continuous_on[OF bounded_bilinear.bounded_linear_right]}, \<^named_theorems>\continuous_intros\) ]))\ subsection \Type of bounded linear functions\ typedef\<^marker>\tag important\ (overloaded) ('a, 'b) blinfun ("(_ \\<^sub>L /_)" [22, 21] 21) = "{f::'a::real_normed_vector\'b::real_normed_vector. bounded_linear f}" morphisms blinfun_apply Blinfun by (blast intro: bounded_linear_intros) declare [[coercion "blinfun_apply :: ('a::real_normed_vector \\<^sub>L'b::real_normed_vector) \ 'a \ 'b"]] lemma bounded_linear_blinfun_apply[bounded_linear_intros]: "bounded_linear g \ bounded_linear (\x. blinfun_apply f (g x))" by (metis blinfun_apply mem_Collect_eq bounded_linear_compose) setup_lifting type_definition_blinfun lemma blinfun_eqI: "(\i. blinfun_apply x i = blinfun_apply y i) \ x = y" by transfer auto lemma bounded_linear_Blinfun_apply: "bounded_linear f \ blinfun_apply (Blinfun f) = f" by (auto simp: Blinfun_inverse) subsection \Type class instantiations\ instantiation blinfun :: (real_normed_vector, real_normed_vector) real_normed_vector begin lift_definition\<^marker>\tag important\ norm_blinfun :: "'a \\<^sub>L 'b \ real" is onorm . lift_definition minus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\f g x. f x - g x" by (rule bounded_linear_sub) definition dist_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b \ real" where "dist_blinfun a b = norm (a - b)" definition [code del]: "(uniformity :: (('a \\<^sub>L 'b) \ ('a \\<^sub>L 'b)) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_blinfun :: "('a \\<^sub>L 'b) set \ bool" where [code del]: "open_blinfun S = (\x\S. \\<^sub>F (x', y) in uniformity. x' = x \ y \ S)" lift_definition uminus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\f x. - f x" by (rule bounded_linear_minus) lift_definition\<^marker>\tag important\ zero_blinfun :: "'a \\<^sub>L 'b" is "\x. 0" by (rule bounded_linear_zero) lift_definition\<^marker>\tag important\ plus_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\f g x. f x + g x" by (metis bounded_linear_add) lift_definition\<^marker>\tag important\ scaleR_blinfun::"real \ 'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" is "\r f x. r *\<^sub>R f x" by (metis bounded_linear_compose bounded_linear_scaleR_right) definition sgn_blinfun :: "'a \\<^sub>L 'b \ 'a \\<^sub>L 'b" where "sgn_blinfun x = scaleR (inverse (norm x)) x" instance apply standard unfolding dist_blinfun_def open_blinfun_def sgn_blinfun_def uniformity_blinfun_def apply (rule refl | (transfer, force simp: onorm_triangle onorm_scaleR onorm_eq_0 algebra_simps))+ done end declare uniformity_Abort[where 'a="('a :: real_normed_vector) \\<^sub>L ('b :: real_normed_vector)", code] lemma norm_blinfun_eqI: assumes "n \ norm (blinfun_apply f x) / norm x" assumes "\x. norm (blinfun_apply f x) \ n * norm x" assumes "0 \ n" shows "norm f = n" by (auto simp: norm_blinfun_def intro!: antisym onorm_bound assms order_trans[OF _ le_onorm] bounded_linear_intros) lemma norm_blinfun: "norm (blinfun_apply f x) \ norm f * norm x" by transfer (rule onorm) lemma norm_blinfun_bound: "0 \ b \ (\x. norm (blinfun_apply f x) \ b * norm x) \ norm f \ b" by transfer (rule onorm_bound) lemma bounded_bilinear_blinfun_apply[bounded_bilinear]: "bounded_bilinear blinfun_apply" proof fix f g::"'a \\<^sub>L 'b" and a b::'a and r::real show "(f + g) a = f a + g a" "(r *\<^sub>R f) a = r *\<^sub>R f a" by (transfer, simp)+ interpret bounded_linear f for f::"'a \\<^sub>L 'b" by (auto intro!: bounded_linear_intros) show "f (a + b) = f a + f b" "f (r *\<^sub>R a) = r *\<^sub>R f a" by (simp_all add: add scaleR) show "\K. \a b. norm (blinfun_apply a b) \ norm a * norm b * K" by (auto intro!: exI[where x=1] norm_blinfun) qed interpretation blinfun: bounded_bilinear blinfun_apply by (rule bounded_bilinear_blinfun_apply) lemmas bounded_linear_apply_blinfun[intro, simp] = blinfun.bounded_linear_left declare blinfun.zero_left [simp] blinfun.zero_right [simp] context bounded_bilinear begin named_theorems bilinear_simps lemmas [bilinear_simps] = add_left add_right diff_left diff_right minus_left minus_right scaleR_left scaleR_right zero_left zero_right sum_left sum_right end instance blinfun :: (real_normed_vector, banach) banach proof fix X::"nat \ 'a \\<^sub>L 'b" assume "Cauchy X" { fix x::'a { fix x::'a assume "norm x \ 1" have "Cauchy (\n. X n x)" proof (rule CauchyI) fix e::real assume "0 < e" from CauchyD[OF \Cauchy X\ \0 < e\] obtain M where M: "\m n. m \ M \ n \ M \ norm (X m - X n) < e" by auto show "\M. \m\M. \n\M. norm (X m x - X n x) < e" proof (safe intro!: exI[where x=M]) fix m n assume le: "M \ m" "M \ n" have "norm (X m x - X n x) = norm ((X m - X n) x)" by (simp add: blinfun.bilinear_simps) also have "\ \ norm (X m - X n) * norm x" by (rule norm_blinfun) also have "\ \ norm (X m - X n) * 1" using \norm x \ 1\ norm_ge_zero by (rule mult_left_mono) also have "\ = norm (X m - X n)" by simp also have "\ < e" using le by fact finally show "norm (X m x - X n x) < e" . qed qed hence "convergent (\n. X n x)" by (metis Cauchy_convergent_iff) } note convergent_norm1 = this define y where "y = x /\<^sub>R norm x" have y: "norm y \ 1" and xy: "x = norm x *\<^sub>R y" by (simp_all add: y_def inverse_eq_divide) have "convergent (\n. norm x *\<^sub>R X n y)" by (intro bounded_bilinear.convergent[OF bounded_bilinear_scaleR] convergent_const convergent_norm1 y) also have "(\n. norm x *\<^sub>R X n y) = (\n. X n x)" by (subst xy) (simp add: blinfun.bilinear_simps) finally have "convergent (\n. X n x)" . } then obtain v where v: "\x. (\n. X n x) \ v x" unfolding convergent_def by metis have "Cauchy (\n. norm (X n))" proof (rule CauchyI) fix e::real assume "e > 0" from CauchyD[OF \Cauchy X\ \0 < e\] obtain M where M: "\m n. m \ M \ n \ M \ norm (X m - X n) < e" by auto show "\M. \m\M. \n\M. norm (norm (X m) - norm (X n)) < e" proof (safe intro!: exI[where x=M]) fix m n assume mn: "m \ M" "n \ M" have "norm (norm (X m) - norm (X n)) \ norm (X m - X n)" by (metis norm_triangle_ineq3 real_norm_def) also have "\ < e" using mn by fact finally show "norm (norm (X m) - norm (X n)) < e" . qed qed then obtain K where K: "(\n. norm (X n)) \ K" unfolding Cauchy_convergent_iff convergent_def by metis have "bounded_linear v" proof fix x y and r::real from tendsto_add[OF v[of x] v [of y]] v[of "x + y", unfolded blinfun.bilinear_simps] tendsto_scaleR[OF tendsto_const[of r] v[of x]] v[of "r *\<^sub>R x", unfolded blinfun.bilinear_simps] show "v (x + y) = v x + v y" "v (r *\<^sub>R x) = r *\<^sub>R v x" by (metis (poly_guards_query) LIMSEQ_unique)+ show "\K. \x. norm (v x) \ norm x * K" proof (safe intro!: exI[where x=K]) fix x have "norm (v x) \ K * norm x" by (rule tendsto_le[OF _ tendsto_mult[OF K tendsto_const] tendsto_norm[OF v]]) (auto simp: norm_blinfun) thus "norm (v x) \ norm x * K" by (simp add: ac_simps) qed qed hence Bv: "\x. (\n. X n x) \ Blinfun v x" by (auto simp: bounded_linear_Blinfun_apply v) have "X \ Blinfun v" proof (rule LIMSEQ_I) fix r::real assume "r > 0" define r' where "r' = r / 2" have "0 < r'" "r' < r" using \r > 0\ by (simp_all add: r'_def) from CauchyD[OF \Cauchy X\ \r' > 0\] obtain M where M: "\m n. m \ M \ n \ M \ norm (X m - X n) < r'" by metis show "\no. \n\no. norm (X n - Blinfun v) < r" proof (safe intro!: exI[where x=M]) fix n assume n: "M \ n" have "norm (X n - Blinfun v) \ r'" proof (rule norm_blinfun_bound) fix x have "eventually (\m. m \ M) sequentially" by (metis eventually_ge_at_top) hence ev_le: "eventually (\m. norm (X n x - X m x) \ r' * norm x) sequentially" proof eventually_elim case (elim m) have "norm (X n x - X m x) = norm ((X n - X m) x)" by (simp add: blinfun.bilinear_simps) also have "\ \ norm ((X n - X m)) * norm x" by (rule norm_blinfun) also have "\ \ r' * norm x" using M[OF n elim] by (simp add: mult_right_mono) finally show ?case . qed have tendsto_v: "(\m. norm (X n x - X m x)) \ norm (X n x - Blinfun v x)" by (auto intro!: tendsto_intros Bv) show "norm ((X n - Blinfun v) x) \ r' * norm x" by (auto intro!: tendsto_upperbound tendsto_v ev_le simp: blinfun.bilinear_simps) qed (simp add: \0 < r'\ less_imp_le) thus "norm (X n - Blinfun v) < r" by (metis \r' < r\ le_less_trans) qed qed thus "convergent X" by (rule convergentI) qed subsection\<^marker>\tag unimportant\ \On Euclidean Space\ lemma Zfun_sum: assumes "finite s" assumes f: "\i. i \ s \ Zfun (f i) F" shows "Zfun (\x. sum (\i. f i x) s) F" using assms by induct (auto intro!: Zfun_zero Zfun_add) lemma norm_blinfun_euclidean_le: fixes a::"'a::euclidean_space \\<^sub>L 'b::real_normed_vector" shows "norm a \ sum (\x. norm (a x)) Basis" apply (rule norm_blinfun_bound) apply (simp add: sum_nonneg) apply (subst euclidean_representation[symmetric, where 'a='a]) apply (simp only: blinfun.bilinear_simps sum_distrib_right) apply (rule order.trans[OF norm_sum sum_mono]) apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm) done lemma tendsto_componentwise1: fixes a::"'a::euclidean_space \\<^sub>L 'b::real_normed_vector" and b::"'c \ 'a \\<^sub>L 'b" assumes "(\j. j \ Basis \ ((\n. b n j) \ a j) F)" shows "(b \ a) F" proof - have "\j. j \ Basis \ Zfun (\x. norm (b x j - a j)) F" using assms unfolding tendsto_Zfun_iff Zfun_norm_iff . hence "Zfun (\x. \j\Basis. norm (b x j - a j)) F" by (auto intro!: Zfun_sum) thus ?thesis unfolding tendsto_Zfun_iff by (rule Zfun_le) (auto intro!: order_trans[OF norm_blinfun_euclidean_le] simp: blinfun.bilinear_simps) qed lift_definition blinfun_of_matrix::"('b::euclidean_space \ 'a::euclidean_space \ real) \ 'a \\<^sub>L 'b" is "\a x. \i\Basis. \j\Basis. ((x \ j) * a i j) *\<^sub>R i" by (intro bounded_linear_intros) lemma blinfun_of_matrix_works: fixes f::"'a::euclidean_space \\<^sub>L 'b::euclidean_space" shows "blinfun_of_matrix (\i j. (f j) \ i) = f" proof (transfer, rule, rule euclidean_eqI) fix f::"'a \ 'b" and x::'a and b::'b assume "bounded_linear f" and b: "b \ Basis" then interpret bounded_linear f by simp have "(\j\Basis. \i\Basis. (x \ i * (f i \ j)) *\<^sub>R j) \ b = (\j\Basis. if j = b then (\i\Basis. (x \ i * (f i \ j))) else 0)" using b by (simp add: inner_sum_left inner_Basis if_distrib cong: if_cong) (simp add: sum.swap) also have "\ = (\i\Basis. (x \ i * (f i \ b)))" using b by (simp add: sum.delta) also have "\ = f x \ b" by (metis (mono_tags, lifting) Linear_Algebra.linear_componentwise linear_axioms) finally show "(\j\Basis. \i\Basis. (x \ i * (f i \ j)) *\<^sub>R j) \ b = f x \ b" . qed lemma blinfun_of_matrix_apply: "blinfun_of_matrix a x = (\i\Basis. \j\Basis. ((x \ j) * a i j) *\<^sub>R i)" by transfer simp lemma blinfun_of_matrix_minus: "blinfun_of_matrix x - blinfun_of_matrix y = blinfun_of_matrix (x - y)" by transfer (auto simp: algebra_simps sum_subtractf) lemma norm_blinfun_of_matrix: "norm (blinfun_of_matrix a) \ (\i\Basis. \j\Basis. \a i j\)" apply (rule norm_blinfun_bound) apply (simp add: sum_nonneg) apply (simp only: blinfun_of_matrix_apply sum_distrib_right) apply (rule order_trans[OF norm_sum sum_mono]) apply (rule order_trans[OF norm_sum sum_mono]) apply (simp add: abs_mult mult_right_mono ac_simps Basis_le_norm) done lemma tendsto_blinfun_of_matrix: assumes "\i j. i \ Basis \ j \ Basis \ ((\n. b n i j) \ a i j) F" shows "((\n. blinfun_of_matrix (b n)) \ blinfun_of_matrix a) F" proof - have "\i j. i \ Basis \ j \ Basis \ Zfun (\x. norm (b x i j - a i j)) F" using assms unfolding tendsto_Zfun_iff Zfun_norm_iff . hence "Zfun (\x. (\i\Basis. \j\Basis. \b x i j - a i j\)) F" by (auto intro!: Zfun_sum) thus ?thesis unfolding tendsto_Zfun_iff blinfun_of_matrix_minus by (rule Zfun_le) (auto intro!: order_trans[OF norm_blinfun_of_matrix]) qed lemma tendsto_componentwise: fixes a::"'a::euclidean_space \\<^sub>L 'b::euclidean_space" and b::"'c \ 'a \\<^sub>L 'b" shows "(\i j. i \ Basis \ j \ Basis \ ((\n. b n j \ i) \ a j \ i) F) \ (b \ a) F" apply (subst blinfun_of_matrix_works[of a, symmetric]) apply (subst blinfun_of_matrix_works[of "b x" for x, symmetric, abs_def]) by (rule tendsto_blinfun_of_matrix) lemma continuous_blinfun_componentwiseI: fixes f:: "'b::t2_space \ 'a::euclidean_space \\<^sub>L 'c::euclidean_space" assumes "\i j. i \ Basis \ j \ Basis \ continuous F (\x. (f x) j \ i)" shows "continuous F f" using assms by (auto simp: continuous_def intro!: tendsto_componentwise) lemma continuous_blinfun_componentwiseI1: fixes f:: "'b::t2_space \ 'a::euclidean_space \\<^sub>L 'c::real_normed_vector" assumes "\i. i \ Basis \ continuous F (\x. f x i)" shows "continuous F f" using assms by (auto simp: continuous_def intro!: tendsto_componentwise1) lemma continuous_on_blinfun_componentwise: fixes f:: "'d::t2_space \ 'e::euclidean_space \\<^sub>L 'f::real_normed_vector" assumes "\i. i \ Basis \ continuous_on s (\x. f x i)" shows "continuous_on s f" using assms by (auto intro!: continuous_at_imp_continuous_on intro!: tendsto_componentwise1 simp: continuous_on_eq_continuous_within continuous_def) lemma bounded_linear_blinfun_matrix: "bounded_linear (\x. (x::_\\<^sub>L _) j \ i)" by (auto intro!: bounded_linearI' bounded_linear_intros) lemma continuous_blinfun_matrix: fixes f:: "'b::t2_space \ 'a::real_normed_vector \\<^sub>L 'c::real_inner" assumes "continuous F f" shows "continuous F (\x. (f x) j \ i)" by (rule bounded_linear.continuous[OF bounded_linear_blinfun_matrix assms]) lemma continuous_on_blinfun_matrix: fixes f::"'a::t2_space \ 'b::real_normed_vector \\<^sub>L 'c::real_inner" assumes "continuous_on S f" shows "continuous_on S (\x. (f x) j \ i)" using assms by (auto simp: continuous_on_eq_continuous_within continuous_blinfun_matrix) lemma continuous_on_blinfun_of_matrix[continuous_intros]: assumes "\i j. i \ Basis \ j \ Basis \ continuous_on S (\s. g s i j)" shows "continuous_on S (\s. blinfun_of_matrix (g s))" using assms by (auto simp: continuous_on intro!: tendsto_blinfun_of_matrix) lemma mult_if_delta: "(if P then (1::'a::comm_semiring_1) else 0) * q = (if P then q else 0)" by auto lemma compact_blinfun_lemma: fixes f :: "nat \ 'a::euclidean_space \\<^sub>L 'b::euclidean_space" assumes "bounded (range f)" shows "\d\Basis. \l::'a \\<^sub>L 'b. \ r::nat\nat. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) i) (l i) < e) sequentially)" by (rule compact_lemma_general[where unproj = "\e. blinfun_of_matrix (\i j. e j \ i)"]) (auto intro!: euclidean_eqI[where 'a='b] bounded_linear_image assms simp: blinfun_of_matrix_works blinfun_of_matrix_apply inner_Basis mult_if_delta sum.delta' scaleR_sum_left[symmetric]) lemma blinfun_euclidean_eqI: "(\i. i \ Basis \ blinfun_apply x i = blinfun_apply y i) \ x = y" apply (auto intro!: blinfun_eqI) apply (subst (2) euclidean_representation[symmetric, where 'a='a]) apply (subst (1) euclidean_representation[symmetric, where 'a='a]) apply (simp add: blinfun.bilinear_simps) done lemma Blinfun_eq_matrix: "bounded_linear f \ Blinfun f = blinfun_of_matrix (\i j. f j \ i)" by (intro blinfun_euclidean_eqI) (auto simp: blinfun_of_matrix_apply bounded_linear_Blinfun_apply inner_Basis if_distrib if_distribR sum.delta' euclidean_representation cong: if_cong) text \TODO: generalize (via \compact_cball\)?\ instance blinfun :: (euclidean_space, euclidean_space) heine_borel proof fix f :: "nat \ 'a \\<^sub>L 'b" assume f: "bounded (range f)" then obtain l::"'a \\<^sub>L 'b" and r where r: "strict_mono r" and l: "\e>0. eventually (\n. \i\Basis. dist (f (r n) i) (l i) < e) sequentially" using compact_blinfun_lemma [OF f] by blast { fix e::real let ?d = "real_of_nat DIM('a) * real_of_nat DIM('b)" assume "e > 0" - hence "e / ?d > 0" by (simp add: DIM_positive) + hence "e / ?d > 0" by (simp) with l have "eventually (\n. \i\Basis. dist (f (r n) i) (l i) < e / ?d) sequentially" by simp moreover { fix n assume n: "\i\Basis. dist (f (r n) i) (l i) < e / ?d" have "norm (f (r n) - l) = norm (blinfun_of_matrix (\i j. (f (r n) - l) j \ i))" unfolding blinfun_of_matrix_works .. also note norm_blinfun_of_matrix also have "(\i\Basis. \j\Basis. \(f (r n) - l) j \ i\) < (\i\(Basis::'b set). e / real_of_nat DIM('b))" proof (rule sum_strict_mono) fix i::'b assume i: "i \ Basis" have "(\j::'a\Basis. \(f (r n) - l) j \ i\) < (\j::'a\Basis. e / ?d)" proof (rule sum_strict_mono) fix j::'a assume j: "j \ Basis" have "\(f (r n) - l) j \ i\ \ norm ((f (r n) - l) j)" by (simp add: Basis_le_norm i) also have "\ < e / ?d" using n i j by (auto simp: dist_norm blinfun.bilinear_simps) finally show "\(f (r n) - l) j \ i\ < e / ?d" by simp qed simp_all also have "\ \ e / real_of_nat DIM('b)" by simp finally show "(\j\Basis. \(f (r n) - l) j \ i\) < e / real_of_nat DIM('b)" by simp qed simp_all also have "\ \ e" by simp finally have "dist (f (r n)) l < e" by (auto simp: dist_norm) } ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" using eventually_elim2 by force } then have *: "((f \ r) \ l) sequentially" unfolding o_def tendsto_iff by simp with r show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" by auto qed subsection\<^marker>\tag unimportant\ \concrete bounded linear functions\ lemma transfer_bounded_bilinear_bounded_linearI: assumes "g = (\i x. (blinfun_apply (f i) x))" shows "bounded_bilinear g = bounded_linear f" proof assume "bounded_bilinear g" then interpret bounded_bilinear f by (simp add: assms) show "bounded_linear f" proof (unfold_locales, safe intro!: blinfun_eqI) fix i show "f (x + y) i = (f x + f y) i" "f (r *\<^sub>R x) i = (r *\<^sub>R f x) i" for r x y by (auto intro!: blinfun_eqI simp: blinfun.bilinear_simps) from _ nonneg_bounded show "\K. \x. norm (f x) \ norm x * K" by (rule ex_reg) (auto intro!: onorm_bound simp: norm_blinfun.rep_eq ac_simps) qed qed (auto simp: assms intro!: blinfun.comp) lemma transfer_bounded_bilinear_bounded_linear[transfer_rule]: "(rel_fun (rel_fun (=) (pcr_blinfun (=) (=))) (=)) bounded_bilinear bounded_linear" by (auto simp: pcr_blinfun_def cr_blinfun_def rel_fun_def OO_def intro!: transfer_bounded_bilinear_bounded_linearI) context bounded_bilinear begin lift_definition prod_left::"'b \ 'a \\<^sub>L 'c" is "(\b a. prod a b)" by (rule bounded_linear_left) declare prod_left.rep_eq[simp] lemma bounded_linear_prod_left[bounded_linear]: "bounded_linear prod_left" by transfer (rule flip) lift_definition prod_right::"'a \ 'b \\<^sub>L 'c" is "(\a b. prod a b)" by (rule bounded_linear_right) declare prod_right.rep_eq[simp] lemma bounded_linear_prod_right[bounded_linear]: "bounded_linear prod_right" by transfer (rule bounded_bilinear_axioms) end lift_definition id_blinfun::"'a::real_normed_vector \\<^sub>L 'a" is "\x. x" by (rule bounded_linear_ident) lemmas blinfun_apply_id_blinfun[simp] = id_blinfun.rep_eq lemma norm_blinfun_id[simp]: "norm (id_blinfun::'a::{real_normed_vector, perfect_space} \\<^sub>L 'a) = 1" by transfer (auto simp: onorm_id) lemma norm_blinfun_id_le: "norm (id_blinfun::'a::real_normed_vector \\<^sub>L 'a) \ 1" by transfer (auto simp: onorm_id_le) lift_definition fst_blinfun::"('a::real_normed_vector \ 'b::real_normed_vector) \\<^sub>L 'a" is fst by (rule bounded_linear_fst) lemma blinfun_apply_fst_blinfun[simp]: "blinfun_apply fst_blinfun = fst" by transfer (rule refl) lift_definition snd_blinfun::"('a::real_normed_vector \ 'b::real_normed_vector) \\<^sub>L 'b" is snd by (rule bounded_linear_snd) lemma blinfun_apply_snd_blinfun[simp]: "blinfun_apply snd_blinfun = snd" by transfer (rule refl) lift_definition blinfun_compose:: "'a::real_normed_vector \\<^sub>L 'b::real_normed_vector \ 'c::real_normed_vector \\<^sub>L 'a \ 'c \\<^sub>L 'b" (infixl "o\<^sub>L" 55) is "(o)" parametric comp_transfer unfolding o_def by (rule bounded_linear_compose) lemma blinfun_apply_blinfun_compose[simp]: "(a o\<^sub>L b) c = a (b c)" by (simp add: blinfun_compose.rep_eq) lemma norm_blinfun_compose: "norm (f o\<^sub>L g) \ norm f * norm g" by transfer (rule onorm_compose) lemma bounded_bilinear_blinfun_compose[bounded_bilinear]: "bounded_bilinear (o\<^sub>L)" by unfold_locales (auto intro!: blinfun_eqI exI[where x=1] simp: blinfun.bilinear_simps norm_blinfun_compose) lemma blinfun_compose_zero[simp]: "blinfun_compose 0 = (\_. 0)" "blinfun_compose x 0 = 0" by (auto simp: blinfun.bilinear_simps intro!: blinfun_eqI) lemma blinfun_bij2: fixes f::"'a \\<^sub>L 'a::euclidean_space" assumes "f o\<^sub>L g = id_blinfun" shows "bij (blinfun_apply g)" proof (rule bijI) show "inj g" using assms by (metis blinfun_apply_id_blinfun blinfun_compose.rep_eq injI inj_on_imageI2) then show "surj g" using blinfun.bounded_linear_right bounded_linear_def linear_inj_imp_surj by blast qed lemma blinfun_bij1: fixes f::"'a \\<^sub>L 'a::euclidean_space" assumes "f o\<^sub>L g = id_blinfun" shows "bij (blinfun_apply f)" proof (rule bijI) show "surj (blinfun_apply f)" by (metis assms blinfun_apply_blinfun_compose blinfun_apply_id_blinfun surjI) then show "inj (blinfun_apply f)" using blinfun.bounded_linear_right bounded_linear_def linear_surj_imp_inj by blast qed lift_definition blinfun_inner_right::"'a::real_inner \ 'a \\<^sub>L real" is "(\)" by (rule bounded_linear_inner_right) declare blinfun_inner_right.rep_eq[simp] lemma bounded_linear_blinfun_inner_right[bounded_linear]: "bounded_linear blinfun_inner_right" by transfer (rule bounded_bilinear_inner) lift_definition blinfun_inner_left::"'a::real_inner \ 'a \\<^sub>L real" is "\x y. y \ x" by (rule bounded_linear_inner_left) declare blinfun_inner_left.rep_eq[simp] lemma bounded_linear_blinfun_inner_left[bounded_linear]: "bounded_linear blinfun_inner_left" by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_inner]) lift_definition blinfun_scaleR_right::"real \ 'a \\<^sub>L 'a::real_normed_vector" is "(*\<^sub>R)" by (rule bounded_linear_scaleR_right) declare blinfun_scaleR_right.rep_eq[simp] lemma bounded_linear_blinfun_scaleR_right[bounded_linear]: "bounded_linear blinfun_scaleR_right" by transfer (rule bounded_bilinear_scaleR) lift_definition blinfun_scaleR_left::"'a::real_normed_vector \ real \\<^sub>L 'a" is "\x y. y *\<^sub>R x" by (rule bounded_linear_scaleR_left) lemmas [simp] = blinfun_scaleR_left.rep_eq lemma bounded_linear_blinfun_scaleR_left[bounded_linear]: "bounded_linear blinfun_scaleR_left" by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_scaleR]) lift_definition blinfun_mult_right::"'a \ 'a \\<^sub>L 'a::real_normed_algebra" is "(*)" by (rule bounded_linear_mult_right) declare blinfun_mult_right.rep_eq[simp] lemma bounded_linear_blinfun_mult_right[bounded_linear]: "bounded_linear blinfun_mult_right" by transfer (rule bounded_bilinear_mult) lift_definition blinfun_mult_left::"'a::real_normed_algebra \ 'a \\<^sub>L 'a" is "\x y. y * x" by (rule bounded_linear_mult_left) lemmas [simp] = blinfun_mult_left.rep_eq lemma bounded_linear_blinfun_mult_left[bounded_linear]: "bounded_linear blinfun_mult_left" by transfer (rule bounded_bilinear.flip[OF bounded_bilinear_mult]) lemmas bounded_linear_function_uniform_limit_intros[uniform_limit_intros] = bounded_linear.uniform_limit[OF bounded_linear_apply_blinfun] bounded_linear.uniform_limit[OF bounded_linear_blinfun_apply] bounded_linear.uniform_limit[OF bounded_linear_blinfun_matrix] subsection \The strong operator topology on continuous linear operators\ text \Let \'a\ and \'b\ be two normed real vector spaces. Then the space of linear continuous operators from \'a\ to \'b\ has a canonical norm, and therefore a canonical corresponding topology (the type classes instantiation are given in \<^file>\Bounded_Linear_Function.thy\). However, there is another topology on this space, the strong operator topology, where \T\<^sub>n\ tends to \T\ iff, for all \x\ in \'a\, then \T\<^sub>n x\ tends to \T x\. This is precisely the product topology where the target space is endowed with the norm topology. It is especially useful when \'b\ is the set of real numbers, since then this topology is compact. We can not implement it using type classes as there is already a topology, but at least we can define it as a topology. Note that there is yet another (common and useful) topology on operator spaces, the weak operator topology, defined analogously using the product topology, but where the target space is given the weak-* topology, i.e., the pullback of the weak topology on the bidual of the space under the canonical embedding of a space into its bidual. We do not define it there, although it could also be defined analogously. \ definition\<^marker>\tag important\ strong_operator_topology::"('a::real_normed_vector \\<^sub>L'b::real_normed_vector) topology" where "strong_operator_topology = pullback_topology UNIV blinfun_apply euclidean" lemma strong_operator_topology_topspace: "topspace strong_operator_topology = UNIV" unfolding strong_operator_topology_def topspace_pullback_topology topspace_euclidean by auto lemma strong_operator_topology_basis: fixes f::"('a::real_normed_vector \\<^sub>L'b::real_normed_vector)" and U::"'i \ 'b set" and x::"'i \ 'a" assumes "finite I" "\i. i \ I \ open (U i)" shows "openin strong_operator_topology {f. \i\I. blinfun_apply f (x i) \ U i}" proof - have "open {g::('a\'b). \i\I. g (x i) \ U i}" by (rule product_topology_basis'[OF assms]) moreover have "{f. \i\I. blinfun_apply f (x i) \ U i} = blinfun_apply-`{g::('a\'b). \i\I. g (x i) \ U i} \ UNIV" by auto ultimately show ?thesis unfolding strong_operator_topology_def by (subst openin_pullback_topology) auto qed lemma strong_operator_topology_continuous_evaluation: "continuous_map strong_operator_topology euclidean (\f. blinfun_apply f x)" proof - have "continuous_map strong_operator_topology euclidean ((\f. f x) o blinfun_apply)" unfolding strong_operator_topology_def apply (rule continuous_map_pullback) using continuous_on_product_coordinates by fastforce then show ?thesis unfolding comp_def by simp qed lemma continuous_on_strong_operator_topo_iff_coordinatewise: "continuous_map T strong_operator_topology f \ (\x. continuous_map T euclidean (\y. blinfun_apply (f y) x))" proof (auto) fix x::"'b" assume "continuous_map T strong_operator_topology f" with continuous_map_compose[OF this strong_operator_topology_continuous_evaluation] have "continuous_map T euclidean ((\z. blinfun_apply z x) o f)" by simp then show "continuous_map T euclidean (\y. blinfun_apply (f y) x)" unfolding comp_def by auto next assume *: "\x. continuous_map T euclidean (\y. blinfun_apply (f y) x)" have "\i. continuous_map T euclidean (\x. blinfun_apply (f x) i)" using * unfolding comp_def by auto then have "continuous_map T euclidean (blinfun_apply o f)" unfolding o_def by (metis (no_types) continuous_map_componentwise_UNIV euclidean_product_topology) show "continuous_map T strong_operator_topology f" unfolding strong_operator_topology_def apply (rule continuous_map_pullback') by (auto simp add: \continuous_map T euclidean (blinfun_apply o f)\) qed lemma strong_operator_topology_weaker_than_euclidean: "continuous_map euclidean strong_operator_topology (\f. f)" by (subst continuous_on_strong_operator_topo_iff_coordinatewise, auto simp add: linear_continuous_on) end diff --git a/src/HOL/Analysis/Brouwer_Fixpoint.thy b/src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy @@ -1,2607 +1,2605 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP *) (* At the moment this is just Brouwer's fixpoint theorem. The proof is from *) (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518 *) (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf". *) (* *) (* The script below is quite messy, but at least we avoid formalizing any *) (* topological machinery; we don't even use barycentric subdivision; this is *) (* the big advantage of Kuhn's proof over the usual Sperner's lemma one. *) (* *) (* (c) Copyright, John Harrison 1998-2008 *) section \Brouwer's Fixed Point Theorem\ theory Brouwer_Fixpoint imports Homeomorphism Derivative begin subsection \Retractions\ lemma retract_of_contractible: assumes "contractible T" "S retract_of T" shows "contractible S" using assms apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with) apply (rule_tac x="r a" in exI) apply (rule_tac x="r \ h" in exI) apply (intro conjI continuous_intros continuous_on_compose) apply (erule continuous_on_subset | force)+ done lemma retract_of_path_connected: "\path_connected T; S retract_of T\ \ path_connected S" by (metis path_connected_continuous_image retract_of_def retraction) lemma retract_of_simply_connected: "\simply_connected T; S retract_of T\ \ simply_connected S" apply (simp add: retract_of_def retraction_def, clarify) apply (rule simply_connected_retraction_gen) apply (force elim!: continuous_on_subset)+ done lemma retract_of_homotopically_trivial: assumes ts: "T retract_of S" and hom: "\f g. \continuous_on U f; f ` U \ S; continuous_on U g; g ` U \ S\ \ homotopic_with_canon (\x. True) U S f g" and "continuous_on U f" "f ` U \ T" and "continuous_on U g" "g ` U \ T" shows "homotopic_with_canon (\x. True) U T f g" proof - obtain r where "r ` S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" using ts by (auto simp: retract_of_def retraction) then obtain k where "Retracts S r T k" unfolding Retracts_def by (metis continuous_on_subset dual_order.trans image_iff image_mono) then show ?thesis apply (rule Retracts.homotopically_trivial_retraction_gen) using assms apply (force simp: hom)+ done qed lemma retract_of_homotopically_trivial_null: assumes ts: "T retract_of S" and hom: "\f. \continuous_on U f; f ` U \ S\ \ \c. homotopic_with_canon (\x. True) U S f (\x. c)" and "continuous_on U f" "f ` U \ T" obtains c where "homotopic_with_canon (\x. True) U T f (\x. c)" proof - obtain r where "r ` S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" using ts by (auto simp: retract_of_def retraction) then obtain k where "Retracts S r T k" unfolding Retracts_def by (metis continuous_on_subset dual_order.trans image_iff image_mono) then show ?thesis apply (rule Retracts.homotopically_trivial_retraction_null_gen) apply (rule TrueI refl assms that | assumption)+ done qed lemma retraction_openin_vimage_iff: "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if retraction: "retraction S T r" and "U \ T" using retraction apply (rule retractionE) apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) using \U \ T\ apply (auto elim: continuous_on_subset) done lemma retract_of_locally_compact: fixes S :: "'a :: {heine_borel,real_normed_vector} set" shows "\ locally compact S; T retract_of S\ \ locally compact T" by (metis locally_compact_closedin closedin_retract) lemma homotopic_into_retract: "\f ` S \ T; g ` S \ T; T retract_of U; homotopic_with_canon (\x. True) S U f g\ \ homotopic_with_canon (\x. True) S T f g" apply (subst (asm) homotopic_with_def) apply (simp add: homotopic_with retract_of_def retraction_def, clarify) apply (rule_tac x="r \ h" in exI) apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ done lemma retract_of_locally_connected: assumes "locally connected T" "S retract_of T" shows "locally connected S" using assms by (auto simp: idempotent_imp_retraction intro!: retraction_openin_vimage_iff elim!: locally_connected_quotient_image retract_ofE) lemma retract_of_locally_path_connected: assumes "locally path_connected T" "S retract_of T" shows "locally path_connected S" using assms by (auto simp: idempotent_imp_retraction intro!: retraction_openin_vimage_iff elim!: locally_path_connected_quotient_image retract_ofE) text \A few simple lemmas about deformation retracts\ lemma deformation_retract_imp_homotopy_eqv: fixes S :: "'a::euclidean_space set" assumes "homotopic_with_canon (\x. True) S S id r" and r: "retraction S T r" shows "S homotopy_eqv T" proof - have "homotopic_with_canon (\x. True) S S (id \ r) id" by (simp add: assms(1) homotopic_with_symD) moreover have "homotopic_with_canon (\x. True) T T (r \ id) id" using r unfolding retraction_def by (metis eq_id_iff homotopic_with_id2 topspace_euclidean_subtopology) ultimately show ?thesis unfolding homotopy_equivalent_space_def by (metis (no_types, lifting) continuous_map_subtopology_eu continuous_on_id' id_def image_id r retraction_def) qed lemma deformation_retract: fixes S :: "'a::euclidean_space set" shows "(\r. homotopic_with_canon (\x. True) S S id r \ retraction S T r) \ T retract_of S \ (\f. homotopic_with_canon (\x. True) S S id f \ f ` S \ T)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: retract_of_def retraction_def) next assume ?rhs then show ?lhs apply (clarsimp simp add: retract_of_def retraction_def) apply (rule_tac x=r in exI, simp) apply (rule homotopic_with_trans, assumption) apply (rule_tac f = "r \ f" and g="r \ id" in homotopic_with_eq) apply (rule_tac Y=S in homotopic_compose_continuous_left) apply (auto simp: homotopic_with_sym) done qed lemma deformation_retract_of_contractible_sing: fixes S :: "'a::euclidean_space set" assumes "contractible S" "a \ S" obtains r where "homotopic_with_canon (\x. True) S S id r" "retraction S {a} r" proof - have "{a} retract_of S" by (simp add: \a \ S\) moreover have "homotopic_with_canon (\x. True) S S id (\x. a)" using assms by (auto simp: contractible_def homotopic_into_contractible image_subset_iff) moreover have "(\x. a) ` S \ {a}" by (simp add: image_subsetI) ultimately show ?thesis using that deformation_retract by metis qed lemma continuous_on_compact_surface_projection_aux: fixes S :: "'a::t2_space set" assumes "compact S" "S \ T" "image q T \ S" and contp: "continuous_on T p" and "\x. x \ S \ q x = x" and [simp]: "\x. x \ T \ q(p x) = q x" and "\x. x \ T \ p(q x) = p x" shows "continuous_on T q" proof - have *: "image p T = image p S" using assms by auto (metis imageI subset_iff) have contp': "continuous_on S p" by (rule continuous_on_subset [OF contp \S \ T\]) have "continuous_on (p ` T) q" by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD) then have "continuous_on T (q \ p)" by (rule continuous_on_compose [OF contp]) then show ?thesis by (rule continuous_on_eq [of _ "q \ p"]) (simp add: o_def) qed lemma continuous_on_compact_surface_projection: fixes S :: "'a::real_normed_vector set" assumes "compact S" and S: "S \ V - {0}" and "cone V" and iff: "\x k. x \ V - {0} \ 0 < k \ (k *\<^sub>R x) \ S \ d x = k" shows "continuous_on (V - {0}) (\x. d x *\<^sub>R x)" proof (rule continuous_on_compact_surface_projection_aux [OF \compact S\ S]) show "(\x. d x *\<^sub>R x) ` (V - {0}) \ S" using iff by auto show "continuous_on (V - {0}) (\x. inverse(norm x) *\<^sub>R x)" by (intro continuous_intros) force show "\x. x \ S \ d x *\<^sub>R x = x" by (metis S zero_less_one local.iff scaleR_one subset_eq) show "d (x /\<^sub>R norm x) *\<^sub>R (x /\<^sub>R norm x) = d x *\<^sub>R x" if "x \ V - {0}" for x using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \cone V\ by (simp add: field_simps cone_def zero_less_mult_iff) show "d x *\<^sub>R x /\<^sub>R norm (d x *\<^sub>R x) = x /\<^sub>R norm x" if "x \ V - {0}" for x proof - have "0 < d x" using local.iff that by blast then show ?thesis by simp qed qed subsection \Kuhn Simplices\ lemma bij_betw_singleton_eq: assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \ A" assumes eq: "(\x. x \ A \ x \ a \ f x = g x)" shows "f a = g a" proof - have "f ` (A - {a}) = g ` (A - {a})" by (intro image_cong) (simp_all add: eq) then have "B - {f a} = B - {g a}" using f g a by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff) moreover have "f a \ B" "g a \ B" using f g a by (auto simp: bij_betw_def) ultimately show ?thesis by auto qed lemma swap_image: "Fun.swap i j f ` A = (if i \ A then (if j \ A then f ` A else f ` ((A - {i}) \ {j})) else (if j \ A then f ` ((A - {j}) \ {i}) else f ` A))" by (auto simp: swap_def cong: image_cong_simp) lemmas swap_apply1 = swap_apply(1) lemmas swap_apply2 = swap_apply(2) lemma pointwise_minimal_pointwise_maximal: fixes s :: "(nat \ nat) set" assumes "finite s" and "s \ {}" and "\x\s. \y\s. x \ y \ y \ x" shows "\a\s. \x\s. a \ x" and "\a\s. \x\s. x \ a" using assms proof (induct s rule: finite_ne_induct) case (insert b s) assume *: "\x\insert b s. \y\insert b s. x \ y \ y \ x" then obtain u l where "l \ s" "\b\s. l \ b" "u \ s" "\b\s. b \ u" using insert by auto with * show "\a\insert b s. \x\insert b s. a \ x" "\a\insert b s. \x\insert b s. x \ a" using *[rule_format, of b u] *[rule_format, of b l] by (metis insert_iff order.trans)+ qed auto lemma kuhn_labelling_lemma: fixes P Q :: "'a::euclidean_space \ bool" assumes "\x. P x \ P (f x)" and "\x. P x \ (\i\Basis. Q i \ 0 \ x\i \ x\i \ 1)" shows "\l. (\x.\i\Basis. l x i \ (1::nat)) \ (\x.\i\Basis. P x \ Q i \ (x\i = 0) \ (l x i = 0)) \ (\x.\i\Basis. P x \ Q i \ (x\i = 1) \ (l x i = 1)) \ (\x.\i\Basis. P x \ Q i \ (l x i = 0) \ x\i \ f x\i) \ (\x.\i\Basis. P x \ Q i \ (l x i = 1) \ f x\i \ x\i)" proof - { fix x i let ?R = "\y. (P x \ Q i \ x \ i = 0 \ y = (0::nat)) \ (P x \ Q i \ x \ i = 1 \ y = 1) \ (P x \ Q i \ y = 0 \ x \ i \ f x \ i) \ (P x \ Q i \ y = 1 \ f x \ i \ x \ i)" { assume "P x" "Q i" "i \ Basis" with assms have "0 \ f x \ i \ f x \ i \ 1" by auto } then have "i \ Basis \ ?R 0 \ ?R 1" by auto } then show ?thesis unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *) by (subst choice_iff[symmetric])+ blast qed subsubsection \The key "counting" observation, somewhat abstracted\ lemma kuhn_counting_lemma: fixes bnd compo compo' face S F defines "nF s == card {f\F. face f s \ compo' f}" assumes [simp, intro]: "finite F" \ \faces\ and [simp, intro]: "finite S" \ \simplices\ and "\f. f \ F \ bnd f \ card {s\S. face f s} = 1" and "\f. f \ F \ \ bnd f \ card {s\S. face f s} = 2" and "\s. s \ S \ compo s \ nF s = 1" and "\s. s \ S \ \ compo s \ nF s = 0 \ nF s = 2" and "odd (card {f\F. compo' f \ bnd f})" shows "odd (card {s\S. compo s})" proof - have "(\s | s \ S \ \ compo s. nF s) + (\s | s \ S \ compo s. nF s) = (\s\S. nF s)" by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong) also have "\ = (\s\S. card {f \ {f\F. compo' f \ bnd f}. face f s}) + (\s\S. card {f \ {f\F. compo' f \ \ bnd f}. face f s})" unfolding sum.distrib[symmetric] by (subst card_Un_disjoint[symmetric]) (auto simp: nF_def intro!: sum.cong arg_cong[where f=card]) also have "\ = 1 * card {f\F. compo' f \ bnd f} + 2 * card {f\F. compo' f \ \ bnd f}" using assms(4,5) by (fastforce intro!: arg_cong2[where f="(+)"] sum_multicount) finally have "odd ((\s | s \ S \ \ compo s. nF s) + card {s\S. compo s})" using assms(6,8) by simp moreover have "(\s | s \ S \ \ compo s. nF s) = (\s | s \ S \ \ compo s \ nF s = 0. nF s) + (\s | s \ S \ \ compo s \ nF s = 2. nF s)" using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+ ultimately show ?thesis by auto qed subsubsection \The odd/even result for faces of complete vertices, generalized\ lemma kuhn_complete_lemma: assumes [simp]: "finite simplices" and face: "\f s. face f s \ (\a\s. f = s - {a})" and card_s[simp]: "\s. s \ simplices \ card s = n + 2" and rl_bd: "\s. s \ simplices \ rl ` s \ {..Suc n}" and bnd: "\f s. s \ simplices \ face f s \ bnd f \ card {s\simplices. face f s} = 1" and nbnd: "\f s. s \ simplices \ face f s \ \ bnd f \ card {s\simplices. face f s} = 2" and odd_card: "odd (card {f. (\s\simplices. face f s) \ rl ` f = {..n} \ bnd f})" shows "odd (card {s\simplices. (rl ` s = {..Suc n})})" proof (rule kuhn_counting_lemma) have finite_s[simp]: "\s. s \ simplices \ finite s" by (metis add_is_0 zero_neq_numeral card_infinite assms(3)) let ?F = "{f. \s\simplices. face f s}" have F_eq: "?F = (\s\simplices. \a\s. {s - {a}})" by (auto simp: face) show "finite ?F" using \finite simplices\ unfolding F_eq by auto show "card {s \ simplices. face f s} = 1" if "f \ ?F" "bnd f" for f using bnd that by auto show "card {s \ simplices. face f s} = 2" if "f \ ?F" "\ bnd f" for f using nbnd that by auto show "odd (card {f \ {f. \s\simplices. face f s}. rl ` f = {..n} \ bnd f})" using odd_card by simp fix s assume s[simp]: "s \ simplices" let ?S = "{f \ {f. \s\simplices. face f s}. face f s \ rl ` f = {..n}}" have "?S = (\a. s - {a}) ` {a\s. rl ` (s - {a}) = {..n}}" using s by (fastforce simp: face) then have card_S: "card ?S = card {a\s. rl ` (s - {a}) = {..n}}" by (auto intro!: card_image inj_onI) { assume rl: "rl ` s = {..Suc n}" then have inj_rl: "inj_on rl s" by (intro eq_card_imp_inj_on) auto moreover obtain a where "rl a = Suc n" "a \ s" by (metis atMost_iff image_iff le_Suc_eq rl) ultimately have n: "{..n} = rl ` (s - {a})" by (auto simp: inj_on_image_set_diff rl) have "{a\s. rl ` (s - {a}) = {..n}} = {a}" using inj_rl \a \ s\ by (auto simp: n inj_on_image_eq_iff[OF inj_rl]) then show "card ?S = 1" unfolding card_S by simp } { assume rl: "rl ` s \ {..Suc n}" show "card ?S = 0 \ card ?S = 2" proof cases assume *: "{..n} \ rl ` s" with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}" by (auto simp: atMost_Suc subset_insert_iff split: if_split_asm) then have "\ inj_on rl s" by (intro pigeonhole) simp then obtain a b where ab: "a \ s" "b \ s" "rl a = rl b" "a \ b" by (auto simp: inj_on_def) then have eq: "rl ` (s - {a}) = rl ` s" by auto with ab have inj: "inj_on rl (s - {a})" by (intro eq_card_imp_inj_on) (auto simp: rl_s card_Diff_singleton_if) { fix x assume "x \ s" "x \ {a, b}" then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})" by (auto simp: eq inj_on_image_set_diff[OF inj]) also have "\ = rl ` (s - {x})" using ab \x \ {a, b}\ by auto also assume "\ = rl ` s" finally have False using \x\s\ by auto } moreover { fix x assume "x \ {a, b}" with ab have "x \ s \ rl ` (s - {x}) = rl ` s" by (simp add: set_eq_iff image_iff Bex_def) metis } ultimately have "{a\s. rl ` (s - {a}) = {..n}} = {a, b}" unfolding rl_s[symmetric] by fastforce with \a \ b\ show "card ?S = 0 \ card ?S = 2" unfolding card_S by simp next assume "\ {..n} \ rl ` s" then have "\x. rl ` (s - {x}) \ {..n}" by auto then show "card ?S = 0 \ card ?S = 2" unfolding card_S by simp qed } qed fact locale kuhn_simplex = fixes p n and base upd and s :: "(nat \ nat) set" assumes base: "base \ {..< n} \ {..< p}" assumes base_out: "\i. n \ i \ base i = p" assumes upd: "bij_betw upd {..< n} {..< n}" assumes s_pre: "s = (\i j. if j \ upd`{..< i} then Suc (base j) else base j) ` {.. n}" begin definition "enum i j = (if j \ upd`{..< i} then Suc (base j) else base j)" lemma s_eq: "s = enum ` {.. n}" unfolding s_pre enum_def[abs_def] .. lemma upd_space: "i < n \ upd i < n" using upd by (auto dest!: bij_betwE) lemma s_space: "s \ {..< n} \ {.. p}" proof - { fix i assume "i \ n" then have "enum i \ {..< n} \ {.. p}" proof (induct i) case 0 then show ?case using base by (auto simp: Pi_iff less_imp_le enum_def) next case (Suc i) with base show ?case by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space) qed } then show ?thesis by (auto simp: s_eq) qed lemma inj_upd: "inj_on upd {..< n}" using upd by (simp add: bij_betw_def) lemma inj_enum: "inj_on enum {.. n}" proof - { fix x y :: nat assume "x \ y" "x \ n" "y \ n" with upd have "upd ` {..< x} \ upd ` {..< y}" by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def) then have "enum x \ enum y" by (auto simp: enum_def fun_eq_iff) } then show ?thesis by (auto simp: inj_on_def) qed lemma enum_0: "enum 0 = base" by (simp add: enum_def[abs_def]) lemma base_in_s: "base \ s" unfolding s_eq by (subst enum_0[symmetric]) auto lemma enum_in: "i \ n \ enum i \ s" unfolding s_eq by auto lemma one_step: assumes a: "a \ s" "j < n" assumes *: "\a'. a' \ s \ a' \ a \ a' j = p'" shows "a j \ p'" proof assume "a j = p'" with * a have "\a'. a' \ s \ a' j = p'" by auto then have "\i. i \ n \ enum i j = p'" unfolding s_eq by auto from this[of 0] this[of n] have "j \ upd ` {..< n}" by (auto simp: enum_def fun_eq_iff split: if_split_asm) with upd \j < n\ show False by (auto simp: bij_betw_def) qed lemma upd_inj: "i < n \ j < n \ upd i = upd j \ i = j" using upd by (auto simp: bij_betw_def inj_on_eq_iff) lemma upd_surj: "upd ` {..< n} = {..< n}" using upd by (auto simp: bij_betw_def) lemma in_upd_image: "A \ {..< n} \ i < n \ upd i \ upd ` A \ i \ A" using inj_on_image_mem_iff[of upd "{..< n}"] upd by (auto simp: bij_betw_def) lemma enum_inj: "i \ n \ j \ n \ enum i = enum j \ i = j" using inj_enum by (auto simp: inj_on_eq_iff) lemma in_enum_image: "A \ {.. n} \ i \ n \ enum i \ enum ` A \ i \ A" using inj_on_image_mem_iff[OF inj_enum] by auto lemma enum_mono: "i \ n \ j \ n \ enum i \ enum j \ i \ j" by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric]) lemma enum_strict_mono: "i \ n \ j \ n \ enum i < enum j \ i < j" using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less) lemma chain: "a \ s \ b \ s \ a \ b \ b \ a" by (auto simp: s_eq enum_mono) lemma less: "a \ s \ b \ s \ a i < b i \ a < b" using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) lemma enum_0_bot: "a \ s \ a = enum 0 \ (\a'\s. a \ a')" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_n_top: "a \ s \ a = enum n \ (\a'\s. a' \ a)" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_Suc: "i < n \ enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))" by (auto simp: fun_eq_iff enum_def upd_inj) lemma enum_eq_p: "i \ n \ n \ j \ enum i j = p" by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric]) lemma out_eq_p: "a \ s \ n \ j \ a j = p" unfolding s_eq by (auto simp: enum_eq_p) lemma s_le_p: "a \ s \ a j \ p" using out_eq_p[of a j] s_space by (cases "j < n") auto lemma le_Suc_base: "a \ s \ a j \ Suc (base j)" unfolding s_eq by (auto simp: enum_def) lemma base_le: "a \ s \ base j \ a j" unfolding s_eq by (auto simp: enum_def) lemma enum_le_p: "i \ n \ j < n \ enum i j \ p" using enum_in[of i] s_space by auto lemma enum_less: "a \ s \ i < n \ enum i < a \ enum (Suc i) \ a" unfolding s_eq by (auto simp: enum_strict_mono enum_mono) lemma ksimplex_0: "n = 0 \ s = {(\x. p)}" using s_eq enum_def base_out by auto lemma replace_0: assumes "j < n" "a \ s" and p: "\x\s - {a}. x j = 0" and "x \ s" shows "x \ a" proof cases assume "x \ a" have "a j \ 0" using assms by (intro one_step[where a=a]) auto with less[OF \x\s\ \a\s\, of j] p[rule_format, of x] \x \ s\ \x \ a\ show ?thesis by auto qed simp lemma replace_1: assumes "j < n" "a \ s" and p: "\x\s - {a}. x j = p" and "x \ s" shows "a \ x" proof cases assume "x \ a" have "a j \ p" using assms by (intro one_step[where a=a]) auto with enum_le_p[of _ j] \j < n\ \a\s\ have "a j < p" by (auto simp: less_le s_eq) with less[OF \a\s\ \x\s\, of j] p[rule_format, of x] \x \ s\ \x \ a\ show ?thesis by auto qed simp end locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t for p n b_s u_s s b_t u_t t begin lemma enum_eq: assumes l: "i \ l" "l \ j" and "j + d \ n" assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}" shows "s.enum l = t.enum (l + d)" using l proof (induct l rule: dec_induct) case base then have s: "s.enum i \ t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \ s.enum ` {i .. j}" using eq by auto from t \i \ j\ \j + d \ n\ have "s.enum i \ t.enum (i + d)" by (auto simp: s.enum_mono) moreover from s \i \ j\ \j + d \ n\ have "t.enum (i + d) \ s.enum i" by (auto simp: t.enum_mono) ultimately show ?case by auto next case (step l) moreover from step.prems \j + d \ n\ have "s.enum l < s.enum (Suc l)" "t.enum (l + d) < t.enum (Suc l + d)" by (simp_all add: s.enum_strict_mono t.enum_strict_mono) moreover have "s.enum (Suc l) \ t.enum ` {i + d .. j + d}" "t.enum (Suc l + d) \ s.enum ` {i .. j}" using step \j + d \ n\ eq by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" using \j + d \ n\ by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) (auto intro!: s.enum_in t.enum_in) then show ?case by simp qed lemma ksimplex_eq_bot: assumes a: "a \ s" "\a'. a' \ s \ a \ a'" assumes b: "b \ t" "\b'. b' \ t \ b \ b'" assumes eq: "s - {a} = t - {b}" shows "s = t" proof cases assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp next assume "n \ 0" have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)" "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)" using \n \ 0\ by (simp_all add: s.enum_Suc t.enum_Suc) moreover have e0: "a = s.enum 0" "b = t.enum 0" using a b by (simp_all add: s.enum_0_bot t.enum_0_bot) moreover { fix j assume "0 < j" "j \ n" moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}" unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum j = t.enum j" using enum_eq[of "1" j n 0] eq by auto } note enum_eq = this then have "s.enum (Suc 0) = t.enum (Suc 0)" using \n \ 0\ by auto moreover { fix j assume "Suc j < n" with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"] have "u_s (Suc j) = u_t (Suc j)" using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"] by (auto simp: fun_eq_iff split: if_split_asm) } then have "\j. 0 < j \ j < n \ u_s j = u_t j" by (auto simp: gr0_conv_Suc) with \n \ 0\ have "u_t 0 = u_s 0" by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto ultimately have "a = b" by simp with assms show "s = t" by auto qed lemma ksimplex_eq_top: assumes a: "a \ s" "\a'. a' \ s \ a' \ a" assumes b: "b \ t" "\b'. b' \ t \ b' \ b" assumes eq: "s - {a} = t - {b}" shows "s = t" proof (cases n) assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp next case (Suc n') have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))" "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))" using Suc by (simp_all add: s.enum_Suc t.enum_Suc) moreover have en: "a = s.enum n" "b = t.enum n" using a b by (simp_all add: s.enum_n_top t.enum_n_top) moreover { fix j assume "j < n" moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}" unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc) ultimately have "s.enum j = t.enum j" using enum_eq[of "0" j n' 0] eq Suc by auto } note enum_eq = this then have "s.enum n' = t.enum n'" using Suc by auto moreover { fix j assume "j < n'" with enum_eq[of j] enum_eq[of "Suc j"] have "u_s j = u_t j" using s.enum_Suc[of j] t.enum_Suc[of j] by (auto simp: Suc fun_eq_iff split: if_split_asm) } then have "\j. j < n' \ u_s j = u_t j" by (auto simp: gr0_conv_Suc) then have "u_t n' = u_s n'" by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc) ultimately have "a = b" by simp with assms show "s = t" by auto qed end inductive ksimplex for p n :: nat where ksimplex: "kuhn_simplex p n base upd s \ ksimplex p n s" lemma finite_ksimplexes: "finite {s. ksimplex p n s}" proof (rule finite_subset) { fix a s assume "ksimplex p n s" "a \ s" then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p n b u s . from s_space \a \ s\ out_eq_p[OF \a \ s\] have "a \ (\f x. if n \ x then p else f x) ` ({..< n} \\<^sub>E {.. p})" by (auto simp: image_iff subset_eq Pi_iff split: if_split_asm intro!: bexI[of _ "restrict a {..< n}"]) } then show "{s. ksimplex p n s} \ Pow ((\f x. if n \ x then p else f x) ` ({..< n} \\<^sub>E {.. p}))" by auto qed (simp add: finite_PiE) lemma ksimplex_card: assumes "ksimplex p n s" shows "card s = Suc n" using assms proof cases case (ksimplex u b) then interpret kuhn_simplex p n u b s . show ?thesis by (simp add: card_image s_eq inj_enum) qed lemma simplex_top_face: assumes "0 < p" "\x\s'. x n = p" shows "ksimplex p n s' \ (\s a. ksimplex p (Suc n) s \ a \ s \ s' = s - {a})" using assms proof safe fix s a assume "ksimplex p (Suc n) s" and a: "a \ s" and na: "\x\s - {a}. x n = p" then show "ksimplex p n (s - {a})" proof cases case (ksimplex base upd) then interpret kuhn_simplex p "Suc n" base upd "s" . have "a n < p" using one_step[of a n p] na \a\s\ s_space by (auto simp: less_le) then have "a = enum 0" using \a \ s\ na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n]) then have s_eq: "s - {a} = enum ` Suc ` {.. n}" - using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident zero_notin_Suc_image in_enum_image subset_eq) + using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident in_enum_image subset_eq) then have "enum 1 \ s - {a}" by auto then have "upd 0 = n" using \a n < p\ \a = enum 0\ na[rule_format, of "enum 1"] by (auto simp: fun_eq_iff enum_Suc split: if_split_asm) then have "bij_betw upd (Suc ` {..< n}) {..< n}" using upd by (subst notIn_Un_bij_betw3[where b=0]) (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) then have "bij_betw (upd\Suc) {..enum 1 \ s - {a}\] \a = enum 0\ by (auto simp: \upd 0 = n\) show ?thesis proof (rule ksimplex.intros, standard) show "bij_betw (upd\Suc) {..< n} {..< n}" by fact show "base(n := p) \ {.. {..i. n\i \ (base(n := p)) i = p" using base base_out by (auto simp: Pi_iff) have "\i. Suc ` {..< i} = {..< Suc i} - {0}" by (auto simp: image_iff Ball_def) arith then have upd_Suc: "\i. i \ n \ (upd\Suc) ` {..< i} = upd ` {..< Suc i} - {n}" using \upd 0 = n\ upd_inj by (auto simp add: image_iff less_Suc_eq_0_disj) have n_in_upd: "\i. n \ upd ` {..< Suc i}" using \upd 0 = n\ by auto define f' where "f' i j = (if j \ (upd\Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j { fix x i assume i [arith]: "i \ n" with upd_Suc have "(upd \ Suc) ` {..a n < p\ \a = enum 0\ \upd 0 = n\ \a n = p - 1\ have "enum (Suc i) x = f' i x" by (auto simp add: f'_def enum_def) } then show "s - {a} = f' ` {.. n}" unfolding s_eq image_comp by (intro image_cong) auto qed qed next assume "ksimplex p n s'" and *: "\x\s'. x n = p" then show "\s a. ksimplex p (Suc n) s \ a \ s \ s' = s - {a}" proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s' . define b where "b = base (n := p - 1)" define u where "u i = (case i of 0 \ n | Suc i \ upd i)" for i have "ksimplex p (Suc n) (s' \ {b})" proof (rule ksimplex.intros, standard) show "b \ {.. {..0 < p\ unfolding lessThan_Suc b_def by (auto simp: PiE_iff) show "\i. Suc n \ i \ b i = p" using base_out by (auto simp: b_def) have "bij_betw u (Suc ` {..< n} \ {0}) ({.. {u 0})" using upd by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def) then show "bij_betw u {.. u`{..< i} then Suc (b j) else b j)" for i j have u_eq: "\i. i \ n \ u ` {..< Suc i} = upd ` {..< i} \ { n }" by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith { fix x have "x \ n \ n \ upd ` {.. {0})" unfolding atMost_Suc_eq_insert_0 by simp also have "\ = (f' \ Suc) ` {.. n} \ {b}" by (auto simp: f'_def) also have "(f' \ Suc) ` {.. n} = s'" using \0 < p\ base_out[of n] unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd) finally show "s' \ {b} = f' ` {.. Suc n}" .. qed moreover have "b \ s'" using * \0 < p\ by (auto simp: b_def) ultimately show ?thesis by auto qed qed lemma ksimplex_replace_0: assumes s: "ksimplex p n s" and a: "a \ s" assumes j: "j < n" and p: "\x\s - {a}. x j = 0" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 1" using s proof cases case (ksimplex b_s u_s) { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t by intro_locales fact+ assume b: "b \ t" "t - {b} = s - {a}" with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t" by (intro ksimplex_eq_top[of a b]) auto } then have "{s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = {s}" using s \a \ s\ by auto then show ?thesis by simp qed lemma ksimplex_replace_1: assumes s: "ksimplex p n s" and a: "a \ s" assumes j: "j < n" and p: "\x\s - {a}. x j = p" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 1" using s proof cases case (ksimplex b_s u_s) { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t by intro_locales fact+ assume b: "b \ t" "t - {b} = s - {a}" with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t" by (intro ksimplex_eq_bot[of a b]) auto } then have "{s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = {s}" using s \a \ s\ by auto then show ?thesis by simp qed lemma card_2_exists: "card s = 2 \ (\x\s. \y\s. x \ y \ (\z\s. z = x \ z = y))" by (auto simp: card_Suc_eq eval_nat_numeral) lemma ksimplex_replace_2: assumes s: "ksimplex p n s" and "a \ s" and "n \ 0" and lb: "\jx\s - {a}. x j \ 0" and ub: "\jx\s - {a}. x j \ p" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 2" using s proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s . from \a \ s\ obtain i where "i \ n" "a = enum i" unfolding s_eq by auto from \i \ n\ have "i = 0 \ i = n \ (0 < i \ i < n)" by linarith then have "\!s'. s' \ s \ ksimplex p n s' \ (\b\s'. s - {a} = s'- {b})" proof (elim disjE conjE) assume "i = 0" define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def) arith+ from rot upd have "bij_betw ?upd {.. ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" proof from \a = enum i\ ub \n \ 0\ \i = 0\ obtain i' where "i' \ n" "enum i' \ enum 0" "enum i' (upd 0) \ p" unfolding s_eq by (auto intro: upd_space simp: enum_inj) then have "enum 1 \ enum i'" "enum i' (upd 0) < p" using enum_le_p[of i' "upd 0"] by (auto simp: enum_inj enum_mono upd_space) then have "enum 1 (upd 0) < p" by (auto simp: le_fun_def intro: le_less_trans) then show "enum (Suc 0) \ {.. {..n \ 0\ by (auto simp: enum_0 enum_Suc PiE_iff extensional_def upd_space) { fix i assume "n \ i" then show "enum (Suc 0) i = p" using \n \ 0\ by (auto simp: enum_eq_p) } show "bij_betw ?upd {..n \ 0\ show ?thesis by (simp only: f'_def enum_def fun_eq_iff image_comp [symmetric]) (auto simp add: upd_inj) qed then have "enum ` Suc ` {..< n} = f' ` {..< n}" by (force simp: enum_inj) also have "Suc ` {..< n} = {.. n} - {0}" by (auto simp: image_iff Ball_def) arith also have "{..< n} = {.. n} - {n}" by auto finally have eq: "s - {a} = f' ` {.. n} - {f' n}" unfolding s_eq \a = enum i\ \i = 0\ by (simp add: inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f']) have "enum 0 < f' 0" using \n \ 0\ by (simp add: enum_strict_mono f'_eq_enum) also have "\ < f' n" using \n \ 0\ b.enum_strict_mono[of 0 n] unfolding b_enum by simp finally have "a \ f' n" using \a = enum i\ \i = 0\ by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b u where "kuhn_simplex p n b u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b u t . { fix x assume "x \ s" "x \ a" then have "x (upd 0) = enum (Suc 0) (upd 0)" by (auto simp: \a = enum i\ \i = 0\ s_eq enum_def enum_inj) } then have eq_upd0: "\x\t-{c}. x (upd 0) = enum (Suc 0) (upd 0)" unfolding eq_sma[symmetric] by auto then have "c (upd 0) \ enum (Suc 0) (upd 0)" using \n \ 0\ by (intro t.one_step[OF \c\t\ ]) (auto simp: upd_space) then have "c (upd 0) < enum (Suc 0) (upd 0) \ c (upd 0) > enum (Suc 0) (upd 0)" by auto then have "t = s \ t = f' ` {..n}" proof (elim disjE conjE) assume *: "c (upd 0) < enum (Suc 0) (upd 0)" interpret st: kuhn_simplex_pair p n base upd s b u t .. { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "c \ x" by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } note top = this have "s = t" using \a = enum i\ \i = 0\ \c \ t\ by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma]) (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) then show ?thesis by simp next assume *: "c (upd 0) > enum (Suc 0) (upd 0)" interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" b u t .. have eq: "f' ` {..n} - {f' n} = t - {c}" using eq_sma eq by simp { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "x \ c" by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } note top = this have "f' ` {..n} = t" using \a = enum i\ \i = 0\ \c \ t\ by (intro st.ksimplex_eq_top[OF _ _ _ _ eq]) (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top) then show ?thesis by simp qed } with ks_f' eq \a \ f' n\ \n \ 0\ show ?thesis apply (intro ex1I[of _ "f' ` {.. n}"]) apply auto [] apply metis done next assume "i = n" from \n \ 0\ obtain n' where n': "n = Suc n'" by (cases n) auto define rot where "rot i = (case i of 0 \ n' | Suc i \ i)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits) arith from rot upd have "bij_betw ?upd {.. ?upd`{..< i} then Suc (b j) else b j)" for i j interpret b: kuhn_simplex p n b "upd \ rot" "f' ` {.. n}" proof { fix i assume "n \ i" then show "b i = p" using base_out[of i] upd_space[of n'] by (auto simp: b_def n') } show "b \ {.. {..n \ 0\ upd_space[of n'] by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n') show "bij_betw ?upd {..n \ 0\ by auto { from \a = enum i\ \n \ 0\ \i = n\ lb upd_space[of n'] obtain i' where "i' \ n" "enum i' \ enum n" "0 < enum i' (upd n')" unfolding s_eq by (auto simp: enum_inj n') moreover have "enum i' (upd n') = base (upd n')" unfolding enum_def using \i' \ n\ \enum i' \ enum n\ by (auto simp: n' upd_inj enum_inj) ultimately have "0 < base (upd n')" by auto } then have benum1: "b.enum (Suc 0) = base" unfolding b.enum_Suc[OF \0] b.enum_0 by (auto simp: b_def rot_def) have [simp]: "\j. Suc j < n \ rot ` {..< Suc j} = {n'} \ {..< j}" by (auto simp: rot_def image_iff Ball_def split: nat.splits) have rot_simps: "\j. rot (Suc j) = j" "rot 0 = n'" by (simp_all add: rot_def) { fix j assume j: "Suc j \ n" then have "b.enum (Suc j) = enum j" by (induct j) (auto simp: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) } note b_enum_eq_enum = this then have "enum ` {..< n} = b.enum ` Suc ` {..< n}" by (auto simp: image_comp intro!: image_cong) also have "Suc ` {..< n} = {.. n} - {0}" by (auto simp: image_iff Ball_def) arith also have "{..< n} = {.. n} - {n}" by auto finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}" unfolding s_eq \a = enum i\ \i = n\ using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"] inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"] by (simp add: comp_def) have "b.enum 0 \ b.enum n" by (simp add: b.enum_mono) also have "b.enum n < enum n" using \n \ 0\ by (simp add: enum_strict_mono b_enum_eq_enum n') finally have "a \ b.enum 0" using \a = enum i\ \i = n\ by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b' u where "kuhn_simplex p n b' u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b' u t . { fix x assume "x \ s" "x \ a" then have "x (upd n') = enum n' (upd n')" by (auto simp: \a = enum i\ n' \i = n\ s_eq enum_def enum_inj in_upd_image) } then have eq_upd0: "\x\t-{c}. x (upd n') = enum n' (upd n')" unfolding eq_sma[symmetric] by auto then have "c (upd n') \ enum n' (upd n')" using \n \ 0\ by (intro t.one_step[OF \c\t\ ]) (auto simp: n' upd_space[unfolded n']) then have "c (upd n') < enum n' (upd n') \ c (upd n') > enum n' (upd n')" by auto then have "t = s \ t = b.enum ` {..n}" proof (elim disjE conjE) assume *: "c (upd n') > enum n' (upd n')" interpret st: kuhn_simplex_pair p n base upd s b' u t .. { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "x \ c" by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } note top = this have "s = t" using \a = enum i\ \i = n\ \c \ t\ by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma]) (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) then show ?thesis by simp next assume *: "c (upd n') < enum n' (upd n')" interpret st: kuhn_simplex_pair p n b "upd \ rot" "f' ` {.. n}" b' u t .. have eq: "f' ` {..n} - {b.enum 0} = t - {c}" using eq_sma eq f' by simp { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "c \ x" by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } note bot = this have "f' ` {..n} = t" using \a = enum i\ \i = n\ \c \ t\ by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq]) (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot) with f' show ?thesis by simp qed } with ks_f' eq \a \ b.enum 0\ \n \ 0\ show ?thesis apply (intro ex1I[of _ "b.enum ` {.. n}"]) apply auto [] apply metis done next assume i: "0 < i" "i < n" define i' where "i' = i - 1" with i have "Suc i' < n" by simp with i have Suc_i': "Suc i' = i" by (simp add: i'_def) let ?upd = "Fun.swap i' i upd" from i upd have "bij_betw ?upd {..< n} {..< n}" by (subst bij_betw_swap_iff) (auto simp: i'_def) define f' where [abs_def]: "f' i j = (if j \ ?upd`{..< i} then Suc (base j) else base j)" for i j interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}" proof show "base \ {.. {.. i" then show "base i = p" by (rule base_out) } show "bij_betw ?upd {.. {..n}" using i by auto { fix j assume "j \ n" moreover have "j < i \ i = j \ i < j" by arith moreover note i ultimately have "enum j = b.enum j \ j \ i" unfolding enum_def[abs_def] b.enum_def[abs_def] by (auto simp: fun_eq_iff swap_image i'_def in_upd_image inj_on_image_set_diff[OF inj_upd]) } note enum_eq_benum = this then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})" by (intro image_cong) auto then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}" unfolding s_eq \a = enum i\ using inj_on_image_set_diff[OF inj_enum Diff_subset \{i} \ {..n}\] inj_on_image_set_diff[OF b.inj_enum Diff_subset \{i} \ {..n}\] by (simp add: comp_def) have "a \ b.enum i" using \a = enum i\ enum_eq_benum i by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b' u where "kuhn_simplex p n b' u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b' u t . have "enum i' \ s - {a}" "enum (i + 1) \ s - {a}" using \a = enum i\ i enum_in by (auto simp: enum_inj i'_def) then obtain l k where l: "t.enum l = enum i'" "l \ n" "t.enum l \ c" and k: "t.enum k = enum (i + 1)" "k \ n" "t.enum k \ c" unfolding eq_sma by (auto simp: t.s_eq) with i have "t.enum l < t.enum k" by (simp add: enum_strict_mono i'_def) with \l \ n\ \k \ n\ have "l < k" by (simp add: t.enum_strict_mono) { assume "Suc l = k" have "enum (Suc (Suc i')) = t.enum (Suc l)" using i by (simp add: k \Suc l = k\ i'_def) then have False using \l < k\ \k \ n\ \Suc i' < n\ by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: if_split_asm) (metis Suc_lessD n_not_Suc_n upd_inj) } with \l < k\ have "Suc l < k" by arith have c_eq: "c = t.enum (Suc l)" proof (rule ccontr) assume "c \ t.enum (Suc l)" then have "t.enum (Suc l) \ s - {a}" using \l < k\ \k \ n\ by (simp add: t.s_eq eq_sma) then obtain j where "t.enum (Suc l) = enum j" "j \ n" "enum j \ enum i" unfolding s_eq \a = enum i\ by auto with i have "t.enum (Suc l) \ t.enum l \ t.enum k \ t.enum (Suc l)" by (auto simp: i'_def enum_mono enum_inj l k) with \Suc l < k\ \k \ n\ show False by (simp add: t.enum_mono) qed { have "t.enum (Suc (Suc l)) \ s - {a}" unfolding eq_sma c_eq t.s_eq using \Suc l < k\ \k \ n\ by (auto simp: t.enum_inj) then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \ n" "j \ i" by (auto simp: s_eq \a = enum i\) moreover have "enum i' < t.enum (Suc (Suc l))" unfolding l(1)[symmetric] using \Suc l < k\ \k \ n\ by (auto simp: t.enum_strict_mono) ultimately have "i' < j" using i by (simp add: enum_strict_mono i'_def) with \j \ i\ \j \ n\ have "t.enum k \ t.enum (Suc (Suc l))" unfolding i'_def by (simp add: enum_mono k eq) then have "k \ Suc (Suc l)" using \k \ n\ \Suc l < k\ by (simp add: t.enum_mono) } with \Suc l < k\ have "Suc (Suc l) = k" by simp then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))" using i by (simp add: k i'_def) also have "\ = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" using \Suc l < k\ \k \ n\ by (simp add: t.enum_Suc l t.upd_inj) finally have "(u l = upd i' \ u (Suc l) = upd (Suc i')) \ (u l = upd (Suc i') \ u (Suc l) = upd i')" using \Suc i' < n\ by (auto simp: enum_Suc fun_eq_iff split: if_split_asm) then have "t = s \ t = b.enum ` {..n}" proof (elim disjE conjE) assume u: "u l = upd i'" have "c = t.enum (Suc l)" unfolding c_eq .. also have "t.enum (Suc l) = enum (Suc i')" using u \l < k\ \k \ n\ \Suc i' < n\ by (simp add: enum_Suc t.enum_Suc l) also have "\ = a" using \a = enum i\ i by (simp add: i'_def) finally show ?thesis using eq_sma \a \ s\ \c \ t\ by auto next assume u: "u l = upd (Suc i')" define B where "B = b.enum ` {..n}" have "b.enum i' = enum i'" using enum_eq_benum[of i'] i by (auto simp: i'_def gr0_conv_Suc) have "c = t.enum (Suc l)" unfolding c_eq .. also have "t.enum (Suc l) = b.enum (Suc i')" using u \l < k\ \k \ n\ \Suc i' < n\ - by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \b.enum i' = enum i'\ swap_apply1) + by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \b.enum i' = enum i'\) (simp add: Suc_i') also have "\ = b.enum i" using i by (simp add: i'_def) finally have "c = b.enum i" . then have "t - {c} = B - {c}" "c \ B" unfolding eq_sma[symmetric] eq B_def using i by auto with \c \ t\ have "t = B" by auto then show ?thesis by (simp add: B_def) qed } with ks_f' eq \a \ b.enum i\ \n \ 0\ \i \ n\ show ?thesis apply (intro ex1I[of _ "b.enum ` {.. n}"]) apply auto [] apply metis done qed then show ?thesis using s \a \ s\ by (simp add: card_2_exists Ex1_def) metis qed text \Hence another step towards concreteness.\ lemma kuhn_simplex_lemma: assumes "\s. ksimplex p (Suc n) s \ rl ` s \ {.. Suc n}" and "odd (card {f. \s a. ksimplex p (Suc n) s \ a \ s \ (f = s - {a}) \ rl ` f = {..n} \ ((\j\n. \x\f. x j = 0) \ (\j\n. \x\f. x j = p))})" shows "odd (card {s. ksimplex p (Suc n) s \ rl ` s = {..Suc n}})" proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq, where bnd="\f. (\j\{..n}. \x\f. x j = 0) \ (\j\{..n}. \x\f. x j = p)"], safe del: notI) have *: "\x y. x = y \ odd (card x) \ odd (card y)" by auto show "odd (card {f. (\s\{s. ksimplex p (Suc n) s}. \a\s. f = s - {a}) \ rl ` f = {..n} \ ((\j\{..n}. \x\f. x j = 0) \ (\j\{..n}. \x\f. x j = p))})" apply (rule *[OF _ assms(2)]) apply (auto simp: atLeast0AtMost) done next fix s assume s: "ksimplex p (Suc n) s" then show "card s = n + 2" by (simp add: ksimplex_card) fix a assume a: "a \ s" then show "rl a \ Suc n" using assms(1) s by (auto simp: subset_eq) let ?S = "{t. ksimplex p (Suc n) t \ (\b\t. s - {a} = t - {b})}" { fix j assume j: "j \ n" "\x\s - {a}. x j = 0" with s a show "card ?S = 1" using ksimplex_replace_0[of p "n + 1" s a j] by (subst eq_commute) simp } { fix j assume j: "j \ n" "\x\s - {a}. x j = p" with s a show "card ?S = 1" using ksimplex_replace_1[of p "n + 1" s a j] by (subst eq_commute) simp } { assume "card ?S \ 2" "\ (\j\{..n}. \x\s - {a}. x j = p)" with s a show "\j\{..n}. \x\s - {a}. x j = 0" using ksimplex_replace_2[of p "n + 1" s a] by (subst (asm) eq_commute) auto } qed subsubsection \Reduced labelling\ definition reduced :: "nat \ (nat \ nat) \ nat" where "reduced n x = (LEAST k. k = n \ x k \ 0)" lemma reduced_labelling: shows "reduced n x \ n" and "\i x (reduced n x) \ 0" proof - show "reduced n x \ n" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto show "\i x (reduced n x) \ 0" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ qed lemma reduced_labelling_unique: "r \ n \ \i r = n \ x r \ 0 \ reduced n x = r" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) (metis le_less not_le)+ lemma reduced_labelling_zero: "j < n \ x j = 0 \ reduced n x \ j" using reduced_labelling[of n x] by auto lemma reduce_labelling_zero[simp]: "reduced 0 x = 0" by (rule reduced_labelling_unique) auto lemma reduced_labelling_nonzero: "j < n \ x j \ 0 \ reduced n x \ j" using reduced_labelling[of n x] by (elim allE[where x=j]) auto lemma reduced_labelling_Suc: "reduced (Suc n) x \ Suc n \ reduced (Suc n) x = reduced n x" using reduced_labelling[of "Suc n" x] by (intro reduced_labelling_unique[symmetric]) auto lemma complete_face_top: assumes "\x\f. \j\n. x j = 0 \ lab x j = 0" and "\x\f. \j\n. x j = p \ lab x j = 1" and eq: "(reduced (Suc n) \ lab) ` f = {..n}" shows "((\j\n. \x\f. x j = 0) \ (\j\n. \x\f. x j = p)) \ (\x\f. x n = p)" proof (safe del: disjCI) fix x j assume j: "j \ n" "\x\f. x j = 0" { fix x assume "x \ f" with assms j have "reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_zero) auto } moreover have "j \ (reduced (Suc n) \ lab) ` f" using j eq by auto ultimately show "x n = p" by force next fix x j assume j: "j \ n" "\x\f. x j = p" and x: "x \ f" have "j = n" proof (rule ccontr) assume "\ ?thesis" { fix x assume "x \ f" with assms j have "reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_nonzero) auto then have "reduced (Suc n) (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover have "n \ (reduced (Suc n) \ lab) ` f" using eq by auto ultimately show False by force qed moreover have "j \ (reduced (Suc n) \ lab) ` f" using j eq by auto ultimately show "x n = p" using j x by auto qed auto text \Hence we get just about the nice induction.\ lemma kuhn_induction: assumes "0 < p" and lab_0: "\x. \j\n. (\j. x j \ p) \ x j = 0 \ lab x j = 0" and lab_1: "\x. \j\n. (\j. x j \ p) \ x j = p \ lab x j = 1" and odd: "odd (card {s. ksimplex p n s \ (reduced n\lab) ` s = {..n}})" shows "odd (card {s. ksimplex p (Suc n) s \ (reduced (Suc n)\lab) ` s = {..Suc n}})" proof - let ?rl = "reduced (Suc n) \ lab" and ?ext = "\f v. \j\n. \x\f. x j = v" let ?ext = "\s. (\j\n. \x\s. x j = 0) \ (\j\n. \x\s. x j = p)" have "\s. ksimplex p (Suc n) s \ ?rl ` s \ {..Suc n}" by (simp add: reduced_labelling subset_eq) moreover have "{s. ksimplex p n s \ (reduced n \ lab) ` s = {..n}} = {f. \s a. ksimplex p (Suc n) s \ a \ s \ f = s - {a} \ ?rl ` f = {..n} \ ?ext f}" proof (intro set_eqI, safe del: disjCI equalityI disjE) fix s assume s: "ksimplex p n s" and rl: "(reduced n \ lab) ` s = {..n}" from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p n u b s . have all_eq_p: "\x\s. x n = p" by (auto simp: out_eq_p) moreover { fix x assume "x \ s" with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] have "?rl x \ n" by (auto intro!: reduced_labelling_nonzero) then have "?rl x = reduced n (lab x)" by (auto intro!: reduced_labelling_Suc) } then have "?rl ` s = {..n}" using rl by (simp cong: image_cong) moreover obtain t a where "ksimplex p (Suc n) t" "a \ t" "s = t - {a}" using s unfolding simplex_top_face[OF \0 < p\ all_eq_p] by auto ultimately show "\t a. ksimplex p (Suc n) t \ a \ t \ s = t - {a} \ ?rl ` s = {..n} \ ?ext s" by auto next fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}" and a: "a \ s" and "?ext (s - {a})" from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p "Suc n" u b s . have all_eq_p: "\x\s. x (Suc n) = p" by (auto simp: out_eq_p) { fix x assume "x \ s - {a}" then have "?rl x \ ?rl ` (s - {a})" by auto then have "?rl x \ n" unfolding rl by auto then have "?rl x = reduced n (lab x)" by (auto intro!: reduced_labelling_Suc) } then show rl': "(reduced n\lab) ` (s - {a}) = {..n}" unfolding rl[symmetric] by (intro image_cong) auto from \?ext (s - {a})\ have all_eq_p: "\x\s - {a}. x n = p" proof (elim disjE exE conjE) fix j assume "j \ n" "\x\s - {a}. x j = 0" with lab_0[rule_format, of j] all_eq_p s_le_p have "\x. x \ s - {a} \ reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_zero) auto moreover have "j \ ?rl ` (s - {a})" using \j \ n\ unfolding rl by auto ultimately show ?thesis by force next fix j assume "j \ n" and eq_p: "\x\s - {a}. x j = p" show ?thesis proof cases assume "j = n" with eq_p show ?thesis by simp next assume "j \ n" { fix x assume x: "x \ s - {a}" have "reduced n (lab x) \ j" proof (rule reduced_labelling_nonzero) show "lab x j \ 0" using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \j \ n\ by auto show "j < n" using \j \ n\ \j \ n\ by simp qed then have "reduced n (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover have "n \ (reduced n\lab) ` (s - {a})" unfolding rl' by auto ultimately show ?thesis by force qed qed show "ksimplex p n (s - {a})" unfolding simplex_top_face[OF \0 < p\ all_eq_p] using s a by auto qed ultimately show ?thesis using assms by (intro kuhn_simplex_lemma) auto qed text \And so we get the final combinatorial result.\ lemma ksimplex_0: "ksimplex p 0 s \ s = {(\x. p)}" proof assume "ksimplex p 0 s" then show "s = {(\x. p)}" by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) next assume s: "s = {(\x. p)}" show "ksimplex p 0 s" proof (intro ksimplex, unfold_locales) show "(\_. p) \ {..<0::nat} \ {..x j. (\j. x j \ p) \ j < n \ x j = 0 \ lab x j = 0" and "\x j. (\j. x j \ p) \ j < n \ x j = p \ lab x j = 1" shows "odd (card {s. ksimplex p n s \ (reduced n\lab) ` s = {..n}})" (is "odd (card (?M n))") using assms proof (induct n) case 0 then show ?case by (simp add: ksimplex_0 cong: conj_cong) next case (Suc n) then have "odd (card (?M n))" by force with Suc show ?case using kuhn_induction[of p n] by (auto simp: comp_def) qed lemma kuhn_lemma: fixes n p :: nat assumes "0 < p" and "\x. (\i p) \ (\i label x i = 1)" and "\x. (\i p) \ (\i label x i = 0)" and "\x. (\i p) \ (\i label x i = 1)" obtains q where "\iir s. (\j r j \ r j \ q j + 1) \ (\j s j \ s j \ q j + 1) \ label r i \ label s i" proof - let ?rl = "reduced n \ label" let ?A = "{s. ksimplex p n s \ ?rl ` s = {..n}}" have "odd (card ?A)" using assms by (intro kuhn_combinatorial[of p n label]) auto then have "?A \ {}" by (rule odd_card_imp_not_empty) then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}" by (auto elim: ksimplex.cases) interpret kuhn_simplex p n b u s by fact show ?thesis proof (intro that[of b] allI impI) fix i assume "i < n" then show "b i < p" using base by auto next fix i assume "i < n" then have "i \ {.. n}" "Suc i \ {.. n}" by auto then obtain u v where u: "u \ s" "Suc i = ?rl u" and v: "v \ s" "i = ?rl v" unfolding rl[symmetric] by blast have "label u i \ label v i" using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] u(2)[symmetric] v(2)[symmetric] \i < n\ by auto moreover have "b j \ u j" "u j \ b j + 1" "b j \ v j" "v j \ b j + 1" if "j < n" for j using that base_le[OF \u\s\] le_Suc_base[OF \u\s\] base_le[OF \v\s\] le_Suc_base[OF \v\s\] by auto ultimately show "\r s. (\j r j \ r j \ b j + 1) \ (\j s j \ s j \ b j + 1) \ label r i \ label s i" by blast qed qed subsubsection \Main result for the unit cube\ lemma kuhn_labelling_lemma': assumes "(\x::nat\real. P x \ P (f x))" and "\x. P x \ (\i::nat. Q i \ 0 \ x i \ x i \ 1)" shows "\l. (\x i. l x i \ (1::nat)) \ (\x i. P x \ Q i \ x i = 0 \ l x i = 0) \ (\x i. P x \ Q i \ x i = 1 \ l x i = 1) \ (\x i. P x \ Q i \ l x i = 0 \ x i \ f x i) \ (\x i. P x \ Q i \ l x i = 1 \ f x i \ x i)" proof - have and_forall_thm: "\P Q. (\x. P x) \ (\x. Q x) \ (\x. P x \ Q x)" by auto have *: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 \ x \ 1 \ x \ y \ x \ 0 \ y \ x" by auto show ?thesis unfolding and_forall_thm apply (subst choice_iff[symmetric])+ apply rule apply rule proof - fix x x' let ?R = "\y::nat. (P x \ Q x' \ x x' = 0 \ y = 0) \ (P x \ Q x' \ x x' = 1 \ y = 1) \ (P x \ Q x' \ y = 0 \ x x' \ (f x) x') \ (P x \ Q x' \ y = 1 \ (f x) x' \ x x')" have "0 \ f x x' \ f x x' \ 1" if "P x" "Q x'" using assms(2)[rule_format,of "f x" x'] that apply (drule_tac assms(1)[rule_format]) apply auto done then have "?R 0 \ ?R 1" by auto then show "\y\1. ?R y" by auto qed qed subsection \Brouwer's fixed point theorem\ text \We start proving Brouwer's fixed point theorem for the unit cube = \cbox 0 One\.\ lemma brouwer_cube: fixes f :: "'a::euclidean_space \ 'a" assumes "continuous_on (cbox 0 One) f" and "f ` cbox 0 One \ cbox 0 One" shows "\x\cbox 0 One. f x = x" proof (rule ccontr) define n where "n = DIM('a)" have n: "1 \ n" "0 < n" "n \ 0" - unfolding n_def by (auto simp: Suc_le_eq DIM_positive) + unfolding n_def by (auto simp: Suc_le_eq) assume "\ ?thesis" then have *: "\ (\x\cbox 0 One. f x - x = 0)" by auto obtain d where d: "d > 0" "\x. x \ cbox 0 One \ d \ norm (f x - x)" apply (rule brouwer_compactness_lemma[OF compact_cbox _ *]) apply (rule continuous_intros assms)+ apply blast done have *: "\x. x \ cbox 0 One \ f x \ cbox 0 One" "\x. x \ (cbox 0 One::'a set) \ (\i\Basis. True \ 0 \ x \ i \ x \ i \ 1)" using assms(2)[unfolded image_subset_iff Ball_def] unfolding cbox_def by auto obtain label :: "'a \ 'a \ nat" where label [rule_format]: "\x. \i\Basis. label x i \ 1" "\x. \i\Basis. x \ cbox 0 One \ x \ i = 0 \ label x i = 0" "\x. \i\Basis. x \ cbox 0 One \ x \ i = 1 \ label x i = 1" "\x. \i\Basis. x \ cbox 0 One \ label x i = 0 \ x \ i \ f x \ i" "\x. \i\Basis. x \ cbox 0 One \ label x i = 1 \ f x \ i \ x \ i" using kuhn_labelling_lemma[OF *] by auto note label = this [rule_format] have lem1: "\x\cbox 0 One. \y\cbox 0 One. \i\Basis. label x i \ label y i \ \f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" proof safe fix x y :: 'a assume x: "x \ cbox 0 One" and y: "y \ cbox 0 One" fix i assume i: "label x i \ label y i" "i \ Basis" have *: "\x y fx fy :: real. x \ fx \ fy \ y \ fx \ x \ y \ fy \ \fx - x\ \ \fy - fx\ + \y - x\" by auto have "\(f x - x) \ i\ \ \(f y - f x)\i\ + \(y - x)\i\" proof (cases "label x i = 0") case True then have fxy: "\ f y \ i \ y \ i \ f x \ i \ x \ i" by (metis True i label(1) label(5) le_antisym less_one not_le_imp_less y) show ?thesis unfolding inner_simps by (rule *) (auto simp: True i label x y fxy) next case False then show ?thesis using label [OF \i \ Basis\] i(1) x y apply (auto simp: inner_diff_left le_Suc_eq) by (metis "*") qed also have "\ \ norm (f y - f x) + norm (y - x)" by (simp add: add_mono i(2) norm_bound_Basis_le) finally show "\f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" unfolding inner_simps . qed have "\e>0. \x\cbox 0 One. \y\cbox 0 One. \z\cbox 0 One. \i\Basis. norm (x - z) < e \ norm (y - z) < e \ label x i \ label y i \ \(f(z) - z)\i\ < d / (real n)" proof - have d': "d / real n / 8 > 0" - using d(1) by (simp add: n_def DIM_positive) + using d(1) by (simp add: n_def) have *: "uniformly_continuous_on (cbox 0 One) f" by (rule compact_uniformly_continuous[OF assms(1) compact_cbox]) obtain e where e: "e > 0" "\x x'. x \ cbox 0 One \ x' \ cbox 0 One \ norm (x' - x) < e \ norm (f x' - f x) < d / real n / 8" using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] unfolding dist_norm by blast show ?thesis proof (intro exI conjI ballI impI) show "0 < min (e / 2) (d / real n / 8)" using d' e by auto fix x y z i assume as: "x \ cbox 0 One" "y \ cbox 0 One" "z \ cbox 0 One" "norm (x - z) < min (e / 2) (d / real n / 8)" "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \ label y i" assume i: "i \ Basis" have *: "\z fz x fx n1 n2 n3 n4 d4 d :: real. \fx - x\ \ n1 + n2 \ \fx - fz\ \ n3 \ \x - z\ \ n4 \ n1 < d4 \ n2 < 2 * d4 \ n3 < d4 \ n4 < d4 \ (8 * d4 = d) \ \fz - z\ < d" by auto show "\(f z - z) \ i\ < d / real n" unfolding inner_simps proof (rule *) show "\f x \ i - x \ i\ \ norm (f y -f x) + norm (y - x)" using as(1) as(2) as(6) i lem1 by blast show "norm (f x - f z) < d / real n / 8" using d' e as by auto show "\f x \ i - f z \ i\ \ norm (f x - f z)" "\x \ i - z \ i\ \ norm (x - z)" unfolding inner_diff_left[symmetric] by (rule Basis_le_norm[OF i])+ have tria: "norm (y - x) \ norm (y - z) + norm (x - z)" using dist_triangle[of y x z, unfolded dist_norm] unfolding norm_minus_commute by auto also have "\ < e / 2 + e / 2" using as(4) as(5) by auto finally show "norm (f y - f x) < d / real n / 8" using as(1) as(2) e(2) by auto have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" using as(4) as(5) by auto with tria show "norm (y - x) < 2 * (d / real n / 8)" by auto qed (use as in auto) qed qed then obtain e where e: "e > 0" "\x y z i. x \ cbox 0 One \ y \ cbox 0 One \ z \ cbox 0 One \ i \ Basis \ norm (x - z) < e \ norm (y - z) < e \ label x i \ label y i \ \(f z - z) \ i\ < d / real n" by blast obtain p :: nat where p: "1 + real n / e \ real p" using real_arch_simple .. have "1 + real n / e > 0" using e(1) n by (simp add: add_pos_pos) then have "p > 0" using p by auto obtain b :: "nat \ 'a" where b: "bij_betw b {..< n} Basis" by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) define b' where "b' = inv_into {..< n} b" then have b': "bij_betw b' Basis {..< n}" using bij_betw_inv_into[OF b] by auto then have b'_Basis: "\i. i \ Basis \ b' i \ {..< n}" unfolding bij_betw_def by (auto simp: set_eq_iff) have bb'[simp]:"\i. i \ Basis \ b (b' i) = i" unfolding b'_def using b by (auto simp: f_inv_into_f bij_betw_def) have b'b[simp]:"\i. i < n \ b' (b i) = i" unfolding b'_def using b by (auto simp: inv_into_f_eq bij_betw_def) have *: "\x :: nat. x = 0 \ x = 1 \ x \ 1" by auto have b'': "\j. j < n \ b j \ Basis" using b unfolding bij_betw_def by auto have q1: "0 < p" "\x. (\i p) \ (\ii\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0 \ (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" unfolding * using \p > 0\ \n > 0\ using label(1)[OF b''] by auto { fix x :: "nat \ nat" and i assume "\i p" "i < n" "x i = p \ x i = 0" then have "(\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ (cbox 0 One::'a set)" using b'_Basis - by (auto simp: cbox_def inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1) } + by (auto simp: cbox_def bij_betw_def zero_le_divide_iff divide_le_eq_1) } note cube = this have q2: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0)" unfolding o_def using cube \p > 0\ by (intro allI impI label(2)) (auto simp: b'') have q3: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" using cube \p > 0\ unfolding o_def by (intro allI impI label(3)) (auto simp: b'') obtain q where q: "\iir s. (\j r j \ r j \ q j + 1) \ (\j s j \ s j \ q j + 1) \ (label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) i \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) i" by (rule kuhn_lemma[OF q1 q2 q3]) define z :: 'a where "z = (\i\Basis. (real (q (b' i)) / real p) *\<^sub>R i)" have "\i\Basis. d / real n \ \(f z - z)\i\" proof (rule ccontr) have "\i\Basis. q (b' i) \ {0..p}" using q(1) b' by (auto intro: less_imp_le simp: bij_betw_def) then have "z \ cbox 0 One" unfolding z_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) then have d_fz_z: "d \ norm (f z - z)" by (rule d) assume "\ ?thesis" then have as: "\i\Basis. \f z \ i - z \ i\ < d / real n" using \n > 0\ by (auto simp: not_le inner_diff) have "norm (f z - z) \ (\i\Basis. \f z \ i - z \ i\)" unfolding inner_diff_left[symmetric] by (rule norm_le_l1) also have "\ < (\(i::'a) \ Basis. d / real n)" by (meson as finite_Basis nonempty_Basis sum_strict_mono) also have "\ = d" using DIM_positive[where 'a='a] by (auto simp: n_def) finally show False using d_fz_z by auto qed then obtain i where i: "i \ Basis" "d / real n \ \(f z - z) \ i\" .. have *: "b' i < n" using i and b'[unfolded bij_betw_def] by auto obtain r s where rs: "\j. j < n \ q j \ r j \ r j \ q j + 1" "\j. j < n \ q j \ s j \ s j \ q j + 1" "(label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) (b' i) \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) (b' i)" using q(2)[rule_format,OF *] by blast have b'_im: "\i. i \ Basis \ b' i < n" using b' unfolding bij_betw_def by auto define r' ::'a where "r' = (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ r (b' i) \ p" apply (rule order_trans) apply (rule rs(1)[OF b'_im,THEN conjunct2]) using q(1)[rule_format,OF b'_im] apply (auto simp: Suc_le_eq) done then have "r' \ cbox 0 One" unfolding r'_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) define s' :: 'a where "s' = (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ s (b' i) \ p" using b'_im q(1) rs(2) by fastforce then have "s' \ cbox 0 One" unfolding s'_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) have "z \ cbox 0 One" unfolding z_def cbox_def using b'_Basis q(1)[rule_format,OF b'_im] \p > 0\ by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) { have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" by (rule sum_mono) (use rs(1)[OF b'_im] in force) also have "\ < e * real p" using p \e > 0\ \p > 0\ by (auto simp: field_simps n_def) finally have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) < e * real p" . } moreover { have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" by (rule sum_mono) (use rs(2)[OF b'_im] in force) also have "\ < e * real p" using p \e > 0\ \p > 0\ by (auto simp: field_simps n_def) finally have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) < e * real p" . } ultimately have "norm (r' - z) < e" and "norm (s' - z) < e" unfolding r'_def s'_def z_def using \p > 0\ apply (rule_tac[!] le_less_trans[OF norm_le_l1]) apply (auto simp: field_simps sum_divide_distrib[symmetric] inner_diff_left) done then have "\(f z - z) \ i\ < d / real n" using rs(3) i unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' by (intro e(2)[OF \r'\cbox 0 One\ \s'\cbox 0 One\ \z\cbox 0 One\]) auto then show False using i by auto qed text \Next step is to prove it for nonempty interiors.\ lemma brouwer_weak: fixes f :: "'a::euclidean_space \ 'a" assumes "compact S" and "convex S" and "interior S \ {}" and "continuous_on S f" and "f ` S \ S" obtains x where "x \ S" and "f x = x" proof - let ?U = "cbox 0 One :: 'a set" have "\Basis /\<^sub>R 2 \ interior ?U" proof (rule interiorI) let ?I = "(\i\Basis. {x::'a. 0 < x \ i} \ {x. x \ i < 1})" show "open ?I" - by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner continuous_on_const continuous_on_id) + by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner) show "\Basis /\<^sub>R 2 \ ?I" by simp show "?I \ cbox 0 One" unfolding cbox_def by force qed then have *: "interior ?U \ {}" by fast have *: "?U homeomorphic S" using homeomorphic_convex_compact[OF convex_box(1) compact_cbox * assms(2,1,3)] . have "\f. continuous_on ?U f \ f ` ?U \ ?U \ (\x\?U. f x = x)" using brouwer_cube by auto then show ?thesis unfolding homeomorphic_fixpoint_property[OF *] using assms by (auto intro: that) qed text \Then the particular case for closed balls.\ lemma brouwer_ball: fixes f :: "'a::euclidean_space \ 'a" assumes "e > 0" and "continuous_on (cball a e) f" and "f ` cball a e \ cball a e" obtains x where "x \ cball a e" and "f x = x" using brouwer_weak[OF compact_cball convex_cball, of a e f] unfolding interior_cball ball_eq_empty using assms by auto text \And finally we prove Brouwer's fixed point theorem in its general version.\ theorem brouwer: fixes f :: "'a::euclidean_space \ 'a" assumes S: "compact S" "convex S" "S \ {}" and contf: "continuous_on S f" and fim: "f ` S \ S" obtains x where "x \ S" and "f x = x" proof - have "\e>0. S \ cball 0 e" using compact_imp_bounded[OF \compact S\] unfolding bounded_pos by auto then obtain e where e: "e > 0" "S \ cball 0 e" by blast have "\x\ cball 0 e. (f \ closest_point S) x = x" proof (rule_tac brouwer_ball[OF e(1)]) show "continuous_on (cball 0 e) (f \ closest_point S)" apply (rule continuous_on_compose) using S compact_eq_bounded_closed continuous_on_closest_point apply blast by (meson S contf closest_point_in_set compact_imp_closed continuous_on_subset image_subsetI) show "(f \ closest_point S) ` cball 0 e \ cball 0 e" by clarsimp (metis S fim closest_point_exists(1) compact_eq_bounded_closed e(2) image_subset_iff mem_cball_0 subsetCE) qed (use assms in auto) then obtain x where x: "x \ cball 0 e" "(f \ closest_point S) x = x" .. have "x \ S" by (metis closest_point_in_set comp_apply compact_imp_closed fim image_eqI S(1) S(3) subset_iff x(2)) then have *: "closest_point S x = x" by (rule closest_point_self) show thesis proof show "closest_point S x \ S" by (simp add: "*" \x \ S\) show "f (closest_point S x) = closest_point S x" using "*" x(2) by auto qed qed subsection \Applications\ text \So we get the no-retraction theorem.\ corollary no_retraction_cball: fixes a :: "'a::euclidean_space" assumes "e > 0" shows "\ (frontier (cball a e) retract_of (cball a e))" proof assume *: "frontier (cball a e) retract_of (cball a e)" have **: "\xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" using scaleR_left_distrib[of 1 1 a] by auto obtain x where x: "x \ {x. norm (a - x) = e}" "2 *\<^sub>R a - x = x" proof (rule retract_fixpoint_property[OF *, of "\x. scaleR 2 a - x"]) show "continuous_on (frontier (cball a e)) ((-) (2 *\<^sub>R a))" by (intro continuous_intros) show "(-) (2 *\<^sub>R a) ` frontier (cball a e) \ frontier (cball a e)" by clarsimp (metis "**" dist_norm norm_minus_cancel) qed (auto simp: dist_norm intro: brouwer_ball[OF assms]) then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" by (auto simp: algebra_simps) then have "a = x" unfolding scaleR_left_distrib[symmetric] by auto then show False using x assms by auto qed corollary contractible_sphere: fixes a :: "'a::euclidean_space" shows "contractible(sphere a r) \ r \ 0" proof (cases "0 < r") case True then show ?thesis unfolding contractible_def nullhomotopic_from_sphere_extension using no_retraction_cball [OF True, of a] by (auto simp: retract_of_def retraction_def) next case False then show ?thesis unfolding contractible_def nullhomotopic_from_sphere_extension - using continuous_on_const less_eq_real_def by auto + using less_eq_real_def by auto qed corollary connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by auto next case equal then show ?thesis by auto next case greater show ?thesis proof assume L: ?lhs have "False" if 1: "DIM('a) = 1" proof - obtain x y where xy: "sphere a r = {x,y}" "x \ y" using sphere_1D_doubleton [OF 1 greater] by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist) then have "finite (sphere a r)" by auto with L \r > 0\ xy show "False" using connected_finite_iff_sing by auto qed with greater show ?rhs by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq) next assume ?rhs then show ?lhs using connected_sphere greater by auto qed qed corollary path_connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "path_connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connected_sphere_eq path_connected_imp_connected by blast next assume R: ?rhs then show ?lhs by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere) qed proposition frontier_subset_retraction: fixes S :: "'a::euclidean_space set" assumes "bounded S" and fros: "frontier S \ T" and contf: "continuous_on (closure S) f" and fim: "f ` S \ T" and fid: "\x. x \ T \ f x = x" shows "S \ T" proof (rule ccontr) assume "\ S \ T" then obtain a where "a \ S" "a \ T" by blast define g where "g \ \z. if z \ closure S then f z else z" have "continuous_on (closure S \ closure(-S)) g" unfolding g_def apply (rule continuous_on_cases) - using fros fid frontier_closures - apply (auto simp: contf continuous_on_id) - done + using fros fid frontier_closures by (auto simp: contf) moreover have "closure S \ closure(- S) = UNIV" using closure_Un by fastforce ultimately have contg: "continuous_on UNIV g" by metis obtain B where "0 < B" and B: "closure S \ ball a B" using \bounded S\ bounded_subset_ballD by blast have notga: "g x \ a" for x unfolding g_def using fros fim \a \ T\ apply (auto simp: frontier_def) using fid interior_subset apply fastforce by (simp add: \a \ S\ closure_def) define h where "h \ (\y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \ g" have "\ (frontier (cball a B) retract_of (cball a B))" by (metis no_retraction_cball \0 < B\) then have "\k. \ retraction (cball a B) (frontier (cball a B)) k" by (simp add: retract_of_def) moreover have "retraction (cball a B) (frontier (cball a B)) h" unfolding retraction_def proof (intro conjI ballI) show "frontier (cball a B) \ cball a B" by force show "continuous_on (cball a B) h" unfolding h_def by (intro continuous_intros) (use contg continuous_on_subset notga in auto) show "h ` cball a B \ frontier (cball a B)" using \0 < B\ by (auto simp: h_def notga dist_norm) show "\x. x \ frontier (cball a B) \ h x = x" apply (auto simp: h_def algebra_simps) apply (simp add: vector_add_divide_simps notga) by (metis (no_types, hide_lams) B add.commute dist_commute dist_norm g_def mem_ball not_less_iff_gr_or_eq subset_eq) qed ultimately show False by simp qed subsubsection \Punctured affine hulls, etc\ lemma rel_frontier_deformation_retract_of_punctured_convex: fixes S :: "'a::euclidean_space set" assumes "convex S" "convex T" "bounded S" and arelS: "a \ rel_interior S" and relS: "rel_frontier S \ T" and affS: "T \ affine hull S" obtains r where "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id r" "retraction (T - {a}) (rel_frontier S) r" proof - have "\d. 0 < d \ (a + d *\<^sub>R l) \ rel_frontier S \ (\e. 0 \ e \ e < d \ (a + e *\<^sub>R l) \ rel_interior S)" if "(a + l) \ affine hull S" "l \ 0" for l apply (rule ray_to_rel_frontier [OF \bounded S\ arelS]) apply (rule that)+ by metis then obtain dd where dd1: "\l. \(a + l) \ affine hull S; l \ 0\ \ 0 < dd l \ (a + dd l *\<^sub>R l) \ rel_frontier S" and dd2: "\l e. \(a + l) \ affine hull S; e < dd l; 0 \ e; l \ 0\ \ (a + e *\<^sub>R l) \ rel_interior S" by metis+ have aaffS: "a \ affine hull S" by (meson arelS subsetD hull_inc rel_interior_subset) have "((\z. z - a) ` (affine hull S - {a})) = ((\z. z - a) ` (affine hull S)) - {0}" by auto moreover have "continuous_on (((\z. z - a) ` (affine hull S)) - {0}) (\x. dd x *\<^sub>R x)" proof (rule continuous_on_compact_surface_projection) show "compact (rel_frontier ((\z. z - a) ` S))" by (simp add: \bounded S\ bounded_translation_minus compact_rel_frontier_bounded) have releq: "rel_frontier ((\z. z - a) ` S) = (\z. z - a) ` rel_frontier S" using rel_frontier_translation [of "-a"] add.commute by simp also have "\ \ (\z. z - a) ` (affine hull S) - {0}" using rel_frontier_affine_hull arelS rel_frontier_def by fastforce finally show "rel_frontier ((\z. z - a) ` S) \ (\z. z - a) ` (affine hull S) - {0}" . show "cone ((\z. z - a) ` (affine hull S))" by (rule subspace_imp_cone) (use aaffS in \simp add: subspace_affine image_comp o_def affine_translation_aux [of a]\) show "(0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)) \ (dd x = k)" if x: "x \ (\z. z - a) ` (affine hull S) - {0}" for k x proof show "dd x = k \ 0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)" using dd1 [of x] that image_iff by (fastforce simp add: releq) next assume k: "0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)" have False if "dd x < k" proof - have "k \ 0" "a + k *\<^sub>R x \ closure S" using k closure_translation [of "-a"] by (auto simp: rel_frontier_def cong: image_cong_simp) then have segsub: "open_segment a (a + k *\<^sub>R x) \ rel_interior S" by (metis rel_interior_closure_convex_segment [OF \convex S\ arelS]) have "x \ 0" and xaffS: "a + x \ affine hull S" using x by auto then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \ rel_frontier S" using dd1 by auto moreover have "a + dd x *\<^sub>R x \ open_segment a (a + k *\<^sub>R x)" using k \x \ 0\ \0 < dd x\ apply (simp add: in_segment) apply (rule_tac x = "dd x / k" in exI) apply (simp add: field_simps that) apply (simp add: vector_add_divide_simps algebra_simps) done ultimately show ?thesis using segsub by (auto simp: rel_frontier_def) qed moreover have False if "k < dd x" using x k that rel_frontier_def by (fastforce simp: algebra_simps releq dest!: dd2) ultimately show "dd x = k" by fastforce qed qed ultimately have *: "continuous_on ((\z. z - a) ` (affine hull S - {a})) (\x. dd x *\<^sub>R x)" by auto have "continuous_on (affine hull S - {a}) ((\x. a + dd x *\<^sub>R x) \ (\z. z - a))" by (intro * continuous_intros continuous_on_compose) with affS have contdd: "continuous_on (T - {a}) ((\x. a + dd x *\<^sub>R x) \ (\z. z - a))" by (blast intro: continuous_on_subset) show ?thesis proof show "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id (\x. a + dd (x - a) *\<^sub>R (x - a))" proof (rule homotopic_with_linear) show "continuous_on (T - {a}) id" by (intro continuous_intros continuous_on_compose) show "continuous_on (T - {a}) (\x. a + dd (x - a) *\<^sub>R (x - a))" using contdd by (simp add: o_def) show "closed_segment (id x) (a + dd (x - a) *\<^sub>R (x - a)) \ T - {a}" if "x \ T - {a}" for x proof (clarsimp simp: in_segment, intro conjI) fix u::real assume u: "0 \ u" "u \ 1" have "a + dd (x - a) *\<^sub>R (x - a) \ T" by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that) then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \ T" using convexD [OF \convex T\] that u by simp have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \ (1 - u + u * d) *\<^sub>R (x - a) = 0" for d by (auto simp: algebra_simps) have "x \ T" "x \ a" using that by auto then have axa: "a + (x - a) \ affine hull S" by (metis (no_types) add.commute affS diff_add_cancel rev_subsetD) then have "\ dd (x - a) \ 0 \ a + dd (x - a) *\<^sub>R (x - a) \ rel_frontier S" using \x \ a\ dd1 by fastforce with \x \ a\ show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \ a" apply (auto simp: iff) using less_eq_real_def mult_le_0_iff not_less u by fastforce qed qed show "retraction (T - {a}) (rel_frontier S) (\x. a + dd (x - a) *\<^sub>R (x - a))" proof (simp add: retraction_def, intro conjI ballI) show "rel_frontier S \ T - {a}" using arelS relS rel_frontier_def by fastforce show "continuous_on (T - {a}) (\x. a + dd (x - a) *\<^sub>R (x - a))" using contdd by (simp add: o_def) show "(\x. a + dd (x - a) *\<^sub>R (x - a)) ` (T - {a}) \ rel_frontier S" apply (auto simp: rel_frontier_def) apply (metis Diff_subset add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def subset_iff) by (metis DiffE add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def rev_subsetD) show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \ rel_frontier S" for x proof - have "x \ a" using that arelS by (auto simp: rel_frontier_def) have False if "dd (x - a) < 1" proof - have "x \ closure S" using x by (auto simp: rel_frontier_def) then have segsub: "open_segment a x \ rel_interior S" by (metis rel_interior_closure_convex_segment [OF \convex S\ arelS]) have xaffS: "x \ affine hull S" using affS relS x by auto then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \ rel_frontier S" using dd1 by (auto simp: \x \ a\) moreover have "a + dd (x - a) *\<^sub>R (x - a) \ open_segment a x" using \x \ a\ \0 < dd (x - a)\ apply (simp add: in_segment) apply (rule_tac x = "dd (x - a)" in exI) apply (simp add: algebra_simps that) done ultimately show ?thesis using segsub by (auto simp: rel_frontier_def) qed moreover have False if "1 < dd (x - a)" using x that dd2 [of "x - a" 1] \x \ a\ closure_affine_hull by (auto simp: rel_frontier_def) ultimately have "dd (x - a) = 1" \ \similar to another proof above\ by fastforce with that show ?thesis by (simp add: rel_frontier_def) qed qed qed qed corollary rel_frontier_retract_of_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" "a \ rel_interior S" shows "rel_frontier S retract_of (affine hull S - {a})" apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S "affine hull S" a]) apply (auto simp: affine_imp_convex rel_frontier_affine_hull retract_of_def assms) done corollary rel_boundary_retract_of_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "a \ rel_interior S" shows "(S - rel_interior S) retract_of (affine hull S - {a})" by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def rel_frontier_retract_of_punctured_affine_hull) lemma homotopy_eqv_rel_frontier_punctured_convex: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ rel_interior S" "convex T" "rel_frontier S \ T" "T \ affine hull S" shows "(rel_frontier S) homotopy_eqv (T - {a})" apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S T]) using assms apply auto using deformation_retract_imp_homotopy_eqv homotopy_equivalent_space_sym by blast lemma homotopy_eqv_rel_frontier_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ rel_interior S" shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})" apply (rule homotopy_eqv_rel_frontier_punctured_convex) using assms rel_frontier_affine_hull by force+ lemma path_connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "path_connected(rel_frontier S)" proof (cases "rel_interior S = {}") case True then show ?thesis by (simp add: \convex S\ convex_imp_path_connected rel_frontier_def) next case False then show ?thesis by (metis aff_dim_affine_hull affine_affine_hull affine_imp_convex all_not_in_conv assms path_connected_punctured_convex rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected) qed lemma connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "connected(rel_frontier S)" by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) subsubsection\Borsuk-style characterization of separation\ lemma continuous_on_Borsuk_map: "a \ s \ continuous_on s (\x. inverse(norm (x - a)) *\<^sub>R (x - a))" by (rule continuous_intros | force)+ lemma Borsuk_map_into_sphere: "(\x. inverse(norm (x - a)) *\<^sub>R (x - a)) ` s \ sphere 0 1 \ (a \ s)" by auto (metis eq_iff_diff_eq_0 left_inverse norm_eq_zero) lemma Borsuk_maps_homotopic_in_path_component: assumes "path_component (- s) a b" shows "homotopic_with_canon (\x. True) s (sphere 0 1) (\x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\x. inverse(norm(x - b)) *\<^sub>R (x - b))" proof - obtain g where "path g" "path_image g \ -s" "pathstart g = a" "pathfinish g = b" using assms by (auto simp: path_component_def) then show ?thesis apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def) apply (rule_tac x = "\z. inverse(norm(snd z - (g \ fst)z)) *\<^sub>R (snd z - (g \ fst)z)" in exI) apply (intro conjI continuous_intros) apply (rule continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+ done qed lemma non_extensible_Borsuk_map: fixes a :: "'a :: euclidean_space" assumes "compact s" and cin: "c \ components(- s)" and boc: "bounded c" and "a \ c" shows "\ (\g. continuous_on (s \ c) g \ g ` (s \ c) \ sphere 0 1 \ (\x \ s. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))" proof - have "closed s" using assms by (simp add: compact_imp_closed) have "c \ -s" using assms by (simp add: in_components_subset) with \a \ c\ have "a \ s" by blast then have ceq: "c = connected_component_set (- s) a" by (metis \a \ c\ cin components_iff connected_component_eq) then have "bounded (s \ connected_component_set (- s) a)" using \compact s\ boc compact_imp_bounded by auto with bounded_subset_ballD obtain r where "0 < r" and r: "(s \ connected_component_set (- s) a) \ ball a r" by blast { fix g assume "continuous_on (s \ c) g" "g ` (s \ c) \ sphere 0 1" and [simp]: "\x. x \ s \ g x = (x - a) /\<^sub>R norm (x - a)" then have [simp]: "\x. x \ s \ c \ norm (g x) = 1" by force have cb_eq: "cball a r = (s \ connected_component_set (- s) a) \ (cball a r - connected_component_set (- s) a)" using ball_subset_cball [of a r] r by auto have cont1: "continuous_on (s \ connected_component_set (- s) a) (\x. a + r *\<^sub>R g x)" apply (rule continuous_intros)+ using \continuous_on (s \ c) g\ ceq by blast have cont2: "continuous_on (cball a r - connected_component_set (- s) a) (\x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" by (rule continuous_intros | force simp: \a \ s\)+ have 1: "continuous_on (cball a r) (\x. if connected_component (- s) a x then a + r *\<^sub>R g x else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" apply (subst cb_eq) apply (rule continuous_on_cases [OF _ _ cont1 cont2]) using ceq cin apply (auto intro: closed_Un_complement_component simp: \closed s\ open_Compl open_connected_component) done have 2: "(\x. a + r *\<^sub>R g x) ` (cball a r \ connected_component_set (- s) a) \ sphere a r " using \0 < r\ by (force simp: dist_norm ceq) have "retraction (cball a r) (sphere a r) (\x. if x \ connected_component_set (- s) a then a + r *\<^sub>R g x else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" using \0 < r\ apply (simp add: retraction_def dist_norm 1 2, safe) apply (force simp: dist_norm abs_if mult_less_0_iff divide_simps \a \ s\) using r by (auto simp: dist_norm norm_minus_commute) then have False using no_retraction_cball [OF \0 < r\, of a, unfolded retract_of_def, simplified, rule_format, of "\x. if x \ connected_component_set (- s) a then a + r *\<^sub>R g x else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"] by blast } then show ?thesis by blast qed subsubsection \Proving surjectivity via Brouwer fixpoint theorem\ lemma brouwer_surjective: fixes f :: "'n::euclidean_space \ 'n" assumes "compact T" and "convex T" and "T \ {}" and "continuous_on T f" and "\x y. \x\S; y\T\ \ x + (y - f y) \ T" and "x \ S" shows "\y\T. f y = x" proof - have *: "\x y. f y = x \ x + (y - f y) = y" by (auto simp add: algebra_simps) show ?thesis unfolding * apply (rule brouwer[OF assms(1-3), of "\y. x + (y - f y)"]) apply (intro continuous_intros) using assms apply auto done qed lemma brouwer_surjective_cball: fixes f :: "'n::euclidean_space \ 'n" assumes "continuous_on (cball a e) f" and "e > 0" and "x \ S" and "\x y. \x\S; y\cball a e\ \ x + (y - f y) \ cball a e" shows "\y\cball a e. f y = x" apply (rule brouwer_surjective) apply (rule compact_cball convex_cball)+ unfolding cball_eq_empty using assms apply auto done subsubsection \Inverse function theorem\ text \See Sussmann: "Multidifferential calculus", Theorem 2.1.1\ lemma sussmann_open_mapping: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "open S" and contf: "continuous_on S f" and "x \ S" and derf: "(f has_derivative f') (at x)" and "bounded_linear g'" "f' \ g' = id" and "T \ S" and x: "x \ interior T" shows "f x \ interior (f ` T)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain B where B: "0 < B" "\x. norm (g' x) \ norm x * B" using bounded_linear.pos_bounded[OF assms(5)] by blast hence *: "1 / (2 * B) > 0" by auto obtain e0 where e0: "0 < e0" "\y. norm (y - x) < e0 \ norm (f y - f x - f' (y - x)) \ 1 / (2 * B) * norm (y - x)" using derf unfolding has_derivative_at_alt using * by blast obtain e1 where e1: "0 < e1" "cball x e1 \ T" using mem_interior_cball x by blast have *: "0 < e0 / B" "0 < e1 / B" using e0 e1 B by auto obtain e where e: "0 < e" "e < e0 / B" "e < e1 / B" using field_lbound_gt_zero[OF *] by blast have lem: "\y\cball (f x) e. f (x + g' (y - f x)) = z" if "z\cball (f x) (e / 2)" for z proof (rule brouwer_surjective_cball) have z: "z \ S" if as: "y \cball (f x) e" "z = x + (g' y - g' (f x))" for y z proof- have "dist x z = norm (g' (f x) - g' y)" unfolding as(2) and dist_norm by auto also have "\ \ norm (f x - y) * B" by (metis B(2) g'.diff) also have "\ \ e * B" by (metis B(1) dist_norm mem_cball real_mult_le_cancel_iff1 that(1)) also have "\ \ e1" using B(1) e(3) pos_less_divide_eq by fastforce finally have "z \ cball x e1" by force then show "z \ S" using e1 assms(7) by auto qed show "continuous_on (cball (f x) e) (\y. f (x + g' (y - f x)))" unfolding g'.diff proof (rule continuous_on_compose2 [OF _ _ order_refl, of _ _ f]) show "continuous_on ((\y. x + (g' y - g' (f x))) ` cball (f x) e) f" by (rule continuous_on_subset[OF contf]) (use z in blast) show "continuous_on (cball (f x) e) (\y. x + (g' y - g' (f x)))" by (intro continuous_intros linear_continuous_on[OF \bounded_linear g'\]) qed next fix y z assume y: "y \ cball (f x) (e / 2)" and z: "z \ cball (f x) e" have "norm (g' (z - f x)) \ norm (z - f x) * B" using B by auto also have "\ \ e * B" by (metis B(1) z dist_norm mem_cball norm_minus_commute real_mult_le_cancel_iff1) also have "\ < e0" using B(1) e(2) pos_less_divide_eq by blast finally have *: "norm (x + g' (z - f x) - x) < e0" by auto have **: "f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto have "norm (f x - (y + (z - f (x + g' (z - f x))))) \ norm (f (x + g' (z - f x)) - z) + norm (f x - y)" using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by (auto simp add: algebra_simps) also have "\ \ 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0(2)[rule_format, OF *] by (simp only: algebra_simps **) auto also have "\ \ 1 / (B * 2) * norm (g' (z - f x)) + e/2" using y by (auto simp: dist_norm) also have "\ \ 1 / (B * 2) * B * norm (z - f x) + e/2" using * B by (auto simp add: field_simps) also have "\ \ 1 / 2 * norm (z - f x) + e/2" by auto also have "\ \ e/2 + e/2" using B(1) \norm (z - f x) * B \ e * B\ by auto finally show "y + (z - f (x + g' (z - f x))) \ cball (f x) e" by (auto simp: dist_norm) qed (use e that in auto) show ?thesis unfolding mem_interior proof (intro exI conjI subsetI) fix y assume "y \ ball (f x) (e / 2)" then have *: "y \ cball (f x) (e / 2)" by auto obtain z where z: "z \ cball (f x) e" "f (x + g' (z - f x)) = y" using lem * by blast then have "norm (g' (z - f x)) \ norm (z - f x) * B" using B by (auto simp add: field_simps) also have "\ \ e * B" by (metis B(1) dist_norm mem_cball norm_minus_commute real_mult_le_cancel_iff1 z(1)) also have "\ \ e1" using e B unfolding less_divide_eq by auto finally have "x + g'(z - f x) \ T" by (metis add_diff_cancel diff_diff_add dist_norm e1(2) mem_cball norm_minus_commute subset_eq) then show "y \ f ` T" using z by auto qed (use e in auto) qed text \Hence the following eccentric variant of the inverse function theorem. This has no continuity assumptions, but we do need the inverse function. We could put \f' \ g = I\ but this happens to fit with the minimal linear algebra theory I've set up so far.\ lemma has_derivative_inverse_strong: fixes f :: "'n::euclidean_space \ 'n" assumes "open S" and "x \ S" and contf: "continuous_on S f" and gf: "\x. x \ S \ g (f x) = x" and derf: "(f has_derivative f') (at x)" and id: "f' \ g' = id" shows "(g has_derivative g') (at (f x))" proof - have linf: "bounded_linear f'" using derf unfolding has_derivative_def by auto then have ling: "bounded_linear g'" unfolding linear_conv_bounded_linear[symmetric] using id right_inverse_linear by blast moreover have "g' \ f' = id" using id linf ling unfolding linear_conv_bounded_linear[symmetric] using linear_inverse_left by auto moreover have *: "\T. \T \ S; x \ interior T\ \ f x \ interior (f ` T)" apply (rule sussmann_open_mapping) apply (rule assms ling)+ apply auto done have "continuous (at (f x)) g" unfolding continuous_at Lim_at proof (rule, rule) fix e :: real assume "e > 0" then have "f x \ interior (f ` (ball x e \ S))" using *[rule_format,of "ball x e \ S"] \x \ S\ by (auto simp add: interior_open[OF open_ball] interior_open[OF assms(1)]) then obtain d where d: "0 < d" "ball (f x) d \ f ` (ball x e \ S)" unfolding mem_interior by blast show "\d>0. \y. 0 < dist y (f x) \ dist y (f x) < d \ dist (g y) (g (f x)) < e" proof (intro exI allI impI conjI) fix y assume "0 < dist y (f x) \ dist y (f x) < d" then have "g y \ g ` f ` (ball x e \ S)" by (metis d(2) dist_commute mem_ball rev_image_eqI subset_iff) then show "dist (g y) (g (f x)) < e" using gf[OF \x \ S\] by (simp add: assms(4) dist_commute image_iff) qed (use d in auto) qed moreover have "f x \ interior (f ` S)" apply (rule sussmann_open_mapping) apply (rule assms ling)+ using interior_open[OF assms(1)] and \x \ S\ apply auto done moreover have "f (g y) = y" if "y \ interior (f ` S)" for y by (metis gf imageE interiorE subsetD that) ultimately show ?thesis using assms by (metis has_derivative_inverse_basic_x open_interior) qed text \A rewrite based on the other domain.\ lemma has_derivative_inverse_strong_x: fixes f :: "'a::euclidean_space \ 'a" assumes "open S" and "g y \ S" and "continuous_on S f" and "\x. x \ S \ g (f x) = x" and "(f has_derivative f') (at (g y))" and "f' \ g' = id" and "f (g y) = y" shows "(g has_derivative g') (at y)" using has_derivative_inverse_strong[OF assms(1-6)] unfolding assms(7) by simp text \On a region.\ theorem has_derivative_inverse_on: fixes f :: "'n::euclidean_space \ 'n" assumes "open S" and derf: "\x. x \ S \ (f has_derivative f'(x)) (at x)" and "\x. x \ S \ g (f x) = x" and "f' x \ g' x = id" and "x \ S" shows "(g has_derivative g'(x)) (at (f x))" proof (rule has_derivative_inverse_strong[where g'="g' x" and f=f]) show "continuous_on S f" unfolding continuous_on_eq_continuous_at[OF \open S\] using derf has_derivative_continuous by blast qed (use assms in auto) end diff --git a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy @@ -1,552 +1,552 @@ (* Title: HOL/Analysis/Cartesian_Euclidean_Space.thy Some material by Jose Divasón, Tim Makarios and L C Paulson *) section \Finite Cartesian Products of Euclidean Spaces\ theory Cartesian_Euclidean_Space imports Convex_Euclidean_Space Derivative begin lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}" by (simp add: subspace_def) lemma sum_mult_product: "sum h {..i\{..j\{..j. j + i * B) {..j. j + i * B) ` {.. {i * B.. (\j. j + i * B) ` {.. cbox a b = {}" "box (vec a) (vec b) = {} \ box a b = {}" by (auto simp: Basis_vec_def mem_box box_eq_empty inner_axis) subsection\Closures and interiors of halfspaces\ lemma interior_halfspace_component_le [simp]: "interior {x. x$k \ a} = {x :: (real^'n). x$k < a}" (is "?LE") and interior_halfspace_component_ge [simp]: "interior {x. x$k \ a} = {x :: (real^'n). x$k > a}" (is "?GE") proof - have "axis k (1::real) \ 0" by (simp add: axis_def vec_eq_iff) moreover have "axis k (1::real) \ x = x$k" for x by (simp add: cart_eq_inner_axis inner_commute) ultimately show ?LE ?GE using interior_halfspace_le [of "axis k (1::real)" a] interior_halfspace_ge [of "axis k (1::real)" a] by auto qed lemma closure_halfspace_component_lt [simp]: "closure {x. x$k < a} = {x :: (real^'n). x$k \ a}" (is "?LE") and closure_halfspace_component_gt [simp]: "closure {x. x$k > a} = {x :: (real^'n). x$k \ a}" (is "?GE") proof - have "axis k (1::real) \ 0" by (simp add: axis_def vec_eq_iff) moreover have "axis k (1::real) \ x = x$k" for x by (simp add: cart_eq_inner_axis inner_commute) ultimately show ?LE ?GE using closure_halfspace_lt [of "axis k (1::real)" a] closure_halfspace_gt [of "axis k (1::real)" a] by auto qed lemma interior_standard_hyperplane: "interior {x :: (real^'n). x$k = a} = {}" proof - have "axis k (1::real) \ 0" by (simp add: axis_def vec_eq_iff) moreover have "axis k (1::real) \ x = x$k" for x by (simp add: cart_eq_inner_axis inner_commute) ultimately show ?thesis using interior_hyperplane [of "axis k (1::real)" a] by force qed lemma matrix_vector_mul_bounded_linear[intro, simp]: "bounded_linear ((*v) A)" for A :: "'a::{euclidean_space,real_algebra_1}^'n^'m" using matrix_vector_mul_linear[of A] by (simp add: linear_conv_bounded_linear linear_matrix_vector_mul_eq) lemma fixes A :: "'a::{euclidean_space,real_algebra_1}^'n^'m" shows matrix_vector_mult_linear_continuous_at [continuous_intros]: "isCont ((*v) A) z" and matrix_vector_mult_linear_continuous_on [continuous_intros]: "continuous_on S ((*v) A)" by (simp_all add: linear_continuous_at linear_continuous_on) subsection\Bounds on components etc.\ relative to operator norm\ lemma norm_column_le_onorm: fixes A :: "real^'n^'m" shows "norm(column i A) \ onorm((*v) A)" proof - have "norm (\ j. A $ j $ i) \ norm (A *v axis i 1)" by (simp add: matrix_mult_dot cart_eq_inner_axis) also have "\ \ onorm ((*v) A)" using onorm [OF matrix_vector_mul_bounded_linear, of A "axis i 1"] by auto finally have "norm (\ j. A $ j $ i) \ onorm ((*v) A)" . then show ?thesis unfolding column_def . qed lemma matrix_component_le_onorm: fixes A :: "real^'n^'m" shows "\A $ i $ j\ \ onorm((*v) A)" proof - have "\A $ i $ j\ \ norm (\ n. (A $ n $ j))" by (metis (full_types, lifting) component_le_norm_cart vec_lambda_beta) also have "\ \ onorm ((*v) A)" by (metis (no_types) column_def norm_column_le_onorm) finally show ?thesis . qed lemma component_le_onorm: fixes f :: "real^'m \ real^'n" shows "linear f \ \matrix f $ i $ j\ \ onorm f" - by (metis linear_matrix_vector_mul_eq matrix_component_le_onorm matrix_vector_mul) + by (metis matrix_component_le_onorm matrix_vector_mul(2)) lemma onorm_le_matrix_component_sum: fixes A :: "real^'n^'m" shows "onorm((*v) A) \ (\i\UNIV. \j\UNIV. \A $ i $ j\)" proof (rule onorm_le) fix x have "norm (A *v x) \ (\i\UNIV. \(A *v x) $ i\)" by (rule norm_le_l1_cart) also have "\ \ (\i\UNIV. \j\UNIV. \A $ i $ j\ * norm x)" proof (rule sum_mono) fix i have "\(A *v x) $ i\ \ \\j\UNIV. A $ i $ j * x $ j\" by (simp add: matrix_vector_mult_def) also have "\ \ (\j\UNIV. \A $ i $ j * x $ j\)" by (rule sum_abs) also have "\ \ (\j\UNIV. \A $ i $ j\ * norm x)" by (rule sum_mono) (simp add: abs_mult component_le_norm_cart mult_left_mono) finally show "\(A *v x) $ i\ \ (\j\UNIV. \A $ i $ j\ * norm x)" . qed finally show "norm (A *v x) \ (\i\UNIV. \j\UNIV. \A $ i $ j\) * norm x" by (simp add: sum_distrib_right) qed lemma onorm_le_matrix_component: fixes A :: "real^'n^'m" assumes "\i j. abs(A$i$j) \ B" shows "onorm((*v) A) \ real (CARD('m)) * real (CARD('n)) * B" proof (rule onorm_le) fix x :: "real^'n::_" have "norm (A *v x) \ (\i\UNIV. \(A *v x) $ i\)" by (rule norm_le_l1_cart) also have "\ \ (\i::'m \UNIV. real (CARD('n)) * B * norm x)" proof (rule sum_mono) fix i have "\(A *v x) $ i\ \ norm(A $ i) * norm x" by (simp add: matrix_mult_dot Cauchy_Schwarz_ineq2) also have "\ \ (\j\UNIV. \A $ i $ j\) * norm x" by (simp add: mult_right_mono norm_le_l1_cart) also have "\ \ real (CARD('n)) * B * norm x" by (simp add: assms sum_bounded_above mult_right_mono) finally show "\(A *v x) $ i\ \ real (CARD('n)) * B * norm x" . qed also have "\ \ CARD('m) * real (CARD('n)) * B * norm x" by simp finally show "norm (A *v x) \ CARD('m) * real (CARD('n)) * B * norm x" . qed lemma rational_approximation: assumes "e > 0" obtains r::real where "r \ \" "\r - x\ < e" using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto proposition matrix_rational_approximation: fixes A :: "real^'n^'m" assumes "e > 0" obtains B where "\i j. B$i$j \ \" "onorm(\x. (A - B) *v x) < e" proof - have "\i j. \q \ \. \q - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"]) then obtain B where B: "\i j. B$i$j \ \" and Bclo: "\i j. \B$i$j - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" by (auto simp: lambda_skolem Bex_def) show ?thesis proof have "onorm ((*v) (A - B)) \ real CARD('m) * real CARD('n) * (e / (2 * real CARD('m) * real CARD('n)))" apply (rule onorm_le_matrix_component) using Bclo by (simp add: abs_minus_commute less_imp_le) also have "\ < e" using \0 < e\ by (simp add: field_split_simps) finally show "onorm ((*v) (A - B)) < e" . qed (use B in auto) qed lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \ (x - ((b \ x) / (b \ b)) *s b) = 0" unfolding inner_simps scalar_mult_eq_scaleR by auto lemma infnorm_cart:"infnorm (x::real^'n) = Sup {\x$i\ |i. i\UNIV}" by (simp add: infnorm_def inner_axis Basis_vec_def) (metis (lifting) inner_axis real_inner_1_right) lemma component_le_infnorm_cart: "\x$i\ \ infnorm (x::real^'n)" using Basis_le_infnorm[of "axis i 1" x] by (simp add: Basis_vec_def axis_eq_axis inner_axis) lemma continuous_component[continuous_intros]: "continuous F f \ continuous F (\x. f x $ i)" unfolding continuous_def by (rule tendsto_vec_nth) lemma continuous_on_component[continuous_intros]: "continuous_on s f \ continuous_on s (\x. f x $ i)" unfolding continuous_on_def by (fast intro: tendsto_vec_nth) lemma continuous_on_vec_lambda[continuous_intros]: "(\i. continuous_on S (f i)) \ continuous_on S (\x. \ i. f i x)" unfolding continuous_on_def by (auto intro: tendsto_vec_lambda) lemma closed_positive_orthant: "closed {x::real^'n. \i. 0 \x$i}" - by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_component) lemma bounded_component_cart: "bounded s \ bounded ((\x. x $ i) ` s)" unfolding bounded_def apply clarify apply (rule_tac x="x $ i" in exI) apply (rule_tac x="e" in exI) apply clarify apply (rule order_trans [OF dist_vec_nth_le], simp) done lemma compact_lemma_cart: fixes f :: "nat \ 'a::heine_borel ^ 'n" assumes f: "bounded (range f)" shows "\l r. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) $ i) (l $ i) < e) sequentially)" (is "?th d") proof - have "\d' \ d. ?th d'" by (rule compact_lemma_general[where unproj=vec_lambda]) - (auto intro!: f bounded_component_cart simp: vec_lambda_eta) + (auto intro!: f bounded_component_cart) then show "?th d" by simp qed instance vec :: (heine_borel, finite) heine_borel proof fix f :: "nat \ 'a ^ 'b" assume f: "bounded (range f)" then obtain l r where r: "strict_mono r" and l: "\e>0. eventually (\n. \i\UNIV. dist (f (r n) $ i) (l $ i) < e) sequentially" using compact_lemma_cart [OF f] by blast let ?d = "UNIV::'b set" { fix e::real assume "e>0" hence "0 < e / (real_of_nat (card ?d))" using zero_less_card_finite divide_pos_pos[of e, of "real_of_nat (card ?d)"] by auto with l have "eventually (\n. \i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))) sequentially" by simp moreover { fix n assume n: "\i. dist (f (r n) $ i) (l $ i) < e / (real_of_nat (card ?d))" have "dist (f (r n)) l \ (\i\?d. dist (f (r n) $ i) (l $ i))" unfolding dist_vec_def using zero_le_dist by (rule L2_set_le_sum) also have "\ < (\i\?d. e / (real_of_nat (card ?d)))" by (rule sum_strict_mono) (simp_all add: n) finally have "dist (f (r n)) l < e" by simp } ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" by (rule eventually_mono) } hence "((f \ r) \ l) sequentially" unfolding o_def tendsto_iff by simp with r show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" by auto qed lemma interval_cart: fixes a :: "real^'n" shows "box a b = {x::real^'n. \i. a$i < x$i \ x$i < b$i}" and "cbox a b = {x::real^'n. \i. a$i \ x$i \ x$i \ b$i}" by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def mem_box Basis_vec_def inner_axis) lemma mem_box_cart: fixes a :: "real^'n" shows "x \ box a b \ (\i. a$i < x$i \ x$i < b$i)" and "x \ cbox a b \ (\i. a$i \ x$i \ x$i \ b$i)" using interval_cart[of a b] by (auto simp add: set_eq_iff less_vec_def less_eq_vec_def) lemma interval_eq_empty_cart: fixes a :: "real^'n" shows "(box a b = {} \ (\i. b$i \ a$i))" (is ?th1) and "(cbox a b = {} \ (\i. b$i < a$i))" (is ?th2) proof - { fix i x assume as:"b$i \ a$i" and x:"x\box a b" hence "a $ i < x $ i \ x $ i < b $ i" unfolding mem_box_cart by auto hence "a$i < b$i" by auto hence False using as by auto } moreover { assume as:"\i. \ (b$i \ a$i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i have "a$i < b$i" using as[THEN spec[where x=i]] by auto hence "a$i < ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i < b$i" unfolding vector_smult_component and vector_add_component by auto } hence "box a b \ {}" using mem_box_cart(1)[of "?x" a b] by auto } ultimately show ?th1 by blast { fix i x assume as:"b$i < a$i" and x:"x\cbox a b" hence "a $ i \ x $ i \ x $ i \ b $ i" unfolding mem_box_cart by auto hence "a$i \ b$i" by auto hence False using as by auto } moreover { assume as:"\i. \ (b$i < a$i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i have "a$i \ b$i" using as[THEN spec[where x=i]] by auto hence "a$i \ ((1/2) *\<^sub>R (a+b)) $ i" "((1/2) *\<^sub>R (a+b)) $ i \ b$i" unfolding vector_smult_component and vector_add_component by auto } hence "cbox a b \ {}" using mem_box_cart(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed lemma interval_ne_empty_cart: fixes a :: "real^'n" shows "cbox a b \ {} \ (\i. a$i \ b$i)" and "box a b \ {} \ (\i. a$i < b$i)" unfolding interval_eq_empty_cart[of a b] by (auto simp add: not_less not_le) (* BH: Why doesn't just "auto" work here? *) lemma subset_interval_imp_cart: fixes a :: "real^'n" shows "(\i. a$i \ c$i \ d$i \ b$i) \ cbox c d \ cbox a b" and "(\i. a$i < c$i \ d$i < b$i) \ cbox c d \ box a b" and "(\i. a$i \ c$i \ d$i \ b$i) \ box c d \ cbox a b" and "(\i. a$i \ c$i \ d$i \ b$i) \ box c d \ box a b" unfolding subset_eq[unfolded Ball_def] unfolding mem_box_cart by (auto intro: order_trans less_le_trans le_less_trans less_imp_le) (* BH: Why doesn't just "auto" work here? *) lemma interval_sing: fixes a :: "'a::linorder^'n" shows "{a .. a} = {a} \ {a<.. cbox a b \ (\i. c$i \ d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th1) and "cbox c d \ box a b \ (\i. c$i \ d$i) --> (\i. a$i < c$i \ d$i < b$i)" (is ?th2) and "box c d \ cbox a b \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th3) and "box c d \ box a b \ (\i. c$i < d$i) --> (\i. a$i \ c$i \ d$i \ b$i)" (is ?th4) using subset_box[of c d a b] by (simp_all add: Basis_vec_def inner_axis) lemma disjoint_interval_cart: fixes a::"real^'n" shows "cbox a b \ cbox c d = {} \ (\i. (b$i < a$i \ d$i < c$i \ b$i < c$i \ d$i < a$i))" (is ?th1) and "cbox a b \ box c d = {} \ (\i. (b$i < a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th2) and "box a b \ cbox c d = {} \ (\i. (b$i \ a$i \ d$i < c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th3) and "box a b \ box c d = {} \ (\i. (b$i \ a$i \ d$i \ c$i \ b$i \ c$i \ d$i \ a$i))" (is ?th4) using disjoint_interval[of a b c d] by (simp_all add: Basis_vec_def inner_axis) lemma Int_interval_cart: fixes a :: "real^'n" shows "cbox a b \ cbox c d = {(\ i. max (a$i) (c$i)) .. (\ i. min (b$i) (d$i))}" unfolding Int_interval by (auto simp: mem_box less_eq_vec_def) (auto simp: Basis_vec_def inner_axis) lemma closed_interval_left_cart: fixes b :: "real^'n" shows "closed {x::real^'n. \i. x$i \ b$i}" - by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_component) lemma closed_interval_right_cart: fixes a::"real^'n" shows "closed {x::real^'n. \i. a$i \ x$i}" - by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_const continuous_on_id continuous_on_component) + by (simp add: Collect_all_eq closed_INT closed_Collect_le continuous_on_component) lemma is_interval_cart: "is_interval (s::(real^'n) set) \ (\a\s. \b\s. \x. (\i. ((a$i \ x$i \ x$i \ b$i) \ (b$i \ x$i \ x$i \ a$i))) \ x \ s)" by (simp add: is_interval_def Ball_def Basis_vec_def inner_axis imp_ex) lemma closed_halfspace_component_le_cart: "closed {x::real^'n. x$i \ a}" by (simp add: closed_Collect_le continuous_on_component) lemma closed_halfspace_component_ge_cart: "closed {x::real^'n. x$i \ a}" by (simp add: closed_Collect_le continuous_on_component) lemma open_halfspace_component_lt_cart: "open {x::real^'n. x$i < a}" by (simp add: open_Collect_less continuous_on_component) lemma open_halfspace_component_gt_cart: "open {x::real^'n. x$i > a}" by (simp add: open_Collect_less continuous_on_component) lemma Lim_component_le_cart: fixes f :: "'a \ real^'n" assumes "(f \ l) net" "\ (trivial_limit net)" "eventually (\x. f x $i \ b) net" shows "l$i \ b" by (rule tendsto_le[OF assms(2) tendsto_const tendsto_vec_nth, OF assms(1, 3)]) lemma Lim_component_ge_cart: fixes f :: "'a \ real^'n" assumes "(f \ l) net" "\ (trivial_limit net)" "eventually (\x. b \ (f x)$i) net" shows "b \ l$i" by (rule tendsto_le[OF assms(2) tendsto_vec_nth tendsto_const, OF assms(1, 3)]) lemma Lim_component_eq_cart: fixes f :: "'a \ real^'n" assumes net: "(f \ l) net" "\ trivial_limit net" and ev:"eventually (\x. f(x)$i = b) net" shows "l$i = b" using ev[unfolded order_eq_iff eventually_conj_iff] and Lim_component_ge_cart[OF net, of b i] and Lim_component_le_cart[OF net, of i b] by auto lemma connected_ivt_component_cart: fixes x :: "real^'n" shows "connected s \ x \ s \ y \ s \ x$k \ a \ a \ y$k \ (\z\s. z$k = a)" using connected_ivt_hyperplane[of s x y "axis k 1" a] by (auto simp add: inner_axis inner_commute) lemma subspace_substandard_cart: "vec.subspace {x. (\i. P i \ x$i = 0)}" unfolding vec.subspace_def by auto lemma closed_substandard_cart: "closed {x::'a::real_normed_vector ^ 'n. \i. P i \ x$i = 0}" proof - { fix i::'n have "closed {x::'a ^ 'n. P i \ x$i = 0}" - by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_const continuous_on_id continuous_on_component) } + by (cases "P i") (simp_all add: closed_Collect_eq continuous_on_component) } thus ?thesis unfolding Collect_all_eq by (simp add: closed_INT) qed subsection "Convex Euclidean Space" lemma Cart_1:"(1::real^'n) = \Basis" using const_vector_cart[of 1] by (simp add: one_vec_def) declare vector_add_ldistrib[simp] vector_ssub_ldistrib[simp] vector_smult_assoc[simp] vector_smult_rneg[simp] declare vector_sadd_rdistrib[simp] vector_sub_rdistrib[simp] lemmas vector_component_simps = vector_minus_component vector_smult_component vector_add_component less_eq_vec_def vec_lambda_beta vector_uminus_component lemma convex_box_cart: assumes "\i. convex {x. P i x}" shows "convex {x. \i. P i (x$i)}" using assms unfolding convex_def by auto lemma convex_positive_orthant_cart: "convex {x::real^'n. (\i. 0 \ x$i)}" by (rule convex_box_cart) (simp add: atLeast_def[symmetric]) lemma unit_interval_convex_hull_cart: "cbox (0::real^'n) 1 = convex hull {x. \i. (x$i = 0) \ (x$i = 1)}" unfolding Cart_1 unit_interval_convex_hull[where 'a="real^'n"] box_real[symmetric] by (rule arg_cong[where f="\x. convex hull x"]) (simp add: Basis_vec_def inner_axis) proposition cube_convex_hull_cart: assumes "0 < d" obtains s::"(real^'n) set" where "finite s" "cbox (x - (\ i. d)) (x + (\ i. d)) = convex hull s" proof - from assms obtain s where "finite s" and "cbox (x - sum ((*\<^sub>R) d) Basis) (x + sum ((*\<^sub>R) d) Basis) = convex hull s" by (rule cube_convex_hull) with that[of s] show thesis by (simp add: const_vector_cart) qed subsection "Derivative" definition\<^marker>\tag important\ "jacobian f net = matrix(frechet_derivative f net)" proposition jacobian_works: "(f::(real^'a) \ (real^'b)) differentiable net \ (f has_derivative (\h. (jacobian f net) *v h)) net" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: frechet_derivative_works has_derivative_linear jacobian_def) next assume ?rhs then show ?lhs by (rule differentiableI) qed text \Component of the differential must be zero if it exists at a local maximum or minimum for that corresponding component\ proposition differential_zero_maxmin_cart: fixes f::"real^'a \ real^'b" assumes "0 < e" "((\y \ ball x e. (f y)$k \ (f x)$k) \ (\y\ball x e. (f x)$k \ (f y)$k))" "f differentiable (at x)" shows "jacobian f (at x) $ k = 0" using differential_zero_maxmin_component[of "axis k 1" e x f] assms vector_cart[of "\j. frechet_derivative f (at x) j $ k"] by (simp add: Basis_vec_def axis_eq_axis inner_axis jacobian_def matrix_def) subsection\<^marker>\tag unimportant\\Routine results connecting the types \<^typ>\real^1\ and \<^typ>\real\\ lemma vec_cbox_1_eq [simp]: shows "vec ` cbox u v = cbox (vec u) (vec v ::real^1)" by (force simp: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box) lemma vec_nth_cbox_1_eq [simp]: fixes u v :: "'a::euclidean_space^1" shows "(\x. x $ 1) ` cbox u v = cbox (u$1) (v$1)" by (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric] mem_box image_iff Bex_def inner_axis) (metis vec_component) lemma vec_nth_1_iff_cbox [simp]: fixes a b :: "'a::euclidean_space" shows "(\x::'a^1. x $ 1) ` S = cbox a b \ S = cbox (vec a) (vec b)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (intro equalityI subsetI) fix x assume "x \ S" then have "x $ 1 \ (\v. v $ (1::1)) ` cbox (vec a) (vec b)" using L by auto then show "x \ cbox (vec a) (vec b)" by (metis (no_types, lifting) imageE vector_one_nth) next fix x :: "'a^1" assume "x \ cbox (vec a) (vec b)" then show "x \ S" by (metis (no_types, lifting) L imageE imageI vec_component vec_nth_cbox_1_eq vector_one_nth) qed qed simp lemma vec_nth_real_1_iff_cbox [simp]: fixes a b :: real shows "(\x::real^1. x $ 1) ` S = {a..b} \ S = cbox (vec a) (vec b)" using vec_nth_1_iff_cbox[of S a b] by simp lemma interval_split_cart: "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" "cbox a b \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" apply (rule_tac[!] set_eqI) unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart unfolding vec_lambda_beta by auto lemmas cartesian_euclidean_space_uniform_limit_intros[uniform_limit_intros] = bounded_linear.uniform_limit[OF blinfun.bounded_linear_right] bounded_linear.uniform_limit[OF bounded_linear_vec_nth] end diff --git a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy @@ -1,7847 +1,7847 @@ section \Complex Path Integrals and Cauchy's Integral Theorem\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ theory Cauchy_Integral_Theorem imports Complex_Transcendental Henstock_Kurzweil_Integration Weierstrass_Theorems Retracts begin lemma leibniz_rule_holomorphic: fixes f::"complex \ 'b::euclidean_space \ complex" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "convex U" shows "(\x. integral (cbox a b) (f x)) holomorphic_on U" using leibniz_rule_field_differentiable[OF assms(1-3) _ assms(4)] by (auto simp: holomorphic_on_def) lemma Ln_measurable [measurable]: "Ln \ measurable borel borel" proof - have *: "Ln (-of_real x) = of_real (ln x) + \ * pi" if "x > 0" for x using that by (subst Ln_minus) (auto simp: Ln_of_real) have **: "Ln (of_real x) = of_real (ln (-x)) + \ * pi" if "x < 0" for x using *[of "-x"] that by simp have cont: "(\x. indicat_real (- \\<^sub>\\<^sub>0) x *\<^sub>R Ln x) \ borel_measurable borel" by (intro borel_measurable_continuous_on_indicator continuous_intros) auto have "(\x. if x \ \\<^sub>\\<^sub>0 then ln (-Re x) + \ * pi else indicator (-\\<^sub>\\<^sub>0) x *\<^sub>R Ln x) \ borel \\<^sub>M borel" (is "?f \ _") by (rule measurable_If_set[OF _ cont]) auto hence "(\x. if x = 0 then Ln 0 else ?f x) \ borel \\<^sub>M borel" by measurable also have "(\x. if x = 0 then Ln 0 else ?f x) = Ln" by (auto simp: fun_eq_iff ** nonpos_Reals_def) finally show ?thesis . qed lemma powr_complex_measurable [measurable]: assumes [measurable]: "f \ measurable M borel" "g \ measurable M borel" shows "(\x. f x powr g x :: complex) \ measurable M borel" using assms by (simp add: powr_def) subsection\<^marker>\tag unimportant\ \Homeomorphisms of arc images\ lemma homeomorphism_arc: fixes g :: "real \ 'a::t2_space" assumes "arc g" obtains h where "homeomorphism {0..1} (path_image g) g h" using assms by (force simp: arc_def homeomorphism_compact path_def path_image_def) lemma homeomorphic_arc_image_interval: fixes g :: "real \ 'a::t2_space" and a::real assumes "arc g" "a < b" shows "(path_image g) homeomorphic {a..b}" proof - have "(path_image g) homeomorphic {0..1::real}" by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc) also have "\ homeomorphic {a..b}" using assms by (force intro: homeomorphic_closed_intervals_real) finally show ?thesis . qed lemma homeomorphic_arc_images: fixes g :: "real \ 'a::t2_space" and h :: "real \ 'b::t2_space" assumes "arc g" "arc h" shows "(path_image g) homeomorphic (path_image h)" proof - have "(path_image g) homeomorphic {0..1::real}" by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc) also have "\ homeomorphic (path_image h)" by (meson assms homeomorphic_def homeomorphism_arc) 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\<^marker>\tag unimportant\ \Piecewise differentiable functions\ definition piecewise_differentiable_on (infixr "piecewise'_differentiable'_on" 50) where "f piecewise_differentiable_on i \ continuous_on i f \ (\S. finite S \ (\x \ i - S. f differentiable (at x within i)))" lemma piecewise_differentiable_on_imp_continuous_on: "f piecewise_differentiable_on S \ continuous_on S f" by (simp add: piecewise_differentiable_on_def) lemma piecewise_differentiable_on_subset: "f piecewise_differentiable_on S \ T \ S \ f piecewise_differentiable_on T" using continuous_on_subset unfolding piecewise_differentiable_on_def apply safe apply (blast elim: continuous_on_subset) by (meson Diff_iff differentiable_within_subset subsetCE) lemma differentiable_on_imp_piecewise_differentiable: fixes a:: "'a::{linorder_topology,real_normed_vector}" shows "f differentiable_on {a..b} \ f piecewise_differentiable_on {a..b}" apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on) apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def) done lemma differentiable_imp_piecewise_differentiable: "(\x. x \ S \ f differentiable (at x within S)) \ f piecewise_differentiable_on S" by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def intro: differentiable_within_subset) lemma piecewise_differentiable_const [iff]: "(\x. z) piecewise_differentiable_on S" by (simp add: differentiable_imp_piecewise_differentiable) lemma piecewise_differentiable_compose: "\f piecewise_differentiable_on S; g piecewise_differentiable_on (f ` S); \x. finite (S \ f-`{x})\ \ (g \ f) piecewise_differentiable_on S" apply (simp add: piecewise_differentiable_on_def, safe) apply (blast intro: continuous_on_compose2) apply (rename_tac A B) apply (rule_tac x="A \ (\x\B. S \ f-`{x})" in exI) apply (blast intro!: differentiable_chain_within) done lemma piecewise_differentiable_affine: fixes m::real assumes "f piecewise_differentiable_on ((\x. m *\<^sub>R x + c) ` S)" shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_differentiable_on S" proof (cases "m = 0") case True then show ?thesis unfolding o_def by (force intro: differentiable_imp_piecewise_differentiable differentiable_const) next case False show ?thesis apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable]) apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+ done qed lemma piecewise_differentiable_cases: fixes c::real assumes "f piecewise_differentiable_on {a..c}" "g piecewise_differentiable_on {c..b}" "a \ c" "c \ b" "f c = g c" shows "(\x. if x \ c then f x else g x) piecewise_differentiable_on {a..b}" proof - obtain S T where st: "finite S" "finite T" and fd: "\x. x \ {a..c} - S \ f differentiable at x within {a..c}" and gd: "\x. x \ {c..b} - T \ g differentiable at x within {c..b}" using assms by (auto simp: piecewise_differentiable_on_def) have finabc: "finite ({a,b,c} \ (S \ T))" by (metis \finite S\ \finite T\ finite_Un finite_insert finite.emptyI) have "continuous_on {a..c} f" "continuous_on {c..b} g" using assms piecewise_differentiable_on_def by auto then have "continuous_on {a..b} (\x. if x \ c then f x else g x)" using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], OF closed_real_atLeastAtMost [of c b], of f g "\x. x\c"] assms by (force simp: ivl_disj_un_two_touch) moreover { fix x assume x: "x \ {a..b} - ({a,b,c} \ (S \ T))" have "(\x. if x \ c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg") proof (cases x c rule: le_cases) case le show ?diff_fg proof (rule differentiable_transform_within [where d = "dist x c"]) have "f differentiable at x" using x le fd [of x] at_within_interior [of x "{a..c}"] by simp then show "f differentiable at x within {a..b}" by (simp add: differentiable_at_withinI) qed (use x le st dist_real_def in auto) next case ge show ?diff_fg proof (rule differentiable_transform_within [where d = "dist x c"]) have "g differentiable at x" using x ge gd [of x] at_within_interior [of x "{c..b}"] by simp then show "g differentiable at x within {a..b}" by (simp add: differentiable_at_withinI) qed (use x ge st dist_real_def in auto) qed } then have "\S. finite S \ (\x\{a..b} - S. (\x. if x \ c then f x else g x) differentiable at x within {a..b})" by (meson finabc) ultimately show ?thesis by (simp add: piecewise_differentiable_on_def) qed lemma piecewise_differentiable_neg: "f piecewise_differentiable_on S \ (\x. -(f x)) piecewise_differentiable_on S" by (auto simp: piecewise_differentiable_on_def continuous_on_minus) lemma piecewise_differentiable_add: assumes "f piecewise_differentiable_on i" "g piecewise_differentiable_on i" shows "(\x. f x + g x) piecewise_differentiable_on i" proof - obtain S T where st: "finite S" "finite T" "\x\i - S. f differentiable at x within i" "\x\i - T. g differentiable at x within i" using assms by (auto simp: piecewise_differentiable_on_def) then have "finite (S \ T) \ (\x\i - (S \ T). (\x. f x + g x) differentiable at x within i)" by auto moreover have "continuous_on i f" "continuous_on i g" using assms piecewise_differentiable_on_def by auto ultimately show ?thesis by (auto simp: piecewise_differentiable_on_def continuous_on_add) qed lemma piecewise_differentiable_diff: "\f piecewise_differentiable_on S; g piecewise_differentiable_on S\ \ (\x. f x - g x) piecewise_differentiable_on S" unfolding diff_conv_add_uminus by (metis piecewise_differentiable_add piecewise_differentiable_neg) lemma continuous_on_joinpaths_D1: "continuous_on {0..1} (g1 +++ g2) \ continuous_on {0..1} g1" apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ ((*)(inverse 2))"]) apply (rule continuous_intros | simp)+ apply (auto elim!: continuous_on_subset simp: joinpaths_def) done lemma continuous_on_joinpaths_D2: "\continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\ \ continuous_on {0..1} g2" apply (rule continuous_on_eq [of _ "(g1 +++ g2) \ (\x. inverse 2*x + 1/2)"]) apply (rule continuous_intros | simp)+ apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def) done lemma piecewise_differentiable_D1: assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" shows "g1 piecewise_differentiable_on {0..1}" proof - obtain S where cont: "continuous_on {0..1} g1" and "finite S" and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" using assms unfolding piecewise_differentiable_on_def by (blast dest!: continuous_on_joinpaths_D1) show ?thesis unfolding piecewise_differentiable_on_def proof (intro exI conjI ballI cont) show "finite (insert 1 (((*)2) ` S))" by (simp add: \finite S\) show "g1 differentiable at x within {0..1}" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x proof (rule_tac d="dist (x/2) (1/2)" in differentiable_transform_within) have "g1 +++ g2 differentiable at (x / 2) within {0..1/2}" by (rule differentiable_subset [OF S [of "x/2"]] | use that in force)+ then show "g1 +++ g2 \ (*) (inverse 2) differentiable at x within {0..1}" using image_affinity_atLeastAtMost_div [of 2 0 "0::real" 1] by (auto intro: differentiable_chain_within) qed (use that in \auto simp: joinpaths_def\) qed qed lemma piecewise_differentiable_D2: assumes "(g1 +++ g2) piecewise_differentiable_on {0..1}" and eq: "pathfinish g1 = pathstart g2" shows "g2 piecewise_differentiable_on {0..1}" proof - have [simp]: "g1 1 = g2 0" using eq by (simp add: pathfinish_def pathstart_def) obtain S where cont: "continuous_on {0..1} g2" and "finite S" and S: "\x. x \ {0..1} - S \ g1 +++ g2 differentiable at x within {0..1}" using assms unfolding piecewise_differentiable_on_def by (blast dest!: continuous_on_joinpaths_D2) show ?thesis unfolding piecewise_differentiable_on_def proof (intro exI conjI ballI cont) show "finite (insert 0 ((\x. 2*x-1)`S))" by (simp add: \finite S\) show "g2 differentiable at x within {0..1}" if "x \ {0..1} - insert 0 ((\x. 2*x-1)`S)" for x proof (rule_tac d="dist ((x+1)/2) (1/2)" in differentiable_transform_within) have x2: "(x + 1) / 2 \ S" using that apply (clarsimp simp: image_iff) by (metis add.commute add_diff_cancel_left' mult_2 field_sum_of_halves) have "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" by (rule differentiable_chain_within differentiable_subset [OF S [of "(x+1)/2"]] | use x2 that in force)+ then show "g1 +++ g2 \ (\x. (x+1) / 2) differentiable at x within {0..1}" by (auto intro: differentiable_chain_within) show "(g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" if "x' \ {0..1}" "dist x' x < dist ((x + 1) / 2) (1/2)" for x' proof - have [simp]: "(2*x'+2)/2 = x'+1" by (simp add: field_split_simps) show ?thesis using that by (auto simp: joinpaths_def) qed qed (use that in \auto simp: joinpaths_def\) qed qed subsection\The concept of continuously differentiable\ text \ John Harrison writes as follows: ``The usual assumption in complex analysis texts is that a path \\\ should be piecewise continuously differentiable, which ensures that the path integral exists at least for any continuous f, since all piecewise continuous functions are integrable. However, our notion of validity is weaker, just piecewise differentiability\ldots{} [namely] continuity plus differentiability except on a finite set\ldots{} [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this can integrate all derivatives.'' "Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" (infix "C1'_differentiable'_on" 50) where "f C1_differentiable_on S \ (\D. (\x \ S. (f has_vector_derivative (D x)) (at x)) \ continuous_on S D)" lemma C1_differentiable_on_eq: "f C1_differentiable_on S \ (\x \ S. f differentiable at x) \ continuous_on S (\x. vector_derivative f (at x))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding C1_differentiable_on_def by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at) next assume ?rhs then show ?lhs using C1_differentiable_on_def vector_derivative_works by fastforce qed lemma C1_differentiable_on_subset: "f C1_differentiable_on T \ S \ T \ f C1_differentiable_on S" unfolding C1_differentiable_on_def continuous_on_eq_continuous_within by (blast intro: continuous_within_subset) lemma C1_differentiable_compose: assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) C1_differentiable_on S" proof - have "\x. x \ S \ g \ f differentiable at x" by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI) moreover have "continuous_on S (\x. vector_derivative (g \ f) (at x))" proof (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) show "continuous_on S (\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))" using fg apply (clarsimp simp add: C1_differentiable_on_eq) apply (rule Limits.continuous_on_scaleR, assumption) by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def) show "\x. x \ S \ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \ f) (at x)" by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at) qed ultimately show ?thesis by (simp add: C1_differentiable_on_eq) qed lemma C1_diff_imp_diff: "f C1_differentiable_on S \ f differentiable_on S" by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on S" - by (auto simp: C1_differentiable_on_eq continuous_on_const) + by (auto simp: C1_differentiable_on_eq) lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" - by (auto simp: C1_differentiable_on_eq continuous_on_const) + by (auto simp: C1_differentiable_on_eq) lemma C1_differentiable_on_add [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_minus [simp, derivative_intros]: "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_diff [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_mult [simp, derivative_intros]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) lemma C1_differentiable_on_scaleR [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ definition\<^marker>\tag important\ piecewise_C1_differentiable_on (infixr "piecewise'_C1'_differentiable'_on" 50) where "f piecewise_C1_differentiable_on i \ continuous_on i f \ (\S. finite S \ (f C1_differentiable_on (i - S)))" lemma C1_differentiable_imp_piecewise: "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) lemma piecewise_C1_imp_differentiable: "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def C1_differentiable_on_def differentiable_def has_vector_derivative_def intro: has_derivative_at_withinI) lemma piecewise_C1_differentiable_compose: assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) piecewise_C1_differentiable_on S" proof - have "continuous_on S (\x. g (f x))" by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" proof - obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" using fg by (auto simp: piecewise_C1_differentiable_on_def) obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" using fg by (auto simp: piecewise_C1_differentiable_on_def) show ?thesis proof (intro exI conjI) show "finite (F \ (\x\G. S \ f-`{x}))" using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" apply (rule C1_differentiable_compose) apply (blast intro: C1_differentiable_on_subset [OF F]) apply (blast intro: C1_differentiable_on_subset [OF G]) by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) qed qed ultimately show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed lemma piecewise_C1_differentiable_on_subset: "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) lemma C1_differentiable_imp_continuous_on: "f C1_differentiable_on S \ continuous_on S f" unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within using differentiable_at_withinI differentiable_imp_continuous_within by blast lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" unfolding C1_differentiable_on_def by auto lemma piecewise_C1_differentiable_affine: fixes m::real assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" proof (cases "m = 0") case True then show ?thesis - unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const) + unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def) next case False have *: "\x. finite (S \ {y. m * y + c = x})" using False not_finite_existsD by fastforce show ?thesis apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) apply (rule * assms derivative_intros | simp add: False vimage_def)+ done qed lemma piecewise_C1_differentiable_cases: fixes c::real assumes "f piecewise_C1_differentiable_on {a..c}" "g piecewise_C1_differentiable_on {c..b}" "a \ c" "c \ b" "f c = g c" shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" proof - obtain S T where st: "f C1_differentiable_on ({a..c} - S)" "g C1_differentiable_on ({c..b} - T)" "finite S" "finite T" using assms by (force simp: piecewise_C1_differentiable_on_def) then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], OF closed_real_atLeastAtMost [of c b], of f g "\x. x\c"] assms by (force simp: ivl_disj_un_two_touch) { fix x assume x: "x \ {a..b} - insert c (S \ T)" have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") proof (cases x c rule: le_cases) case le show ?diff_fg apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) next case ge show ?diff_fg apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) qed } then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" by auto moreover { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" if "a < x" "x < c" "x \ S" for x proof - have f: "f differentiable at x" by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) show ?thesis using that apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] f) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" if "c < x" "x < b" "x \ T" for x proof - have g: "g differentiable at x" by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) show ?thesis using that apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] g) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed ultimately have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" by (rule continuous_on_subset [OF continuous_on_open_Un], auto) } note * = this have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" apply (rule_tac x="{a,b,c} \ S \ T" in exI) using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) with cab show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed lemma piecewise_C1_differentiable_neg: "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def by (auto intro!: continuous_on_minus C1_differentiable_on_minus) lemma piecewise_C1_differentiable_add: assumes "f piecewise_C1_differentiable_on i" "g piecewise_C1_differentiable_on i" shows "(\x. f x + g x) piecewise_C1_differentiable_on i" proof - obtain S t where st: "finite S" "finite t" "f C1_differentiable_on (i-S)" "g C1_differentiable_on (i-t)" using assms by (auto simp: piecewise_C1_differentiable_on_def) then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) moreover have "continuous_on i f" "continuous_on i g" using assms piecewise_C1_differentiable_on_def by auto ultimately show ?thesis by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) qed lemma piecewise_C1_differentiable_diff: "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ \ (\x. f x - g x) piecewise_C1_differentiable_on S" unfolding diff_conv_add_uminus by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) lemma piecewise_C1_differentiable_D1: fixes g1 :: "real \ 'a::real_normed_field" assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" shows "g1 piecewise_C1_differentiable_on {0..1}" proof - obtain S where "finite S" and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x proof (rule differentiable_transform_within) show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" using that g12D apply (simp only: joinpaths_def) by (rule differentiable_chain_at derivative_intros | force)+ show "\x'. \dist x' x < dist (x/2) (1/2)\ \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" using that by (auto simp: dist_real_def joinpaths_def) qed (use that in \auto simp: dist_real_def\) have [simp]: "vector_derivative (g1 \ (*) 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x apply (subst vector_derivative_chain_at) using that apply (rule derivative_eq_intros g1D | simp)+ done have "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" using co12 by (rule continuous_on_subset) force then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 \ (*)2) (at x))" proof (rule continuous_on_eq [OF _ vector_derivative_at]) show "(g1 +++ g2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" if "x \ {0..1/2} - insert (1/2) S" for x proof (rule has_vector_derivative_transform_within) show "(g1 \ (*) 2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" using that by (force intro: g1D differentiable_chain_at simp: vector_derivative_works [symmetric]) show "\x'. \dist x' x < dist x (1/2)\ \ (g1 \ (*) 2) x' = (g1 +++ g2) x'" using that by (auto simp: dist_norm joinpaths_def) qed (use that in \auto simp: dist_norm\) qed have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" apply (rule continuous_intros)+ using coDhalf apply (simp add: scaleR_conv_of_real image_set_diff image_image) done then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g1" using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) apply (rule_tac x="insert 1 (((*)2)`S)" in exI) apply (simp add: g1D con_g1) done qed lemma piecewise_C1_differentiable_D2: fixes g2 :: "real \ 'a::real_normed_field" assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" shows "g2 piecewise_C1_differentiable_on {0..1}" proof - obtain S where "finite S" and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x proof (rule differentiable_transform_within) show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" using g12D that apply (simp only: joinpaths_def) apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) apply (rule differentiable_chain_at derivative_intros | force)+ done show "\x'. dist x' x < dist ((x + 1) / 2) (1/2) \ (g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" using that by (auto simp: dist_real_def joinpaths_def field_simps) qed (use that in \auto simp: dist_norm\) have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x using that by (auto simp: vector_derivative_chain_at field_split_simps g2D) have "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" using co12 by (rule continuous_on_subset) force then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g2 \ (\x. 2*x-1)) (at x))" proof (rule continuous_on_eq [OF _ vector_derivative_at]) show "(g1 +++ g2 has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) (at x)" if "x \ {1 / 2..1} - insert (1 / 2) S" for x proof (rule_tac f="g2 \ (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) show "(g2 \ (\x. 2 * x - 1) has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) (at x)" using that by (force intro: g2D differentiable_chain_at simp: vector_derivative_works [symmetric]) show "\x'. \dist x' x < dist (3 / 4) ((x + 1) / 2)\ \ (g2 \ (\x. 2 * x - 1)) x' = (g1 +++ g2) x'" using that by (auto simp: dist_norm joinpaths_def add_divide_distrib) qed (use that in \auto simp: dist_norm\) qed have [simp]: "((\x. (x+1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` S))) = ({1/2..1} - insert (1/2) S)" apply (simp add: image_set_diff inj_on_def image_image) apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) done have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) \ (\x. (x+1)/2))" by (rule continuous_intros | simp add: coDhalf)+ then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) (\x. vector_derivative g2 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g2" using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` S)" in exI) apply (simp add: g2D con_g2) done qed subsection \Valid paths, and their start and finish\ definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" where "closed_path g \ g 0 = g 1" text\In particular, all results for paths apply\ lemma valid_path_imp_path: "valid_path g \ path g" by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" by (metis connected_path_image valid_path_imp_path) lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" by (metis compact_path_image valid_path_imp_path) lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" by (metis bounded_path_image valid_path_imp_path) lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" by (metis closed_path_image valid_path_imp_path) lemma valid_path_compose: assumes "valid_path g" and der: "\x. x \ path_image g \ f field_differentiable (at x)" and con: "continuous_on (path_image g) (deriv f)" shows "valid_path (f \ g)" proof - obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto have "f \ g differentiable at t" when "t\{0..1} - S" for t proof (rule differentiable_chain_at) show "g differentiable at t" using \valid_path g\ by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) next have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis then show "f differentiable at (g t)" using der[THEN field_differentiable_imp_differentiable] by auto qed moreover have "continuous_on ({0..1} - S) (\x. vector_derivative (f \ g) (at x))" proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], rule continuous_intros) show "continuous_on ({0..1} - S) (\x. vector_derivative g (at x))" using g_diff C1_differentiable_on_eq by auto next have "continuous_on {0..1} (\x. deriv f (g x))" using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def by blast then show "continuous_on ({0..1} - S) (\x. deriv f (g x))" using continuous_on_subset by blast next show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" when "t \ {0..1} - S" for t proof (rule vector_derivative_chain_at_general[symmetric]) show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) next have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis then show "f field_differentiable at (g t)" using der by auto qed qed ultimately have "f \ g C1_differentiable_on {0..1} - S" using C1_differentiable_on_eq by blast moreover have "path (f \ g)" apply (rule path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]]) using der by (simp add: continuous_at_imp_continuous_on field_differentiable_imp_continuous_at) ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def using \finite S\ by auto qed lemma valid_path_uminus_comp[simp]: fixes g::"real \ 'a ::real_normed_field" shows "valid_path (uminus \ g) \ valid_path g" proof show "valid_path g \ valid_path (uminus \ g)" for g::"real \ 'a" by (auto intro!: valid_path_compose derivative_intros simp add: deriv_linear[of "-1",simplified]) then show "valid_path g" when "valid_path (uminus \ g)" by (metis fun.map_comp group_add_class.minus_comp_minus id_comp that) qed lemma valid_path_offset[simp]: shows "valid_path (\t. g t - z) \ valid_path g" proof show *: "valid_path (g::real\'a) \ valid_path (\t. g t - z)" for g z unfolding valid_path_def by (fastforce intro:derivative_intros C1_differentiable_imp_piecewise piecewise_C1_differentiable_diff) show "valid_path (\t. g t - z) \ valid_path g" using *[of "\t. g t - z" "-z",simplified] . qed subsection\Contour Integrals along a path\ text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ text\piecewise differentiable function on [0,1]\ definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" (infixr "has'_contour'_integral" 50) where "(f has_contour_integral i) g \ ((\x. f(g x) * vector_derivative g (at x within {0..1})) has_integral i) {0..1}" definition\<^marker>\tag important\ contour_integrable_on (infixr "contour'_integrable'_on" 50) where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" definition\<^marker>\tag important\ contour_integral where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" unfolding contour_integrable_on_def contour_integral_def by blast lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) using has_integral_unique by blast lemma has_contour_integral_eqpath: "\(f has_contour_integral y) p; f contour_integrable_on \; contour_integral p f = contour_integral \ f\ \ (f has_contour_integral y) \" using contour_integrable_on_def contour_integral_unique by auto lemma has_contour_integral_integral: "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" by (metis contour_integral_unique contour_integrable_on_def) lemma has_contour_integral_unique: "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" using has_integral_unique by (auto simp: has_contour_integral_def) lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" using contour_integrable_on_def by blast text\Show that we can forget about the localized derivative.\ lemma has_integral_localized_vector_derivative: "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" proof - have *: "{a..b} - {a,b} = interior {a..b}" by (simp add: atLeastAtMost_diff_ends) show ?thesis apply (rule has_integral_spike_eq [of "{a,b}"]) apply (auto simp: at_within_interior [of _ "{a..b}"]) done qed lemma integrable_on_localized_vector_derivative: "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" by (simp add: integrable_on_def has_integral_localized_vector_derivative) lemma has_contour_integral: "(f has_contour_integral i) g \ ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) lemma contour_integrable_on: "f contour_integrable_on g \ (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) subsection\<^marker>\tag unimportant\ \Reversing a path\ lemma valid_path_imp_reverse: assumes "valid_path g" shows "valid_path(reversepath g)" proof - obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) then have "finite ((-) 1 ` S)" by auto moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" unfolding reversepath_def apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) using S - by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq continuous_on_const elim!: continuous_on_subset)+ + by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq elim!: continuous_on_subset)+ ultimately show ?thesis using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) qed lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" using valid_path_imp_reverse by force lemma has_contour_integral_reversepath: assumes "valid_path g" and f: "(f has_contour_integral i) g" shows "(f has_contour_integral (-i)) (reversepath g)" proof - { fix S x assume xs: "g C1_differentiable_on ({0..1} - S)" "x \ (-) 1 ` S" "0 \ x" "x \ 1" have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - vector_derivative g (at (1 - x) within {0..1})" proof - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" using xs by (force simp: has_vector_derivative_def C1_differentiable_on_def) have "(g \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" by (simp add: o_def) show ?thesis using xs by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) qed } note * = this obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) {0..1}" using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] by (simp add: has_integral_neg) then show ?thesis using S apply (clarsimp simp: reversepath_def has_contour_integral_def) apply (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) apply (auto simp: *) done qed lemma contour_integrable_reversepath: "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" using has_contour_integral_reversepath contour_integrable_on_def by blast lemma contour_integrable_reversepath_eq: "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" using contour_integrable_reversepath valid_path_reversepath by fastforce lemma contour_integral_reversepath: assumes "valid_path g" shows "contour_integral (reversepath g) f = - (contour_integral g f)" proof (cases "f contour_integrable_on g") case True then show ?thesis by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) next case False then have "\ f contour_integrable_on (reversepath g)" by (simp add: assms contour_integrable_reversepath_eq) with False show ?thesis by (simp add: not_integrable_contour_integral) qed subsection\<^marker>\tag unimportant\ \Joining two paths together\ lemma valid_path_join: assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" shows "valid_path(g1 +++ g2)" proof - have "g1 1 = g2 0" using assms by (auto simp: pathfinish_def pathstart_def) moreover have "(g1 \ (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" apply (rule piecewise_C1_differentiable_compose) using assms apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) done moreover have "(g2 \ (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" apply (rule piecewise_C1_differentiable_compose) using assms unfolding valid_path_def piecewise_C1_differentiable_on_def by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) ultimately show ?thesis apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) apply (rule piecewise_C1_differentiable_cases) apply (auto simp: o_def) done qed lemma valid_path_join_D1: fixes g1 :: "real \ 'a::real_normed_field" shows "valid_path (g1 +++ g2) \ valid_path g1" unfolding valid_path_def by (rule piecewise_C1_differentiable_D1) lemma valid_path_join_D2: fixes g2 :: "real \ 'a::real_normed_field" shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" unfolding valid_path_def by (rule piecewise_C1_differentiable_D2) lemma valid_path_join_eq [simp]: fixes g2 :: "real \ 'a::real_normed_field" shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast lemma has_contour_integral_join: assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" "valid_path g1" "valid_path g2" shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" proof - obtain s1 s2 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" using assms by (auto simp: has_contour_integral) have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) using s1 apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) done moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) using s2 apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) done ultimately show ?thesis apply (simp add: has_contour_integral) apply (rule has_integral_combine [where c = "1/2"], auto) done qed lemma contour_integrable_joinI: assumes "f contour_integrable_on g1" "f contour_integrable_on g2" "valid_path g1" "valid_path g2" shows "f contour_integrable_on (g1 +++ g2)" using assms by (meson has_contour_integral_join contour_integrable_on_def) lemma contour_integrable_joinD1: assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" shows "f contour_integrable_on g1" proof - obtain s1 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" using assms apply (auto simp: contour_integrable_on) apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) done then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g1: "\0 < z; z < 1; z \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = 2 *\<^sub>R vector_derivative g1 (at z)" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) using s1 apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) done show ?thesis using s1 apply (auto simp: contour_integrable_on) apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g1) done qed lemma contour_integrable_joinD2: assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" shows "f contour_integrable_on g2" proof - obtain s2 where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" using assms apply (auto simp: contour_integrable_on) apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) apply (simp add: image_affinity_atLeastAtMost_diff) done then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g2: "\0 < z; z < 1; z \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = 2 *\<^sub>R vector_derivative g2 (at z)" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) using s2 apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left vector_derivative_works add_divide_distrib) done show ?thesis using s2 apply (auto simp: contour_integrable_on) apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g2) done qed lemma contour_integrable_join [simp]: shows "\valid_path g1; valid_path g2\ \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast lemma contour_integral_join [simp]: shows "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" by (auto simp: shiftpath_def) lemma valid_path_shiftpath [intro]: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "valid_path(shiftpath a g)" using assms apply (auto simp: valid_path_def shiftpath_alt_def) apply (rule piecewise_C1_differentiable_cases) apply (auto simp: algebra_simps) apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) done lemma has_contour_integral_shiftpath: assumes f: "(f has_contour_integral i) g" "valid_path g" and a: "a \ {0..1}" shows "(f has_contour_integral i) (shiftpath a g)" proof - obtain s where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" using assms by (auto simp: has_contour_integral) then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + integral {0..a} (\x. f (g x) * vector_derivative g (at x))" apply (rule has_integral_unique) apply (subst add.commute) apply (subst integral_combine) using assms * integral_unique by auto { fix x have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" unfolding shiftpath_def apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) apply (intro derivative_eq_intros | simp)+ using g apply (drule_tac x="x+a" in bspec) using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) done } note vd1 = this { fix x have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" unfolding shiftpath_def apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) apply (intro derivative_eq_intros | simp)+ using g apply (drule_tac x="x+a-1" in bspec) using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) done } note vd2 = this have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" using * a by (fastforce intro: integrable_subinterval_real) have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" apply (rule integrable_subinterval_real) using * a by auto have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" apply (rule has_integral_spike_finite [where S = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) using s apply blast using a apply (auto simp: algebra_simps vd1) apply (force simp: shiftpath_def add.commute) using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) done moreover have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" apply (rule has_integral_spike_finite [where S = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) using s apply blast using a apply (auto simp: algebra_simps vd2) apply (force simp: shiftpath_def add.commute) using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) apply (simp add: algebra_simps) done ultimately show ?thesis using a by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) qed lemma has_contour_integral_shiftpath_D: assumes "(f has_contour_integral i) (shiftpath a g)" "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "(f has_contour_integral i) g" proof - obtain s where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) { fix x assume x: "0 < x" "x < 1" "x \ s" then have gx: "g differentiable at x" using g by auto have "vector_derivative g (at x within {0..1}) = vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" apply (rule vector_derivative_at_within_ivl [OF has_vector_derivative_transform_within_open [where f = "(shiftpath (1 - a) (shiftpath a g))" and S = "{0<..<1}-s"]]) using s g assms x apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) done } note vd = this have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" using assms by (auto intro!: has_contour_integral_shiftpath) show ?thesis apply (simp add: has_contour_integral_def) apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) using s assms vd apply (auto simp: Path_Connected.shiftpath_shiftpath) done qed lemma has_contour_integral_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast lemma contour_integrable_on_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto lemma contour_integral_shiftpath: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "contour_integral (shiftpath a g) f = contour_integral g f" using assms by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) subsection\<^marker>\tag unimportant\ \More about straight-line paths\ lemma has_vector_derivative_linepath_within: "(linepath a b has_vector_derivative (b - a)) (at x within s)" apply (simp add: linepath_def has_vector_derivative_def algebra_simps) apply (rule derivative_eq_intros | simp)+ done lemma vector_derivative_linepath_within: "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified]) apply (auto simp: has_vector_derivative_linepath_within) done lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" by (simp add: has_vector_derivative_linepath_within vector_derivative_at) lemma valid_path_linepath [iff]: "valid_path (linepath a b)" apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) apply (rule_tac x="{}" in exI) apply (simp add: differentiable_on_def differentiable_def) using has_vector_derivative_def has_vector_derivative_linepath_within apply (fastforce simp add: continuous_on_eq_continuous_within) done lemma has_contour_integral_linepath: shows "(f has_contour_integral i) (linepath a b) \ ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" - by (simp add: has_contour_integral vector_derivative_linepath_at) + by (simp add: has_contour_integral) lemma linepath_in_path: shows "x \ {0..1} \ linepath a b x \ closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_in_convex_hull: fixes x::real assumes a: "a \ convex hull s" and b: "b \ convex hull s" and x: "0\x" "x\1" shows "linepath a b x \ convex hull s" apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) using x apply (auto simp: linepath_image_01 [symmetric]) done lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" by (simp add: linepath_def) lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" by (simp add: linepath_def) lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" by (simp add: has_contour_integral_linepath) lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" using has_contour_integral_unique by blast lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" using has_contour_integral_trivial contour_integral_unique by blast lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) lemma bounded_linear_linepath: assumes "bounded_linear f" shows "f (linepath a b x) = linepath (f a) (f b) x" proof - interpret f: bounded_linear f by fact show ?thesis by (simp add: linepath_def f.add f.scale) qed lemma bounded_linear_linepath': assumes "bounded_linear f" shows "f \ linepath a b = linepath (f a) (f b)" using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) lemma cnj_linepath: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma cnj_linepath': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) subsection\Relation to subpath construction\ lemma valid_path_subpath: fixes g :: "real \ 'a :: real_normed_vector" assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "valid_path(subpath u v g)" proof (cases "v=u") case True then show ?thesis unfolding valid_path_def subpath_def by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) next case False have "(g \ (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" apply (rule piecewise_C1_differentiable_compose) apply (simp add: C1_differentiable_imp_piecewise) apply (simp add: image_affinity_atLeastAtMost) using assms False apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) apply (subst Int_commute) apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) done then show ?thesis by (auto simp: o_def valid_path_def subpath_def) qed lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" by (simp add: has_contour_integral subpath_def) lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" using has_contour_integral_subpath_refl contour_integrable_on_def by blast lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" - by (simp add: has_contour_integral_subpath_refl contour_integral_unique) + by (simp add: contour_integral_unique) lemma has_contour_integral_subpath: assumes f: "f contour_integrable_on g" and g: "valid_path g" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) (subpath u v g)" proof (cases "v=u") case True then show ?thesis using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) next case False obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) {0..1}" using f uv apply (simp add: contour_integrable_on subpath_def has_contour_integral) apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) apply (simp_all add: has_integral_integral) apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps False) done { fix x have "x \ {0..1} \ x \ (\t. (v-u) *\<^sub>R t + u) -` s \ vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) apply (intro derivative_eq_intros | simp)+ apply (cut_tac s [of "(v - u) * x + u"]) using uv mult_left_le [of x "v-u"] apply (auto simp: vector_derivative_works) done } note vd = this show ?thesis apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) using fs assms apply (simp add: False subpath_def has_contour_integral) apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) done qed lemma contour_integrable_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "f contour_integrable_on (subpath u v g)" apply (cases u v rule: linorder_class.le_cases) apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) apply (subst reversepath_subpath [symmetric]) apply (rule contour_integrable_reversepath) using assms apply (blast intro: valid_path_subpath) apply (simp add: contour_integrable_on_def) using assms apply (blast intro: has_contour_integral_subpath) done lemma has_integral_contour_integral_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(((\x. f(g x) * vector_derivative g (at x))) has_integral contour_integral (subpath u v g) f) {u..v}" using assms apply (auto simp: has_integral_integrable_integral) apply (rule integrable_on_subcbox [where a=u and b=v and S = "{0..1}", simplified]) apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) done lemma contour_integral_subcontour_integral: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "contour_integral (subpath u v g) f = integral {u..v} (\x. f(g x) * vector_derivative g (at x))" using assms has_contour_integral_subpath contour_integral_unique by blast lemma contour_integral_subpath_combine_less: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" "u {0..1}" "v \ {0..1}" "w \ {0..1}" shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = contour_integral (subpath u w g) f" proof (cases "u\v \ v\w \ u\w") case True have *: "subpath v u g = reversepath(subpath u v g) \ subpath w u g = reversepath(subpath u w g) \ subpath w v g = reversepath(subpath v w g)" by (auto simp: reversepath_subpath) have "u < v \ v < w \ u < w \ w < v \ v < u \ u < w \ v < w \ w < u \ w < u \ u < v \ w < v \ v < u" using True assms by linarith with assms show ?thesis using contour_integral_subpath_combine_less [of f g u v w] contour_integral_subpath_combine_less [of f g u w v] contour_integral_subpath_combine_less [of f g v u w] contour_integral_subpath_combine_less [of f g v w u] contour_integral_subpath_combine_less [of f g w u v] contour_integral_subpath_combine_less [of f g w v u] apply simp apply (elim disjE) apply (auto simp: * contour_integral_reversepath contour_integrable_subpath - valid_path_reversepath valid_path_subpath algebra_simps) + valid_path_subpath algebra_simps) done next case False then show ?thesis - apply (auto simp: contour_integral_subpath_refl) + apply (auto) using assms - by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath) + by (metis eq_neg_iff_add_eq_0 contour_integral_reversepath reversepath_subpath valid_path_subpath) qed lemma contour_integral_integral: "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) lemma contour_integral_cong: assumes "g = g'" "\x. x \ path_image g \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral using assms by (intro integral_cong) (auto simp: path_image_def) text \Contour integral along a segment on the real axis\ lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed text\Cauchy's theorem where there's a primitive\ lemma contour_integral_primitive_lemma: fixes f :: "complex \ complex" and g :: "real \ complex" assumes "a \ b" and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) has_integral (f(g b) - f(g a))) {a..b}" proof - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {a..b} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "a < x" and b: "x < b" and xk: "x \ k" then have "g differentiable at x within {a..b}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } note * = this show ?thesis apply (rule fundamental_theorem_of_calculus_interior_strong) using k assms cfg * apply (auto simp: at_within_Icc_at) done qed lemma contour_integral_primitive: assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "valid_path g" "path_image g \ s" shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" using assms apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) done corollary Cauchy_theorem_primitive: assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" shows "(f' has_contour_integral 0) g" using assms by (metis diff_self contour_integral_primitive) text\Existence of path integral for continuous function\ lemma contour_integrable_continuous_linepath: assumes "continuous_on (closed_segment a b) f" shows "f contour_integrable_on (linepath a b)" proof - have "continuous_on {0..1} ((\x. f x * (b - a)) \ linepath a b)" apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) apply (rule continuous_intros | simp add: assms)+ done then show ?thesis apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) apply (rule integrable_continuous [of 0 "1::real", simplified]) apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) apply (auto simp: vector_derivative_linepath_within) done qed lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" by (rule has_derivative_imp_has_field_derivative) (rule derivative_intros | simp)+ lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" apply (rule contour_integral_unique) using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] apply (auto simp: field_simps has_field_der_id) done lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" - by (simp add: continuous_on_const contour_integrable_continuous_linepath) + by (simp add: contour_integrable_continuous_linepath) lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" - by (simp add: continuous_on_id contour_integrable_continuous_linepath) + by (simp add: contour_integrable_continuous_linepath) subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ lemma has_contour_integral_neg: "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" by (simp add: has_integral_neg has_contour_integral_def) lemma has_contour_integral_add: "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" by (simp add: has_integral_add has_contour_integral_def algebra_simps) lemma has_contour_integral_diff: "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" by (simp add: has_integral_diff has_contour_integral_def algebra_simps) lemma has_contour_integral_lmul: "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" apply (simp add: has_contour_integral_def) apply (drule has_integral_mult_right) apply (simp add: algebra_simps) done lemma has_contour_integral_rmul: "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" apply (drule has_contour_integral_lmul) apply (simp add: mult.commute) done lemma has_contour_integral_div: "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) lemma has_contour_integral_eq: "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" apply (simp add: path_image_def has_contour_integral_def) by (metis (no_types, lifting) image_eqI has_integral_eq) lemma has_contour_integral_bound_linepath: assumes "(f has_contour_integral i) (linepath a b)" "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" shows "norm i \ B * norm(b - a)" proof - { fix x::real assume x: "0 \ x" "x \ 1" have "norm (f (linepath a b x)) * norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) } note * = this have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" apply (rule has_integral_bound [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) using assms * unfolding has_contour_integral_def apply (auto simp: norm_mult) done then show ?thesis by (auto simp: content_real) qed (*UNUSED lemma has_contour_integral_bound_linepath_strong: fixes a :: real and f :: "complex \ real" assumes "(f has_contour_integral i) (linepath a b)" "finite k" "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" shows "norm i \ B*norm(b - a)" *) lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" unfolding has_contour_integral_linepath by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" by (simp add: has_contour_integral_def) lemma has_contour_integral_is_0: "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto lemma has_contour_integral_sum: "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) subsection\<^marker>\tag unimportant\ \Operations on path integrals\ lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) lemma contour_integral_neg: "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) lemma contour_integral_add: "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = contour_integral g f1 + contour_integral g f2" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) lemma contour_integral_diff: "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = contour_integral g f1 - contour_integral g f2" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) lemma contour_integral_lmul: shows "f contour_integrable_on g \ contour_integral g (\x. c * f x) = c*contour_integral g f" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) lemma contour_integral_rmul: shows "f contour_integrable_on g \ contour_integral g (\x. f x * c) = contour_integral g f * c" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) lemma contour_integral_div: shows "f contour_integrable_on g \ contour_integral g (\x. f x / c) = contour_integral g f / c" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) lemma contour_integral_eq: "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" apply (simp add: contour_integral_def) using has_contour_integral_eq by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) lemma contour_integral_eq_0: "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" by (simp add: has_contour_integral_is_0 contour_integral_unique) lemma contour_integral_bound_linepath: shows "\f contour_integrable_on (linepath a b); 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" apply (rule has_contour_integral_bound_linepath [of f]) apply (auto simp: has_contour_integral_integral) done lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" by (simp add: contour_integral_unique has_contour_integral_0) lemma contour_integral_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) lemma contour_integrable_eq: "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_eq) subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ lemma contour_integrable_neg: "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" using has_contour_integral_neg contour_integrable_on_def by blast lemma contour_integrable_add: "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" using has_contour_integral_add contour_integrable_on_def by fastforce lemma contour_integrable_diff: "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" using has_contour_integral_diff contour_integrable_on_def by fastforce lemma contour_integrable_lmul: "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" using has_contour_integral_lmul contour_integrable_on_def by fastforce lemma contour_integrable_rmul: "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" using has_contour_integral_rmul contour_integrable_on_def by fastforce lemma contour_integrable_div: "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" using has_contour_integral_div contour_integrable_on_def by fastforce lemma contour_integrable_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ (\x. sum (\a. f a x) s) contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_sum) subsection\<^marker>\tag unimportant\ \Reversing a path integral\ lemma has_contour_integral_reverse_linepath: "(f has_contour_integral i) (linepath a b) \ (f has_contour_integral (-i)) (linepath b a)" using has_contour_integral_reversepath valid_path_linepath by fastforce lemma contour_integral_reverse_linepath: "continuous_on (closed_segment a b) f \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" apply (rule contour_integral_unique) apply (rule has_contour_integral_reverse_linepath) by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) (* Splitting a path integral in a flat way.*) lemma has_contour_integral_split: assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "(f has_contour_integral (i + j)) (linepath a b)" proof (cases "k = 0 \ k = 1") case True then show ?thesis using assms by auto next case False then have k: "0 < k" "k < 1" "complex_of_real k \ 1" using assms by auto have c': "c = k *\<^sub>R (b - a) + a" by (metis diff_add_cancel c) have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" by (simp add: algebra_simps c') { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" using False apply (simp add: c' algebra_simps) apply (simp add: real_vector.scale_left_distrib [symmetric] field_split_simps) done have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" using k has_integral_affinity01 [OF *, of "inverse k" "0"] apply (simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) apply (auto dest: has_integral_cmul [where c = "inverse k"]) done } note fi = this { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" using k apply (simp add: c' field_simps) apply (simp add: scaleR_conv_of_real divide_simps) apply (simp add: field_simps) done have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) done } note fj = this show ?thesis using f k apply (simp add: has_contour_integral_linepath) apply (simp add: linepath_def) apply (rule has_integral_combine [OF _ _ fi fj], simp_all) done qed lemma continuous_on_closed_segment_transform: assumes f: "continuous_on (closed_segment a b) f" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "continuous_on (closed_segment a c) f" proof - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" using c by (simp add: algebra_simps) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) then show "continuous_on (closed_segment a c) f" by (rule continuous_on_subset [OF f]) qed lemma contour_integral_split: assumes f: "continuous_on (closed_segment a b) f" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" proof - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" using c by (simp add: algebra_simps) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) moreover have "closed_segment c b \ closed_segment a b" by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) ultimately have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" by (auto intro: continuous_on_subset [OF f]) show ?thesis by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) qed lemma contour_integral_split_linepath: assumes f: "continuous_on (closed_segment a b) f" and c: "c \ closed_segment a b" shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) text\The special case of midpoints used in the main quadrisection\ lemma has_contour_integral_midpoint: assumes "(f has_contour_integral i) (linepath a (midpoint a b))" "(f has_contour_integral j) (linepath (midpoint a b) b)" shows "(f has_contour_integral (i + j)) (linepath a b)" apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) using assms apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) done lemma contour_integral_midpoint: "continuous_on (closed_segment a b) f \ contour_integral (linepath a b) f = contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) done text\A couple of special case lemmas that are useful below\ lemma triangle_linear_has_chain_integral: "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) apply (auto intro!: derivative_eq_intros) done lemma has_chain_integral_chain_integral3: "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" apply (subst contour_integral_unique [symmetric], assumption) apply (drule has_contour_integral_integrable) apply (simp add: valid_path_join) done lemma has_chain_integral_chain_integral4: "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" apply (subst contour_integral_unique [symmetric], assumption) apply (drule has_contour_integral_integrable) apply (simp add: valid_path_join) done subsection\Reversing the order in a double path integral\ text\The condition is stronger than needed but it's often true in typical situations\ lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" by (auto simp: cbox_Pair_eq) lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" by (auto simp: cbox_Pair_eq) proposition contour_integral_swap: assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" and vp: "valid_path g" "valid_path h" and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" shows "contour_integral g (\w. contour_integral h (f w)) = contour_integral h (\z. contour_integral g (\w. f w z))" proof - have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) \ (\t. (g x, h t))" by (rule ext) simp have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\t. (g t, h x))" by (rule ext) simp have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have "\y. y \ {0..1} \ continuous_on {0..1} (\x. f (g x) (h y))" by (subst fgh2) (rule fcon_im2 gcon continuous_intros | simp)+ then have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" using continuous_on_mult gvcon integrable_continuous_real by blast have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) \ fst" by auto then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" apply (rule ssubst) apply (rule continuous_intros | simp add: gvcon)+ done have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) \ snd" by auto then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" apply (rule ssubst) apply (rule continuous_intros | simp add: hvcon)+ done have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w))" by auto then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" apply (rule ssubst) apply (rule gcon hcon continuous_intros | simp)+ apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) done have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" unfolding contour_integrable_on apply (rule integrable_continuous_real) apply (rule continuous_on_mult [OF _ hvcon]) apply (subst fgh1) apply (rule fcon_im1 hcon continuous_intros | simp)+ done qed also have "\ = integral {0..1} (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" unfolding contour_integral_integral apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ unfolding integral_mult_left [symmetric] apply (simp only: mult_ac) done also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" unfolding contour_integral_integral apply (rule integral_cong) unfolding integral_mult_left [symmetric] apply (simp add: algebra_simps) done finally show ?thesis by (simp add: contour_integral_integral) qed subsection\<^marker>\tag unimportant\ \The key quadrisection step\ lemma norm_sum_half: assumes "norm(a + b) \ e" shows "norm a \ e/2 \ norm b \ e/2" proof - have "e \ norm (- a - b)" by (simp add: add.commute assms norm_minus_commute) thus ?thesis using norm_triangle_ineq4 order_trans by fastforce qed lemma norm_sum_lemma: assumes "e \ norm (a + b + c + d)" shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" proof - have "e \ norm ((a + b) + (c + d))" using assms by (simp add: algebra_simps) then show ?thesis by (auto dest!: norm_sum_half) qed lemma Cauchy_theorem_quadrisection: assumes f: "continuous_on (convex hull {a,b,c}) f" and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" and e: "e * K^2 \ norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" shows "\a' b' c'. a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" (is "\x y z. ?\ x y z") proof - note divide_le_eq_numeral1 [simp del] define a' where "a' = midpoint b c" define b' where "b' = midpoint c a" define c' where "c' = midpoint a b" have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ have fcont': "continuous_on (closed_segment c' b') f" "continuous_on (closed_segment a' c') f" "continuous_on (closed_segment b' a') f" unfolding a'_def b'_def c'_def by (rule continuous_on_subset [OF f], metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ let ?pathint = "\x y. contour_integral(linepath x y) f" have *: "?pathint a b + ?pathint b c + ?pathint c a = (?pathint a c' + ?pathint c' b' + ?pathint b' a) + (?pathint a' c' + ?pathint c' b + ?pathint b a') + (?pathint a' c + ?pathint c b' + ?pathint b' a') + (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" by (simp add: fcont' contour_integral_reverse_linepath) (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" by (metis left_diff_distrib mult.commute norm_mult_numeral1) have [simp]: "\x y. cmod (x - y) = cmod (y - x)" by (simp add: norm_minus_commute) consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" using assms unfolding * by (blast intro: that dest!: norm_sum_lemma) then show ?thesis proof cases case 1 then have "?\ a c' b'" using assms apply (clarsimp simp: c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 2 then have "?\ a' c' b" using assms apply (clarsimp simp: a'_def c'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 3 then have "?\ a' c b'" using assms apply (clarsimp simp: a'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 4 then have "?\ a' b' c'" using assms apply (clarsimp simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast qed qed subsection\<^marker>\tag unimportant\ \Cauchy's theorem for triangles\ lemma triangle_points_closer: fixes a::complex shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ \ norm(x - y) \ norm(a - b) \ norm(x - y) \ norm(b - c) \ norm(x - y) \ norm(c - a)" using simplex_extremal_le [of "{a,b,c}"] by (auto simp: norm_minus_commute) lemma holomorphic_point_small_triangle: assumes x: "x \ S" and f: "continuous_on S f" and cd: "f field_differentiable (at x within S)" and e: "0 < e" shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ x \ convex hull {a,b,c} \ convex hull {a,b,c} \ S \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f) \ e*(dist a b + dist b c + dist c a)^2" (is "\k>0. \a b c. _ \ ?normle a b c") proof - have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ \ a \ e*(x + y + z)^2" by (simp add: algebra_simps power2_eq_square) have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" for x::real and a b c by linarith have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" if "convex hull {a, b, c} \ S" for a b c using segments_subset_convex_hull that by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] { fix f' a b c d assume d: "0 < d" and f': "\y. \cmod (y - x) \ d; y \ S\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" and xc: "x \ convex hull {a, b, c}" and S: "convex hull {a, b, c} \ S" have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF S]) apply (simp add: field_simps) done { fix y assume yc: "y \ convex hull {a,b,c}" have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" proof (rule f') show "cmod (y - x) \ d" by (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) qed (use S yc in blast) also have "\ \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" by (simp add: yc e xc disj_le [OF triangle_points_closer]) finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . } note cm_le = this have "?normle a b c" unfolding dist_norm pa apply (rule le_of_3) using f' xc S e apply simp_all apply (intro norm_triangle_le add_mono path_bound) apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ done } note * = this show ?thesis using cd e apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) apply (clarify dest!: spec mp) using * unfolding dist_norm apply blast done qed text\Hence the most basic theorem for a triangle.\ locale Chain = fixes x0 At Follows assumes At0: "At x0 0" and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" begin primrec f where "f 0 = x0" | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" lemma At: "At (f n) n" proof (induct n) case 0 show ?case by (simp add: At0) next case (Suc n) show ?case by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) qed lemma Follows: "Follows (f(Suc n)) (f n)" by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) declare f.simps(2) [simp del] end lemma Chain3: assumes At0: "At x0 y0 z0 0" and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" obtains f g h where "f 0 = x0" "g 0 = y0" "h 0 = z0" "\n. At (f n) (g n) (h n) n" "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" proof - interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" apply unfold_locales using At0 AtSuc by auto show ?thesis apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) using three.At three.Follows apply simp_all apply (simp_all add: split_beta') done qed proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle: assumes "f holomorphic_on (convex hull {a,b,c})" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof - have contf: "continuous_on (convex hull {a,b,c}) f" by (metis assms holomorphic_on_imp_continuous_on) let ?pathint = "\x y. contour_integral(linepath x y) f" { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" define e where "e = norm y / K^2" have K1: "K \ 1" by (simp add: K_def max.coboundedI1) then have K: "K > 0" by linarith have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" by (simp_all add: K_def) have e: "e > 0" unfolding e_def using ynz K1 by simp define At where "At x y z n \ convex hull {x,y,z} \ convex hull {a,b,c} \ dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" for x y z n have At0: "At a b c 0" using fy by (simp add: At_def e_def has_chain_integral_chain_integral3) { fix x y z n assume At: "At x y z n" then have contf': "continuous_on (convex hull {x,y,z}) f" using contf At_def continuous_on_subset by metis have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" using At Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] apply (simp add: At_def algebra_simps) apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) done } note AtSuc = this obtain fa fb fc where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" and dist: "\n. dist (fa n) (fb n) \ K/2^n" "\n. dist (fb n) (fc n) \ K/2^n" "\n. dist (fc n) (fa n) \ K/2^n" and no: "\n. norm(?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" apply (rule Chain3 [of At, OF At0 AtSuc]) apply (auto simp: At_def) done obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" proof (rule bounded_closed_nest) show "\n. closed (convex hull {fa n, fb n, fc n})" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) show "\m n. m \ n \ convex hull {fa n, fb n, fc n} \ convex hull {fa m, fb m, fc m}" by (erule transitive_stepwise_le) (auto simp: conv_le) qed (fastforce intro: finite_imp_bounded_convex_hull)+ then have xin: "x \ convex hull {a,b,c}" using assms f0 by blast then have fx: "f field_differentiable at x within (convex hull {a,b,c})" using assms holomorphic_on_def by blast { fix k n assume k: "0 < k" and le: "\x' y' z'. \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; x \ convex hull {x',y',z'}; convex hull {x',y',z'} \ convex hull {a,b,c}\ \ cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" and Kk: "K / k < 2 ^ n" have "K / 2 ^ n < k" using Kk k by (auto simp: field_simps) then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" using dist [of n] k by linarith+ have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ (3 * K / 2 ^ n)\<^sup>2" using dist [of n] e K by (simp add: abs_le_square_iff [symmetric]) have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" by linarith have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" using ynz dle e mult_le_cancel_left_pos by blast also have "\ < cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" using no [of n] e K apply (simp add: e_def field_simps) apply (simp only: zero_less_norm_iff [symmetric]) done finally have False using le [OF DD x cosb] by auto } then have ?thesis using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e apply clarsimp apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]], force+) done } moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) segments_subset_convex_hull(3) segments_subset_convex_hull(5)) ultimately show ?thesis using has_contour_integral_integral by fastforce qed subsection\<^marker>\tag unimportant\ \Version needing function holomorphic in interior only\ lemma Cauchy_theorem_flat_lemma: assumes f: "continuous_on (convex hull {a,b,c}) f" and c: "c - a = k *\<^sub>R (b - a)" and k: "0 \ k" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ show ?thesis proof (cases "k \ 1") case True show ?thesis by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) next case False then show ?thesis using fabc c apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) apply (metis closed_segment_commute fabc(3)) apply (auto simp: k contour_integral_reverse_linepath) done qed qed lemma Cauchy_theorem_flat: assumes f: "continuous_on (convex hull {a,b,c}) f" and c: "c - a = k *\<^sub>R (b - a)" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "0 \ k") case True with assms show ?thesis by (blast intro: Cauchy_theorem_flat_lemma) next case False have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + contour_integral (linepath c b) f = 0" apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) using False c apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) done ultimately show ?thesis apply (auto simp: contour_integral_reverse_linepath) using add_eq_0_iff by force qed lemma Cauchy_theorem_triangle_interior: assumes contf: "continuous_on (convex hull {a,b,c}) f" and holf: "f holomorphic_on interior (convex hull {a,b,c})" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using contf continuous_on_subset segments_subset_convex_hull by metis+ have "bounded (f ` (convex hull {a,b,c}))" by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" by (auto simp: dest!: bounded_pos [THEN iffD1]) have "bounded (convex hull {a,b,c})" by (simp add: bounded_convex_hull) then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" using bounded_pos_less by blast then have diff_2C: "norm(x - y) \ 2*C" if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y proof - have "cmod x \ C" using x by (meson Cno not_le not_less_iff_gr_or_eq) hence "cmod (x - y) \ C + C" using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) thus "cmod (x - y) \ 2 * C" by (metis mult_2) qed have contf': "continuous_on (convex hull {b,a,c}) f" using contf by (simp add: insert_commute) { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" by (rule has_chain_integral_chain_integral3 [OF fy]) have ?thesis proof (cases "c=a \ a=b \ b=c") case True then show ?thesis using Cauchy_theorem_flat [OF contf, of 0] using has_chain_integral_chain_integral3 [OF fy] ynz by (force simp: fabc contour_integral_reverse_linepath) next case False then have car3: "card {a, b, c} = Suc (DIM(complex))" by auto { assume "interior(convex hull {a,b,c}) = {}" then have "collinear{a,b,c}" using interior_convex_hull_eq_empty [OF car3] by (simp add: collinear_3_eq_affine_dependent) with False obtain d where "c \ a" "a \ b" "b \ c" "c - b = d *\<^sub>R (a - b)" by (auto simp: collinear_3 collinear_lemma) then have "False" using False Cauchy_theorem_flat [OF contf'] pi_eq_y ynz by (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) } then obtain d where d: "d \ interior (convex hull {a, b, c})" by blast { fix d1 assume d1_pos: "0 < d1" and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ \ cmod (f x' - f x) < cmod y / (24 * C)" define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x let ?pathint = "\x y. contour_integral(linepath x y) f" have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) then have eCB: "24 * e * C * B \ cmod y" using \C>0\ \B>0\ by (simp add: field_simps) have e_le_d1: "e * (4 * C) \ d1" using e \C>0\ by (simp add: field_simps) have "shrink a \ interior(convex hull {a,b,c})" "shrink b \ interior(convex hull {a,b,c})" "shrink c \ interior(convex hull {a,b,c})" using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) then have fhp0: "(f has_contour_integral 0) (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal) then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" by (simp add: has_chain_integral_chain_integral3) have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" "f contour_integrable_on linepath (shrink b) (shrink c)" "f contour_integrable_on linepath (shrink c) (shrink a)" using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" by (simp add: algebra_simps) have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" using False \C>0\ diff_2C [of b a] ynz by (auto simp: field_split_simps hull_inc) have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u apply (cases "x=0", simp add: \0) using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast { fix u v assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" have shr_uv: "shrink u \ interior(convex hull {a,b,c})" "shrink v \ interior(convex hull {a,b,c})" using d e uv by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" apply (rule order_trans [OF _ eCB]) using e \B>0\ diff_2C [of u v] uv by (auto simp: field_simps) { fix x::real assume x: "0\x" "x\1" have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) using uv x d interior_subset apply (auto simp: hull_inc intro!: less_C) done have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" apply (simp only: ll norm_mult scaleR_diff_right) using \e>0\ cmod_less_4C apply (force intro: norm_triangle_lt less_le_trans [OF _ e_le_d1]) done have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" using x uv shr_uv cmod_less_dt by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) also have "\ \ cmod y / cmod (v - u) / 12" using False uv \C>0\ diff_2C [of v u] ynz by (auto simp: field_split_simps hull_inc) finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" by simp then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" using uv False by (auto simp: field_simps) have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ B * (cmod y / 24 / C / B * 2 * C) + 2 * C * (cmod y / 24 / C)" apply (rule add_mono [OF mult_mono]) using By_uv e \0 < B\ \0 < C\ x apply (simp_all add: cmod_fuv cmod_shr cmod_12_le) apply (simp add: field_simps) done also have "\ \ cmod y / 6" by simp finally have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / 6" . } note cmod_diff_le = this have f_uv: "continuous_on (closed_segment u v) f" by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" by (simp add: algebra_simps) have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ (B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)) * content (cbox 0 (1::real))" apply (rule has_integral_bound [of _ "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" _ 0 1]) using ynz \0 < B\ \0 < C\ apply (simp_all del: le_divide_eq_numeral1) apply (simp add: has_integral_diff has_contour_integral_linepath [symmetric] has_contour_integral_integral fpi_uv f_uv contour_integrable_continuous_linepath) apply (auto simp: ** norm_triangle_le norm_mult cmod_diff_le simp del: le_divide_eq_numeral1) done also have "\ \ norm y / 6" by simp finally have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" . } note * = this have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) moreover have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) moreover have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) ultimately have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) \ norm y / 6 + norm y / 6 + norm y / 6" by (metis norm_triangle_le add_mono) also have "\ = norm y / 2" by simp finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - (?pathint a b + ?pathint b c + ?pathint c a)) \ norm y / 2" by (simp add: algebra_simps) then have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) then have "False" using pi_eq_y ynz by auto } moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" by (simp add: contf compact_convex_hull compact_uniformly_continuous) ultimately have "False" unfolding uniformly_continuous_on_def by (force simp: ynz \0 < C\ dist_norm) then show ?thesis .. qed } moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" using fabc contour_integrable_continuous_linepath by auto ultimately show ?thesis using has_contour_integral_integral by fastforce qed subsection\<^marker>\tag unimportant\ \Version allowing finite number of exceptional points\ proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle_cofinite: assumes "continuous_on (convex hull {a,b,c}) f" and "finite S" and "(\x. x \ interior(convex hull {a,b,c}) - S \ f field_differentiable (at x))" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" using assms proof (induction "card S" arbitrary: a b c S rule: less_induct) case (less S a b c) show ?case proof (cases "S={}") case True with less show ?thesis by (fastforce simp: holomorphic_on_def field_differentiable_at_within Cauchy_theorem_triangle_interior) next case False then obtain d S' where d: "S = insert d S'" "d \ S'" by (meson Set.set_insert all_not_in_conv) then show ?thesis proof (cases "d \ convex hull {a,b,c}") case False show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof (rule less.hyps) show "\x. x \ interior (convex hull {a, b, c}) - S' \ f field_differentiable at x" using False d interior_subset by (auto intro!: less.prems) qed (use d less.prems in auto) next case True have *: "convex hull {a, b, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" proof (rule less.hyps) show "\x. x \ interior (convex hull {a, b, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have *: "convex hull {b, c, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" proof (rule less.hyps) show "\x. x \ interior (convex hull {b, c, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have *: "convex hull {c, a, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" proof (rule less.hyps) show "\x. x \ interior (convex hull {c, a, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have "f contour_integrable_on linepath a b" using less.prems abd contour_integrable_joinD1 contour_integrable_on_def by blast moreover have "f contour_integrable_on linepath b c" using less.prems bcd contour_integrable_joinD1 contour_integrable_on_def by blast moreover have "f contour_integrable_on linepath c a" using less.prems cad contour_integrable_joinD1 contour_integrable_on_def by blast ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" by auto { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" have cont_ad: "continuous_on (closed_segment a d) f" by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) have cont_bd: "continuous_on (closed_segment b d) f" by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) have cont_cd: "continuous_on (closed_segment c d) f" by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" using has_chain_integral_chain_integral3 [OF abd] has_chain_integral_chain_integral3 [OF bcd] has_chain_integral_chain_integral3 [OF cad] by (simp_all add: algebra_simps add_eq_0_iff) then have ?thesis using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce } then show ?thesis using fpi contour_integrable_on_def by blast qed qed qed subsection\<^marker>\tag unimportant\ \Cauchy's theorem for an open starlike set\ lemma starlike_convex_subset: assumes S: "a \ S" "closed_segment b c \ S" and subs: "\x. x \ S \ closed_segment a x \ S" shows "convex hull {a,b,c} \ S" using S apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) done lemma triangle_contour_integrals_starlike_primitive: assumes contf: "continuous_on S f" and S: "a \ S" "open S" and x: "x \ S" and subs: "\y. y \ S \ closed_segment a y \ S" and zer: "\b c. closed_segment b c \ S \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" proof - let ?pathint = "\x y. contour_integral(linepath x y) f" { fix e y assume e: "0 < e" and bxe: "ball x e \ S" and close: "cmod (y - x) < e" have y: "y \ S" using bxe close by (force simp: dist_norm norm_minus_commute) have cont_ayf: "continuous_on (closed_segment a y) f" using contf continuous_on_subset subs y by blast have xys: "closed_segment x y \ S" apply (rule order_trans [OF _ bxe]) using close by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) have "?pathint a y - ?pathint a x = ?pathint x y" using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force } note [simp] = this { fix e::real assume e: "0 < e" have cont_atx: "continuous (at x) f" using x S contf continuous_on_eq_continuous_at by blast then obtain d1 where d1: "d1>0" and d1_less: "\y. cmod (y - x) < d1 \ cmod (f y - f x) < e/2" unfolding continuous_at Lim_at dist_norm using e by (drule_tac x="e/2" in spec) force obtain d2 where d2: "d2>0" "ball x d2 \ S" using \open S\ x by (auto simp: open_contains_ball) have dpos: "min d1 d2 > 0" using d1 d2 by simp { fix y assume yx: "y \ x" and close: "cmod (y - x) < min d1 d2" have y: "y \ S" using d2 close by (force simp: dist_norm norm_minus_commute) have "closed_segment x y \ S" using close d2 by (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) then have fxy: "f contour_integrable_on linepath x y" by (metis contour_integrable_continuous_linepath continuous_on_subset [OF contf]) then obtain i where i: "(f has_contour_integral i) (linepath x y)" by (auto simp: contour_integrable_on_def) then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" proof (rule has_contour_integral_bound_linepath) show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" by (meson close d1_less le_less_trans less_imp_le min.strict_boundedE segment_bound1) qed (use e in simp) also have "\ < e * cmod (y - x)" by (simp add: e yx) finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using i yx by (simp add: contour_integral_unique divide_less_eq) } then have "\d>0. \y. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using dpos by blast } then have *: "(\y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \x\ 0" by (simp add: Lim_at dist_norm inverse_eq_divide) show ?thesis apply (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right) apply (rule Lim_transform [OF * tendsto_eventually]) using \open S\ x apply (force simp: dist_norm open_contains_ball inverse_eq_divide [symmetric] eventually_at) done qed (** Existence of a primitive.*) lemma holomorphic_starlike_primitive: fixes f :: "complex \ complex" assumes contf: "continuous_on S f" and S: "starlike S" and os: "open S" and k: "finite k" and fcd: "\x. x \ S - k \ f field_differentiable at x" shows "\g. \x \ S. (g has_field_derivative f x) (at x)" proof - obtain a where a: "a\S" and a_cs: "\x. x\S \ closed_segment a x \ S" using S by (auto simp: starlike_def) { fix x b c assume "x \ S" "closed_segment b c \ S" then have abcs: "convex hull {a, b, c} \ S" by (simp add: a a_cs starlike_convex_subset) then have "continuous_on (convex hull {a, b, c}) f" by (simp add: continuous_on_subset [OF contf]) then have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" using abcs interior_subset by (force intro: fcd Cauchy_theorem_triangle_cofinite [OF _ k]) } note 0 = this show ?thesis apply (intro exI ballI) apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) apply (metis a_cs) apply (metis has_chain_integral_chain_integral3 0) done qed lemma Cauchy_theorem_starlike: "\open S; starlike S; finite k; continuous_on S f; \x. x \ S - k \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) lemma Cauchy_theorem_starlike_simple: "\open S; starlike S; f holomorphic_on S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) apply (simp_all add: holomorphic_on_imp_continuous_on) apply (metis at_within_open holomorphic_on_def) done subsection\Cauchy's theorem for a convex set\ text\For a convex set we can avoid assuming openness and boundary analyticity\ lemma triangle_contour_integrals_convex_primitive: assumes contf: "continuous_on S f" and S: "a \ S" "convex S" and x: "x \ S" and zer: "\b c. \b \ S; c \ S\ \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within S)" proof - let ?pathint = "\x y. contour_integral(linepath x y) f" { fix y assume y: "y \ S" have cont_ayf: "continuous_on (closed_segment a y) f" using S y by (meson contf continuous_on_subset convex_contains_segment) have xys: "closed_segment x y \ S" (*?*) using convex_contains_segment S x y by auto have "?pathint a y - ?pathint a x = ?pathint x y" using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force } note [simp] = this { fix e::real assume e: "0 < e" have cont_atx: "continuous (at x within S) f" using x S contf by (simp add: continuous_on_eq_continuous_within) then obtain d1 where d1: "d1>0" and d1_less: "\y. \y \ S; cmod (y - x) < d1\ \ cmod (f y - f x) < e/2" unfolding continuous_within Lim_within dist_norm using e by (drule_tac x="e/2" in spec) force { fix y assume yx: "y \ x" and close: "cmod (y - x) < d1" and y: "y \ S" have fxy: "f contour_integrable_on linepath x y" using convex_contains_segment S x y by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) then obtain i where i: "(f has_contour_integral i) (linepath x y)" by (auto simp: contour_integrable_on_def) then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" proof (rule has_contour_integral_bound_linepath) show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" by (meson assms(3) close convex_contains_segment d1_less le_less_trans less_imp_le segment_bound1 subset_iff x y) qed (use e in simp) also have "\ < e * cmod (y - x)" by (simp add: e yx) finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using i yx by (simp add: contour_integral_unique divide_less_eq) } then have "\d>0. \y\S. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using d1 by blast } then have *: "((\y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \ 0) (at x within S)" by (simp add: Lim_within dist_norm inverse_eq_divide) show ?thesis apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) apply (rule Lim_transform [OF * tendsto_eventually]) using linordered_field_no_ub apply (force simp: inverse_eq_divide [symmetric] eventually_at) done qed lemma contour_integral_convex_primitive: assumes "convex S" "continuous_on S f" "\a b c. \a \ S; b \ S; c \ S\ \ (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (cases "S={}") case False with assms that show ?thesis by (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) qed auto lemma holomorphic_convex_primitive: fixes f :: "complex \ complex" assumes "convex S" "finite K" and contf: "continuous_on S f" and fd: "\x. x \ interior S - K \ f field_differentiable at x" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (rule contour_integral_convex_primitive [OF \convex S\ contf Cauchy_theorem_triangle_cofinite]) have *: "convex hull {a, b, c} \ S" if "a \ S" "b \ S" "c \ S" for a b c by (simp add: \convex S\ hull_minimal that) show "continuous_on (convex hull {a, b, c}) f" if "a \ S" "b \ S" "c \ S" for a b c by (meson "*" contf continuous_on_subset that) show "f field_differentiable at x" if "a \ S" "b \ S" "c \ S" "x \ interior (convex hull {a, b, c}) - K" for a b c x by (metis "*" DiffD1 DiffD2 DiffI fd interior_mono subsetCE that) qed (use assms in \force+\) lemma holomorphic_convex_primitive': fixes f :: "complex \ complex" assumes "convex S" and "open S" and "f holomorphic_on S" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (rule holomorphic_convex_primitive) fix x assume "x \ interior S - {}" with assms show "f field_differentiable at x" by (auto intro!: holomorphic_on_imp_differentiable_at simp: interior_open) qed (use assms in \auto intro: holomorphic_on_imp_continuous_on\) corollary\<^marker>\tag unimportant\ Cauchy_theorem_convex: "\continuous_on S f; convex S; finite K; \x. x \ interior S - K \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) corollary Cauchy_theorem_convex_simple: "\f holomorphic_on S; convex S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: holomorphic_on_imp_continuous_on) using at_within_interior holomorphic_on_def interior_subset by fastforce text\In particular for a disc\ corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc: "\finite K; continuous_on (cball a e) f; \x. x \ ball a e - K \ f field_differentiable at x; valid_path g; path_image g \ cball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (auto intro: Cauchy_theorem_convex) corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc_simple: "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (simp add: Cauchy_theorem_convex_simple) subsection\<^marker>\tag unimportant\ \Generalize integrability to local primitives\ lemma contour_integral_local_primitive_lemma: fixes f :: "complex\complex" shows "\g piecewise_differentiable_on {a..b}; \x. x \ s \ (f has_field_derivative f' x) (at x within s); \x. x \ {a..b} \ g x \ s\ \ (\x. f' (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b}" apply (cases "cbox a b = {}", force) apply (simp add: integrable_on_def) apply (rule exI) apply (rule contour_integral_primitive_lemma, assumption+) using atLeastAtMost_iff by blast lemma contour_integral_local_primitive_any: fixes f :: "complex \ complex" assumes gpd: "g piecewise_differentiable_on {a..b}" and dh: "\x. x \ s \ \d h. 0 < d \ (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" and gs: "\x. x \ {a..b} \ g x \ s" shows "(\x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" proof - { fix x assume x: "a \ x" "x \ b" obtain d h where d: "0 < d" and h: "(\y. norm(y - g x) < d \ (h has_field_derivative f y) (at y within s))" using x gs dh by (metis atLeastAtMost_iff) have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast then obtain e where e: "e>0" and lessd: "\x'. x' \ {a..b} \ \x' - x\ < e \ cmod (g x' - g x) < d" using x d apply (auto simp: dist_norm continuous_on_iff) apply (drule_tac x=x in bspec) using x apply simp apply (drule_tac x=d in spec, auto) done have "\d>0. \u v. u \ x \ x \ v \ {u..v} \ ball x d \ (u \ v \ a \ u \ v \ b) \ (\x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" apply (rule_tac x=e in exI) using e apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) done } then show ?thesis by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) qed lemma contour_integral_local_primitive: fixes f :: "complex \ complex" assumes g: "valid_path g" "path_image g \ s" and dh: "\x. x \ s \ \d h. 0 < d \ (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" shows "f contour_integrable_on g" using g apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def has_integral_localized_vector_derivative integrable_on_def [symmetric]) using contour_integral_local_primitive_any [OF _ dh] by (meson image_subset_iff piecewise_C1_imp_differentiable) text\In particular if a function is holomorphic\ lemma contour_integrable_holomorphic: assumes contf: "continuous_on s f" and os: "open s" and k: "finite k" and g: "valid_path g" "path_image g \ s" and fcd: "\x. x \ s - k \ f field_differentiable at x" shows "f contour_integrable_on g" proof - { fix z assume z: "z \ s" obtain d where "d>0" and d: "ball z d \ s" using \open s\ z by (auto simp: open_contains_ball) then have contfb: "continuous_on (ball z d) f" using contf continuous_on_subset by blast obtain h where "\y\ball z d. (h has_field_derivative f y) (at y within ball z d)" by (metis holomorphic_convex_primitive [OF convex_ball k contfb fcd] d interior_subset Diff_iff subsetD) then have "\y\ball z d. (h has_field_derivative f y) (at y within s)" by (metis open_ball at_within_open d os subsetCE) then have "\h. (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" by (force simp: dist_norm norm_minus_commute) then have "\d h. 0 < d \ (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" using \0 < d\ by blast } then show ?thesis by (rule contour_integral_local_primitive [OF g]) qed lemma contour_integrable_holomorphic_simple: assumes fh: "f holomorphic_on S" and os: "open S" and g: "valid_path g" "path_image g \ S" shows "f contour_integrable_on g" apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) apply (simp add: fh holomorphic_on_imp_continuous_on) using fh by (simp add: field_differentiable_def holomorphic_on_open os) lemma continuous_on_inversediff: fixes z:: "'a::real_normed_field" shows "z \ S \ continuous_on S (\w. 1 / (w - z))" by (rule continuous_intros | force)+ lemma contour_integrable_inversediff: "\valid_path g; z \ path_image g\ \ (\w. 1 / (w-z)) contour_integrable_on g" apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) done text\Key fact that path integral is the same for a "nearby" path. This is the main lemma for the homotopy form of Cauchy's theorem and is also useful if we want "without loss of generality" to assume some nice properties of a path (e.g. smoothness). It can also be used to define the integrals of analytic functions over arbitrary continuous paths. This is just done for winding numbers now. \ text\A technical definition to avoid duplication of similar proofs, for paths joined at the ends versus looping paths\ definition linked_paths :: "bool \ (real \ 'a) \ (real \ 'a::topological_space) \ bool" where "linked_paths atends g h == (if atends then pathstart h = pathstart g \ pathfinish h = pathfinish g else pathfinish g = pathstart g \ pathfinish h = pathstart h)" text\This formulation covers two cases: \<^term>\g\ and \<^term>\h\ share their start and end points; \<^term>\g\ and \<^term>\h\ both loop upon themselves.\ lemma contour_integral_nearby: assumes os: "open S" and p: "path p" "path_image p \ S" shows "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ linked_paths atends g h \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" proof - have "\z. \e. z \ path_image p \ 0 < e \ ball z e \ S" using open_contains_ball os p(2) by blast then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ S" by metis define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" have "compact (path_image p)" by (metis p(1) compact_path_image) moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" using ee by auto ultimately have "\D \ cover. finite D \ path_image p \ \D" by (simp add: compact_eq_Heine_Borel cover_def) then obtain D where D: "D \ cover" "finite D" "path_image p \ \D" by blast then obtain k where k: "k \ {0..1}" "finite k" and D_eq: "D = ((\z. ball z (ee z / 3)) \ p) ` k" apply (simp add: cover_def path_image_def image_comp) apply (blast dest!: finite_subset_image [OF \finite D\]) done then have kne: "k \ {}" using D by auto have pi: "\i. i \ k \ p i \ path_image p" using k by (auto simp: path_image_def) then have eepi: "\i. i \ k \ 0 < ee((p i))" by (metis ee) define e where "e = Min((ee \ p) ` k)" have fin_eep: "finite ((ee \ p) ` k)" using k by blast have "0 < e" using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) have "uniformly_continuous_on {0..1} p" using p by (simp add: path_def compact_uniformly_continuous) then obtain d::real where d: "d>0" and de: "\x x'. \x' - x\ < d \ x\{0..1} \ x'\{0..1} \ cmod (p x' - p x) < e/3" unfolding uniformly_continuous_on_def dist_norm real_norm_def by (metis divide_pos_pos \0 < e\ zero_less_numeral) then obtain N::nat where N: "N>0" "inverse N < d" using real_arch_inverse [of d] by auto show ?thesis proof (intro exI conjI allI; clarify?) show "e/3 > 0" using \0 < e\ by simp fix g h assume g: "valid_path g" and ghp: "\t\{0..1}. cmod (g t - p t) < e / 3 \ cmod (h t - p t) < e / 3" and h: "valid_path h" and joins: "linked_paths atends g h" { fix t::real assume t: "0 \ t" "t \ 1" then obtain u where u: "u \ k" and ptu: "p t \ ball(p u) (ee(p u) / 3)" using \path_image p \ \D\ D_eq by (force simp: path_image_def) then have ele: "e \ ee (p u)" using fin_eep by (simp add: e_def) have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" using ghp t by auto with ele have "cmod (g t - p t) < ee (p u) / 3" "cmod (h t - p t) < ee (p u) / 3" by linarith+ then have "g t \ ball(p u) (ee(p u))" "h t \ ball(p u) (ee(p u))" using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u by (force simp: dist_norm ball_def norm_minus_commute)+ then have "g t \ S" "h t \ S" using ee u k by (auto simp: path_image_def ball_def) } then have ghs: "path_image g \ S" "path_image h \ S" by (auto simp: path_image_def) moreover { fix f assume fhols: "f holomorphic_on S" then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple by blast+ have contf: "continuous_on S f" by (simp add: fhols holomorphic_on_imp_continuous_on) { fix z assume z: "z \ path_image p" have "f holomorphic_on ball z (ee z)" using fhols ee z holomorphic_on_subset by blast then have "\ff. (\w \ ball z (ee z). (ff has_field_derivative f w) (at w))" using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) } then obtain ff where ff: "\z w. \z \ path_image p; w \ ball z (ee z)\ \ (ff z has_field_derivative f w) (at w)" by metis { fix n assume n: "n \ N" then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" proof (induct n) case 0 show ?case by simp next case (Suc n) obtain t where t: "t \ k" and "p (n/N) \ ball(p t) (ee(p t) / 3)" using \path_image p \ \D\ [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems by (force simp: path_image_def) then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" by (simp add: dist_norm) have e3le: "e/3 \ ee (p t) / 3" using fin_eep t by (simp add: e_def) { fix x assume x: "n/N \ x" "x \ (1 + n)/N" then have nN01: "0 \ n/N" "(1 + n)/N \ 1" using Suc.prems by auto then have x01: "0 \ x" "x \ 1" using x by linarith+ have "cmod (p t - p x) < ee (p t) / 3 + e/3" proof (rule norm_diff_triangle_less [OF ptu de]) show "\real n / real N - x\ < d" using x N by (auto simp: field_simps) qed (use x01 Suc.prems in auto) then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" using e3le eepi [OF t] by simp have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " apply (rule norm_diff_triangle_less [OF ptx]) using ghp x01 by (simp add: norm_minus_commute) also have "\ \ ee (p t)" using e3le eepi [OF t] by simp finally have gg: "cmod (p t - g x) < ee (p t)" . have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " apply (rule norm_diff_triangle_less [OF ptx]) using ghp x01 by (simp add: norm_minus_commute) also have "\ \ ee (p t)" using e3le eepi [OF t] by simp finally have "cmod (p t - g x) < ee (p t)" "cmod (p t - h x) < ee (p t)" using gg by auto } note ptgh_ee = this have "closed_segment (g (real n / real N)) (h (real n / real N)) = path_image (linepath (h (n/N)) (g (n/N)))" by (simp add: closed_segment_commute) also have pi_hgn: "\ \ ball (p t) (ee (p t))" using ptgh_ee [of "n/N"] Suc.prems by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) finally have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ S" using ee pi t by blast have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) \ ball (p t) (ee (p t))" using ptgh_ee [of "(1+n)/N"] Suc.prems by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ S" using \N>0\ Suc.prems ee pi t by (auto simp: Path_Connected.path_image_join field_simps) have pi_subset_ball: "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) \ ball (p t) (ee (p t))" apply (intro subset_path_image_join pi_hgn pi_ghn') using \N>0\ Suc.prems apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) done have pi0: "(f has_contour_integral 0) (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) apply (metis ff open_ball at_within_open pi t) using Suc.prems pi_subset_ball apply (simp_all add: valid_path_join valid_path_subpath g h) done have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" using Suc.prems by (simp add: contour_integrable_subpath g fpa) have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" using gh_n's by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" using gh_ns by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + contour_integral (subpath ((Suc n) / N) (n/N) h) f + contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" using contour_integral_unique [OF pi0] Suc.prems by (simp add: g h fpa valid_path_subpath contour_integrable_subpath fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. \hn - gn = ghn - gh0; gd + ghn' + he + hgn = (0::complex); hn - he = hn'; gn + gd = gn'; hgn = -ghn\ \ hn' - gn' = ghn' - gh0" by (auto simp: algebra_simps) have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) also have "\ = contour_integral (subpath 0 ((Suc n) / N) h) f" using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) finally have pi0_eq: "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = contour_integral (subpath 0 ((Suc n) / N) h) f" . show ?case apply (rule * [OF Suc.hyps eq0 pi0_eq]) using Suc.prems apply (simp_all add: g h fpa contour_integral_subpath_combine contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath continuous_on_subset [OF contf gh_ns]) done qed } note ind = this have "contour_integral h f = contour_integral g f" using ind [OF order_refl] N joins by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) } ultimately show "path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f)" by metis qed qed lemma assumes "open S" "path p" "path_image p \ S" shows contour_integral_nearby_ends: "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ pathstart h = pathstart g \ pathfinish h = pathfinish g \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" and contour_integral_nearby_loops: "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ pathfinish g = pathstart g \ pathfinish h = pathstart h \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" using contour_integral_nearby [OF assms, where atends=True] using contour_integral_nearby [OF assms, where atends=False] unfolding linked_paths_def by simp_all lemma C1_differentiable_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ p C1_differentiable_on S" by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) lemma valid_path_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ valid_path p" by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) lemma valid_path_subpath_trivial [simp]: fixes g :: "real \ 'a::euclidean_space" shows "z \ g x \ valid_path (subpath x x g)" by (simp add: subpath_def valid_path_polynomial_function) lemma contour_integral_bound_exists: assumes S: "open S" and g: "valid_path g" and pag: "path_image g \ S" shows "\L. 0 < L \ (\f B. f holomorphic_on S \ (\z \ S. norm(f z) \ B) \ norm(contour_integral g f) \ L*B)" proof - have "path g" using g by (simp add: valid_path_imp_path) then obtain d::real and p where d: "0 < d" and p: "polynomial_function p" "path_image p \ S" and pi: "\f. f holomorphic_on S \ contour_integral g f = contour_integral p f" using contour_integral_nearby_ends [OF S \path g\ pag] apply clarify apply (drule_tac x=g in spec) apply (simp only: assms) apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) done then obtain p' where p': "polynomial_function p'" "\x. (p has_vector_derivative (p' x)) (at x)" by (blast intro: has_vector_derivative_polynomial_function that) then have "bounded(p' ` {0..1})" using continuous_on_polymonial_function by (force simp: intro!: compact_imp_bounded compact_continuous_image) then obtain L where L: "L>0" and nop': "\x. \0 \ x; x \ 1\ \ norm (p' x) \ L" by (force simp: bounded_pos) { fix f B assume f: "f holomorphic_on S" and B: "\z. z\S \ cmod (f z) \ B" then have "f contour_integrable_on p \ valid_path p" using p S by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) moreover have "cmod (vector_derivative p (at x)) * cmod (f (p x)) \ L * B" if "0 \ x" "x \ 1" for x proof (rule mult_mono) show "cmod (vector_derivative p (at x)) \ L" by (metis nop' p'(2) that vector_derivative_at) show "cmod (f (p x)) \ B" by (metis B atLeastAtMost_iff imageI p(2) path_defs(4) subset_eq that) qed (use \L>0\ in auto) ultimately have "cmod (contour_integral g f) \ L * B" apply (simp only: pi [OF f]) apply (simp only: contour_integral_integral) apply (rule order_trans [OF integral_norm_bound_integral]) apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) done } then show ?thesis by (force simp: L contour_integral_integral) qed text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ subsection \Winding Numbers\ definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where "winding_number_prop \ z e p n \ valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t \ {0..1}. norm(\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" lemma winding_number: assumes "path \" "z \ path_image \" "0 < e" shows "\p. winding_number_prop \ z e p (winding_number \ z)" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain d where d: "d>0" and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ (\t \ {0..1}. norm(h t - \ t) < d/2)" using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" have "\n. \e > 0. \p. winding_number_prop \ z e p n" proof (rule_tac x=nn in exI, clarify) fix e::real assume e: "e>0" obtain p where p: "polynomial_function p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto simp: intro!: holomorphic_intros) then show "\p. winding_number_prop \ z e p nn" apply (rule_tac x=p in exI) using pi_eq [of h p] h p d apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) done qed then show ?thesis unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) qed lemma winding_number_unique: assumes \: "path \" "z \ path_image \" and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "winding_number_prop \ z e p n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by (auto simp: winding_number_prop_def) also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed (*NB not winding_number_prop here due to the loop in p*) lemma winding_number_unique_loop: assumes \: "path \" "z \ path_image \" and loop: "pathfinish \ = pathstart \" and pi: "\e. e>0 \ \p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by auto also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed proposition winding_number_valid_path: assumes "valid_path \" "z \ path_image \" shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" by (rule winding_number_unique) (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) proposition has_contour_integral_winding_number: assumes \: "valid_path \" "z \ path_image \" shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" by (simp add: winding_number_valid_path) lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" by (simp add: path_image_subpath winding_number_valid_path) lemma winding_number_join: assumes \1: "path \1" "z \ path_image \1" and \2: "path \2" "z \ path_image \2" and "pathfinish \1 = pathstart \2" shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" proof (rule winding_number_unique) show "\p. winding_number_prop (\1 +++ \2) z e p (winding_number \1 z + winding_number \2 z)" if "e > 0" for e proof - obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" using \0 < e\ \1 winding_number by blast moreover obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" using \0 < e\ \2 winding_number by blast ultimately have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" using assms apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) apply (auto simp: joinpaths_def) done then show ?thesis by blast qed qed (use assms in \auto simp: not_in_path_image_join\) lemma winding_number_reversepath: assumes "path \" "z \ path_image \" shows "winding_number(reversepath \) z = - (winding_number \ z)" proof (rule winding_number_unique) show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" using assms apply (simp add: winding_number_prop_def contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) apply (auto simp: reversepath_def) done then show ?thesis by blast qed qed (use assms in auto) lemma winding_number_shiftpath: assumes \: "path \" "z \ path_image \" and "pathfinish \ = pathstart \" "a \ {0..1}" shows "winding_number(shiftpath a \) z = winding_number \ z" proof (rule winding_number_unique_loop) show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then show ?thesis apply (rule_tac x="shiftpath a p" in exI) using assms that apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) apply (simp add: shiftpath_def) done qed qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) lemma winding_number_split_linepath: assumes "c \ closed_segment a b" "z \ closed_segment a b" shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" proof - have "z \ closed_segment a c" "z \ closed_segment c b" using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ then show ?thesis using assms by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) qed lemma winding_number_cong: "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) lemma winding_number_constI: assumes "c\z" "\t. \0\t; t\1\ \ g t = c" shows "winding_number g z = 0" proof - have "winding_number g z = winding_number (linepath c c) z" apply (rule winding_number_cong) using assms unfolding linepath_def by auto moreover have "winding_number (linepath c c) z =0" apply (rule winding_number_trivial) using assms by auto ultimately show ?thesis by auto qed lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" unfolding winding_number_def proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) fix n e g assume "0 < e" and g: "winding_number_prop p z e g n" then show "\r. winding_number_prop (\w. p w - z) 0 e r n" by (rule_tac x="\t. g t - z" in exI) (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) next fix n e g assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" then show "\r. winding_number_prop p z e r n" apply (rule_tac x="\t. g t + z" in exI) apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) apply (force simp: algebra_simps) done qed subsubsection\<^marker>\tag unimportant\ \Some lemmas about negating a path\ lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast lemma has_contour_integral_negatepath: assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" shows "(f has_contour_integral i) (uminus \ \)" proof - obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) then have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) show "negligible S" by (simp add: \finite S\ negligible_finite) show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = - (f (- \ x) * vector_derivative \ (at x within {0..1}))" if "x \ {0..1} - S" for x proof - have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" proof (rule vector_derivative_within_cbox) show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" using that unfolding o_def by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) qed (use that in auto) then show ?thesis by simp qed qed then show ?thesis by (simp add: has_contour_integral_def) qed lemma winding_number_negatepath: assumes \: "valid_path \" and 0: "0 \ path_image \" shows "winding_number(uminus \ \) 0 = winding_number \ 0" proof - have "(/) 1 contour_integrable_on \" using "0" \ contour_integrable_inversediff by fastforce then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" by (rule has_contour_integral_integral) then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" using has_contour_integral_neg by auto then show ?thesis using assms apply (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) apply (simp add: contour_integral_unique has_contour_integral_negatepath) done qed lemma contour_integrable_negatepath: assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" shows "f contour_integrable_on (uminus \ \)" by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) (* A combined theorem deducing several things piecewise.*) lemma winding_number_join_pos_combined: "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ lemma Re_winding_number: "\valid_path \; z \ path_image \\ \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) lemma winding_number_pos_le: assumes \: "valid_path \" "z \ path_image \" and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 \ Re(winding_number \ z)" proof - have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x using ge by (simp add: Complex.Im_divide algebra_simps x) let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have hi: "(?vd has_integral ?int z) (cbox 0 1)" unfolding box_real apply (subst has_contour_integral [symmetric]) using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) have "0 \ Im (?int z)" proof (rule has_integral_component_nonneg [of \, simplified]) show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" by (force simp: ge0) show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" by (rule has_integral_spike_interior [OF hi]) simp qed then show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt_lemma: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" shows "0 < Re(winding_number \ z)" proof - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have hi: "(?vd has_integral ?int z) (cbox 0 1)" unfolding box_real apply (subst has_contour_integral [symmetric]) using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) show "((\x. if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e) has_integral ?int z) {0..1}" by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) show "\x. 0 \ x \ x \ 1 \ e \ Im (if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e)" by (simp add: ge) qed (use has_integral_const_real [of _ 0 1] in auto) with e show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 < Re (winding_number \ z)" proof - have bm: "bounded ((\w. w - z) ` (path_image \))" using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" using bounded_pos [THEN iffD1, OF bm] by blast { fix x::real assume x: "0 < x" "x < 1" then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] by (simp add: path_image_def power2_eq_square mult_mono') with x have "\ x \ z" using \ using path_image_def by fastforce then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" using B ge [OF x] B2 e apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) apply (auto simp: divide_left_mono divide_right_mono) done then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) } note * = this show ?thesis using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) qed subsection\The winding number is an integer\ text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, Also on page 134 of Serge Lang's book with the name title, etc.\ lemma exp_fg: fixes z::complex assumes g: "(g has_vector_derivative g') (at x within s)" and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" and z: "g x \ z" shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" proof - have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" using assms unfolding has_vector_derivative_def scaleR_conv_of_real by (auto intro!: derivative_eq_intros) show ?thesis apply (rule has_vector_derivative_eq_rhs) using z apply (auto intro!: derivative_eq_intros * [unfolded o_def] g) done qed lemma winding_number_exp_integral: fixes z::complex assumes \: "\ piecewise_C1_differentiable_on {a..b}" and ab: "a \ b" and z: "z \ \ ` {a..b}" shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" (is "?thesis1") "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" (is "?thesis2") proof - let ?D\ = "\x. vector_derivative \ (at x)" have [simp]: "\x. a \ x \ x \ b \ \ x \ z" using z by force have cong: "continuous_on {a..b} \" using \ by (simp add: piecewise_C1_differentiable_on_def) obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" using \ by (force simp: piecewise_C1_differentiable_on_def) have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) moreover have "{a<.. {a..b} - k" by force ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) { fix w assume "w \ z" have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" by (auto simp: dist_norm intro!: continuous_intros) moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" by (auto simp: intro!: derivative_eq_intros) ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) } then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" by meson have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" unfolding integrable_on_def [symmetric] proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) show "\d h. 0 < d \ (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" if "w \ - {z}" for w apply (rule_tac x="norm(w - z)" in exI) using that inverse_eq_divide has_field_derivative_at_within h by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) qed simp have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" unfolding box_real [symmetric] divide_inverse_commute by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) with ab show ?thesis1 by (simp add: divide_inverse_commute integral_def integrable_on_def) { fix t assume t: "t \ {a..b}" have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" using z by (auto intro!: continuous_intros simp: dist_norm) have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] by simp (auto simp: ball_def dist_norm that) { fix x D assume x: "x \ k" "a < x" "x < b" then have "x \ interior ({a..b} - k)" using open_subset_interior [OF \] by fastforce then have con: "isCont ?D\ x" using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" by (rule continuous_at_imp_continuous_within) have gdx: "\ differentiable at x" using x by (simp add: g_diff_at) have "\d. \x \ k; a < x; x < b; (\ has_vector_derivative d) (at x); a \ t; t \ b\ \ ((\x. integral {a..x} (\x. ?D\ x / (\ x - z))) has_vector_derivative d / (\ x - z)) (at x within {a..b})" apply (rule has_vector_derivative_eq_rhs) apply (rule integral_has_vector_derivative_continuous_at [where S = "{}", simplified]) apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ done then have "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) (at x within {a..b})" using x gdx t apply (clarsimp simp add: differentiable_iff_scaleR) apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_withinI) apply (simp_all add: has_vector_derivative_def [symmetric]) done } note * = this have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) using t apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous_1 [OF vg_int] simp add: ab)+ done } with ab show ?thesis2 by (simp add: divide_inverse_commute integral_def) qed lemma winding_number_exp_2pi: "\path p; z \ path_image p\ \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) lemma integer_winding_number_eq: assumes \: "path \" and z: "z \ path_image \" shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" proof - obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \" "pathfinish p = pathfinish \" and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto then have wneq: "winding_number \ z = winding_number p z" using eq winding_number_valid_path by force have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" using eq by (simp add: exp_eq_1 complex_is_Int_iff) have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" using p winding_number_exp_integral(2) [of p 0 1 z] apply (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) by (metis path_image_def pathstart_def pathstart_in_path_image) then have "winding_number p z \ \ \ pathfinish p = pathstart p" using p wneq iff by (auto simp: path_defs) then show ?thesis using p eq by (auto simp: winding_number_valid_path) qed theorem integer_winding_number: "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" by (metis integer_winding_number_eq) text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) We can thus bound the winding number of a path that doesn't intersect a given ray. \ lemma winding_number_pos_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" proof - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" using z by (auto simp: path_image_def) have [simp]: "z \ \ ` {0..1}" using path_image_def z by auto have gpd: "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast define r where "r = (w - z) / (\ 0 - z)" have [simp]: "r \ 0" using w z by (auto simp: r_def) have cont: "continuous_on {0..1} (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) have "Arg2pi r \ 2*pi" by (simp add: Arg2pi less_eq_real_def) also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" using 1 apply (simp add: winding_number_valid_path [OF \ z] contour_integral_integral) apply (simp add: Complex.Re_divide field_simps power2_eq_square) done finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by (simp add: Arg2pi_ge_0 cont IVT') then obtain t where t: "t \ {0..1}" and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by blast define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" have iArg: "Arg2pi r = Im i" using eqArg by (simp add: i_def) have gpdt: "\ piecewise_C1_differentiable_on {0..t}" by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) have "exp (- i) * (\ t - z) = \ 0 - z" unfolding i_def apply (rule winding_number_exp_integral [OF gpdt]) using t z unfolding path_image_def by force+ then have *: "\ t - z = exp i * (\ 0 - z)" by (simp add: exp_minus field_simps) then have "(w - z) = r * (\ 0 - z)" by (simp add: r_def) then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" apply simp apply (subst Complex_Transcendental.Arg2pi_eq [of r]) apply (simp add: iArg) using * apply (simp add: exp_eq_polar field_simps) done with t show ?thesis by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) qed lemma winding_number_big_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" proof - { assume "Re (winding_number \ z) \ - 1" then have "Re (winding_number (reversepath \) z) \ 1" by (simp add: \ valid_path_imp_path winding_number_reversepath z) moreover have "valid_path (reversepath \)" using \ valid_path_imp_reverse by auto moreover have "z \ path_image (reversepath \)" by (simp add: z) ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" using winding_number_pos_meets w by blast then have ?thesis by simp } then show ?thesis using assms by (simp add: abs_if winding_number_pos_meets split: if_split_asm) qed lemma winding_number_less_1: fixes z::complex shows "\valid_path \; z \ path_image \; w \ z; \a::real. 0 < a \ z + a*(w - z) \ path_image \\ \ Re(winding_number \ z) < 1" by (auto simp: not_less dest: winding_number_big_meets) text\One way of proving that WN=1 for a loop.\ lemma winding_number_eq_1: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" shows "winding_number \ z = 1" proof - have "winding_number \ z \ Ints" by (simp add: \ integer_winding_number loop valid_path_imp_path z) then show ?thesis using 0 2 by (auto simp: Ints_def) qed subsection\Continuity of winding number and invariance on connected sets\ lemma continuous_at_winding_number: fixes z::complex assumes \: "path \" and z: "z \ path_image \" shows "continuous (at z) (winding_number \)" proof - obtain e where "e>0" and cbg: "cball z e \ - path_image \" using open_contains_cball [of "- path_image \"] z by (force simp: closed_def [symmetric] closed_path_image [OF \]) then have ppag: "path_image \ \ - cball z (e/2)" by (force simp: cball_def dist_norm) have oc: "open (- cball z (e / 2))" by (simp add: closed_def [symmetric]) obtain d where "d>0" and pi_eq: "\h1 h2. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ \ path_image h1 \ - cball z (e / 2) \ path_image h2 \ - cball z (e / 2) \ (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [OF oc \ ppag] by metis obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \ \ pathfinish p = pathfinish \" and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) { fix w assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" then have wnotp: "w \ path_image p" using cbg \d>0\ \e>0\ apply (simp add: path_image_def cball_def dist_norm, clarify) apply (frule pg) apply (drule_tac c="\ x" in subsetD) apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) done have wnotg: "w \ path_image \" using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) { fix k::real assume k: "k>0" then obtain q where q: "valid_path q" "w \ path_image q" "pathstart q = pathstart \ \ pathfinish q = pathfinish \" and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k by (force simp: min_divide_distrib_right winding_number_prop_def) have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) apply (frule pg) apply (frule qg) using p q \d>0\ e2 apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) done then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" by (simp add: pi qi) } note pip = this have "path p" using p by (simp add: valid_path_imp_path) then have "winding_number p w = winding_number \ w" apply (rule winding_number_unique [OF _ wnotp]) apply (rule_tac x=p in exI) apply (simp add: p wnotp min_divide_distrib_right pip winding_number_prop_def) done } note wnwn = this obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" using p open_contains_cball [of "- path_image p"] by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) obtain L where "L>0" and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ cmod (contour_integral p f) \ L * B" using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force { fix e::real and w::complex assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" then have [simp]: "w \ path_image p" using cbp p(2) \0 < pe\ by (force simp: dist_norm norm_minus_commute path_image_def cball_def) have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = contour_integral p (\x. 1/(x - w) - 1/(x - z))" by (simp add: p contour_integrable_inversediff contour_integral_diff) { fix x assume pe: "3/4 * pe < cmod (z - x)" have "cmod (w - x) < pe/4 + cmod (z - x)" by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" using norm_diff_triangle_le by blast also have "\ < pe/4 + cmod (w - x)" using w by (simp add: norm_minus_commute) finally have "pe/2 < cmod (w - x)" using pe by auto then have "(pe/2)^2 < cmod (w - x) ^ 2" apply (rule power_strict_mono) using \pe>0\ by auto then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" by (simp add: power_divide) have "8 * L * cmod (w - z) < e * pe\<^sup>2" using w \L>0\ by (simp add: field_simps) also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" using pe2 \e>0\ by (simp add: power2_eq_square) also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" using wx apply (rule mult_strict_left_mono) using pe2 e not_less_iff_gr_or_eq by fastforce finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" by simp also have "\ \ e * cmod (w - x) * cmod (z - x)" using e by simp finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" apply (cases "x=z \ x=w") using pe \pe>0\ w \L>0\ apply (force simp: norm_minus_commute) using wx w(2) \L>0\ pe pe2 Lwz apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) done } note L_cmod_le = this have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" apply (rule L) using \pe>0\ w apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) using \pe>0\ w \L>0\ apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) done have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" apply simp apply (rule le_less_trans [OF *]) using \L>0\ e apply (force simp: field_simps) done then have "cmod (winding_number p w - winding_number p z) < e" using pi_ge_two e by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) } note cmod_wn_diff = this then have "isCont (winding_number p) z" apply (simp add: continuous_at_eps_delta, clarify) apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) using \pe>0\ \L>0\ apply (simp add: dist_norm cmod_wn_diff) done then show ?thesis apply (rule continuous_transform_within [where d = "min d e / 2"]) apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) done qed corollary continuous_on_winding_number: "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" shows "winding_number \ constant_on S" proof - 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 simp flip: of_int_diff) qed have cont: "continuous_on S (\w. winding_number \ w)" using continuous_on_winding_number [OF \] sg by (meson continuous_on_subset disjoint_eq_subset_Compl) show ?thesis using "*" zero_less_one by (blast intro: continuous_discrete_range_constant [OF cs cont]) 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 (metis constant_on_def) lemma open_winding_number_levelsets: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "open {z. z \ path_image \ \ winding_number \ z = k}" proof - have opn: "open (- path_image \)" by (simp add: closed_path_image \ open_Compl) { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" obtain e where e: "e>0" "ball z e \ - path_image \" using open_contains_ball [of "- path_image \"] opn z by blast have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" apply (rule_tac x=e in exI) using e apply (simp add: dist_norm ball_def norm_minus_commute) apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where S = "ball z e"]) done } then show ?thesis by (auto simp: open_dist) qed subsection\Winding number is zero "outside" a curve\ proposition winding_number_zero_in_outside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" shows "winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto obtain w::complex where w: "w \ ball 0 (B + 1)" by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) have "- ball 0 (B + 1) \ outside (path_image \)" apply (rule outside_subset_convex) using B subset_ball by auto then have wout: "w \ outside (path_image \)" using w by blast moreover have "winding_number \ constant_on outside (path_image \)" using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) ultimately have "winding_number \ z = winding_number \ w" by (metis (no_types, hide_lams) constant_on_def z) also have "\ = 0" proof - have wnot: "w \ path_image \" using wout by (simp add: outside_def) { fix e::real assume "0" "pathfinish p = pathfinish \" and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force have pip: "path_image p \ ball 0 (B + 1)" using B apply (clarsimp simp add: path_image_def dist_norm ball_def) apply (frule (1) pg1) apply (fastforce dest: norm_add_less) done then have "w \ path_image p" using w by blast then have "\p. valid_path p \ w \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" apply (rule_tac x=p in exI) apply (simp add: p valid_path_polynomial_function) apply (intro conjI) using pge apply (simp add: norm_minus_commute) apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) apply (rule holomorphic_intros | simp add: dist_norm)+ using mem_ball_0 w apply blast using p apply (simp_all add: valid_path_polynomial_function loop pip) done } then show ?thesis by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) qed finally show ?thesis . qed corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" by (rule winding_number_zero_in_outside) (auto simp: pathfinish_def pathstart_def path_polynomial_function) corollary\<^marker>\tag unimportant\ winding_number_zero_outside: "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) lemma winding_number_zero_at_infinity: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "\B. \z. B \ norm z \ winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto then show ?thesis apply (rule_tac x="B+1" in exI, clarify) apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) apply (meson less_add_one mem_cball_0 not_le order_trans) using ball_subset_cball by blast qed lemma winding_number_zero_point: "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ \ \z. z \ s \ winding_number \ z = 0" using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside by (fastforce simp add: compact_path_image) text\If a path winds round a set, it winds rounds its inside.\ lemma winding_number_around_inside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" shows "winding_number \ w = winding_number \ z" proof - have ssb: "s \ inside(path_image \)" proof fix x :: complex assume "x \ s" hence "x \ path_image \" by (meson disjoint_iff_not_equal s_disj) thus "x \ inside (path_image \)" using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) qed show ?thesis apply (rule winding_number_eq [OF \ loop w]) using z apply blast apply (simp add: cls connected_with_inside cos) apply (simp add: Int_Un_distrib2 s_disj, safe) by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) qed text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ lemma winding_number_subpath_continuous: assumes \: "valid_path \" and z: "z \ path_image \" shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" proof - have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = winding_number (subpath 0 x \) z" if x: "0 \ x" "x \ 1" for x proof - have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" using assms x apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) done also have "\ = winding_number (subpath 0 x \) z" apply (subst winding_number_valid_path) using assms x apply (simp_all add: path_image_subpath valid_path_subpath) by (force simp: path_image_def) finally show ?thesis . qed show ?thesis apply (rule continuous_on_eq [where f = "\x. 1 / (2*pi*\) * integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) apply (rule continuous_intros)+ apply (rule indefinite_integral_continuous_1) apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) using assms apply (simp add: *) done qed lemma winding_number_ivt_pos: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) apply (rule winding_number_subpath_continuous [OF \ z]) using assms apply (auto simp: path_image_def image_def) done lemma winding_number_ivt_neg: assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) apply (rule winding_number_subpath_continuous [OF \ z]) using assms apply (auto simp: path_image_def image_def) done lemma winding_number_ivt_abs: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] by force lemma winding_number_lt_half_lemma: assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" shows "Re(winding_number \ z) < 1/2" proof - { assume "Re(winding_number \ z) \ 1/2" then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" using winding_number_ivt_pos [OF \ z, of "1/2"] by auto have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" using winding_number_exp_2pi [of "subpath 0 t \" z] apply (simp add: t \ valid_path_imp_path) using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) have "b < a \ \ 0" proof - have "\ 0 \ {c. b < a \ c}" by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) thus ?thesis by blast qed moreover have "b < a \ \ t" proof - have "\ t \ {c. b < a \ c}" by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) thus ?thesis by blast qed ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az by (simp add: inner_diff_right)+ then have False by (simp add: gt inner_mult_right mult_less_0_iff) } then show ?thesis by force qed lemma winding_number_lt_half: assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" shows "\Re (winding_number \ z)\ < 1/2" proof - have "z \ path_image \" using assms by auto with assms show ?thesis apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) done qed lemma winding_number_le_half: assumes \: "valid_path \" and z: "z \ path_image \" and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" shows "\Re (winding_number \ z)\ \ 1/2" proof - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" have "isCont (winding_number \) z" by (metis continuous_at_winding_number valid_path_imp_path \ z) then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse apply (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square anz) apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) done have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" by simp then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp then have wnz_12': "\Re (winding_number \ z')\ > 1/2" by linarith moreover have "\Re (winding_number \ z')\ < 1/2" apply (rule winding_number_lt_half [OF \ *]) using azb \d>0\ pag - apply (auto simp: add_strict_increasing anz field_split_simps algebra_simps dest!: subsetD) + apply (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) done ultimately have False by simp } then show ?thesis by force qed lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" using separating_hyperplane_closed_point [of "closed_segment a b" z] apply auto apply (simp add: closed_segment_def) apply (drule less_imp_le) apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) apply (auto simp: segment) done text\ Positivity of WN for a linepath.\ lemma winding_number_linepath_pos_lt: assumes "0 < Im ((b - a) * cnj (b - z))" shows "0 < Re(winding_number(linepath a b) z)" proof - have z: "z \ path_image (linepath a b)" using assms by (simp add: closed_segment_def) (force simp: algebra_simps) show ?thesis apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) apply (simp add: linepath_def algebra_simps) done qed subsection\Cauchy's integral formula, again for a convex enclosing set\ lemma Cauchy_integral_formula_weak: assumes s: "convex s" and "finite k" and conf: "continuous_on s f" and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" and z: "z \ interior s - k" and vpg: "valid_path \" and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - obtain f' where f': "(f has_field_derivative f') (at z)" using fcd [OF z] by (auto simp: field_differentiable_def) have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x proof (cases "x = z") case True then show ?thesis apply (simp add: continuous_within) apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) using has_field_derivative_at_within has_field_derivative_iff f' apply (fastforce simp add:)+ done next case False then have dxz: "dist x z > 0" by auto have cf: "continuous (at x within s) f" using conf continuous_on_eq_continuous_within that by blast have "continuous (at x within s) (\w. (f w - f z) / (w - z))" by (rule cf continuous_intros | simp add: False)+ then show ?thesis apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) apply (force simp: dist_commute) done qed have fink': "finite (insert z k)" using \finite k\ by blast have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \" apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) using c apply (force simp: continuous_on_eq_continuous_within) apply (rename_tac w) apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) apply (simp_all add: dist_pos_lt dist_commute) apply (metis less_irrefl) apply (rule derivative_intros fcd | simp)+ done show ?thesis apply (rule has_contour_integral_eq) using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] apply (auto simp: ac_simps divide_simps) done qed theorem Cauchy_integral_formula_convex_simple: "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; pathfinish \ = pathstart \\ \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" apply (rule Cauchy_integral_formula_weak [where k = "{}"]) using holomorphic_on_imp_continuous_on by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) subsection\Homotopy forms of Cauchy's theorem\ lemma Cauchy_theorem_homotopic: assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" proof - have pathsf: "linked_paths atends g h" using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) obtain k :: "real \ real \ complex" where contk: "continuous_on ({0..1} \ {0..1}) k" and ks: "k ` ({0..1} \ {0..1}) \ s" and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) { fix t::real assume t: "t \ {0..1}" have pak: "path (k \ (\u. (t, u)))" unfolding path_def apply (rule continuous_intros continuous_on_subset [OF contk])+ using t by force have pik: "path_image (k \ Pair t) \ s" using ks t by (auto simp: path_image_def) obtain e where "e>0" and e: "\g h. \valid_path g; valid_path h; \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; linked_paths atends g h\ \ contour_integral h f = contour_integral g f" using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis obtain d where "d>0" and d: "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) { fix t1 t2 assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" using \e > 0\ apply (rule_tac y = k1 in norm_triangle_half_l) apply (auto simp: norm_minus_commute intro: order_less_trans) done have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f" apply (rule_tac x="e/4" in exI) using t t1 t2 ltd \e > 0\ apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) done } then have "\e. 0 < e \ (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e \ (\d. 0 < d \ (\g1 g2. valid_path g1 \ valid_path g2 \ (\u \ {0..1}. norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f)))" by (rule_tac x=d in exI) (simp add: \d > 0\) } then obtain ee where ee: "\t. t \ {0..1} \ ee t > 0 \ (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t \ (\d. 0 < d \ (\g1 g2. valid_path g1 \ valid_path g2 \ (\u \ {0..1}. norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f)))" by metis note ee_rule = ee [THEN conjunct2, rule_format] define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" proof (rule compactE [OF compact_interval]) show "{0..1} \ \C" using ee [THEN conjunct1] by (auto simp: C_def dist_norm) qed (use C_def in auto) define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" have kk01: "kk \ {0..1}" by (auto simp: kk_def) define e where "e = Min (ee ` kk)" have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" using C' by (auto simp: kk_def C_def) have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" by (simp add: kk_def ee) moreover have "finite kk" using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force ultimately have "e > 0" using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) then obtain N::nat where "N > 0" and N: "1/N < e/3" by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) have e_le_ee: "\i. i \ kk \ e \ ee i" using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x using C' subsetD [OF C'01 that] unfolding C'_eq by blast have [OF order_refl]: "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j \ contour_integral j f = contour_integral g f)" if "n \ N" for n using that proof (induct n) case 0 show ?case using ee_rule [of 0 0 0] apply clarsimp apply (rule_tac x=d in exI, safe) by (metis diff_self vpg norm_zero) next case (Suc n) then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" using plus [of "n/N"] by blast then have nN_less: "\n/N - t\ < ee t" by (simp add: dist_norm del: less_divide_eq_numeral1) have n'N_less: "\real (Suc n) / real N - t\ < ee t" using t N \N > 0\ e_le_ee [of t] by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast obtain d1 where "d1 > 0" and d1: "\g1 g2. \valid_path g1; valid_path g2; \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; linked_paths atends g1 g2\ \ contour_integral g2 f = contour_integral g1 f" using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce have "n \ N" using Suc.prems by auto with Suc.hyps obtain d2 where "d2 > 0" and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ \ contour_integral j f = contour_integral g f" by auto have "continuous_on {0..1} (k \ (\u. (n/N, u)))" apply (rule continuous_intros continuous_on_subset [OF contk])+ using N01 by auto then have pkn: "path (\u. k (n/N, u))" by (simp add: path_def) have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) obtain p where "polynomial_function p" and psf: "pathstart p = pathstart (\u. k (n/N, u))" "pathfinish p = pathfinish (\u. k (n/N, u))" and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" using path_approx_polynomial_function [OF pkn min12] by blast then have vpp: "valid_path p" using valid_path_polynomial_function by blast have lpa: "linked_paths atends g p" by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) show ?case proof (intro exI; safe) fix j assume "valid_path j" "linked_paths atends g j" and "\u\{0..1}. cmod (j u - k (real (Suc n) / real N, u)) < min d1 d2" then have "contour_integral j f = contour_integral p f" using pk_le N01(1) ksf by (force intro!: vpp d1 simp add: linked_paths_def psf) also have "... = contour_integral g f" using pk_le by (force intro!: vpp d2 lpa) finally show "contour_integral j f = contour_integral g f" . qed (simp add: \0 < d1\ \0 < d2\) qed then obtain d where "0 < d" "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ linked_paths atends g j \ contour_integral j f = contour_integral g f" using \N>0\ by auto then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" using \N>0\ vph by fastforce then show ?thesis by (simp add: pathsf) qed proposition Cauchy_theorem_homotopic_paths: assumes hom: "homotopic_paths s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" using Cauchy_theorem_homotopic [of True s g h] assms by simp proposition Cauchy_theorem_homotopic_loops: assumes hom: "homotopic_loops s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" using Cauchy_theorem_homotopic [of False s g h] assms by simp lemma has_contour_integral_newpath: "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ \ (f has_contour_integral y) g" using has_contour_integral_integral contour_integral_unique by auto lemma Cauchy_theorem_null_homotopic: "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) using contour_integrable_holomorphic_simple apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) by (simp add: Cauchy_theorem_homotopic_loops) subsection\<^marker>\tag unimportant\ \More winding number properties\ text\including the fact that it's +-1 inside a simple closed curve.\ lemma winding_number_homotopic_paths: assumes "homotopic_paths (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_paths_imp_subset [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_paths (-{z}) g p" and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_paths (-{z}) h q" using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have "homotopic_paths (- {z}) g p" by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) moreover have "homotopic_paths (- {z}) h q" by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) ultimately have "homotopic_paths (- {z}) p q" by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_homotopic_loops: assumes "homotopic_loops (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_loops_imp_subset [OF assms] by auto moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" using homotopic_loops_imp_loop [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_loops (-{z}) g p" and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_loops (-{z}) h q" using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have gp: "homotopic_loops (- {z}) g p" by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) have hq: "homotopic_loops (- {z}) h q" by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" proof (rule Cauchy_theorem_homotopic_loops) show "homotopic_loops (- {z}) p q" by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) qed (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_paths_linear_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) lemma winding_number_loops_linear_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) lemma winding_number_nearby_paths_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) lemma winding_number_nearby_loops_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) lemma winding_number_subpath_combine: "\path g; z \ path_image g; u \ {0..1}; v \ {0..1}; w \ {0..1}\ \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = winding_number (subpath u w g) z" apply (rule trans [OF winding_number_join [THEN sym] winding_number_homotopic_paths [OF homotopic_join_subpaths]]) using path_image_subpath_subset by auto subsection\Partial circle path\ definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" lemma pathstart_part_circlepath [simp]: "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" by (metis part_circlepath_def pathstart_def pathstart_linepath) lemma pathfinish_part_circlepath [simp]: "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" by (metis part_circlepath_def pathfinish_def pathfinish_linepath) lemma reversepath_part_circlepath[simp]: "reversepath (part_circlepath z r s t) = part_circlepath z r t s" unfolding part_circlepath_def reversepath_def linepath_def by (auto simp:algebra_simps) lemma has_vector_derivative_part_circlepath [derivative_intros]: "((part_circlepath z r s t) has_vector_derivative (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) (at x within X)" apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) apply (rule has_vector_derivative_real_field) apply (rule derivative_eq_intros | simp)+ done lemma differentiable_part_circlepath: "part_circlepath c r a b differentiable at x within A" using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast lemma vector_derivative_part_circlepath: "vector_derivative (part_circlepath z r s t) (at x) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" using has_vector_derivative_part_circlepath vector_derivative_at by blast lemma vector_derivative_part_circlepath01: "\0 \ x; x \ 1\ \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" using has_vector_derivative_part_circlepath by (auto simp: vector_derivative_at_within_ivl) lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" apply (simp add: valid_path_def) apply (rule C1_differentiable_imp_piecewise) apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath intro!: continuous_intros) done lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" by (simp add: valid_path_imp_path) proposition path_image_part_circlepath: assumes "s \ t" shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" proof - { fix z::real assume "0 \ z" "z \ 1" with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" apply (rule_tac x="(1 - z) * s + z * t" in exI) apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) apply (rule conjI) using mult_right_mono apply blast using affine_ineq by (metis "mult.commute") } moreover { fix z assume "s \ z" "z \ t" then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" apply (rule_tac x="(z - s)/(t - s)" in image_eqI) apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) apply (auto simp: field_split_simps) done } ultimately show ?thesis by (fastforce simp add: path_image_def part_circlepath_def) qed lemma path_image_part_circlepath': "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" proof - have "path_image (part_circlepath z r s t) = (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" by (simp add: image_image path_image_def part_circlepath_def) also have "linepath s t ` {0..1} = closed_segment s t" by (rule linepath_image_01) finally show ?thesis by (simp add: cis_conv_exp) qed lemma path_image_part_circlepath_subset: "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) lemma in_path_image_part_circlepath: assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" shows "norm(w - z) = r" proof - have "w \ {c. dist z c = r}" by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) thus ?thesis by (simp add: dist_norm norm_minus_commute) qed lemma path_image_part_circlepath_subset': assumes "r \ 0" shows "path_image (part_circlepath z r s t) \ sphere z r" proof (cases "s \ t") case True thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp next case False thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all qed lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma has_contour_integral_part_circlepath_iff: assumes "a < b" shows "(f has_contour_integral I) (part_circlepath c r a b) \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" proof - have "(f has_contour_integral I) (part_circlepath c r a b) \ ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) (at x within {0..1})) has_integral I) {0..1}" unfolding has_contour_integral_def .. also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * cis (linepath a b x)) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_part_circlepath01) (simp_all add: cis_conv_exp) also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * r * \ * exp (\ * linepath (of_real a) (of_real b) x) * vector_derivative (linepath (of_real a) (of_real b)) (at x within {0..1})) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_linepath_within) (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) (linepath (of_real a) (of_real b))" by (simp add: has_contour_integral_def) also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) finally show ?thesis . qed lemma contour_integrable_part_circlepath_iff: assumes "a < b" shows "f contour_integrable_on (part_circlepath c r a b) \ (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (auto simp: contour_integrable_on_def integrable_on_def has_contour_integral_part_circlepath_iff) lemma contour_integral_part_circlepath_eq: assumes "a < b" shows "contour_integral (part_circlepath c r a b) f = integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" proof (cases "f contour_integrable_on part_circlepath c r a b") case True hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with True show ?thesis using has_contour_integral_part_circlepath_iff[OF assms] contour_integral_unique has_integral_integrable_integral by blast next case False hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with False show ?thesis by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma contour_integral_part_circlepath_reverse: "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all lemma contour_integral_part_circlepath_reverse': "b < a \ contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (rule contour_integral_part_circlepath_reverse) lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" proof (cases "w = 0") case True then show ?thesis by auto next case False have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" apply (simp add: norm_mult finite_int_iff_bounded_le) apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) apply (auto simp: field_split_simps le_floor_iff) done have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" by blast show ?thesis apply (subst exp_Ln [OF False, symmetric]) apply (simp add: exp_eq) using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) done qed lemma finite_bounded_log2: fixes a::complex assumes "a \ 0" shows "finite {z. norm z \ b \ exp(a*z) = w}" proof - have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" by (rule finite_imageI [OF finite_bounded_log]) show ?thesis by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) qed lemma has_contour_integral_bound_part_circlepath_strong: assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" and "finite k" and le: "0 \ B" "0 < r" "s \ t" and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod i \ B * r * (t - s)" proof - consider "s = t" | "s < t" using \s \ t\ by linarith then show ?thesis proof cases case 1 with fi [unfolded has_contour_integral] have "i = 0" by (simp add: vector_derivative_part_circlepath) with assms show ?thesis by simp next case 2 have [simp]: "\r\ = r" using \r > 0\ by linarith have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y proof - define w where "w = (y - z)/of_real r / exp(\ * of_real s)" have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" apply (rule finite_vimageI [OF finite_bounded_log2]) using \s < t\ apply (auto simp: inj_of_real) done show ?thesis apply (simp add: part_circlepath_def linepath_def vimage_def) apply (rule finite_subset [OF _ fin]) using le apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) done qed then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" by (rule finite_finite_vimage_IntI [OF \finite k\]) have **: "((\x. if (part_circlepath z r s t x) \ k then 0 else f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" by (auto intro!: B [unfolded path_image_def image_def, simplified]) show ?thesis apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) using assms apply force apply (simp add: norm_mult vector_derivative_part_circlepath) using le * "2" \r > 0\ by auto qed qed lemma has_contour_integral_bound_part_circlepath: "\(f has_contour_integral i) (part_circlepath z r s t); 0 \ B; 0 < r; s \ t; \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ \ norm i \ B*r*(t - s)" by (auto intro: has_contour_integral_bound_part_circlepath_strong) lemma contour_integrable_continuous_part_circlepath: "continuous_on (path_image (part_circlepath z r s t)) f \ f contour_integrable_on (part_circlepath z r s t)" apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) apply (rule integrable_continuous_real) apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) done proposition winding_number_part_circlepath_pos_less: assumes "s < t" and no: "norm(w - z) < r" shows "0 < Re (winding_number(part_circlepath z r s t) w)" proof - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) note valid_path_part_circlepath moreover have " w \ path_image (part_circlepath z r s t)" using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) moreover have "0 < r * (t - s) * (r - cmod (w - z))" using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) ultimately show ?thesis apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) apply (rule mult_left_mono)+ using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) using assms \0 < r\ by auto qed lemma simple_path_part_circlepath: "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" proof (cases "r = 0 \ s = t") case True then show ?thesis unfolding part_circlepath_def simple_path_def by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ next case False then have "r \ 0" "s \ t" by auto have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" by (simp add: algebra_simps) have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" by auto have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] intro: exI [where x = "-n" for n]) have 1: "\s - t\ \ 2 * pi" if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" proof (rule ccontr) assume "\ \s - t\ \ 2 * pi" then have *: "\n. t - s \ of_int n * \s - t\" using False that [of "2*pi / \t - s\"] by (simp add: abs_minus_commute divide_simps) show False using * [of 1] * [of "-1"] by auto qed have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n proof - have "t-s = 2 * (real_of_int n * pi)/x" using that by (simp add: field_simps) then show ?thesis by (metis abs_minus_commute) qed have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" by force show ?thesis using False apply (simp add: simple_path_def) apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) apply (subst abs_away) apply (auto simp: 1) apply (rule ccontr) apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) done qed lemma arc_part_circlepath: assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" shows "arc (part_circlepath z r s t)" proof - have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n proof (rule ccontr) assume "x \ y" have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" by (force simp: field_simps) have "\real_of_int n\ < \y - x\" using assms \x \ y\ by (simp add: st abs_mult field_simps) then show False using assms x y st by (auto dest: of_int_lessD) qed show ?thesis using assms apply (simp add: arc_def) apply (simp add: part_circlepath_def inj_on_def exp_eq) apply (blast intro: *) done qed subsection\Special case of one complete circle\ definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" where "circlepath z r \ part_circlepath z r 0 (2*pi)" lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" by (simp add: circlepath_def) lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" proof - have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = z + of_real r * exp (2 * pi * \ * x + pi * \)" by (simp add: divide_simps) (simp add: algebra_simps) also have "\ = z - r * exp (2 * pi * \ * x)" by (simp add: exp_add) finally show ?thesis by (simp add: circlepath path_image_def sphere_def dist_norm) qed lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] by (simp add: add.commute) lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" using circlepath_add1 [of z r "x-1/2"] by (simp add: add.commute) lemma path_image_circlepath_minus_subset: "path_image (circlepath z (-r)) \ path_image (circlepath z r)" apply (simp add: path_image_def image_def circlepath_minus, clarify) apply (case_tac "xa \ 1/2", force) apply (force simp: circlepath_add_half)+ done lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" using path_image_circlepath_minus_subset by fastforce lemma has_vector_derivative_circlepath [derivative_intros]: "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) (at x within X)" apply (simp add: circlepath_def scaleR_conv_of_real) apply (rule derivative_eq_intros) apply (simp add: algebra_simps) done lemma vector_derivative_circlepath: "vector_derivative (circlepath z r) (at x) = 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" using has_vector_derivative_circlepath vector_derivative_at by blast lemma vector_derivative_circlepath01: "\0 \ x; x \ 1\ \ vector_derivative (circlepath z r) (at x within {0..1}) = 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" using has_vector_derivative_circlepath by (auto simp: vector_derivative_at_within_ivl) lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" by (simp add: circlepath_def) lemma path_circlepath [simp]: "path (circlepath z r)" by (simp add: valid_path_imp_path) lemma path_image_circlepath_nonneg: assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" proof - have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x proof (cases "x = z") case True then show ?thesis by force next case False define w where "w = x - z" then have "w \ 0" by (simp add: False) have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" using cis_conv_exp complex_eq_iff by auto show ?thesis apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) apply (rule_tac x="t / (2*pi)" in image_eqI) apply (simp add: field_simps \w \ 0\) using False ** apply (auto simp: w_def) done qed show ?thesis unfolding circlepath path_image_def sphere_def dist_norm by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) qed lemma path_image_circlepath [simp]: "path_image (circlepath z r) = sphere z \r\" using path_image_circlepath_minus by (force simp: path_image_circlepath_nonneg abs_if) lemma has_contour_integral_bound_circlepath_strong: "\(f has_contour_integral i) (circlepath z r); finite k; 0 \ B; 0 < r; \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ \ norm i \ B*(2*pi*r)" unfolding circlepath_def by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) lemma has_contour_integral_bound_circlepath: "\(f has_contour_integral i) (circlepath z r); 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ \ norm i \ B*(2*pi*r)" by (auto intro: has_contour_integral_bound_circlepath_strong) lemma contour_integrable_continuous_circlepath: "continuous_on (path_image (circlepath z r)) f \ f contour_integrable_on (circlepath z r)" by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" by (simp add: circlepath_def simple_path_part_circlepath) lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" by (simp add: sphere_def dist_norm norm_minus_commute) lemma contour_integral_circlepath: assumes "r > 0" shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" proof (rule contour_integral_unique) show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" unfolding has_contour_integral_def using assms apply (subst has_integral_cong) apply (simp add: vector_derivative_circlepath01) using has_integral_const_real [of _ 0 1] apply (force simp: circlepath) done qed lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" apply (rule winding_number_unique_loop) apply (simp_all add: sphere_def valid_path_imp_path) apply (rule_tac x="circlepath z r" in exI) apply (simp add: sphere_def contour_integral_circlepath) done proposition winding_number_circlepath: assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" proof (cases "w = z") case True then show ?thesis using assms winding_number_circlepath_centre by auto next case False have [simp]: "r > 0" using assms le_less_trans norm_ge_zero by blast define r' where "r' = norm(w - z)" have "r' < r" by (simp add: assms r'_def) have disjo: "cball z r' \ sphere z r = {}" using \r' < r\ by (force simp: cball_def sphere_def) have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" proof (rule winding_number_around_inside [where s = "cball z r'"]) show "winding_number (circlepath z r) z \ 0" by (simp add: winding_number_circlepath_centre) show "cball z r' \ path_image (circlepath z r) = {}" by (simp add: disjo less_eq_real_def) qed (auto simp: r'_def dist_norm norm_minus_commute) also have "\ = 1" by (simp add: winding_number_circlepath_centre) finally show ?thesis . qed text\ Hence the Cauchy formula for points inside a circle.\ theorem Cauchy_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" proof - have "r > 0" using assms le_less_trans norm_ge_zero by blast have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) (circlepath z r)" proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) show "\x. x \ interior (cball z r) - {} \ f field_differentiable at x" using holf holomorphic_on_imp_differentiable_at by auto have "w \ sphere z r" by simp (metis dist_commute dist_norm not_le order_refl wz) then show "path_image (circlepath z r) \ cball z r - {w}" using \r > 0\ by (auto simp add: cball_def sphere_def) qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) then show ?thesis by (simp add: winding_number_circlepath assms) qed corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: assumes "f holomorphic_on cball z r" "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) lemma no_bounded_connected_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" shows "winding_number g z = 0" apply (rule winding_number_zero_in_outside) apply (simp_all add: assms) by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) lemma no_bounded_path_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" shows "winding_number g z = 0" apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) by (simp add: bounded_subset nb path_component_subset_connected_component) subsection\ Uniform convergence of path integral\ text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ proposition contour_integral_uniform_limit: assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" and ul_f: "uniform_limit (path_image \) f l F" and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and \: "valid_path \" and [simp]: "\ trivial_limit F" shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" proof - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) { fix e::real assume "0 < e" then have "0 < e / (\B\ + 1)" by simp then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" using ul_f [unfolded uniform_limit_iff dist_norm] by auto with ev_fint obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" using eventually_happens [OF eventually_conj] by (fastforce simp: contour_integrable_on path_image_def) have Ble: "B * e / (\B\ + 1) \ e" using \0 \ B\ \0 < e\ by (simp add: field_split_simps) have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" proof (intro exI conjI ballI) show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" if "x \ {0..1}" for x apply (rule order_trans [OF _ Ble]) using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le]) done qed (rule inta) } then show lintg: "l contour_integrable_on \" unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) { fix e::real define B' where "B' = B + 1" have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) assume "0 < e" then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B' / 2"] B' by (simp add: field_simps) have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t proof - have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto also have "\ < e" by (simp add: B' \0 < e\ mult_imp_div_pos_less) finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . then show ?thesis by (simp add: left_diff_distrib [symmetric] norm_mult) qed have le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ \ cmod (integral {0..1} (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" apply (rule le_less_trans [OF integral_norm_bound_integral ie]) apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) apply (blast intro: *)+ done have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) done } then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" by (rule tendstoI) qed corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" and "uniform_limit (sphere z r) f l F" and "\ trivial_limit F" "0 < r" shows "l contour_integrable_on (circlepath z r)" "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ lemma Cauchy_next_derivative: assumes "continuous_on (path_image \) f'" and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" and k: "k \ 0" and "open s" and \: "valid_path \" and w: "w \ s - path_image \" shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) (at w)" (is "?thes2") proof - have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w using open_contains_ball by blast have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" by (metis norm_of_nat of_nat_Suc) have cint: "\x. \x \ w; cmod (x - w) < d\ \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" apply (rule contour_integrable_div [OF contour_integrable_diff]) using int w d by (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) contour_integrable_on \" unfolding eventually_at apply (rule_tac x=d in exI) apply (simp add: \d > 0\ dist_norm field_simps cint) done have bim_g: "bounded (image f' (path_image \))" by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" by (force simp: bounded_pos path_image_def) have twom: "\\<^sub>F n in at w. \x\path_image \. cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" if "0 < e" for e proof - have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" for u x proof - define ff where [abs_def]: "ff n w = (if n = 0 then inverse(x - w)^k else if n = 1 then k / (x - w)^(Suc k) else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" if "z \ ball w (d/2)" "i \ 1" for i z proof - have "z \ path_image \" using \x \ path_image \\ d that ball_divide_subset_numeral by blast then have xz[simp]: "x \ z" using \x \ path_image \\ by blast then have neq: "x * x + z * z \ x * (z * 2)" by (blast intro: dest!: sum_sqs_eq) with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" by (simp add: algebra_simps) show ?thesis using \i \ 1\ apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ done qed { fix a::real and b::real assume ab: "a > 0" "b > 0" then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" by (subst mult_le_cancel_left_pos) (use \k \ 0\ in \auto simp: divide_simps\) with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" by (simp add: field_simps) } note canc = this have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" if "v \ ball w (d/2)" for v proof - have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) have "d/2 \ cmod (x - v)" using d x that using lessd d x by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) then have "d \ cmod (x - v) * 2" by (simp add: field_split_simps) then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" using \0 < d\ order_less_imp_le power_mono by blast have "x \ v" using that using \x \ path_image \\ ball_divide_subset_numeral d by fastforce then show ?thesis using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) - using dpow_le apply (simp add: algebra_simps field_split_simps mult_less_0_iff) + using dpow_le apply (simp add: field_split_simps) done qed have ub: "u \ ball w (d/2)" using uwd by (simp add: dist_commute dist_norm) have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] by (simp add: ff_def \0 < d\) then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" by (simp add: field_simps) then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) / (cmod (u - w) * real k) \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) also have "\ < e" using uw_less \0 < d\ by (simp add: mult_ac divide_simps) finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) / cmod ((u - w) * real k) < e" by (simp add: norm_mult) have "x \ u" using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) show ?thesis apply (rule le_less_trans [OF _ e]) using \k \ 0\ \x \ u\ \u \ w\ apply (simp add: field_simps norm_divide [symmetric]) done qed show ?thesis unfolding eventually_at apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) done qed have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" unfolding uniform_limit_iff dist_norm proof clarify fix e::real assume "0 < e" have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" and x: "0 \ x" "x \ 1" for u x proof (cases "(f' (\ x)) = 0") case True then show ?thesis by (simp add: \0 < e\) next case False have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k))" by (simp add: field_simps) also have "\ = cmod (f' (\ x)) * cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k)" by (simp add: norm_mult) also have "\ < cmod (f' (\ x)) * (e/C)" using False mult_strict_left_mono [OF ec] by force also have "\ \ e" using C by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) finally show ?thesis . qed show "\\<^sub>F n in at w. \x\path_image \. cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def by (force intro: * elim: eventually_mono) qed show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = (f u - f w) / (u - w) / k" if "dist u w < d" for u proof - have u: "u \ s - path_image \" by (metis subsetD d dist_commute mem_ball that) show ?thesis apply (rule contour_integral_unique) apply (simp add: diff_divide_distrib algebra_simps) apply (intro has_contour_integral_diff has_contour_integral_div) using u w apply (simp_all add: field_simps int) done qed show ?thes2 apply (simp add: has_field_derivative_iff del: power_Suc) apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) apply (simp add: \k \ 0\ **) done qed lemma Cauchy_next_derivative_circlepath: assumes contf: "continuous_on (path_image (circlepath z r)) f" and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" and k: "k \ 0" and w: "w \ ball z r" shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" (is "?thes1") and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" (is "?thes2") proof - have "r > 0" using w using ball_eq_empty by fastforce have wim: "w \ ball z r - path_image (circlepath z r)" using w by (auto simp: dist_norm) show ?thes1 ?thes2 by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; auto simp: vector_derivative_circlepath norm_mult)+ qed text\ In particular, the first derivative formula.\ lemma Cauchy_derivative_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" (is "?thes1") and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" (is "?thes2") proof - have [simp]: "r \ 0" using w using ball_eq_empty by fastforce have f: "continuous_on (path_image (circlepath z r)) f" by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) have int: "\w. dist z w < r \ ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) show ?thes1 apply (simp add: power2_eq_square) apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) apply (blast intro: int) done have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" apply (simp add: power2_eq_square) apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) apply (blast intro: int) done then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) show ?thes2 by simp (rule fder) qed subsection\Existence of all higher derivatives\ proposition derivative_is_holomorphic: assumes "open S" and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" shows "f' holomorphic_on S" proof - have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z proof - obtain r where "r > 0" and r: "cball z r \ S" using open_contains_cball \z \ S\ \open S\ by blast then have holf_cball: "f holomorphic_on cball z r" apply (simp add: holomorphic_on_def) using field_differentiable_at_within field_differentiable_def fder by blast then have "continuous_on (path_image (circlepath z r)) f" using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" by (auto intro: continuous_intros)+ have contf_cball: "continuous_on (cball z r) f" using holf_cball by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) have holf_ball: "f holomorphic_on ball z r" using holf_cball using ball_subset_cball holomorphic_on_subset by blast { fix w assume w: "w \ ball z r" have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) (at w)" by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) (circlepath z r)" by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) (circlepath z r)" by (simp add: algebra_simps) then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" by (simp add: f'_eq) } note * = this show ?thesis apply (rule exI) apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) apply (simp_all add: \0 < r\ * dist_norm) done qed show ?thesis by (simp add: holomorphic_on_open [OF \open S\] *) qed lemma holomorphic_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" using analytic_on_holomorphic holomorphic_deriv by auto lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" by (induction n) (auto simp: holomorphic_deriv) lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" unfolding analytic_on_def using holomorphic_higher_deriv by blast lemma has_field_derivative_higher_deriv: "\f holomorphic_on S; open S; x \ S\ \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) lemma valid_path_compose_holomorphic: assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" shows "valid_path (f \ g)" proof (rule valid_path_compose[OF \valid_path g\]) fix x assume "x \ path_image g" then show "f field_differentiable at x" using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast next have "deriv f holomorphic_on S" using holomorphic_deriv holo \open S\ by auto then show "continuous_on (path_image g) (deriv f)" using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto qed subsection\Morera's theorem\ lemma Morera_local_triangle_ball: assumes "\z. z \ S \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ (\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" shows "f analytic_on S" proof - { fix z assume "z \ S" with assms obtain e a where "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" and 0: "\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" by fastforce have az: "dist a z < e" using mem_ball z by blast have sb_ball: "ball z (e - dist a z) \ ball a e" by (simp add: dist_commute ball_subset_ball_iff) have "\e>0. f holomorphic_on ball z e" proof (intro exI conjI) have sub_ball: "\y. dist a y < e \ closed_segment a y \ ball a e" by (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) show "f holomorphic_on ball z (e - dist a z)" apply (rule holomorphic_on_subset [OF _ sb_ball]) apply (rule derivative_is_holomorphic[OF open_ball]) apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) apply (simp_all add: 0 \0 < e\ sub_ball) done qed (simp add: az) } then show ?thesis by (simp add: analytic_on_def) qed lemma Morera_local_triangle: assumes "\z. z \ S \ \t. open t \ z \ t \ continuous_on t f \ (\a b c. convex hull {a,b,c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" shows "f analytic_on S" proof - { fix z assume "z \ S" with assms obtain t where "open t" and z: "z \ t" and contf: "continuous_on t f" and 0: "\a b c. convex hull {a,b,c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" by force then obtain e where "e>0" and e: "ball z e \ t" using open_contains_ball by blast have [simp]: "continuous_on (ball z e) f" using contf using continuous_on_subset e by blast have eq0: "\b c. closed_segment b c \ ball z e \ contour_integral (linepath z b) f + contour_integral (linepath b c) f + contour_integral (linepath c z) f = 0" by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ (\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" using \e > 0\ eq0 by force } then show ?thesis by (simp add: Morera_local_triangle_ball) qed proposition Morera_triangle: "\continuous_on S f; open S; \a b c. convex hull {a,b,c} \ S \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0\ \ f analytic_on S" using Morera_local_triangle by blast subsection\Combining theorems for higher derivatives including Leibniz rule\ lemma higher_deriv_linear [simp]: "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" by (induction n) auto lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" by (induction n) auto lemma higher_deriv_ident [simp]: "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" apply (induction n, simp) apply (metis higher_deriv_linear lambda_one) done lemma higher_deriv_id [simp]: "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" by (simp add: id_def) lemma has_complex_derivative_funpow_1: "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" apply (induction n, auto) apply (simp add: id_def) by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) lemma higher_deriv_uminus: assumes "f holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" using Suc.prems assms has_field_derivative_higher_deriv by auto have "((deriv ^^ n) (\w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)" apply (rule has_field_derivative_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) apply (rule derivative_eq_intros | rule * refl assms)+ apply (auto simp add: Suc) done then show ?case by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_add: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" using Suc.prems assms has_field_derivative_higher_deriv by auto have "((deriv ^^ n) (\w. f w + g w) has_field_derivative deriv ((deriv ^^ n) f) z + deriv ((deriv ^^ n) g) z) (at z)" apply (rule has_field_derivative_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) apply (rule derivative_eq_intros | rule * refl assms)+ apply (auto simp add: Suc) done then show ?case by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_diff: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) apply (subst higher_deriv_add) using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) done lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" by (cases k) simp_all lemma higher_deriv_mult: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" shows "(deriv ^^ n) (\w. f w * g w) z = (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" using Suc.prems assms has_field_derivative_higher_deriv by auto have sumeq: "(\i = 0..n. of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" apply (simp add: bb algebra_simps sum.distrib) apply (subst (4) sum_Suc_reindex) apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong) done have "((deriv ^^ n) (\w. f w * g w) has_field_derivative (\i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z)) (at z)" apply (rule has_field_derivative_transform_within_open [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) apply (simp add: algebra_simps) apply (rule DERIV_cong [OF DERIV_sum]) apply (rule DERIV_cmult) apply (auto intro: DERIV_mult * sumeq \open S\ Suc.prems Suc.IH [symmetric]) done then show ?case unfolding funpow.simps o_apply by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_transform_within_open: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" and fg: "\w. w \ S \ f w = g w" shows "(deriv ^^ i) f z = (deriv ^^ i) g z" using z by (induction i arbitrary: z) (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) lemma higher_deriv_compose_linear: fixes z::complex assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" and fg: "\w. w \ S \ u * w \ T" shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have holo0: "f holomorphic_on (*) u ` S" by (meson fg f holomorphic_on_subset image_subset_iff) have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) have holo1: "(\w. f (u * w)) holomorphic_on S" apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) apply (rule holo0 holomorphic_intros)+ done have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" apply (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems]) apply (rule holomorphic_higher_deriv [OF holo1 S]) apply (simp add: Suc.IH) done also have "\ = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" apply (rule deriv_cmult) apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=T, unfolded o_def]) - apply (simp add: analytic_on_linear) + apply (simp) apply (simp add: analytic_on_open f holomorphic_higher_deriv T) apply (blast intro: fg) done also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "(*) u", unfolded o_def]) apply (rule derivative_intros) using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv T apply blast - apply (simp add: deriv_linear) + apply (simp) done finally show ?case by simp qed lemma higher_deriv_add_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_add show ?thesis by (auto simp: analytic_at_two) qed lemma higher_deriv_diff_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_diff show ?thesis by (auto simp: analytic_at_two) qed lemma higher_deriv_uminus_at: "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" using higher_deriv_uminus by (auto simp: analytic_at) lemma higher_deriv_mult_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w * g w) z = (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_mult show ?thesis by (auto simp: analytic_at_two) qed text\ Nonexistence of isolated singularities and a stronger integral formula.\ proposition no_isolated_singularity: fixes z::complex assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" shows "f holomorphic_on S" proof - { fix z assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" have "f field_differentiable at z" proof (cases "z \ K") case False then show ?thesis by (blast intro: cdf \z \ S\) next case True with finite_set_avoid [OF K, of z] obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" by blast obtain e where "e>0" and e: "ball z e \ S" using S \z \ S\ by (force simp: open_contains_ball) have fde: "continuous_on (ball z (min d e)) f" by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c by (simp add: hull_minimal continuous_on_subset [OF fde]) have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ \ f field_differentiable at x" for a b c x by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" apply (rule contour_integral_convex_primitive [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) using cont fd by auto then have "f holomorphic_on ball z (min d e)" by (metis open_ball at_within_open derivative_is_holomorphic) then show ?thesis unfolding holomorphic_on_def by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) qed } with holf S K show ?thesis by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) qed lemma no_isolated_singularity': fixes z::complex assumes f: "\z. z \ K \ (f \ f z) (at z within S)" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" shows "f holomorphic_on S" proof (rule no_isolated_singularity[OF _ assms(2-)]) show "continuous_on S f" unfolding continuous_on_def proof fix z assume z: "z \ S" show "(f \ f z) (at z within S)" proof (cases "z \ K") case False from holf have "continuous_on (S - K) f" by (rule holomorphic_on_imp_continuous_on) with z False have "(f \ f z) (at z within (S - K))" by (simp add: continuous_on_def) also from z K S False have "at z within (S - K) = at z within S" by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) finally show "(f \ f z) (at z within S)" . qed (insert assms z, simp_all) qed qed proposition Cauchy_integral_formula_convex: assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" and z: "z \ interior S" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have *: "\x. x \ interior S \ f field_differentiable at x" unfolding holomorphic_on_open [symmetric] field_differentiable_def using no_isolated_singularity [where S = "interior S"] by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd field_differentiable_at_within field_differentiable_def holomorphic_onI holomorphic_on_imp_differentiable_at open_interior) show ?thesis by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) qed text\ Formula for higher derivatives.\ lemma Cauchy_has_contour_integral_higher_derivative_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) (circlepath z r)" using w proof (induction k arbitrary: w) case 0 then show ?case using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) next case (Suc k) have [simp]: "r > 0" using w using ball_eq_empty by fastforce have f: "continuous_on (path_image (circlepath z r)) f" by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] by (auto simp: contour_integrable_on_def) then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" by (rule contour_integral_unique) have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" using Suc.prems assms has_field_derivative_higher_deriv by auto then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" by (force simp: field_differentiable_def) have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) also have "\ = of_nat (Suc k) * X" by (simp only: con) finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" by (metis deriv_cmult dnf_diff) then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" by (simp add: field_simps) then show ?case using of_nat_eq_0_iff X by fastforce qed lemma Cauchy_higher_derivative_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" (is "?thes1") and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" (is "?thes2") proof - have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) (circlepath z r)" using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] by simp show ?thes1 using * using contour_integrable_on_def by blast show ?thes2 unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) qed corollary Cauchy_contour_integral_circlepath: assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) lemma Cauchy_contour_integral_circlepath_2: assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" using Cauchy_contour_integral_circlepath [OF assms, of 1] by (simp add: power2_eq_square) subsection\A holomorphic function is analytic, i.e. has local power series\ theorem holomorphic_power_series: assumes holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" proof - \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" proof have "cball z ((r + dist w z) / 2) \ ball z r" using w by (simp add: dist_commute field_sum_of_halves subset_eq) then show "f holomorphic_on cball z ((r + dist w z) / 2)" by (rule holomorphic_on_subset [OF holf]) have "r > 0" using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) then show "0 < (r + dist w z) / 2" by simp (use zero_le_dist [of w z] in linarith) qed (use w in \auto simp: dist_commute\) then have holf: "f holomorphic_on ball z r" using ball_subset_cball holomorphic_on_subset by blast have contf: "continuous_on (cball z r) f" by (simp add: holfc holomorphic_on_imp_continuous_on) have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" proof show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) qed (use w in \auto simp: dist_norm norm_minus_commute\) have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" unfolding uniform_limit_iff dist_norm proof clarify fix e::real assume "0 < e" have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto obtain n where n: "((r - k) / r) ^ n < e / B * k" using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force have "norm ((\k N" and r: "r = dist z u" for N u proof - have N: "((r - k) / r) ^ N < e / B * k" apply (rule le_less_trans [OF power_decreasing n]) using \n \ N\ k by auto have u [simp]: "(u \ z) \ (u \ w)" using \0 < r\ r w by auto have wzu_not1: "(w - z) / (u - z) \ 1" by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" using \0 < B\ apply (auto simp: geometric_sum [OF wzu_not1]) apply (simp add: field_simps norm_mult [symmetric]) done also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" by (simp add: algebra_simps) also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" by (simp add: norm_mult norm_power norm_minus_commute) also have "\ \ (((r - k)/r)^N) * B" using \0 < r\ w k apply (simp add: divide_simps) apply (rule mult_mono [OF power_mono]) apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) done also have "\ < e * k" using \0 < B\ N by (simp add: divide_simps) also have "\ \ e * norm (u - w)" using r kle \0 < e\ by (simp add: dist_commute dist_norm) finally show ?thesis by (simp add: field_split_simps norm_divide del: power_Suc) qed with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. norm ((\k\<^sub>F x in sequentially. contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" apply (rule eventuallyI) apply (subst contour_integral_sum, simp) using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) apply (simp only: contour_integral_lmul cint algebra_simps) done have cic: "\u. (\y. \k0 < r\ by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf]) have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) sums contour_integral (circlepath z r) (\u. f u/(u - w))" unfolding sums_def apply (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul] cic) using \0 < r\ apply auto done then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) sums (2 * of_real pi * \ * f w)" using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" by (rule sums_divide) then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) sums f w" by (simp add: field_simps) then show ?thesis by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) qed subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ text\ These weak Liouville versions don't even need the derivative formula.\ lemma Liouville_weak_0: assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" shows "f z = 0" proof (rule ccontr) assume fz: "f z \ 0" with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" by (auto simp: dist_norm) define R where "R = 1 + \B\ + norm z" have "R > 0" unfolding R_def proof - have "0 \ cmod z + \B\" by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) then show "0 < 1 + \B\ + cmod z" by linarith qed have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" apply (rule Cauchy_integral_circlepath) using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ done have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x unfolding R_def by (rule B) (use norm_triangle_ineq4 [of x z] in auto) with \R > 0\ fz show False using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) qed proposition Liouville_weak: assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" shows "f z = l" using Liouville_weak_0 [of "\z. f z - l"] - by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero) + by (simp add: assms holomorphic_on_diff LIM_zero) proposition Liouville_weak_inverse: assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" obtains z where "f z = 0" proof - { assume f: "\z. f z \ 0" have 1: "(\x. 1 / f x) holomorphic_on UNIV" - by (simp add: holomorphic_on_divide holomorphic_on_const assms f) + by (simp add: holomorphic_on_divide assms f) have 2: "((\x. 1 / f x) \ 0) at_infinity" apply (rule tendstoI [OF eventually_mono]) apply (rule_tac B="2/e" in unbounded) - apply (simp add: dist_norm norm_divide field_split_simps mult_ac) + apply (simp add: dist_norm norm_divide field_split_simps) done have False using Liouville_weak_0 [OF 1 2] f by simp } then show ?thesis using that by blast qed text\ In particular we get the Fundamental Theorem of Algebra.\ theorem fundamental_theorem_of_algebra: fixes a :: "nat \ complex" assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" obtains z where "(\i\n. a i * z^i) = 0" using assms proof (elim disjE bexE) assume "a 0 = 0" then show ?thesis by (auto simp: that [of 0]) next fix i assume i: "i \ {1..n}" and nz: "a i \ 0" have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" by (rule holomorphic_intros)+ show thesis proof (rule Liouville_weak_inverse [OF 1]) show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B using i polyfun_extremal nz by force qed (use that in auto) qed subsection\Weierstrass convergence theorem\ lemma holomorphic_uniform_limit: assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" and ulim: "uniform_limit (cball z r) f g F" and F: "\ trivial_limit F" obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) next case equal then show ?thesis by (force simp: holomorphic_on_def intro: that) next case greater have contg: "continuous_on (cball z r) g" using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast have "path_image (circlepath z r) \ cball z r" using \0 < r\ by auto then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" by (intro continuous_intros continuous_on_subset [OF contg]) have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" if w: "w \ ball z r" for w proof - define d where "d = (r - norm(w - z))" have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" apply (rule eventually_mono [OF cont]) using w apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) done have ul_less: "uniform_limit (sphere z r) (\n x. f n x / (x - w)) (\x. g x / (x - w)) F" using greater \0 < d\ apply (clarsimp simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps) apply (rule_tac e1="e * d" in eventually_mono [OF uniform_limitD [OF ulim]]) apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ done have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" proof (rule Lim_transform_eventually) show "\\<^sub>F x in F. contour_integral (circlepath z r) (\u. f x u / (u - w)) = 2 * of_real pi * \ * f x w" apply (rule eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]]) using w\0 < d\ d_def by auto qed (auto simp: cif_tends_cig) have "\e. 0 < e \ \\<^sub>F n in F. dist (f n w) (g w) < e" by (rule eventually_mono [OF uniform_limitD [OF ulim]]) (use w in auto) then have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" by (rule tendsto_mult_left [OF tendstoI]) then have "((\u. g u / (u - w)) has_contour_integral 2 * of_real pi * \ * g w) (circlepath z r)" using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w by fastforce then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" using has_contour_integral_div [where c = "2 * of_real pi * \"] by (force simp: field_simps) then show ?thesis by (simp add: dist_norm) qed show ?thesis using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] by (fastforce simp add: holomorphic_on_open contg intro: that) qed text\ Version showing that the limit is the limit of the derivatives.\ proposition has_complex_derivative_uniform_limit: fixes z::complex assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" and ulim: "uniform_limit (cball z r) f g F" and F: "\ trivial_limit F" and "0 < r" obtains g' where "continuous_on (cball z r) g" "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" proof - let ?conint = "contour_integral (circlepath z r)" have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] ulim F]; auto simp: holomorphic_on_open field_differentiable_def)+ then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" using DERIV_deriv_iff_has_field_derivative by (fastforce simp add: holomorphic_on_open) then have derg: "\x. x \ ball z r \ deriv g x = g' x" by (simp add: DERIV_imp_deriv) have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w proof - have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" if cont_fn: "continuous_on (cball z r) (f n)" and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n proof - have hol_fn: "f n holomorphic_on ball z r" using fnd by (force simp: holomorphic_on_open) have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" using DERIV_unique [OF fnd] w by blast show ?thesis by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] field_split_simps) qed define d where "d = (r - norm(w - z))^2" have "d > 0" using w by (simp add: dist_commute dist_norm d_def) have dle: "d \ cmod ((y - w)\<^sup>2)" if "r = cmod (z - y)" for y proof - have "w \ ball z (cmod (z - y))" using that w by fastforce then have "cmod (w - z) \ cmod (z - y)" by (simp add: dist_complex_def norm_minus_commute) moreover have "cmod (z - y) - cmod (w - z) \ cmod (y - w)" by (metis diff_add_cancel diff_add_eq_diff_diff_swap norm_minus_commute norm_triangle_ineq2) ultimately show ?thesis using that by (simp add: d_def norm_power power_mono) qed have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) have 2: "uniform_limit (sphere z r) (\n x. f n x / (x - w)\<^sup>2) (\x. g x / (x - w)\<^sup>2) F" unfolding uniform_limit_iff proof clarify fix e::real assume "0 < e" with \r > 0\ show "\\<^sub>F n in F. \x\sphere z r. dist (f n x / (x - w)\<^sup>2) (g x / (x - w)\<^sup>2) < e" apply (simp add: norm_divide field_split_simps sphere_def dist_norm) apply (rule eventually_mono [OF uniform_limitD [OF ulim], of "e*d"]) apply (simp add: \0 < d\) apply (force simp: dist_norm dle intro: less_le_trans) done qed have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" using Lim_null by (force intro!: tendsto_mult_right_zero) have "((\n. f' n w - g' w) \ 0) F" apply (rule Lim_transform_eventually [OF tendsto_0]) apply (force simp: divide_simps intro: eq_f' eventually_mono [OF cont]) done then show ?thesis using Lim_null by blast qed obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" by (blast intro: tends_f'n_g' g') then show ?thesis using g using that by blast qed subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ lemma holomorphic_uniform_sequence: assumes S: "open S" and hol_fn: "\n. (f n) holomorphic_on S" and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" shows "g holomorphic_on S" proof - have "\f'. (g has_field_derivative f') (at z)" if "z \ S" for z proof - obtain r where "0 < r" and r: "cball z r \ S" and ul: "uniform_limit (cball z r) f g sequentially" using ulim_g [OF \z \ S\] by blast have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" proof (intro eventuallyI conjI) show "continuous_on (cball z r) (f x)" for x using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast show "f x holomorphic_on ball z r" for x by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) qed show ?thesis apply (rule holomorphic_uniform_limit [OF *]) using \0 < r\ centre_in_ball ul apply (auto simp: holomorphic_on_open) done qed with S show ?thesis by (simp add: holomorphic_on_open) qed lemma has_complex_derivative_uniform_sequence: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ ((f n) has_field_derivative f' n x) (at x)" and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" shows "\g'. \x \ S. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" proof - have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ S" for z proof - obtain r where "0 < r" and r: "cball z r \ S" and ul: "uniform_limit (cball z r) f g sequentially" using ulim_g [OF \z \ S\] by blast have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" proof (intro eventuallyI conjI ballI) show "continuous_on (cball z r) (f x)" for x by (meson S continuous_on_subset hfd holomorphic_on_imp_continuous_on holomorphic_on_open r) show "w \ ball z r \ (f x has_field_derivative f' x w) (at w)" for w x using ball_subset_cball hfd r by blast qed show ?thesis by (rule has_complex_derivative_uniform_limit [OF *, of g]) (use \0 < r\ ul in \force+\) qed show ?thesis by (rule bchoice) (blast intro: y) qed subsection\On analytic functions defined by a series\ lemma series_and_derivative_comparison: fixes S :: "complex set" assumes S: "open S" and h: "summable h" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\\<^sub>F n in sequentially. \x\S. norm (f n x) \ h n" obtains g g' where "\x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" proof - obtain g where g: "uniform_limit S (\n x. \id>0. cball x d \ S \ uniform_limit (cball x d) (\n x. \i S" for x proof - obtain d where "d>0" and d: "cball x d \ S" using open_contains_cball [of "S"] \x \ S\ S by blast show ?thesis proof (intro conjI exI) show "uniform_limit (cball x d) (\n x. \id > 0\ d in auto) qed have "\x. x \ S \ (\n. \i g x" by (metis tendsto_uniform_limitI [OF g]) moreover have "\g'. \x\S. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" by (rule has_complex_derivative_uniform_sequence [OF S]) (auto intro: * hfd DERIV_sum)+ ultimately show ?thesis by (metis sums_def that) qed text\A version where we only have local uniform/comparative convergence.\ lemma series_and_derivative_comparison_local: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ (\\<^sub>F n in sequentially. \y\ball x d \ S. norm (f n y) \ h n)" shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" proof - have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" if "z \ S" for z proof - obtain d h where "0 < d" "summable h" and le_h: "\\<^sub>F n in sequentially. \y\ball z d \ S. norm (f n y) \ h n" using to_g \z \ S\ by meson then obtain r where "r>0" and r: "ball z r \ ball z d \ S" using \z \ S\ S by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) have 1: "open (ball z d \ S)" by (simp add: open_Int S) have 2: "\n x. x \ ball z d \ S \ (f n has_field_derivative f' n x) (at x)" by (auto simp: hfd) obtain g g' where gg': "\x \ ball z d \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) then have "(\n. f' n z) sums g' z" by (meson \0 < r\ centre_in_ball contra_subsetD r) moreover have "(\n. f n z) sums (\n. f n z)" using summable_sums centre_in_ball \0 < d\ \summable h\ le_h by (metis (full_types) Int_iff gg' summable_def that) moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" proof (rule has_field_derivative_transform_within) show "\x. dist x z < r \ g x = (\n. f n x)" by (metis subsetD dist_commute gg' mem_ball r sums_unique) qed (use \0 < r\ gg' \z \ S\ \0 < d\ in auto) ultimately show ?thesis by auto qed then show ?thesis by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson qed text\Sometimes convenient to compare with a complex series of positive reals. (?)\ lemma series_and_derivative_comparison_complex: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" apply (rule series_and_derivative_comparison_local [OF S hfd], assumption) apply (rule ex_forward [OF to_g], assumption) apply (erule exE) apply (rule_tac x="Re \ h" in exI) apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff) done text\Sometimes convenient to compare with a complex series of positive reals. (?)\ lemma series_differentiable_comparison_complex: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ f n field_differentiable (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" obtains g where "\x \ S. ((\n. f n x) sums g x) \ g field_differentiable (at x)" proof - have hfd': "\n x. x \ S \ (f n has_field_derivative deriv (f n) x) (at x)" using hfd field_differentiable_derivI by blast have "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. deriv (f n) x) sums g' x) \ (g has_field_derivative g' x) (at x)" by (metis series_and_derivative_comparison_complex [OF S hfd' to_g]) then show ?thesis using field_differentiable_def that by blast qed text\In particular, a power series is analytic inside circle of convergence.\ lemma power_series_and_derivative_0: fixes a :: "nat \ complex" and r::real assumes "summable (\n. a n * r^n)" shows "\g g'. \z. cmod z < r \ ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" proof (cases "0 < r") case True have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" by (rule derivative_eq_intros | simp)+ have y_le: "\cmod (z - y) * 2 < r - cmod z\ \ cmod y \ cmod (of_real r + of_real (cmod z)) / 2" for z y using \r > 0\ apply (auto simp: algebra_simps norm_mult norm_divide norm_power simp flip: of_real_add) using norm_triangle_ineq2 [of y z] apply (simp only: diff_le_eq norm_minus_commute mult_2) done have "summable (\n. a n * complex_of_real r ^ n)" using assms \r > 0\ by simp moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" using \r > 0\ by (simp flip: of_real_add) ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" by (rule power_series_conv_imp_absconv_weak) have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" apply (rule series_and_derivative_comparison_complex [OF open_ball der]) apply (rule_tac x="(r - norm z)/2" in exI) apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) using \r > 0\ apply (auto simp: sum eventually_sequentially norm_mult norm_power dist_norm intro!: mult_left_mono power_mono y_le) done then show ?thesis by (simp add: ball_def) next case False then show ?thesis apply (simp add: not_less) using less_le_trans norm_not_less_zero by blast qed proposition\<^marker>\tag unimportant\ power_series_and_derivative: fixes a :: "nat \ complex" and r::real assumes "summable (\n. a n * r^n)" obtains g g' where "\z \ ball w r. ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" using power_series_and_derivative_0 [OF assms] apply clarify apply (rule_tac g="(\z. g(z - w))" in that) using DERIV_shift [where z="-w"] apply (auto simp: norm_minus_commute Ball_def dist_norm) done proposition\<^marker>\tag unimportant\ power_series_holomorphic: assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" shows "f holomorphic_on ball z r" proof - have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w proof - have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" proof - have wz: "cmod (w - z) < r" using w by (auto simp: field_split_simps dist_norm norm_minus_commute) then have "0 \ r" by (meson less_eq_real_def norm_ge_zero order_trans) show ?thesis using w by (simp add: dist_norm \0\r\ flip: of_real_add) qed have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" using assms [OF inb] by (force simp: summable_def dist_norm) obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ (\n. a n * (u - z) ^ n) sums g u \ (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" by (rule power_series_and_derivative [OF sum, of z]) fastforce have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u proof - have less: "cmod (z - u) * 2 < cmod (z - w) + r" using that dist_triangle2 [of z u w] by (simp add: dist_norm [symmetric] algebra_simps) show ?thesis apply (rule sums_unique2 [of "\n. a n*(u - z)^n"]) using gg' [of u] less w apply (auto simp: assms dist_norm) done qed have "(f has_field_derivative g' w) (at w)" by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"]) (use w gg' [of w] in \(force simp: dist_norm)+\) then show ?thesis .. qed then show ?thesis by (simp add: holomorphic_on_open) qed corollary holomorphic_iff_power_series: "f holomorphic_on ball z r \ (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" apply (intro iffI ballI holomorphic_power_series, assumption+) apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) done lemma power_series_analytic: "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" by (force simp: analytic_on_open intro!: power_series_holomorphic) lemma analytic_iff_power_series: "f analytic_on ball z r \ (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" by (simp add: analytic_on_open holomorphic_iff_power_series) subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ lemma holomorphic_fun_eq_on_ball: "\f holomorphic_on ball z r; g holomorphic_on ball z r; w \ ball z r; \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ \ f w = g w" apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) apply (auto simp: holomorphic_iff_power_series) done lemma holomorphic_fun_eq_0_on_ball: "\f holomorphic_on ball z r; w \ ball z r; \n. (deriv ^^ n) f z = 0\ \ f w = 0" apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) apply (auto simp: holomorphic_iff_power_series) done lemma holomorphic_fun_eq_0_on_connected: assumes holf: "f holomorphic_on S" and "open S" and cons: "connected S" and der: "\n. (deriv ^^ n) f z = 0" and "z \ S" "w \ S" shows "f w = 0" proof - have *: "ball x e \ (\n. {w \ S. (deriv ^^ n) f w = 0})" if "\u. (deriv ^^ u) f x = 0" "ball x e \ S" for x e proof - have "\x' n. dist x x' < e \ (deriv ^^ n) f x' = 0" apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) apply (rule holomorphic_on_subset [OF holf]) using that apply simp_all by (metis funpow_add o_apply) with that show ?thesis by auto qed have 1: "openin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" apply (rule open_subset, force) using \open S\ apply (simp add: open_contains_ball Ball_def) apply (erule all_forward) using "*" by auto blast+ have 2: "closedin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" using assms by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) obtain e where "e>0" and e: "ball w e \ S" using openE [OF \open S\ \w \ S\] . then have holfb: "f holomorphic_on ball w e" using holf holomorphic_on_subset by blast have 3: "(\n. {w \ S. (deriv ^^ n) f w = 0}) = S \ f w = 0" using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) show ?thesis using cons der \z \ S\ apply (simp add: connected_clopen) apply (drule_tac x="\n. {w \ S. (deriv ^^ n) f w = 0}" in spec) apply (auto simp: 1 2 3) done qed lemma holomorphic_fun_eq_on_connected: assumes "f holomorphic_on S" "g holomorphic_on S" and "open S" "connected S" and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" and "z \ S" "w \ S" shows "f w = g w" proof (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" S z, simplified]) show "(\x. f x - g x) holomorphic_on S" by (intro assms holomorphic_intros) show "\n. (deriv ^^ n) (\x. f x - g x) z = 0" using assms higher_deriv_diff by auto qed (use assms in auto) lemma holomorphic_fun_eq_const_on_connected: assumes holf: "f holomorphic_on S" and "open S" and cons: "connected S" and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" and "z \ S" "w \ S" shows "f w = f z" proof (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" S z, simplified]) show "(\w. f w - f z) holomorphic_on S" by (intro assms holomorphic_intros) show "\n. (deriv ^^ n) (\w. f w - f z) z = 0" by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) qed (use assms in auto) subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ lemma pole_lemma: assumes holf: "f holomorphic_on S" and a: "a \ interior S" shows "(\z. if z = a then deriv f a else (f z - f a) / (z - a)) holomorphic_on S" (is "?F holomorphic_on S") proof - have F1: "?F field_differentiable (at u within S)" if "u \ S" "u \ a" for u proof - have fcd: "f field_differentiable at u within S" using holf holomorphic_on_def by (simp add: \u \ S\) have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within S" by (rule fcd derivative_intros | simp add: that)+ have "0 < dist a u" using that dist_nz by blast then show ?thesis by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ S\) qed have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ S" for e proof - have holfb: "f holomorphic_on ball a e" by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) have 2: "?F holomorphic_on ball a e - {a}" apply (simp add: holomorphic_on_def flip: field_differentiable_def) using mem_ball that apply (auto intro: F1 field_differentiable_within_subset) done have "isCont (\z. if z = a then deriv f a else (f z - f a) / (z - a)) x" if "dist a x < e" for x proof (cases "x=a") case True then have "f field_differentiable at a" using holfb \0 < e\ holomorphic_on_imp_differentiable_at by auto with True show ?thesis by (auto simp: continuous_at has_field_derivative_iff simp flip: DERIV_deriv_iff_field_differentiable elim: rev_iffD1 [OF _ LIM_equal]) next case False with 2 that show ?thesis by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) qed then have 1: "continuous_on (ball a e) ?F" by (clarsimp simp: continuous_on_eq_continuous_at) have "?F holomorphic_on ball a e" by (auto intro: no_isolated_singularity [OF 1 2]) with that show ?thesis by (simp add: holomorphic_on_open field_differentiable_def [symmetric] field_differentiable_at_within) qed show ?thesis proof fix x assume "x \ S" show "?F field_differentiable at x within S" proof (cases "x=a") case True then show ?thesis using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) next case False with F1 \x \ S\ show ?thesis by blast qed qed qed lemma pole_theorem: assumes holg: "g holomorphic_on S" and a: "a \ interior S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) holomorphic_on S" using pole_lemma [OF holg a] by (rule holomorphic_transform) (simp add: eq field_split_simps) lemma pole_lemma_open: assumes "f holomorphic_on S" "open S" shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on S" proof (cases "a \ S") case True with assms interior_eq pole_lemma show ?thesis by fastforce next case False with assms show ?thesis apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) apply (rule derivative_intros | force)+ done qed lemma pole_theorem_open: assumes holg: "g holomorphic_on S" and S: "open S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) holomorphic_on S" using pole_lemma_open [OF holg S] by (rule holomorphic_transform) (auto simp: eq divide_simps) lemma pole_theorem_0: assumes holg: "g holomorphic_on S" and a: "a \ interior S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f holomorphic_on S" using pole_theorem [OF holg a eq] by (rule holomorphic_transform) (auto simp: eq field_split_simps) lemma pole_theorem_open_0: assumes holg: "g holomorphic_on S" and S: "open S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f holomorphic_on S" using pole_theorem_open [OF holg S eq] by (rule holomorphic_transform) (auto simp: eq field_split_simps) lemma pole_theorem_analytic: assumes g: "g analytic_on S" and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" (is "?F analytic_on S") unfolding analytic_on_def proof fix x assume "x \ S" with g obtain e where "0 < e" and e: "g holomorphic_on ball x e" by (auto simp add: analytic_on_def) obtain d where "0 < d" and d: "\w. w \ ball x d - {a} \ g w = (w - a) * f w" using \x \ S\ eq by blast have "?F holomorphic_on ball x (min d e)" using d e \x \ S\ by (fastforce simp: holomorphic_on_subset subset_ball intro!: pole_theorem_open) then show "\e>0. ?F holomorphic_on ball x e" using \0 < d\ \0 < e\ not_le by fastforce qed lemma pole_theorem_analytic_0: assumes g: "g analytic_on S" and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" and [simp]: "f a = deriv g a" "g a = 0" shows "f analytic_on S" proof - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" by auto show ?thesis using pole_theorem_analytic [OF g eq] by simp qed lemma pole_theorem_analytic_open_superset: assumes g: "g analytic_on S" and "S \ T" "open T" and eq: "\z. z \ T - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" proof (rule pole_theorem_analytic [OF g]) fix z assume "z \ S" then obtain e where "0 < e" and e: "ball z e \ T" using assms openE by blast then show "\d>0. \w\ball z d - {a}. g w = (w - a) * f w" using eq by auto qed lemma pole_theorem_analytic_open_superset_0: assumes g: "g analytic_on S" "S \ T" "open T" "\z. z \ T - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f analytic_on S" proof - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" by auto have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" by (rule pole_theorem_analytic_open_superset [OF g]) then show ?thesis by simp qed subsection\General, homology form of Cauchy's theorem\ text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ lemma contour_integral_continuous_on_linepath_2D: assumes "open U" and cont_dw: "\w. w \ U \ F w contour_integrable_on (linepath a b)" and cond_uu: "continuous_on (U \ U) (\(x,y). F x y)" and abu: "closed_segment a b \ U" shows "continuous_on U (\w. contour_integral (linepath a b) (F w))" proof - have *: "\d>0. \x'\U. dist x' w < d \ dist (contour_integral (linepath a b) (F x')) (contour_integral (linepath a b) (F w)) \ \" if "w \ U" "0 < \" "a \ b" for w \ proof - obtain \ where "\>0" and \: "cball w \ \ U" using open_contains_cball \open U\ \w \ U\ by force let ?TZ = "cball w \ \ closed_segment a b" have "uniformly_continuous_on ?TZ (\(x,y). F x y)" proof (rule compact_uniformly_continuous) show "continuous_on ?TZ (\(x,y). F x y)" by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \ abu in blast) show "compact ?TZ" by (simp add: compact_Times) qed then obtain \ where "\>0" and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) using \0 < \\ \a \ b\ by auto have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" for x1 x2 x1' x2' using \ [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm) have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' proof - have "(\x. F x' x - F w x) contour_integrable_on linepath a b" by (simp add: \w \ U\ cont_dw contour_integrable_diff that) then have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) using \0 < \\ \0 < \\ that apply (auto simp: norm_minus_commute) done also have "\ = \" using \a \ b\ by simp finally show ?thesis . qed show ?thesis apply (rule_tac x="min \ \" in exI) using \0 < \\ \0 < \\ apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ U\ intro: le_ee) done qed show ?thesis proof (cases "a=b") case True then show ?thesis by simp next case False show ?thesis by (rule continuous_onI) (use False in \auto intro: *\) qed qed text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ lemma Cauchy_integral_formula_global_weak: assumes "open U" and holf: "f holomorphic_on U" and z: "z \ U" and \: "polynomial_function \" and pasz: "path_image \ \ U - {z}" and loop: "pathfinish \ = pathstart \" and zero: "\w. w \ U \ winding_number \ w = 0" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" using has_vector_derivative_polynomial_function [OF \] by blast then have "bounded(path_image \')" by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" using bounded_pos by force define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" have "path \" "valid_path \" using \ by (auto simp: path_polynomial_function valid_path_polynomial_function) then have ov: "open v" by (simp add: v_def open_winding_number_levelsets loop) have uv_Un: "U \ v = UNIV" using pasz zero by (auto simp: v_def) have conf: "continuous_on U f" by (metis holf holomorphic_on_imp_continuous_on) have hol_d: "(d y) holomorphic_on U" if "y \ U" for y proof - have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U" by (simp add: holf pole_lemma_open \open U\) then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \open U\ by fastforce then have "continuous_on U (d y)" apply (simp add: d_def continuous_on_eq_continuous_at \open U\, clarify) using * holomorphic_on_def by (meson field_differentiable_within_open field_differentiable_imp_continuous_at \open U\) moreover have "d y holomorphic_on U - {y}" proof - have "\w. w \ U - {y} \ (\w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w" apply (rule_tac d="dist w y" and f = "\w. (f w - f y)/(w - y)" in field_differentiable_transform_within) apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) using \open U\ holf holomorphic_on_imp_differentiable_at by blast then show ?thesis unfolding field_differentiable_def by (simp add: d_def holomorphic_on_open \open U\ open_delete) qed ultimately show ?thesis by (rule no_isolated_singularity) (auto simp: \open U\) qed have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y proof (rule contour_integrable_holomorphic_simple [where S = "U-{y}"]) show "(\x. (f x - f y) / (x - y)) holomorphic_on U - {y}" by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) show "path_image \ \ U - {y}" using pasz that by blast qed (auto simp: \open U\ open_delete \valid_path \\) define h where "h z = (if z \ U then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z have U: "((d z) has_contour_integral h z) \" if "z \ U" for z proof - have "d z holomorphic_on U" by (simp add: hol_d that) with that show ?thesis apply (simp add: h_def) by (meson Diff_subset \open U\ \valid_path \\ contour_integrable_holomorphic_simple has_contour_integral_integral pasz subset_trans) qed have V: "((\w. f w / (w - z)) has_contour_integral h z) \" if z: "z \ v" for z proof - have 0: "0 = (f z) * 2 * of_real (2 * pi) * \ * winding_number \ z" using v_def z by auto then have "((\x. 1 / (x - z)) has_contour_integral 0) \" using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" using has_contour_integral_lmul by fastforce then have "((\x. f z / (x - z)) has_contour_integral 0) \" by (simp add: field_split_simps) moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" using z apply (auto simp: v_def) apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) done ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" by (rule has_contour_integral_add) have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" if "z \ U" using * by (auto simp: divide_simps has_contour_integral_eq) moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" if "z \ U" apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]]) using U pasz \valid_path \\ that apply (auto intro: holomorphic_on_imp_continuous_on hol_d) apply (rule continuous_intros conf holomorphic_intros holf assms | force)+ done ultimately show ?thesis using z by (simp add: h_def) qed have znot: "z \ path_image \" using pasz by blast obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - U \ d0 \ dist x y" using separate_compact_closed [of "path_image \" "-U"] pasz \open U\ by (fastforce simp add: \path \\ compact_path_image) obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ U" apply (rule that [of "d0/2"]) using \0 < d0\ apply (auto simp: dist_norm dest: d0) done have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" apply (rule_tac x=x in exI) apply (rule_tac x="x'-x" in exI) apply (force simp: dist_norm) done then have 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" apply (clarsimp simp add: mem_interior) using \0 < dd\ apply (rule_tac x="dd/2" in exI, auto) done obtain T where "compact T" and subt: "path_image \ \ interior T" and T: "T \ U" apply (rule that [OF _ 1]) apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) apply (rule order_trans [OF _ dd]) using \0 < dd\ by fastforce obtain L where "L>0" and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ cmod (contour_integral \ f) \ L * B" using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] by blast have "bounded(f ` T)" by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" by (auto simp: bounded_pos) obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" using \compact T\ bounded_pos compact_imp_bounded by force have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y proof - have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp with le have ybig: "norm y > C" by force with C have "y \ T" by force then have ynot: "y \ path_image \" using subt interior_subset by blast have [simp]: "winding_number \ y = 0" apply (rule winding_number_zero_outside [of _ "cball 0 C"]) using ybig interior_subset subt apply (force simp: loop \path \\ dist_norm intro!: C)+ done have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) have holint: "(\w. f w / (w - y)) holomorphic_on interior T" apply (rule holomorphic_on_divide) using holf holomorphic_on_subset interior_subset T apply blast apply (rule holomorphic_intros)+ using \y \ T\ interior_subset by auto have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z proof - have "D * L / e + cmod z \ cmod y" using le C [of z] z using interior_subset by force then have DL2: "D * L / e \ cmod (z - y)" using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) also have "\ \ D * (e / L / D)" apply (rule mult_mono) using that D interior_subset apply blast using \L>0\ \e>0\ \D>0\ DL2 - apply (auto simp: norm_divide field_split_simps algebra_simps) + apply (auto simp: norm_divide field_split_simps) done finally show ?thesis . qed have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" by (simp add: dist_norm) also have "\ \ L * (D * (e / L / D))" by (rule L [OF holint leD]) also have "\ = e" using \L>0\ \0 < D\ by auto finally show ?thesis . qed then have "(h \ 0) at_infinity" by (meson Lim_at_infinityI) moreover have "h holomorphic_on UNIV" proof - have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" if "x \ U" "z \ U" "x \ z" for x z using that conf apply (simp add: split_def continuous_on_eq_continuous_at \open U\) apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ done have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" by (rule continuous_intros)+ have open_uu_Id: "open (U \ U - Id)" apply (rule open_Diff) apply (simp add: open_Times \open U\) using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] apply (auto simp: Id_fstsnd_eq algebra_simps) done have con_derf: "continuous (at z) (deriv f)" if "z \ U" for z apply (rule continuous_on_interior [of U]) apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open U\) by (simp add: interior_open that \open U\) have tendsto_f': "((\(x,y). if y = x then deriv f (x) else (f (y) - f (x)) / (y - x)) \ deriv f x) (at (x, x) within U \ U)" if "x \ U" for x proof (rule Lim_withinI) fix e::real assume "0 < e" obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] by (metis UNIV_I dist_norm) obtain k2 where "k2>0" and k2: "ball x k2 \ U" by (blast intro: openE [OF \open U\] \x \ U\) have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" for x' z' proof - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w apply (drule segment_furthest_le [where y=x]) by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) have "closed_segment x' z' \ U" by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" by (rule has_contour_integral_div) have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] \e > 0\ \z' \ x'\ apply (auto simp: norm_divide divide_simps derf_le) done also have "\ \ e" using \0 < e\ by simp finally show ?thesis . qed show "\d>0. \xa\U \ U. 0 < dist xa (x, x) \ dist xa (x, x) < d \ dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" apply (rule_tac x="min k1 k2" in exI) using \k1>0\ \k2>0\ \e>0\ apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) done qed have con_pa_f: "continuous_on (path_image \) f" by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T) have le_B: "\T. T \ {0..1} \ cmod (vector_derivative \ (at T)) \ B" apply (rule B) using \' using path_image_def vector_derivative_at by fastforce have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" by (simp add: V) have cond_uu: "continuous_on (U \ U) (\(x,y). d x y)" apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) using con_ff apply (auto simp: continuous_within) done have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w proof - have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) have **: "\x. \x \ U; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) apply (rule \open U\ derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ done show ?thesis unfolding d_def apply (rule no_isolated_singularity [OF * _ \open U\, where K = "{w}"]) apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \open U\ **) done qed { fix a b assume abu: "closed_segment a b \ U" then have "\w. w \ U \ (\z. d z w) contour_integrable_on (linepath a b)" by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) then have cont_cint_d: "continuous_on U (\w. contour_integral (linepath a b) (\z. d z w))" apply (rule contour_integral_continuous_on_linepath_2D [OF \open U\ _ _ abu]) apply (auto intro: continuous_on_swap_args cond_uu) done have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" proof (rule continuous_on_compose) show "continuous_on {0..1} \" using \path \\ path_def by blast show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" using pasz unfolding path_image_def by (auto intro!: continuous_on_subset [OF cont_cint_d]) qed have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" apply (simp add: contour_integrable_on) apply (rule integrable_continuous_real) apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) using pf\' by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" using abu by (force simp: h_def intro: contour_integral_eq) also have "\ = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" apply (rule contour_integral_swap) apply (rule continuous_on_subset [OF cond_uu]) using abu pasz \valid_path \\ apply (auto intro!: continuous_intros) by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) finally have cint_h_eq: "contour_integral (linepath a b) h = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . note cint_cint cint_h_eq } note cint_h = this have conthu: "continuous_on U h" proof (simp add: continuous_on_sequentially, clarify) fix a x assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" by (meson U contour_integrable_on_def eventuallyI) obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" unfolding uniform_limit_iff dist_norm proof clarify fix ee::real assume "0 < ee" show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" proof - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) using dd pasz \valid_path \\ apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) done then obtain kk where "kk>0" and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" for w z using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) show ?thesis using ax unfolding lim_sequentially eventually_sequentially apply (drule_tac x="min dd kk" in spec) using \dd > 0\ \kk > 0\ apply (fastforce simp: kk dist_norm) done qed qed have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" by (simp add: h_def x) then show "(h \ a) \ h x" by (simp add: h_def x au o_def) qed show ?thesis proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) fix z0 consider "z0 \ v" | "z0 \ U" using uv_Un by blast then show "h field_differentiable at z0" proof cases assume "z0 \ v" then show ?thesis using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ by (auto simp: field_differentiable_def v_def) next assume "z0 \ U" then obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c proof - have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" using hol_dw holomorphic_on_imp_continuous_on \open U\ by (auto intro!: contour_integrable_holomorphic_simple) have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" using that e segments_subset_convex_hull by fastforce+ have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) apply (rule holomorphic_on_subset [OF hol_dw]) using e abc_subset by auto have "contour_integral \ (\x. contour_integral (linepath a b) (\z. d z x) + (contour_integral (linepath b c) (\z. d z x) + contour_integral (linepath c a) (\z. d z x))) = 0" apply (rule contour_integral_eq_0) using abc pasz U apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ done then show ?thesis by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) qed show ?thesis using e \e > 0\ by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic Morera_triangle continuous_on_subset [OF conthu] *) qed qed qed ultimately have [simp]: "h z = 0" for z by (meson Liouville_weak) have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" by (metis mult.commute has_contour_integral_lmul) then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" by (simp add: field_split_simps) moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) show ?thesis using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) qed theorem Cauchy_integral_formula_global: assumes S: "open S" and holf: "f holomorphic_on S" and z: "z \ S" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" and zero: "\w. w \ S \ winding_number \ w = 0" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have "path \" using vpg by (blast intro: valid_path_imp_path) have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) obtain d where "d>0" and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; pathstart h = pathstart g \ pathfinish h = pathfinish g\ \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis obtain p where polyp: "polynomial_function p" and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast then have ploop: "pathfinish p = pathstart p" using loop by auto have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast have [simp]: "z \ path_image \" using pasz by blast have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" using pf ps led d [OF vpg vpp] \d > 0\ by auto have wn_eq: "winding_number p z = winding_number \ z" using vpp paps by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) have "winding_number p w = winding_number \ w" if "w \ S" for w proof - have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) have "w \ path_image p" "w \ path_image \" using paps pasz that by auto then show ?thesis using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) qed then have wn0: "\w. w \ S \ winding_number p w = 0" by (simp add: zero) show ?thesis using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) qed theorem Cauchy_theorem_global: assumes S: "open S" and holf: "f holomorphic_on S" and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" and pas: "path_image \ \ S" and zero: "\w. w \ S \ winding_number \ w = 0" shows "(f has_contour_integral 0) \" proof - obtain z where "z \ S" and znot: "z \ path_image \" proof - have "compact (path_image \)" using compact_valid_path_image vpg by blast then have "path_image \ \ S" by (metis (no_types) compact_open path_image_nonempty S) with pas show ?thesis by (blast intro: that) qed then have pasz: "path_image \ \ S - {z}" using pas by blast have hol: "(\w. (w - z) * f w) holomorphic_on S" by (rule holomorphic_intros holf)+ show ?thesis using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] by (auto simp: znot elim!: has_contour_integral_eq) qed corollary Cauchy_theorem_global_outside: assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" "\w. w \ S \ w \ outside(path_image \)" shows "(f has_contour_integral 0) \" by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) lemma simply_connected_imp_winding_number_zero: assumes "simply_connected S" "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" shows "winding_number g z = 0" proof - have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" by (rule winding_number_homotopic_paths) also have "\ = 0" using assms by (force intro: winding_number_trivial) finally show ?thesis . qed lemma Cauchy_theorem_simply_connected: assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" "path_image g \ S" "pathfinish g = pathstart g" shows "(f has_contour_integral 0) g" using assms apply (simp add: simply_connected_eq_contractible_path) apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] homotopic_paths_imp_homotopic_loops) using valid_path_imp_path by blast proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: assumes A: "convex A" "open A" and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" and z0: "z0 \ A" obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" proof - note f' = holomorphic_derivI [OF f(1) A(2)] obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" proof (rule holomorphic_convex_primitive' [OF A]) show "(\x. deriv f x / f x) holomorphic_on A" by (intro holomorphic_intros f A) qed (auto simp: A at_within_open[of _ A]) define h where "h = (\x. -g z0 + ln (f z0) + g x)" from g and A have g_holo: "g holomorphic_on A" by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) hence h_holo: "h holomorphic_on A" by (auto simp: h_def intro!: holomorphic_intros) have "\c. \x\A. f x / exp (h x) - 1 = c" proof (rule has_field_derivative_zero_constant, goal_cases) case (2 x) note [simp] = at_within_open[OF _ \open A\] from 2 and z0 and f show ?case by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') qed fact+ then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" by blast from c[OF z0] and z0 and f have "c = 0" by (simp add: h_def) with c have "\x. x \ A \ exp (h x) = f x" by simp from that[OF h_holo this] show ?thesis . qed end diff --git a/src/HOL/Analysis/Change_Of_Vars.thy b/src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy +++ b/src/HOL/Analysis/Change_Of_Vars.thy @@ -1,3474 +1,3474 @@ (* Title: HOL/Analysis/Change_Of_Vars.thy Authors: LC Paulson, based on material from HOL Light *) section\Change of Variables Theorems\ theory Change_Of_Vars imports Vitali_Covering_Theorem Determinants begin subsection \Measurable Shear and Stretch\ proposition fixes a :: "real^'n" assumes "m \ n" and ab_ne: "cbox a b \ {}" and an: "0 \ a$n" shows measurable_shear_interval: "(\x. \ i. if i = m then x$m + x$n else x$i) ` (cbox a b) \ lmeasurable" (is "?f ` _ \ _") and measure_shear_interval: "measure lebesgue ((\x. \ i. if i = m then x$m + x$n else x$i) ` cbox a b) = measure lebesgue (cbox a b)" (is "?Q") proof - have lin: "linear ?f" by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" by (simp add: lin measurable_linear_image_interval) let ?c = "\ i. if i = m then b$m + b$n else b$i" let ?mn = "axis m 1 - axis n (1::real)" have eq1: "measure lebesgue (cbox a ?c) = measure lebesgue (?f ` cbox a b) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a$m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m})" proof (rule measure_Un3_negligible) show "cbox a ?c \ {x. ?mn \ x \ a$m} \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "negligible {x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ a $ m}) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible ((?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ a $ m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ b$m}) \ {x. ?mn \ x = b$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m})) \ {x. ?mn \ x = b$m}" using \m \ n\ ab_ne apply (auto simp: algebra_simps mem_box_cart inner_axis') apply (drule_tac x=m in spec)+ apply simp done ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) show "?f ` cbox a b \ cbox a ?c \ {x. ?mn \ x \ a $ m} \ cbox a ?c \ {x. ?mn \ x \ b$m} = cbox a ?c" (is "?lhs = _") proof show "?lhs \ cbox a ?c" by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) show "cbox a ?c \ ?lhs" apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) apply (auto simp: mem_box_cart split: if_split_asm) done qed qed (fact fab) let ?d = "\ i. if i = m then a $ m - b $ m else 0" have eq2: "measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a $ m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m}) = measure lebesgue (cbox a (\ i. if i = m then a $ m + b $ n else b $ i))" proof (rule measure_translate_add[of "cbox a ?c \ {x. ?mn \ x \ a$m}" "cbox a ?c \ {x. ?mn \ x \ b$m}" "(\ i. if i = m then a$m - b$m else 0)" "cbox a (\ i. if i = m then a$m + b$n else b$i)"]) show "(cbox a ?c \ {x. ?mn \ x \ a$m}) \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "\x. \x $ n + a $ m \ x $ m\ \ x \ (+) (\ i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \ x $ m}" using \m \ n\ by (rule_tac x="x - (\ i. if i = m then a$m - b$m else 0)" in image_eqI) (simp_all add: mem_box_cart) then have imeq: "(+) ?d ` {x. b $ m \ ?mn \ x} = {x. a $ m \ ?mn \ x}" using \m \ n\ by (auto simp: mem_box_cart inner_axis' algebra_simps) have "\x. \0 \ a $ n; x $ n + a $ m \ x $ m; \i. i \ m \ a $ i \ x $ i \ x $ i \ b $ i\ \ a $ m \ x $ m" using \m \ n\ by force then have "(+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {x. a $ m \ ?mn \ x}" using an ab_ne apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq) apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib) by (metis (full_types) add_mono mult_2_right) then show "cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs") using an \m \ n\ apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force) apply (drule_tac x=n in spec)+ by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1)) have "negligible{x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x})) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}))" by (rule negligible_subset) qed have ac_ne: "cbox a ?c \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have ax_ne: "cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\ i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove if_distrib [of "\u. u - z" for z] prod.remove) show ?Q using eq1 eq2 eq3 by (simp add: algebra_simps) qed proposition fixes S :: "(real^'n) set" assumes "S \ lmeasurable" shows measurable_stretch: "((\x. \ k. m k * x$k) ` S) \ lmeasurable" (is "?f ` S \ _") and measure_stretch: "measure lebesgue ((\x. \ k. m k * x$k) ` S) = \prod m UNIV\ * measure lebesgue S" (is "?MEQ") proof - have "(?f ` S) \ lmeasurable \ ?MEQ" proof (cases "\k. m k \ 0") case True have m0: "0 < \prod m UNIV\" using True by simp have "(indicat_real (?f ` S) has_integral \prod m UNIV\ * measure lebesgue S) UNIV" proof (clarsimp simp add: has_integral_alt [where i=UNIV]) fix e :: "real" assume "e > 0" have "(indicat_real S has_integral (measure lebesgue S)) UNIV" using assms lmeasurable_iff_has_integral by blast then obtain B where "B>0" and B: "\a b. ball 0 B \ cbox a b \ \z. (indicat_real S has_integral z) (cbox a b) \ \z - measure lebesgue S\ < e / \prod m UNIV\" by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \e > 0\) show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (indicat_real (?f ` S) has_integral z) (cbox a b) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" proof (intro exI conjI allI) let ?C = "Max (range (\k. \m k\)) * B" show "?C > 0" using True \B > 0\ by (simp add: Max_gr_iff) show "ball 0 ?C \ cbox u v \ (\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" for u v proof assume uv: "ball 0 ?C \ cbox u v" with \?C > 0\ have cbox_ne: "cbox u v \ {}" using centre_in_ball by blast let ?\ = "\k. u$k / m k" let ?\ = "\k. v$k / m k" have invm0: "\k. inverse (m k) \ 0" using True by auto have "ball 0 B \ (\x. \ k. x $ k / m k) ` ball 0 ?C" proof clarsimp fix x :: "real^'n" assume x: "norm x < B" have [simp]: "\Max (range (\k. \m k\))\ = Max (range (\k. \m k\))" by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI) have "norm (\ k. m k * x $ k) \ norm (Max (range (\k. \m k\)) *\<^sub>R x)" by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono) also have "\ < ?C" using x by simp (metis \B > 0\ \?C > 0\ mult.commute real_mult_less_iff1 zero_less_mult_pos) finally have "norm (\ k. m k * x $ k) < ?C" . then show "x \ (\x. \ k. x $ k / m k) ` ball 0 ?C" using stretch_Galois [of "inverse \ m"] True by (auto simp: image_iff field_simps) qed then have Bsub: "ball 0 B \ cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))" using cbox_ne uv image_stretch_interval_cart [of "inverse \ m" u v, symmetric] by (force simp: field_simps) obtain z where zint: "(indicat_real S has_integral z) (cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" and zless: "\z - measure lebesgue S\ < e / \prod m UNIV\" using B [OF Bsub] by blast have ind: "indicat_real (?f ` S) = (\x. indicator S (\ k. x$k / m k))" using True stretch_Galois [of m] by (force simp: indicator_def) show "\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e" proof (simp add: ind, intro conjI exI) have "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) ((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" using True has_integral_stretch_cart [OF zint, of "inverse \ m"] by (simp add: field_simps prod_dividef) moreover have "((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))) = cbox u v" using True image_stretch_interval_cart [of "inverse \ m" u v, symmetric] image_stretch_interval_cart [of "\k. 1" u v, symmetric] \cbox u v \ {}\ by (simp add: field_simps image_comp o_def) ultimately show "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) (cbox u v)" by simp have "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ = \prod m UNIV\ * \z - measure lebesgue S\" by (metis (no_types, hide_lams) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib') also have "\ < e" using zless True by (simp add: field_simps) finally show "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ < e" . qed qed qed qed then show ?thesis by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable) next case False then obtain k where "m k = 0" and prm: "prod m UNIV = 0" by auto have nfS: "negligible (?f ` S)" by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \m k = 0\ in auto) then have "(?f ` S) \ lmeasurable" by (simp add: negligible_iff_measure) with nfS show ?thesis by (simp add: prm negligible_iff_measure0) qed then show "(?f ` S) \ lmeasurable" ?MEQ by metis+ qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "linear f" "S \ lmeasurable" shows measurable_linear_image: "(f ` S) \ lmeasurable" and measure_linear_image: "measure lebesgue (f ` S) = \det (matrix f)\ * measure lebesgue S" (is "?Q f S") proof - have "\S \ lmeasurable. (f ` S) \ lmeasurable \ ?Q f S" proof (rule induct_linear_elementary [OF \linear f\]; intro ballI) fix f g and S :: "(real,'n) vec set" assume "linear f" and "linear g" and f [rule_format]: "\S \ lmeasurable. f ` S \ lmeasurable \ ?Q f S" and g [rule_format]: "\S \ lmeasurable. g ` S \ lmeasurable \ ?Q g S" and S: "S \ lmeasurable" then have gS: "g ` S \ lmeasurable" by blast show "(f \ g) ` S \ lmeasurable \ ?Q (f \ g) S" using f [OF gS] g [OF S] matrix_compose [OF \linear g\ \linear f\] by (simp add: o_def image_comp abs_mult det_mul) next fix f :: "real^'n::_ \ real^'n::_" and i and S :: "(real^'n::_) set" assume "linear f" and 0: "\x. f x $ i = 0" and "S \ lmeasurable" then have "\ inj f" by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) have detf: "det (matrix f) = 0" using \\ inj f\ det_nz_iff_inj[OF \linear f\] by blast show "f ` S \ lmeasurable \ ?Q f S" proof show "f ` S \ lmeasurable" using lmeasurable_iff_indicator_has_integral \linear f\ \\ inj f\ negligible_UNIV negligible_linear_singular_image by blast have "measure lebesgue (f ` S) = 0" by (meson \\ inj f\ \linear f\ negligible_imp_measure0 negligible_linear_singular_image) also have "\ = \det (matrix f)\ * measure lebesgue S" by (simp add: detf) finally show "?Q f S" . qed next fix c and S :: "(real^'n::_) set" assume "S \ lmeasurable" show "(\a. \ i. c i * a $ i) ` S \ lmeasurable \ ?Q (\a. \ i. c i * a $ i) S" proof show "(\a. \ i. c i * a $ i) ` S \ lmeasurable" by (simp add: \S \ lmeasurable\ measurable_stretch) show "?Q (\a. \ i. c i * a $ i) S" by (simp add: measure_stretch [OF \S \ lmeasurable\, of c] axis_def matrix_def det_diagonal) qed next fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i" have lin: "linear ?h" by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) have meq: "measure lebesgue ((\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i) ` cbox a b) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have him: "?h ` (cbox a b) \ {}" by blast have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)" by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis swap_id_eq)+ show ?thesis using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\i. (b - a)$i", symmetric] by (simp add: eq content_cbox_cart False) qed have "(\ i j. if Fun.swap m n id i = j then 1 else 0) = (\ i j. if j = Fun.swap m n id i then 1 else (0::real))" by (auto intro!: Cart_lambda_cong) then have "matrix ?h = transpose(\ i j. mat 1 $ i $ Fun.swap m n id j)" by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def) then have 1: "\det (matrix ?h)\ = 1" by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) show "?h ` S \ lmeasurable \ ?Q ?h S" proof show "?h ` S \ lmeasurable" "?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force+ qed next fix m n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. if i = m then v $ m + v $ n else v $ i" have lin: "linear ?h" by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) consider "m < n" | " n < m" using \m \ n\ less_linear by blast then have 1: "det(matrix ?h) = 1" proof cases assume "m < n" have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j < i\ axis_def \m < n\ by auto with \m < n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) next assume "n < m" have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j > i\ axis_def \m > n\ by auto with \m > n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) qed have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have ne: "(+) (\ i. if i = n then - a $ n else 0) ` cbox a b \ {}" by auto let ?v = "\ i. if i = n then - a $ n else 0" have "?h ` cbox a b = (+) (\ i. if i = m \ i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)" using \m \ n\ unfolding image_comp o_def by (force simp: vec_eq_iff) then have "measure lebesgue (?h ` (cbox a b)) = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` (+) ?v ` cbox a b)" by (rule ssubst) (rule measure_translation) also have "\ = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" by (metis (no_types, lifting) cbox_translation) also have "\ = measure lebesgue ((+) (\ i. if i = n then - a $ n else 0) ` cbox a b)" apply (subst measure_shear_interval) using \m \ n\ ne apply auto apply (simp add: cbox_translation) by (metis cbox_borel cbox_translation measure_completion sets_lborel) also have "\ = measure lebesgue (cbox a b)" by (rule measure_translation) finally show ?thesis . qed show "?h ` S \ lmeasurable \ ?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force qed with assms show "(f ` S) \ lmeasurable" "?Q f S" by metis+ qed lemma fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes f: "orthogonal_transformation f" and S: "S \ lmeasurable" shows measurable_orthogonal_image: "f ` S \ lmeasurable" and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S" proof - have "linear f" by (simp add: f orthogonal_transformation_linear) then show "f ` S \ lmeasurable" by (metis S measurable_linear_image) show "measure lebesgue (f ` S) = measure lebesgue S" by (simp add: measure_linear_image \linear f\ S f) qed proposition measure_semicontinuous_with_hausdist_explicit: assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" obtains d where "d > 0" "\T. \T \ lmeasurable; \y. y \ T \ \x. x \ S \ dist x y < d\ \ measure lebesgue T < measure lebesgue S + e" proof (cases "S = {}") case True with that \e > 0\ show ?thesis by force next case False then have frS: "frontier S \ {}" using \bounded S\ frontier_eq_empty not_bounded_UNIV by blast have "S \ lmeasurable" by (simp add: \bounded S\ measurable_Jordan neg) have null: "(frontier S) \ null_sets lebesgue" by (metis neg negligible_iff_null_sets) have "frontier S \ lmeasurable" and mS0: "measure lebesgue (frontier S) = 0" using neg negligible_imp_measurable negligible_iff_measure by blast+ with \e > 0\ sets_lebesgue_outer_open obtain U where "open U" and U: "frontier S \ U" "U - frontier S \ lmeasurable" "emeasure lebesgue (U - frontier S) < e" by (metis fmeasurableD) with null have "U \ lmeasurable" by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel) have "measure lebesgue (U - frontier S) = measure lebesgue U" using mS0 by (simp add: \U \ lmeasurable\ fmeasurableD measure_Diff_null_set null) with U have mU: "measure lebesgue U < e" by (simp add: emeasure_eq_measure2 ennreal_less_iff) show ?thesis proof have "U \ UNIV" using \U \ lmeasurable\ by auto then have "- U \ {}" by blast with \open U\ \frontier S \ U\ show "setdist (frontier S) (- U) > 0" by (auto simp: \bounded S\ open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS) fix T assume "T \ lmeasurable" and T: "\t. t \ T \ \y. y \ S \ dist y t < setdist (frontier S) (- U)" then have "measure lebesgue T - measure lebesgue S \ measure lebesgue (T - S)" by (simp add: \S \ lmeasurable\ measure_diff_le_measure_setdiff) also have "\ \ measure lebesgue U" proof - have "T - S \ U" proof clarify fix x assume "x \ T" and "x \ S" then obtain y where "y \ S" and y: "dist y x < setdist (frontier S) (- U)" using T by blast have "closed_segment x y \ frontier S \ {}" using connected_Int_frontier \x \ S\ \y \ S\ by blast then obtain z where z: "z \ closed_segment x y" "z \ frontier S" by auto with y have "dist z x < setdist(frontier S) (- U)" by (auto simp: dist_commute dest!: dist_in_closed_segment) with z have False if "x \ -U" using setdist_le_dist [OF \z \ frontier S\ that] by auto then show "x \ U" by blast qed then show ?thesis by (simp add: \S \ lmeasurable\ \T \ lmeasurable\ \U \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Diff) qed finally have "measure lebesgue T - measure lebesgue S \ measure lebesgue U" . with mU show "measure lebesgue T < measure lebesgue S + e" by linarith qed qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" and bounded: "\x. x \ S \ \det (matrix (f' x))\ \ B" shows measurable_bounded_differentiable_image: "f ` S \ lmeasurable" and measure_bounded_differentiable_image: "measure lebesgue (f ` S) \ B * measure lebesgue S" (is "?M") proof - have "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ B * measure lebesgue S" proof (cases "B < 0") case True then have "S = {}" by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI) then show ?thesis by auto next case False then have "B \ 0" by arith let ?\ = "measure lebesgue" have f_diff: "f differentiable_on S" using deriv by (auto simp: differentiable_on_def differentiable_def) have eps: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * ?\ S" (is "?ME") if "e > 0" for e proof - have eps_d: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * (?\ S + d)" (is "?MD") if "d > 0" for d proof - obtain T where T: "open T" "S \ T" and TS: "(T-S) \ lmeasurable" and "emeasure lebesgue (T-S) < ennreal d" using S \d > 0\ sets_lebesgue_outer_open by blast then have "?\ (T-S) < d" by (metis emeasure_eq_measure2 ennreal_leI not_less) with S T TS have "T \ lmeasurable" and Tless: "?\ T < ?\ S + d" by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D) have "\r. 0 < r \ r < d \ ball x r \ T \ f ` (S \ ball x r) \ lmeasurable \ ?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" if "x \ S" "d > 0" for x d proof - have lin: "linear (f' x)" and lim0: "((\y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \ 0) (at x within S)" using deriv \x \ S\ by (auto simp: has_derivative_within bounded_linear.linear field_simps) have bo: "bounded (f' x ` ball 0 1)" by (simp add: bounded_linear_image linear_linear lin) have neg: "negligible (frontier (f' x ` ball 0 1))" using deriv has_derivative_linear \x \ S\ by (auto intro!: negligible_convex_frontier [OF convex_linear_image]) have 0: "0 < e * unit_ball_vol (real CARD('n))" using \e > 0\ by simp obtain k where "k > 0" and k: "\U. \U \ lmeasurable; \y. y \ U \ \z. z \ f' x ` ball 0 1 \ dist z y < k\ \ ?\ U < ?\ (f' x ` ball 0 1) + e * unit_ball_vol (CARD('n))" using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast obtain l where "l > 0" and l: "ball x l \ T" using \x \ S\ \open T\ \S \ T\ openE by blast obtain \ where "0 < \" and \: "\y. \y \ S; y \ x; dist y x < \\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x) < k" using lim0 \k > 0\ by (simp add: Lim_within) (auto simp add: field_simps) define r where "r \ min (min l (\/2)) (min 1 (d/2))" show ?thesis proof (intro exI conjI) show "r > 0" "r < d" using \l > 0\ \\ > 0\ \d > 0\ by (auto simp: r_def) have "r \ l" by (auto simp: r_def) with l show "ball x r \ T" by auto have ex_lessK: "\x' \ ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k" if "y \ S" and "dist x y < r" for y proof (cases "y = x") case True with lin linear_0 \k > 0\ that show ?thesis by (rule_tac x=0 in bexI) (auto simp: linear_0) next case False then show ?thesis proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" by (simp add: lin linear_scale) then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)" by (simp add: dist_norm) also have "\ = norm (f' x (y - x) - (f y - f x)) / r" using \r > 0\ by (simp add: divide_simps scale_right_diff_distrib [symmetric]) also have "\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x)" using that \r > 0\ False by (simp add: algebra_simps field_split_simps dist_norm norm_minus_commute mult_right_mono) also have "\ < k" using that \0 < \\ by (simp add: dist_commute r_def \ [OF \y \ S\ False]) finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" . show "(y - x) /\<^sub>R r \ ball 0 1" using that \r > 0\ by (simp add: dist_norm divide_simps norm_minus_commute) qed qed let ?rfs = "(\x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \ ball x r)" have rfs_mble: "?rfs \ lmeasurable" proof (rule bounded_set_imp_lmeasurable) have "f differentiable_on S \ ball x r" using f_diff by (auto simp: fmeasurableD differentiable_on_subset) with S show "?rfs \ sets lebesgue" by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue) let ?B = "(\(x, y). x + y) ` (f' x ` ball 0 1 \ ball 0 k)" have "bounded ?B" by (simp add: bounded_plus [OF bo]) moreover have "?rfs \ ?B" apply (auto simp: dist_norm image_iff dest!: ex_lessK) by (metis (no_types, hide_lams) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball) ultimately show "bounded (?rfs)" by (rule bounded_subset) qed then have "(\x. r *\<^sub>R x) ` ?rfs \ lmeasurable" by (simp add: measurable_linear_image) with \r > 0\ have "(+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) then have "(+) (f x) ` (+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" using measurable_translation by blast then show fsb: "f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) have "?\ (f ` (S \ ball x r)) = ?\ (?rfs) * r ^ CARD('n)" using \r > 0\ fsb by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp) also have "\ \ (\det (matrix (f' x))\ * unit_ball_vol (CARD('n)) + e * unit_ball_vol (CARD('n))) * r ^ CARD('n)" proof - have "?\ (?rfs) < ?\ (f' x ` ball 0 1) + e * unit_ball_vol (CARD('n))" using rfs_mble by (force intro: k dest!: ex_lessK) then have "?\ (?rfs) < \det (matrix (f' x))\ * unit_ball_vol (CARD('n)) + e * unit_ball_vol (CARD('n))" by (simp add: lin measure_linear_image [of "f' x"] content_ball) with \r > 0\ show ?thesis by auto qed also have "\ \ (B + e) * ?\ (ball x r)" using bounded [OF \x \ S\] \r > 0\ by (simp add: content_ball algebra_simps) finally show "?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" . qed qed then obtain r where r0d: "\x d. \x \ S; d > 0\ \ 0 < r x d \ r x d < d" and rT: "\x d. \x \ S; d > 0\ \ ball x (r x d) \ T" and r: "\x d. \x \ S; d > 0\ \ (f ` (S \ ball x (r x d))) \ lmeasurable \ ?\ (f ` (S \ ball x (r x d))) \ (B + e) * ?\ (ball x (r x d))" by metis obtain C where "countable C" and Csub: "C \ {(x,r x t) |x t. x \ S \ 0 < t}" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible(S - (\i \ C. ball (fst i) (snd i)))" apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \ S \ 0 < t}" fst snd]) apply auto by (metis dist_eq_0_iff r0d) let ?UB = "(\(x,s) \ C. ball x s)" have eq: "f ` (S \ ?UB) = (\(x,s) \ C. f ` (S \ ball x s))" by auto have mle: "?\ (\(x,s) \ K. f ` (S \ ball x s)) \ (B + e) * (?\ S + d)" (is "?l \ ?r") if "K \ C" and "finite K" for K proof - have gt0: "b > 0" if "(a, b) \ K" for a b using Csub that \K \ C\ r0d by auto have inj: "inj_on (\(x, y). ball x y) K" by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) have disjnt: "disjoint ((\(x, y). ball x y) ` K)" using pwC that apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) by (metis subsetD fst_conv snd_conv) have "?l \ (\i\K. ?\ (case i of (x, s) \ f ` (S \ ball x s)))" proof (rule measure_UNION_le [OF \finite K\], clarify) fix x r assume "(x,r) \ K" then have "x \ S" using Csub \K \ C\ by auto show "f ` (S \ ball x r) \ sets lebesgue" by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) qed also have "\ \ (\(x,s) \ K. (B + e) * ?\ (ball x s))" apply (rule sum_mono) using Csub r \K \ C\ by auto also have "\ = (B + e) * (\(x,s) \ K. ?\ (ball x s))" by (simp add: prod.case_distrib sum_distrib_left) also have "\ = (B + e) * sum ?\ ((\(x, y). ball x y) ` K)" using \B \ 0\ \e > 0\ by (simp add: inj sum.reindex prod.case_distrib) also have "\ = (B + e) * ?\ (\(x,s) \ K. ball x s)" using \B \ 0\ \e > 0\ that by (subst measure_Union') (auto simp: disjnt measure_Union') also have "\ \ (B + e) * ?\ T" using \B \ 0\ \e > 0\ that apply simp apply (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) using Csub rT by force+ also have "\ \ (B + e) * (?\ S + d)" using \B \ 0\ \e > 0\ Tless by simp finally show ?thesis . qed have fSUB_mble: "(f ` (S \ ?UB)) \ lmeasurable" unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: fmeasurable_UN_bound [OF \countable C\ _ mle]) have fSUB_meas: "?\ (f ` (S \ ?UB)) \ (B + e) * (?\ S + d)" (is "?MUB") unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: measure_UN_bound [OF \countable C\ _ mle]) have neg: "negligible ((f ` (S \ ?UB) - f ` S) \ (f ` S - f ` (S \ ?UB)))" proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]]) show "f differentiable_on S - (\i\C. ball (fst i) (snd i))" by (meson DiffE differentiable_on_subset subsetI f_diff) qed force show "f ` S \ lmeasurable" by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg]) show ?MD using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp qed show "f ` S \ lmeasurable" using eps_d [of 1] by simp show ?ME proof (rule field_le_epsilon) fix \ :: real assume "0 < \" then show "?\ (f ` S) \ (B + e) * ?\ S + \" using eps_d [of "\ / (B+e)"] \e > 0\ \B \ 0\ by (auto simp: divide_simps mult_ac) qed qed show ?thesis proof (cases "?\ S = 0") case True with eps have "?\ (f ` S) = 0" by (metis mult_zero_right not_le zero_less_measure_iff) then show ?thesis using eps [of 1] by (simp add: True) next case False have "?\ (f ` S) \ B * ?\ S" proof (rule field_le_epsilon) fix e :: real assume "e > 0" then show "?\ (f ` S) \ B * ?\ S + e" using eps [of "e / ?\ S"] False by (auto simp: algebra_simps zero_less_measure_iff) qed with eps [of 1] show ?thesis by auto qed qed then show "f ` S \ lmeasurable" ?M by blast+ qed lemma m_diff_image_weak: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" proof - let ?\ = "measure lebesgue" have aint_S: "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int unfolding absolutely_integrable_on_def by auto define m where "m \ integral S (\x. \det (matrix (f' x))\)" have *: "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" if "e > 0" for e proof - define T where "T \ \n. {x \ S. n * e \ \det (matrix (f' x))\ \ \det (matrix (f' x))\ < (Suc n) * e}" have meas_t: "T n \ lmeasurable" for n proof - have *: "(\x. \det (matrix (f' x))\) \ borel_measurable (lebesgue_on S)" using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def) have [intro]: "x \ sets (lebesgue_on S) \ x \ sets lebesgue" for x using S sets_restrict_space_subset by blast have "{x \ S. real n * e \ \det (matrix (f' x))\} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space) then have 1: "{x \ S. real n * e \ \det (matrix (f' x))\} \ lmeasurable" using S by (simp add: fmeasurableI2) have "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space) then have 2: "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ lmeasurable" using S by (simp add: fmeasurableI2) show ?thesis using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong) qed have aint_T: "\k. (\x. \det (matrix (f' x))\) absolutely_integrable_on T k" using set_integrable_subset [OF aint_S] meas_t T_def by blast have Seq: "S = (\n. T n)" apply (auto simp: T_def) apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) using that apply auto using of_int_floor_le pos_le_divide_eq apply blast by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) have meas_ft: "f ` T n \ lmeasurable" for n proof (rule measurable_bounded_differentiable_image) show "T n \ lmeasurable" by (simp add: meas_t) next fix x :: "(real,'n) vec" assume "x \ T n" show "(f has_derivative f' x) (at x within T n)" by (metis (no_types, lifting) \x \ T n\ deriv has_derivative_within_subset mem_Collect_eq subsetI T_def) show "\det (matrix (f' x))\ \ (Suc n) * e" using \x \ T n\ T_def by auto next show "(\x. \det (matrix (f' x))\) integrable_on T n" using aint_T absolutely_integrable_on_def by blast qed have disT: "disjoint (range T)" unfolding disjoint_def proof clarsimp show "T m \ T n = {}" if "T m \ T n" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) with \e > 0\ show ?case unfolding T_def proof (clarsimp simp add: Collect_conj_eq [symmetric]) fix x assume "e > 0" "m < n" "n * e \ \det (matrix (f' x))\" "\det (matrix (f' x))\ < (1 + real m) * e" then have "n < 1 + real m" by (metis (no_types, hide_lams) less_le_trans mult.commute not_le real_mult_le_cancel_iff2) then show "False" using less.hyps by linarith qed qed auto qed have injT: "inj_on T ({n. T n \ {}})" unfolding inj_on_def proof clarsimp show "m = n" if "T m = T n" "T n \ {}" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) have False if "T n \ T m" "x \ T n" for x using \e > 0\ \m < n\ that apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) by (metis add.commute less_le_trans nat_less_real_le not_le real_mult_le_cancel_iff2) then show ?case using less.prems by blast qed auto qed have sum_eq_Tim: "(\k\n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \ real" and n proof (subst sum.reindex_nontrivial) fix i j assume "i \ {..n}" "j \ {..n}" "i \ j" "T i = T j" with that injT [unfolded inj_on_def] show "f (T i) = 0" by simp metis qed (use atMost_atLeast0 in auto) let ?B = "m + e * ?\ S" have "(\k\n. ?\ (f ` T k)) \ ?B" for n proof - have "(\k\n. ?\ (f ` T k)) \ (\k\n. ((k+1) * e) * ?\(T k))" proof (rule sum_mono [OF measure_bounded_differentiable_image]) show "(f has_derivative f' x) (at x within T k)" if "x \ T k" for k x using that unfolding T_def by (blast intro: deriv has_derivative_within_subset) show "(\x. \det (matrix (f' x))\) integrable_on T k" for k using absolutely_integrable_on_def aint_T by blast show "\det (matrix (f' x))\ \ real (k + 1) * e" if "x \ T k" for k x using T_def that by auto qed (use meas_t in auto) also have "\ \ (\k\n. (k * e) * ?\(T k)) + (\k\n. e * ?\(T k))" by (simp add: algebra_simps sum.distrib) also have "\ \ ?B" proof (rule add_mono) have "(\k\n. real k * e * ?\ (T k)) = (\k\n. integral (T k) (\x. k * e))" by (simp add: lmeasure_integral [OF meas_t] flip: integral_mult_right integral_mult_left) also have "\ \ (\k\n. integral (T k) (\x. (abs (det (matrix (f' x))))))" proof (rule sum_mono) fix k assume "k \ {..n}" show "integral (T k) (\x. k * e) \ integral (T k) (\x. \det (matrix (f' x))\)" proof (rule integral_le [OF integrable_on_const [OF meas_t]]) show "(\x. \det (matrix (f' x))\) integrable_on T k" using absolutely_integrable_on_def aint_T by blast next fix x assume "x \ T k" show "k * e \ \det (matrix (f' x))\" using \x \ T k\ T_def by blast qed qed also have "\ = sum (\T. integral T (\x. \det (matrix (f' x))\)) (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = integral (\k\n. T k) (\x. \det (matrix (f' x))\)" proof (rule integral_unique [OF has_integral_Union, symmetric]) fix S assume "S \ T ` {..n}" then show "((\x. \det (matrix (f' x))\) has_integral integral S (\x. \det (matrix (f' x))\)) S" using absolutely_integrable_on_def aint_T by blast next show "pairwise (\S S'. negligible (S \ S')) (T ` {..n})" using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible) qed auto also have "\ \ m" unfolding m_def proof (rule integral_subset_le) have "(\x. \det (matrix (f' x))\) absolutely_integrable_on (\k\n. T k)" apply (rule set_integrable_subset [OF aint_S]) apply (intro measurable meas_t fmeasurableD) apply (force simp: Seq) done then show "(\x. \det (matrix (f' x))\) integrable_on (\k\n. T k)" using absolutely_integrable_on_def by blast qed (use Seq int in auto) finally show "(\k\n. real k * e * ?\ (T k)) \ m" . next have "(\k\n. ?\ (T k)) = sum ?\ (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = ?\ (\k\n. T k)" using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric]) also have "\ \ ?\ S" using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable) finally have "(\k\n. ?\ (T k)) \ ?\ S" . then show "(\k\n. e * ?\ (T k)) \ e * ?\ S" by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that) qed finally show "(\k\n. ?\ (f ` T k)) \ ?B" . qed moreover have "measure lebesgue (\k\n. f ` T k) \ (\k\n. ?\ (f ` T k))" for n by (simp add: fmeasurableD meas_ft measure_UNION_le) ultimately have B_ge_m: "?\ (\k\n. (f ` T k)) \ ?B" for n by (meson order_trans) have "(\n. f ` T n) \ lmeasurable" by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m]) moreover have "?\ (\n. f ` T n) \ m + e * ?\ S" by (rule measure_countable_Union_le [OF meas_ft B_ge_m]) ultimately show "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" by (auto simp: Seq image_Union) qed show ?thesis proof show "f ` S \ lmeasurable" using * linordered_field_no_ub by blast let ?x = "m - ?\ (f ` S)" have False if "?\ (f ` S) > integral S (\x. \det (matrix (f' x))\)" proof - have ml: "m < ?\ (f ` S)" using m_def that by blast then have "?\ S \ 0" using "*"(2) bgauge_existence_lemma by fastforce with ml have 0: "0 < - (m - ?\ (f ` S))/2 / ?\ S" using that zero_less_measure_iff by force then show ?thesis using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm) qed then show "?\ (f ` S) \ integral S (\x. \det (matrix (f' x))\)" by fastforce qed qed theorem fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows measurable_differentiable_image: "f ` S \ lmeasurable" and measure_differentiable_image: "measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" (is "?M") proof - let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" apply (auto simp: mem_box_cart) apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) then have Seq: "S = (\n. ?I n)" by auto have fIn: "f ` ?I n \ lmeasurable" and mfIn: "?\ (f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" (is ?MN) for n proof - have In: "?I n \ lmeasurable" by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int) moreover have "\x. x \ ?I n \ (f has_derivative f' x) (at x within ?I n)" by (meson Int_iff deriv has_derivative_within_subset subsetI) moreover have int_In: "(\x. \det (matrix (f' x))\) integrable_on ?I n" proof - have "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int absolutely_integrable_integrable_bound by force then have "(\x. \det (matrix (f' x))\) absolutely_integrable_on ?I n" by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) then show ?thesis using absolutely_integrable_on_def by blast qed ultimately have "f ` ?I n \ lmeasurable" "?\ (f ` ?I n) \ integral (?I n) (\x. \det (matrix (f' x))\)" using m_diff_image_weak by metis+ moreover have "integral (?I n) (\x. \det (matrix (f' x))\) \ integral S (\x. \det (matrix (f' x))\)" by (simp add: int_In int integral_subset_le) ultimately show "f ` ?I n \ lmeasurable" ?MN by auto qed have "?I k \ ?I n" if "k \ n" for k n by (rule Int_mono) (use that in \auto simp: subset_interval_imp_cart\) then have "(\k\n. f ` ?I k) = f ` ?I n" for n by (fastforce simp add:) with mfIn have "?\ (\k\n. f ` ?I k) \ integral S (\x. \det (matrix (f' x))\)" for n by simp then have "(\n. f ` ?I n) \ lmeasurable" "?\ (\n. f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+ then show "f ` S \ lmeasurable" ?M by (metis Seq image_UN)+ qed lemma borel_measurable_simple_function_limit_increasing: fixes f :: "'a::euclidean_space \ real" shows "(f \ borel_measurable lebesgue \ (\x. 0 \ f x)) \ (\g. (\n x. 0 \ g n x \ g n x \ f x) \ (\n x. g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \ (\x. (\n. g n x) \ f x))" (is "?lhs = ?rhs") proof assume f: ?lhs have leb_f: "{x. a \ f x \ f x < b} \ sets lebesgue" for a b proof - have "{x. a \ f x \ f x < b} = {x. f x < b} - {x. f x < a}" by auto also have "\ \ sets lebesgue" using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto finally show ?thesis . qed have "g n x \ f x" if inc_g: "\n x. 0 \ g n x \ g n x \ g (Suc n) x" and meas_g: "\n. g n \ borel_measurable lebesgue" and fin: "\n. finite(range (g n))" and lim: "\x. (\n. g n x) \ f x" for g n x proof - have "\r>0. \N. \n\N. dist (g n x) (f x) \ r" if "g n x > f x" proof - have g: "g n x \ g (N + n) x" for N by (rule transitive_stepwise_le) (use inc_g in auto) have "\na\N. g n x - f x \ dist (g na x) (f x)" for N apply (rule_tac x="N+n" in exI) using g [of N] by (auto simp: dist_norm) with that show ?thesis using diff_gt_0_iff_gt by blast qed with lim show ?thesis apply (auto simp: lim_sequentially) by (meson less_le_not_le not_le_imp_less) qed moreover let ?\ = "\n k. indicator {y. k/2^n \ f y \ f y < (k+1)/2^n}" let ?g = "\n x. (\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x)" have "\g. (\n x. 0 \ g n x \ g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \(\x. (\n. g n x) \ f x)" proof (intro exI allI conjI) show "0 \ ?g n x" for n x proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg) fix k::real assume "k \ \" and k: "\k\ \ 2 ^ (2*n)" show "0 \ k/2^n * ?\ n k x" using f \k \ \\ apply (auto simp: indicator_def field_split_simps Ints_def) apply (drule spec [where x=x]) using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] by linarith qed show "?g n x \ ?g (Suc n) x" for n x proof - have "?g n x = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * (indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x + indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x))" by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x)" by (simp add: comm_monoid_add_class.sum.distrib algebra_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \ f y \ f y < (2 * k+1)/2 ^ Suc n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2 * k+1) + 1)/2 ^ Suc n} x)" by (force simp: field_simps indicator_def intro: sum.cong) also have "\ \ (\k | k \ \ \ \k\ \ 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x))" (is "?a + _ \ ?b") proof - have *: "\sum f I \ sum h I; a + sum h I \ b\ \ a + sum f I \ b" for I a b f and h :: "real\real" by linarith let ?h = "\k. (2*k+1)/2 ^ Suc n * (indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2*k+1) + 1)/2 ^ Suc n} x)" show ?thesis proof (rule *) show "(\k | k \ \ \ \k\ \ 2 ^ (2*n). 2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < (2 * k+1 + 1)/2 ^ Suc n} x) \ sum ?h {k \ \. \k\ \ 2 ^ (2*n)}" by (rule sum_mono) (simp add: indicator_def field_split_simps) next have \: "?a = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have \: "sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have 0: "(*) 2 ` {k \ \. P k} \ (\x. 2 * x + 1) ` {k \ \. P k} = {}" for P :: "real \ bool" proof - have "2 * i \ 2 * j + 1" for i j :: int by arith thus ?thesis unfolding Ints_def by auto (use of_int_eq_iff in fastforce) qed have "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" unfolding \ \ using finite_abs_int_segment [of "2 ^ (2*n)"] by (subst sum_Un) (auto simp: 0) also have "\ \ ?b" proof (rule sum_mono2) show "finite {k::real. k \ \ \ \k\ \ 2 ^ (2 * Suc n)}" by (rule finite_abs_int_segment) show "(*) 2 ` {k::real. k \ \ \ \k\ \ 2^(2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2^(2*n)} \ {k \ \. \k\ \ 2 ^ (2 * Suc n)}" apply auto using one_le_power [of "2::real" "2*n"] by linarith have *: "\x \ (S \ T) - U; \x. x \ S \ x \ U; \x. x \ T \ x \ U\ \ P x" for S T U P by blast have "0 \ b" if "b \ \" "f x * (2 * 2^n) < b + 1" for b proof - have "0 \ f x * (2 * 2^n)" by (simp add: f) also have "\ < b+1" by (simp add: that) finally show "0 \ b" using \b \ \\ by (auto simp: elim!: Ints_cases) qed then show "0 \ b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \ f y \ f y < (b + 1)/2 ^ Suc n} x" if "b \ {k \ \. \k\ \ 2 ^ (2 * Suc n)} - ((*) 2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)})" for b using that by (simp add: indicator_def divide_simps) qed finally show "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} \ ?b" . qed qed finally show ?thesis . qed show "?g n \ borel_measurable lebesgue" for n apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum) using leb_f sets_restrict_UNIV by auto show "finite (range (?g n))" for n proof - have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" for x proof (cases "\k. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ f x \ f x < (k+1)/2^n") case True then show ?thesis by (blast intro: indicator_sum_eq) next case False then have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) = 0" by auto then show ?thesis by force qed then have "range (?g n) \ ((\k. (k/2^n)) ` {k. k \ \ \ \k\ \ 2 ^ (2*n)})" by auto moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (intro finite_imageI finite_abs_int_segment) ultimately show ?thesis by (rule finite_subset) qed show "(\n. ?g n x) \ f x" for x proof (clarsimp simp add: lim_sequentially) fix e::real assume "e > 0" obtain N1 where N1: "2 ^ N1 > abs(f x)" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by fastforce have "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) < e" if "N1 + N2 \ n" for n proof - let ?m = "real_of_int \2^n * f x\" have "\?m\ \ 2^n * 2^N1" using N1 apply (simp add: f) by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power) also have "\ \ 2 ^ (2*n)" by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff power_add power_increasing_iff semiring_norm(76)) finally have m_le: "\?m\ \ 2 ^ (2*n)" . have "?m/2^n \ f x" "f x < (?m + 1)/2^n" by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos) then have eq: "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) = dist (?m/2^n) (f x)" by (subst indicator_sum_eq [of ?m]) (auto simp: m_le) have "\2^n\ * \?m/2^n - f x\ = \2^n * (?m/2^n - f x)\" by (simp add: abs_mult) also have "\ < 2 ^ N2 * e" using N2 by (simp add: divide_simps mult.commute) linarith also have "\ \ \2^n\ * e" using that \e > 0\ by auto finally have "dist (?m/2^n) (f x) < e" by (simp add: dist_norm) then show ?thesis using eq by linarith qed then show "\no. \n\no. dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k * ?\ n k x/2^n) (f x) < e" by force qed qed ultimately show ?rhs by metis next assume RHS: ?rhs with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq] show ?lhs by (blast intro: order_trans) qed subsection\Borel measurable Jacobian determinant\ lemma lemma_partial_derivatives0: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim0: "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" shows "f x = 0" proof - interpret linear f by fact have "dim {x. f x = 0} \ DIM('a)" by (rule dim_subset_UNIV) moreover have False if less: "dim {x. f x = 0} < DIM('a)" proof - obtain d where "d \ 0" and d: "\y. f y = 0 \ d \ y = 0" using orthogonal_to_subspace_exists [OF less] orthogonal_def by (metis (mono_tags, lifting) mem_Collect_eq span_base) then obtain k where "k > 0" and k: "\e. e > 0 \ \y. y \ S - {0} \ norm y < e \ k * norm y \ \d \ y\" using lb by blast have "\h. \n. ((h n \ S \ h n \ 0 \ k * norm (h n) \ \d \ h n\) \ norm (h n) < 1 / real (Suc n)) \ norm (h (Suc n)) < norm (h n)" proof (rule dependent_nat_choice) show "\y. (y \ S \ y \ 0 \ k * norm y \ \d \ y\) \ norm y < 1 / real (Suc 0)" by simp (metis DiffE insertCI k not_less not_one_le_zero) qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto) then obtain \ where \: "\n. \ n \ S - {0}" and kd: "\n. k * norm(\ n) \ \d \ \ n\" and norm_lt: "\n. norm(\ n) < 1/(Suc n)" by force let ?\ = "\n. \ n /\<^sub>R norm (\ n)" have com: "\g. (\n. g n \ sphere (0::'a) 1) \ \l \ sphere 0 1. \\::nat\nat. strict_mono \ \ (g \ \) \ l" using compact_sphere compact_def by metis moreover have "\n. ?\ n \ sphere 0 1" using \ by auto ultimately obtain l::'a and \::"nat\nat" where l: "l \ sphere 0 1" and "strict_mono \" and to_l: "(?\ \ \) \ l" by meson moreover have "continuous (at l) (\x. (\d \ x\ - k))" by (intro continuous_intros) ultimately have lim_dl: "((\x. (\d \ x\ - k)) \ (?\ \ \)) \ (\d \ l\ - k)" by (meson continuous_imp_tendsto) have "\\<^sub>F i in sequentially. 0 \ ((\x. \d \ x\ - k) \ ((\n. \ n /\<^sub>R norm (\ n)) \ \)) i" using \ kd by (auto simp: field_split_simps) then have "k \ \d \ l\" using tendsto_lowerbound [OF lim_dl, of 0] by auto moreover have "d \ l = 0" proof (rule d) show "f l = 0" proof (rule LIMSEQ_unique [of "f \ ?\ \ \"]) have "isCont f l" using \linear f\ linear_continuous_at linear_conv_bounded_linear by blast then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ f l" unfolding comp_assoc using to_l continuous_imp_tendsto by blast have "\ \ 0" using norm_lt LIMSEQ_norm_0 by metis with \strict_mono \\ have "(\ \ \) \ 0" by (metis LIMSEQ_subseq_LIMSEQ) with lim0 \ have "((\x. f x /\<^sub>R norm x) \ (\ \ \)) \ 0" by (force simp: tendsto_at_iff_sequentially) then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ 0" by (simp add: o_def scale) qed qed ultimately show False using \k > 0\ by auto qed ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis using dim_eq_full by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis mem_Collect_eq module_hom_zero span_base span_raw_def) qed lemma lemma_partial_derivatives: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim: "((\x. f (x - a) /\<^sub>R norm (x - a)) \ 0) (at a within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x \ S - {a}. norm(a - x) < e \ k * norm(a - x) \ \v \ (x - a)\)" shows "f x = 0" proof - have "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within (\x. x-a) ` S)" using lim by (simp add: Lim_within dist_norm) then show ?thesis proof (rule lemma_partial_derivatives0 [OF \linear f\]) fix v :: "'a" assume v: "v \ 0" show "\k>0. \e>0. \x. x \ (\x. x - a) ` S - {0} \ norm x < e \ k * norm x \ \v \ x\" using lb [OF v] by (force simp: norm_minus_commute) qed qed proposition borel_measurable_partial_derivatives: fixes f :: "real^'m::{finite,wellorder} \ real^'n" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. (matrix(f' x)$m$n)) \ borel_measurable (lebesgue_on S)" proof - have contf: "continuous_on S f" using continuous_on_eq_continuous_within f has_derivative_continuous by blast have "{x \ S. (matrix (f' x)$m$n) \ b} \ sets lebesgue" for b proof (rule sets_negligible_symdiff) let ?T = "{x \ S. \e>0. \d>0. \A. A$m$n < b \ (\i j. A$i$j \ \) \ (\y \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x))}" let ?U = "S \ (\e \ {e \ \. e > 0}. \A \ {A. A$m$n < b \ (\i j. A$i$j \ \)}. \d \ {d \ \. 0 < d}. S \ (\y \ S. {x \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x)}))" have "?T = ?U" proof (intro set_eqI iffI) fix x assume xT: "x \ ?T" then show "x \ ?U" proof (clarsimp simp add:) fix q :: real assume "q \ \" "q > 0" then obtain d A where "d > 0" and A: "A $ m $ n < b" "\i j. A $ i $ j \ \" "\y. \y\S; norm (y - x) < d\ \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)" using xT by auto then obtain \ where "d > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with A show "\A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\s. s \ \ \ 0 < s \ (\y\S. norm (y - x) < s \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)))" by force qed next fix x assume xU: "x \ ?U" then show "x \ ?T" proof clarsimp fix e :: "real" assume "e > 0" then obtain \ where \: "e > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with xU obtain A r where "x \ S" and Ar: "A $ m $ n < b" "\i j. A $ i $ j \ \" "r \ \" "r > 0" and "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ \ * norm (y - x)" by (auto simp: split: if_split_asm) then have "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)" by (meson \e > \\ less_eq_real_def mult_right_mono norm_ge_zero order_trans) then show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" using \x \ S\ Ar by blast qed qed moreover have "?U \ sets lebesgue" proof - have coQ: "countable {e \ \. 0 < e}" using countable_Collect countable_rat by blast have ne: "{e \ \. (0::real) < e} \ {}" using zero_less_one Rats_1 by blast have coA: "countable {A. A $ m $ n < b \ (\i j. A $ i $ j \ \)}" proof (rule countable_subset) show "countable {A. \i j. A $ i $ j \ \}" using countable_vector [OF countable_vector, of "\i j. \"] by (simp add: countable_rat) qed blast have *: "\U \ {} \ closedin (top_of_set S) (S \ \ U)\ \ closedin (top_of_set S) (S \ \ U)" for U by fastforce have eq: "{x::(real,'m)vec. P x \ (Q x \ R x)} = {x. P x \ \ Q x} \ {x. P x \ R x}" for P Q R by auto have sets: "S \ (\y\S. {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}) \ sets lebesgue" for e A d proof - have clo: "closedin (top_of_set S) {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}" for y proof - have cont1: "continuous_on S (\x. norm (y - x))" and cont2: "continuous_on S (\x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" by (force intro: contf continuous_intros)+ have clo1: "closedin (top_of_set S) {x \ S. d \ norm (y - x)}" using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def) have clo2: "closedin (top_of_set S) {x \ S. norm (f y - f x - (A *v y - A *v x)) \ e * norm (y - x)}" using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def) show ?thesis by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2) qed show ?thesis by (rule lebesgue_closedin [of S]) (force intro: * S clo)+ qed show ?thesis by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto qed ultimately show "?T \ sets lebesgue" by simp let ?M = "(?T - {x \ S. matrix (f' x) $ m $ n \ b} \ ({x \ S. matrix (f' x) $ m $ n \ b} - ?T))" let ?\ = "\x v. \\>0. \e>0. \y \ S-{x}. norm (x - y) < e \ \v \ (y - x)\ < \ * norm (x - y)" have nN: "negligible {x \ S. \v\0. ?\ x v}" unfolding negligible_eq_zero_density proof clarsimp fix x v and r e :: "real" assume "x \ S" "v \ 0" "r > 0" "e > 0" and Theta [rule_format]: "?\ x v" moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0" by (simp add: \v \ 0\ \e > 0\) ultimately obtain d where "d > 0" and dless: "\y. \y \ S - {x}; norm (x - y) < d\ \ \v \ (y - x)\ < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)" by metis let ?W = "ball x (min d r) \ {y. \v \ (y - x)\ < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}" have "open {x. \v \ (x - a)\ < b}" for a b by (intro open_Collect_less continuous_intros) show "\d>0. d \ r \ (\U. {x' \ S. \v\0. ?\ x' v} \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d))" proof (intro exI conjI) show "0 < min d r" "min d r \ r" using \r > 0\ \d > 0\ by auto show "{x' \ S. \v. v \ 0 \ (\\>0. \e>0. \z\S - {x'}. norm (x' - z) < e \ \v \ (z - x')\ < \ * norm (x' - z))} \ ball x (min d r) \ ?W" proof (clarsimp simp: dist_norm norm_minus_commute) fix y w assume "y \ S" "w \ 0" and less [rule_format]: "\\>0. \e>0. \z\S - {y}. norm (y - z) < e \ \w \ (z - y)\ < \ * norm (y - z)" and d: "norm (y - x) < d" and r: "norm (y - x) < r" show "\v \ (y - x)\ < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" proof (cases "y = x") case True with \r > 0\ \d > 0\ \e > 0\ \v \ 0\ show ?thesis by simp next case False have "\v \ (y - x)\ < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)" apply (rule dless) using False \y \ S\ d by (auto simp: norm_minus_commute) also have "\ \ norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" using d r \e > 0\ by (simp add: field_simps norm_minus_commute mult_left_mono) finally show ?thesis . qed qed show "?W \ lmeasurable" by (simp add: fmeasurable_Int_fmeasurable borel_open) obtain k::'m where True by metis obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis define b where "b \ norm v" have "b > 0" using \v \ 0\ by (auto simp: b_def) obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)" by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv) let ?v = "\ i. min d r / CARD('m)" let ?v' = "\ i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r" let ?x' = "inv T x" let ?W' = "(ball ?x' (min d r) \ {y. \(y - ?x')$k\ < e * min d r / (2 * CARD('m) ^ CARD('m))})" have abs: "x - e \ y \ y \ x + e \ abs(y - x) \ e" for x y e::real by auto have "?W = T ` ?W'" proof - have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)" by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f) have 2: "{y. \v \ (y - x)\ < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} = T ` {y. \y $ k - ?x' $ k\ < e * min d r / (2 * real CARD('m) ^ CARD('m))}" proof - have *: "\T (b *\<^sub>R axis k 1) \ (y - x)\ = b * \inv T y $ k - ?x' $ k\" for y proof - have "\T (b *\<^sub>R axis k 1) \ (y - x)\ = \(b *\<^sub>R axis k 1) \ inv T (y - x)\" by (metis (no_types, hide_lams) b_def eqb invT orthogonal_transformation_def v) also have "\ = b * \(axis k 1) \ inv T (y - x)\" using \b > 0\ by (simp add: abs_mult) also have "\ = b * \inv T y $ k - ?x' $ k\" using orthogonal_transformation_linear [OF invT] by (simp add: inner_axis' linear_diff) finally show ?thesis by simp qed show ?thesis using v b_def [symmetric] using \b > 0\ by (simp add: * bij_image_Collect_eq [OF \bij T\] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right) qed show ?thesis using \b > 0\ by (simp add: image_Int \inj T\ 1 2 b_def [symmetric]) qed moreover have "?W' \ lmeasurable" by (auto intro: fmeasurable_Int_fmeasurable) ultimately have "measure lebesgue ?W = measure lebesgue ?W'" by (metis measure_orthogonal_image T) also have "\ \ measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))" proof (rule measure_mono_fmeasurable) show "?W' \ cbox (?x' - ?v') (?x' + ?v')" apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff) by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component) qed auto also have "\ \ e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))" proof - have "cbox (?x' - ?v) (?x' + ?v) \ {}" using \r > 0\ \d > 0\ by (auto simp: interval_eq_empty_cart divide_less_0_iff) with \r > 0\ \d > 0\ \e > 0\ show ?thesis apply (simp add: content_cbox_if_cart mem_box_cart) apply (auto simp: prod_nonneg) - apply (simp add: abs if_distrib prod.delta_remove prod_constant field_simps power_diff split: if_split_asm) + apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm) done qed also have "\ \ e/2 * measure lebesgue (cball ?x' (min d r))" proof (rule mult_left_mono [OF measure_mono_fmeasurable]) have *: "norm (?x' - y) \ min d r" if y: "\i. \?x' $ i - y $ i\ \ min d r / real CARD('m)" for y proof - have "norm (?x' - y) \ (\i\UNIV. \(?x' - y) $ i\)" by (rule norm_le_l1_cart) also have "\ \ real CARD('m) * (min d r / real CARD('m))" by (rule sum_bounded_above) (use y in auto) finally show ?thesis by simp qed show "cbox (?x' - ?v) (?x' + ?v) \ cball ?x' (min d r)" apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *) by (simp add: abs_diff_le_iff abs_minus_commute) qed (use \e > 0\ in auto) also have "\ < e * content (cball ?x' (min d r))" using \r > 0\ \d > 0\ \e > 0\ by auto also have "\ = e * content (ball x (min d r))" using \r > 0\ \d > 0\ by (simp add: content_cball content_ball) finally show "measure lebesgue ?W < e * content (ball x (min d r))" . qed qed have *: "(\x. (x \ S) \ (x \ T \ x \ U)) \ (T - U) \ (U - T) \ S" for S T U :: "(real,'m) vec set" by blast have MN: "?M \ {x \ S. \v\0. ?\ x v}" proof (rule *) fix x assume x: "x \ {x \ S. \v\0. ?\ x v}" show "(x \ ?T) \ (x \ {x \ S. matrix (f' x) $ m $ n \ b})" proof (cases "x \ S") case True then have x: "\ ?\ x v" if "v \ 0" for v using x that by force show ?thesis proof (rule iffI; clarsimp) assume b: "\e>0. \d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" (is "\e>0. \d>0. \A. ?\ e d A") then have "\k. \d>0. \A. ?\ (1 / Suc k) d A" by (metis (no_types, hide_lams) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff) then obtain \ A where \: "\k. \ k > 0" and Ab: "\k. A k $ m $ n < b" and A: "\k y. \y \ S; norm (y - x) < \ k\ \ norm (f y - f x - A k *v (y - x)) \ 1/(Suc k) * norm (y - x)" by metis have "\i j. \a. (\n. A n $ i $ j) \ a" proof (intro allI) fix i j have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n by (metis cart_eq_inner_axis matrix_vector_mul_component) let ?CA = "{x. Cauchy (\n. (A n) *v x)}" have "subspace ?CA" unfolding subspace_def convergent_eq_Cauchy [symmetric] by (force simp: algebra_simps intro: tendsto_intros) then have CA_eq: "?CA = span ?CA" by (metis span_eq_iff) also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" using dim_subset_UNIV[of ?CA] by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" using less by (force intro: orthogonal_to_subspace_exists [of ?CA]) with x [OF \d \ 0\] obtain \ where "\ > 0" and \: "\e. e > 0 \ \y \ S - {x}. norm (x - y) < e \ \ * norm (x - y) \ \d \ (y - x)\" by (fastforce simp: not_le Bex_def) obtain \ z where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \x: "\ \ x" and z: "(\n. (\ n - x) /\<^sub>R norm (\ n - x)) \ z" proof - have "\\. (\i. (\ i \ S - {x} \ \ * norm(\ i - x) \ \d \ (\ i - x)\ \ norm(\ i - x) < 1/Suc i) \ norm(\(Suc i) - x) < norm(\ i - x))" proof (rule dependent_nat_choice) show "\y. y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1 / Suc 0" using \ [of 1] by (auto simp: dist_norm norm_minus_commute) next fix y i assume "y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1/Suc i" then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0" by auto then obtain y' where "y' \ S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))" "\ * norm (x - y') \ \d \ (y' - x)\" using \ by metis with \ show "\y'. (y' \ S - {x} \ \ * norm (y' - x) \ \d \ (y' - x)\ \ norm (y' - x) < 1/(Suc (Suc i))) \ norm (y' - x) < norm (y - x)" by (auto simp: dist_norm norm_minus_commute) qed then obtain \ where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \conv: "\i. norm(\ i - x) < 1/(Suc i)" by blast let ?f = "\i. (\ i - x) /\<^sub>R norm (\ i - x)" have "?f i \ sphere 0 1" for i using \Sx by auto then obtain l \ where "l \ sphere 0 1" "strict_mono \" and l: "(?f \ \) \ l" using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson show thesis proof show "(\ \ \) i \ S - {x}" "\ * norm ((\ \ \) i - x) \ \d \ ((\ \ \) i - x)\" for i using \Sx \le by auto have "\ \ x" proof (clarsimp simp add: LIMSEQ_def dist_norm) fix r :: "real" assume "r > 0" with real_arch_invD obtain no where "no \ 0" "real no > 1/r" by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2) with \conv show "\no. \n\no. norm (\ n - x) < r" by (metis \r > 0\ add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc) qed with \strict_mono \\ show "(\ \ \) \ x" by (metis LIMSEQ_subseq_LIMSEQ) show "(\n. ((\ \ \) n - x) /\<^sub>R norm ((\ \ \) n - x)) \ l" using l by (auto simp: o_def) qed qed have "isCont (\x. (\d \ x\ - \)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\y. \d \ ((\ y - x) /\<^sub>R norm (\ y - x))\ - \) \ \d \ z\ - \" by auto moreover have "\\<^sub>F i in sequentially. 0 \ \d \ ((\ i - x) /\<^sub>R norm (\ i - x))\ - \" proof (rule eventuallyI) fix n show "0 \ \d \ ((\ n - x) /\<^sub>R norm (\ n - x))\ - \" using \le [of n] \Sx by (auto simp: abs_mult divide_simps) qed ultimately have "\ \ \d \ z\" using tendsto_lowerbound [where a=0] by fastforce have "Cauchy (\n. (A n) *v z)" proof (clarsimp simp add: Cauchy_def) fix \ :: "real" assume "0 < \" then obtain N::nat where "N > 0" and N: "\/2 > 1/N" by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse) show "\M. \m\M. \n\M. dist (A m *v z) (A n *v z) < \" proof (intro exI allI impI) fix i j assume ij: "N \ i" "N \ j" let ?V = "\i k. A i *v ((\ k - x) /\<^sub>R norm (\ k - x))" have "\\<^sub>F k in sequentially. dist (\ k) x < min (\ i) (\ j)" using \x [unfolded tendsto_iff] by (meson min_less_iff_conj \) then have even: "\\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \ 0" proof (rule eventually_mono, clarsimp) fix p assume p: "dist (\ p) x < \ i" "dist (\ p) x < \ j" let ?C = "\k. f (\ p) - f x - A k *v (\ p - x)" have "norm ((A i - A j) *v (\ p - x)) = norm (?C j - ?C i)" by (simp add: algebra_simps) also have "\ \ norm (?C j) + norm (?C i)" using norm_triangle_ineq4 by blast also have "\ \ 1/(Suc j) * norm (\ p - x) + 1/(Suc i) * norm (\ p - x)" by (metis A Diff_iff \Sx dist_norm p add_mono) also have "\ \ 1/N * norm (\ p - x) + 1/N * norm (\ p - x)" apply (intro add_mono mult_right_mono) using ij \N > 0\ by (auto simp: field_simps) also have "\ = 2 / N * norm (\ p - x)" by simp finally have no_le: "norm ((A i - A j) *v (\ p - x)) \ 2 / N * norm (\ p - x)" . have "norm (?V i p - ?V j p) = norm ((A i - A j) *v ((\ p - x) /\<^sub>R norm (\ p - x)))" by (simp add: algebra_simps) also have "\ = norm ((A i - A j) *v (\ p - x)) / norm (\ p - x)" by (simp add: divide_inverse matrix_vector_mult_scaleR) also have "\ \ 2 / N" using no_le by (auto simp: field_split_simps) finally show "norm (?V i p - ?V j p) \ 2 / N" . qed have "isCont (\w. (norm(A i *v w - A j *v w) - 2 / N)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\w. norm (A i *v ((\ w - x) /\<^sub>R norm (\ w - x)) - A j *v ((\ w - x) /\<^sub>R norm (\ w - x))) - 2 / N) \ norm (A i *v z - A j *v z) - 2 / N" by auto have "dist (A i *v z) (A j *v z) \ 2 / N" using tendsto_upperbound [OF lim even] by (auto simp: dist_norm) with N show "dist (A i *v z) (A j *v z) < \" by linarith qed qed then have "d \ z = 0" using CA_eq d orthogonal_def by auto then show False using \0 < \\ \\ \ \d \ z\\ by auto qed ultimately show ?thesis using dim_eq_full by fastforce qed finally have "?CA = UNIV" . then have "Cauchy (\n. (A n) *v axis j 1)" by auto then obtain L where "(\n. A n *v axis j 1) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) then have "(\x. (A x *v axis j 1) $ i) \ L $ i" by (rule tendsto_vec_nth) then show "\a. (\n. A n $ i $ j) \ a" by (force simp: vax) qed then obtain B where B: "\i j. (\n. A n $ i $ j) \ B $ i $ j" by (auto simp: lambda_skolem) have lin_df: "linear (f' x)" and lim_df: "((\y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \ 0) (at x within S)" using \x \ S\ assms by (auto simp: has_derivative_within linear_linear) moreover interpret linear "f' x" by fact have "(matrix (f' x) - B) *v w = 0" for w proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"]) show "linear ((*v) (matrix (f' x) - B))" by (rule matrix_vector_mul_linear) have "((\y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" using tendsto_minus [OF lim_df] by (simp add: algebra_simps field_split_simps) then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \ 0) (at x within S)" proof (rule Lim_transform) have "((\y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \ 0) (at x within S)" proof (clarsimp simp add: Lim_within dist_norm) fix e :: "real" assume "e > 0" then obtain q::nat where "q \ 0" and qe2: "1/q < e/2" by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral) let ?g = "\p. sum (\i. sum (\j. abs((A p - B)$i$j)) UNIV) UNIV" have "(\k. onorm (\y. (A k - B) *v y)) \ 0" proof (rule Lim_null_comparison) show "\\<^sub>F k in sequentially. norm (onorm (\y. (A k - B) *v y)) \ ?g k" proof (rule eventually_sequentiallyI) fix k :: "nat" assume "0 \ k" have "0 \ onorm ((*v) (A k - B))" using matrix_vector_mul_bounded_linear by (rule onorm_pos_le) then show "norm (onorm ((*v) (A k - B))) \ (\i\UNIV. \j\UNIV. \(A k - B) $ i $ j\)" by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) qed next show "?g \ 0" using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum) qed with \e > 0\ obtain p where "\n. n \ p \ \onorm ((*v) (A n - B))\ < e/2" unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral) then have pqe2: "\onorm ((*v) (A (p + q) - B))\ < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*) using le_add1 by blast show "\d>0. \y\S. y \ x \ norm (y - x) < d \ inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" proof (intro exI, safe) show "0 < \(p + q)" by (simp add: \) next fix y assume y: "y \ S" "norm (y - x) < \(p + q)" and "y \ x" have *: "\norm(b - c) < e - d; norm(y - x - b) \ d\ \ norm(y - x - c) < e" for b c d e x and y:: "real^'n" using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)" proof (rule *) show "norm (f y - f x - A (p + q) *v (y - x)) \ norm (y - x) / (Suc (p + q))" using A [OF y] by simp have "norm (A (p + q) *v (y - x) - B *v (y - x)) \ onorm(\x. (A(p + q) - B) *v x) * norm(y - x)" by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm) also have "\ < (e/2) * norm (y - x)" using \y \ x\ pqe2 by auto also have "\ \ (e - 1 / (Suc (p + q))) * norm (y - x)" proof (rule mult_right_mono) have "1 / Suc (p + q) \ 1 / q" using \q \ 0\ by (auto simp: field_split_simps) also have "\ < e/2" using qe2 by auto finally show "e / 2 \ e - 1 / real (Suc (p + q))" by linarith qed auto finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))" by (simp add: algebra_simps) qed then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" using \y \ x\ by (simp add: field_split_simps algebra_simps) qed qed then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" by (simp add: algebra_simps diff lin_df matrix_vector_mul_linear scalar_mult_eq_scaleR) qed qed (use x in \simp; auto simp: not_less\) ultimately have "f' x = (*v) B" by (force simp: algebra_simps scalar_mult_eq_scaleR) show "matrix (f' x) $ m $ n \ b" proof (rule tendsto_upperbound [of "\i. (A i $ m $ n)" _ sequentially]) show "(\i. A i $ m $ n) \ matrix (f' x) $ m $ n" by (simp add: B \f' x = (*v) B\) show "\\<^sub>F i in sequentially. A i $ m $ n \ b" by (simp add: Ab less_eq_real_def) qed auto next fix e :: "real" assume "x \ S" and b: "matrix (f' x) $ m $ n \ b" and "e > 0" then obtain d where "d>0" and d: "\y. y\S \ 0 < dist y x \ dist y x < d \ norm (f y - f x - f' x (y - x)) / (norm (y - x)) < e/2" using f [OF \x \ S\] by (simp add: Deriv.has_derivative_at_within Lim_within) (auto simp add: field_simps dest: spec [of _ "e/2"]) let ?A = "matrix(f' x) - (\ i j. if i = m \ j = n then e / 4 else 0)" obtain B where BRats: "\i j. B$i$j \ \" and Bo_e6: "onorm((*v) (?A - B)) < e/6" using matrix_rational_approximation \e > 0\ by (metis zero_less_divide_iff zero_less_numeral) show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" proof (intro exI conjI ballI allI impI) show "d>0" by (rule \d>0\) show "B $ m $ n < b" proof - have "\matrix ((*v) (?A - B)) $ m $ n\ \ onorm ((*v) (?A - B))" using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis then show ?thesis using b Bo_e6 by simp qed show "B $ i $ j \ \" for i j using BRats by auto show "norm (f y - f x - B *v (y - x)) \ e * norm (y - x)" if "y \ S" and y: "norm (y - x) < d" for y proof (cases "y = x") case True then show ?thesis by simp next case False have *: "norm(d' - d) \ e/2 \ norm(y - (x + d')) < e/2 \ norm(y - x - d) \ e" for d d' e and x y::"real^'n" using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp show ?thesis proof (rule *) have split246: "\norm y \ e / 6; norm(x - y) \ e / 4\ \ norm x \ e/2" if "e > 0" for e and x y :: "real^'n" using norm_triangle_le [of y "x-y" "e/2"] \e > 0\ by simp have "linear (f' x)" using True f has_derivative_linear by blast then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" by (simp add: matrix_vector_mult_diff_rdistrib) also have "\ \ (e * norm (y - x)) / 2" proof (rule split246) have "norm ((?A - B) *v (y - x)) / norm (y - x) \ onorm(\x. (?A - B) *v x)" by (rule le_onorm) auto also have "\ < e/6" by (rule Bo_e6) finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . then show "norm ((?A - B) *v (y - x)) \ e * norm (y - x) / 6" by (simp add: field_split_simps False) have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x))" by (simp add: algebra_simps) also have "\ = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))" proof - have "(\j\UNIV. (if i = m \ j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i proof (cases "i=m") case True then show ?thesis by (auto simp: if_distrib [of "\z. z * _"] cong: if_cong) next case False then show ?thesis by (simp add: axis_def) qed then have "(\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)" by (auto simp: vec_eq_iff matrix_vector_mult_def) then show ?thesis by metis qed also have "\ \ e * norm (y - x) / 4" using \e > 0\ apply (simp add: norm_mult abs_mult) by (metis component_le_norm_cart vector_minus_component) finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \ e * norm (y - x) / 4" . show "0 < e * norm (y - x)" by (simp add: False \e > 0\) qed finally show "norm (f' x (y - x) - B *v (y - x)) \ (e * norm (y - x)) / 2" . show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2" using False d [OF \y \ S\] y by (simp add: dist_norm field_simps) qed qed qed qed qed auto qed show "negligible ?M" using negligible_subset [OF nN MN] . qed then show ?thesis by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms) qed theorem borel_measurable_det_Jacobian: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ borel_measurable (lebesgue_on S)" unfolding det_def by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S]) text\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) theorem borel_measurable_lebesgue_on_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" proof - have "{x. (if x \ S then f x else 0) \ T} \ sets lebesgue \ {x \ S. f x \ T} \ sets lebesgue" if "T \ sets borel" for T proof (cases "0 \ T") case True then have "{x \ S. f x \ T} = {x. (if x \ S then f x else 0) \ T} \ S" "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T} \ -S" by auto then show ?thesis by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un) next case False then have "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T}" by auto then show ?thesis by auto qed then show ?thesis unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric] by blast qed lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" proof - obtain T N N' where "S = T \ N" "N \ N'" "N' \ null_sets lborel" "T \ sets borel" using sets_completionE [OF assms] by auto then show thesis by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) qed lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" shows "(\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue) \ f \ borel_measurable (lebesgue_on S) \ (\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue)" (is "?lhs \ _ \ ?rhs") unfolding borel_measurable_lebesgue_on_preimage_borel [OF S] proof (intro iffI allI conjI impI, safe) fix V :: "'b set" assume *: "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "V \ sets borel" then have V: "V \ sets lebesgue" by simp have "{x \ S. f x \ V} = {x \ S. f x \ T \ V}" using fim by blast also have "{x \ S. f x \ T \ V} \ sets lebesgue" using T V * le_inf_iff by blast finally show "{x \ S. f x \ V} \ sets lebesgue" . next fix U :: "'b set" assume "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" "negligible U" "U \ T" then show "{x \ S. f x \ U} \ sets lebesgue" using negligible_imp_sets by blast next fix U :: "'b set" assume 1 [rule_format]: "(\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" and 2 [rule_format]: "\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "U \ sets lebesgue" "U \ T" then obtain C N where C: "C \ sets borel \ negligible N \ C \ N = U" using sets_lebesgue_almost_borel by metis then have "{x \ S. f x \ C} \ sets lebesgue" by (blast intro: 1) moreover have "{x \ S. f x \ N} \ sets lebesgue" using C \U \ T\ by (blast intro: 2) moreover have "{x \ S. f x \ C \ N} = {x \ S. f x \ C} \ {x \ S. f x \ N}" by auto ultimately show "{x \ S. f x \ U} \ sets lebesgue" using C by auto qed subsection\Simplest case of Sard's theorem (we don't need continuity of derivative)\ lemma Sard_lemma00: fixes P :: "'b::euclidean_space set" assumes "a \ 0" and a: "a *\<^sub>R i \ 0" and i: "i \ Basis" and P: "P \ {x. a *\<^sub>R i \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (DIM('b) - 1)" proof - have "a > 0" using assms by simp let ?v = "(\j\Basis. (if j = i then e else m) *\<^sub>R j)" show thesis proof have "- e \ x \ i" "x \ i \ e" if "t \ P" "norm (x - t) \ e" for x t using \a > 0\ that Basis_le_norm [of i "x-t"] P i by (auto simp: inner_commute algebra_simps) moreover have "- m \ x \ j" "x \ j \ m" if "norm x \ m" "t \ P" "norm (x - t) \ e" "j \ Basis" and "j \ i" for x t j using that Basis_le_norm [of j x] by auto ultimately show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ cbox (-?v) ?v" by (auto simp: mem_box) have *: "\k\Basis. - ?v \ k \ ?v \ k" using \0 \ m\ \0 \ e\ by (auto simp: inner_Basis) have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)" by (metis DIM_positive Suc_pred power_Suc) show "measure lebesgue (cbox (-?v) ?v) \ 2 * e * (2 * m) ^ (DIM('b) - 1)" using \i \ Basis\ by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2) qed blast qed text\As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\ lemma Sard_lemma0: fixes P :: "(real^'n::{finite,wellorder}) set" assumes "a \ 0" and P: "P \ {x. a \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis have Tinv [simp]: "T (inv T x) = x" for x by (simp add: T orthogonal_transformation_surj surj_f_inv_f) obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ T-`P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) have "norm a *\<^sub>R axis k 1 \ x = 0" if "T x \ P" for x proof - have "a \ T x = 0" using P that by blast then show ?thesis by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) qed then show "T -` P \ {x. norm a *\<^sub>R axis k 1 \ x = 0}" by auto qed (use assms T in auto) show thesis proof show "T ` S \ lmeasurable" using S measurable_orthogonal_image T by blast have "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" proof clarsimp fix x t assume "norm x \ m" "t \ P" "norm (x - t) \ e" then have "norm (inv T x) \ m" using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) moreover have "\t\T -` P. norm (inv T x - t) \ e" proof have "T (inv T x - inv T t) = x - t" using T linear_diff orthogonal_transformation_def by (metis (no_types, hide_lams) Tinv) then have "norm (inv T x - inv T t) = norm (x - t)" by (metis T orthogonal_transformation_norm) then show "norm (inv T x - inv T t) \ e" using \norm (x - t) \ e\ by linarith next show "inv T t \ T -` P" using \t \ P\ by force qed ultimately show "x \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" by force qed then show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` S" using image_mono [OF subS] by (rule order_trans) show "measure lebesgue (T ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" using mS T by (simp add: S measure_orthogonal_image) qed qed text\As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\ lemma Sard_lemma1: fixes P :: "(real^'n::{finite,wellorder}) set" assumes P: "dim P < CARD('n)" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm(z - w) \ m \ (\t \ P. norm(z - w - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" by (rule Sard_lemma0 [OF _ _ \0 \ m\ \0 \ e\]) show thesis proof show "(+)w ` S \ lmeasurable" by (metis measurable_translation S) show "{z. norm (z - w) \ m \ (\t\P. norm (z - w - t) \ e)} \ (+)w ` S" using subS by force show "measure lebesgue ((+)w ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" by (metis measure_translation mS) qed qed lemma Sard_lemma2: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" (is "?m \ ?n") and "B > 0" "bounded S" and derS: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" and B: "\x. x \ S \ onorm(f' x) \ B" shows "negligible(f ` S)" proof - have lin_f': "\x. x \ S \ linear(f' x)" using derS has_derivative_linear by blast show ?thesis proof (clarsimp simp add: negligible_outer_le) fix e :: "real" assume "e > 0" obtain c where csub: "S \ cbox (- (vec c)) (vec c)" and "c > 0" proof - obtain b where b: "\x. x \ S \ norm x \ b" using \bounded S\ by (auto simp: bounded_iff) show thesis proof have "- \b\ - 1 \ x $ i \ x $ i \ \b\ + 1" if "x \ S" for x i using component_le_norm_cart [of x i] b [OF that] by auto then show "S \ cbox (- vec (\b\ + 1)) (vec (\b\ + 1))" by (auto simp: mem_box_cart) qed auto qed then have box_cc: "box (- (vec c)) (vec c) \ {}" and cbox_cc: "cbox (- (vec c)) (vec c) \ {}" by (auto simp: interval_eq_empty_cart) obtain d where "d > 0" "d \ B" and d: "(d * 2) * (4 * B) ^ (?n - 1) \ e / (2*c) ^ ?m / ?m ^ ?m" apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"]) using \B > 0\ \c > 0\ \e > 0\ by (simp_all add: divide_simps min_mult_distrib_right) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. y \ S \ norm(y - x) < r \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)))" for x proof (cases "x \ S") case True then obtain r where "r > 0" and "\y. \y \ S; norm (y - x) < r\ \ norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using derS \d > 0\ by (force simp: has_derivative_within_alt) then show ?thesis by (rule_tac x="min r (1/2)" in exI) simp next case False then show ?thesis by (rule_tac x="1/2" in exI) simp qed then obtain r where r12: "\x. 0 < r x \ r x \ 1/2" and r: "\x y. \x \ S; y \ S; norm(y - x) < r x\ \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)" by metis then have ga: "gauge (\x. ball x (r x))" by (auto simp: gauge_def) obtain \ where \: "countable \" and sub_cc: "\\ \ cbox (- vec c) (vec c)" and cbox: "\K. K \ \ \ interior K \ {} \ (\u v. K = cbox u v)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ ball x (r x)" and close: "\u v. cbox u v \ \ \ \n. \i::'m. v $ i - u $ i = 2*c / 2^n" and covers: "S \ \\" apply (rule covering_lemma [OF csub box_cc ga]) apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric]) done let ?\ = "measure lebesgue" have "\T. T \ lmeasurable \ f ` (K \ S) \ T \ ?\ T \ e / (2*c) ^ ?m * ?\ K" if "K \ \" for K proof - obtain u v where uv: "K = cbox u v" using cbox \K \ \\ by blast then have uv_ne: "cbox u v \ {}" using cbox that by fastforce obtain x where x: "x \ S \ cbox u v" "cbox u v \ ball x (r x)" using \K \ \\ covered uv by blast then have "dim (range (f' x)) < ?n" using rank_dim_range [of "matrix (f' x)"] x rank[of x] by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') then obtain T where T: "T \ lmeasurable" and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" (is "_ \ ?DVU") apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) using \B > 0\ \d > 0\ by simp_all show ?thesis proof (intro exI conjI) have "f ` (K \ S) \ {z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))}" unfolding uv proof (clarsimp simp: mult.assoc, intro conjI) fix y assume y: "y \ cbox u v" and "y \ S" then have "norm (y - x) < r x" by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2)) then have le_dyx: "norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using r [of x y] x \y \ S\ by blast have yx_le: "norm (y - x) \ norm (v - u)" proof (rule norm_le_componentwise_cart) show "norm ((y - x) $ i) \ norm ((v - u) $ i)" for i using x y by (force simp: mem_box_cart dest!: spec [where x=i]) qed have *: "\norm(y - x - z) \ d; norm z \ B; d \ B\ \ norm(y - x) \ 2 * B" for x y z :: "real^'n::_" and d B using norm_triangle_ineq2 [of "y - x" z] by auto show "norm (f y - f x) \ 2 * (B * norm (v - u))" proof (rule * [OF le_dyx]) have "norm (f' x (y - x)) \ onorm (f' x) * norm (y - x)" using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) also have "\ \ B * norm (v - u)" proof (rule mult_mono) show "onorm (f' x) \ B" using B x by blast qed (use \B > 0\ yx_le in auto) finally show "norm (f' x (y - x)) \ B * norm (v - u)" . show "d * norm (y - x) \ B * norm (v - u)" using \B > 0\ by (auto intro: mult_mono [OF \d \ B\ yx_le]) qed show "\t. norm (f y - f x - f' x t) \ d * norm (v - u)" apply (rule_tac x="y-x" in exI) using \d > 0\ yx_le le_dyx mult_left_mono [where c=d] by (meson order_trans real_mult_le_cancel_iff2) qed with subT show "f ` (K \ S) \ T" by blast show "?\ T \ e / (2*c) ^ ?m * ?\ K" proof (rule order_trans [OF measT]) have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n" using \c > 0\ apply (simp add: algebra_simps power_mult_distrib) by (metis Suc_pred power_Suc zero_less_card_finite) also have "\ \ (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n" by (rule mult_right_mono [OF d]) auto also have "\ \ e / (2*c) ^ ?m * ?\ K" proof - have "u \ ball (x) (r x)" "v \ ball x (r x)" using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+ moreover have "r x \ 1/2" using r12 by auto ultimately have "norm (v - u) \ 1" using norm_triangle_half_r [of x u 1 v] by (metis (no_types, hide_lams) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball) then have "norm (v - u) ^ ?n \ norm (v - u) ^ ?m" by (simp add: power_decreasing [OF mlen]) also have "\ \ ?\ K * real (?m ^ ?m)" proof - obtain n where n: "\i. v$i - u$i = 2 * c / 2^n" using close [of u v] \K \ \\ uv by blast have "norm (v - u) ^ ?m \ (\i\UNIV. \(v - u) $ i\) ^ ?m" by (intro norm_le_l1_cart power_mono) auto also have "\ \ (\i\UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)" by (simp add: n field_simps \c > 0\ less_eq_real_def) also have "\ = ?\ K * real (?m ^ ?m)" by (simp add: uv uv_ne content_cbox_cart) finally show ?thesis . qed finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \ ?\ K" by (simp add: field_split_simps) show ?thesis using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \c > 0\ \e > 0\ by auto qed finally show "?DVU \ e / (2*c) ^ ?m * ?\ K" . qed qed (use T in auto) qed then obtain g where meas_g: "\K. K \ \ \ g K \ lmeasurable" and sub_g: "\K. K \ \ \ f ` (K \ S) \ g K" and le_g: "\K. K \ \ \ ?\ (g K) \ e / (2*c)^?m * ?\ K" by metis have le_e: "?\ (\i\\. g i) \ e" if "\ \ \" "finite \" for \ proof - have "?\ (\i\\. g i) \ (\i\\. ?\ (g i))" using meas_g \\ \ \\ by (auto intro: measure_UNION_le [OF \finite \\]) also have "\ \ (\K\\. e / (2*c) ^ ?m * ?\ K)" using \\ \ \\ sum_mono [OF le_g] by (meson le_g subsetCE sum_mono) also have "\ = e / (2*c) ^ ?m * (\K\\. ?\ K)" by (simp add: sum_distrib_left) also have "\ \ e" proof - have "\ division_of \\" proof (rule division_ofI) show "K \ \\" "K \ {}" "\a b. K = cbox a b" if "K \ \" for K using \K \ \\ covered cbox \\ \ \\ by (auto simp: Union_upper) show "interior K \ interior L = {}" if "K \ \" and "L \ \" and "K \ L" for K L by (metis (mono_tags, lifting) \\ \ \\ pairwiseD djointish pairwise_subset that) qed (use that in auto) then have "sum ?\ \ \ ?\ (\\)" by (simp add: content_division) also have "\ \ ?\ (cbox (- vec c) (vec c) :: (real, 'm) vec set)" proof (rule measure_mono_fmeasurable) show "\\ \ cbox (- vec c) (vec c)" by (meson Sup_subset_mono sub_cc order_trans \\ \ \\) qed (use \\ division_of \\\ lmeasurable_division in auto) also have "\ = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)" by simp also have "\ \ (2 ^ ?m * c ^ ?m)" using \c > 0\ by (simp add: content_cbox_if_cart) finally have "sum ?\ \ \ (2 ^ ?m * c ^ ?m)" . then show ?thesis using \e > 0\ \c > 0\ by (auto simp: field_split_simps mult_less_0_iff) qed finally show ?thesis . qed show "\T. f ` S \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "f ` S \ \ (g ` \)" using covers sub_g by force show "\ (g ` \) \ lmeasurable" by (rule fmeasurable_UN_bound [OF \countable \\ meas_g le_e]) show "?\ (\ (g ` \)) \ e" by (rule measure_UN_bound [OF \countable \\ meas_g le_e]) qed qed qed theorem baby_Sard: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" and der: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" shows "negligible(f ` S)" proof - let ?U = "\n. {x \ S. norm(x) \ n \ onorm(f' x) \ real n}" have "\x. x \ S \ \n. norm x \ real n \ onorm (f' x) \ real n" by (meson linear order_trans real_arch_simple) then have eq: "S = (\n. ?U n)" by auto have "negligible (f ` ?U n)" for n proof (rule Sard_lemma2 [OF mlen]) show "0 < real n + 1" by auto show "bounded (?U n)" using bounded_iff by blast show "(f has_derivative f' x) (at x within ?U n)" if "x \ ?U n" for x using der that by (force intro: has_derivative_subset) qed (use rank in auto) then show ?thesis by (subst eq) (simp add: image_Union negligible_Union_nat) qed subsection\A one-way version of change-of-variables not assuming injectivity. \ lemma integral_on_image_ubound_weak: fixes f :: "real^'n::{finite,wellorder} \ real" assumes S: "S \ sets lebesgue" and f: "f \ borel_measurable (lebesgue_on (g ` S))" and nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and det_int_fg: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" and meas_gim: "\T. \T \ g ` S; T \ sets lebesgue\ \ {x \ S. g x \ T} \ sets lebesgue" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. \det (matrix (g' x))\" have cont_g: "continuous_on S g" using der_g has_derivative_continuous_on by blast have [simp]: "space (lebesgue_on S) = S" by (simp add: S) have gS_in_sets_leb: "g ` S \ sets lebesgue" apply (rule differentiable_image_in_sets_lebesgue) using der_g by (auto simp: S differentiable_def differentiable_on_def) obtain h where nonneg_h: "\n x. 0 \ h n x" and h_le_f: "\n x. x \ S \ h n (g x) \ f (g x)" and h_inc: "\n x. h n x \ h (Suc n) x" and h_meas: "\n. h n \ borel_measurable lebesgue" and fin_R: "\n. finite(range (h n))" and lim: "\x. x \ g ` S \ (\n. h n x) \ f x" proof - let ?f = "\x. if x \ g ` S then f x else 0" have "?f \ borel_measurable lebesgue \ (\x. 0 \ ?f x)" by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric]) then show ?thesis apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing) apply (rename_tac h) by (rule_tac h=h in that) (auto split: if_split_asm) qed have h_lmeas: "{t. h n (g t) = y} \ S \ sets lebesgue" for y n proof - have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV" by simp then have "((h n) -`{y} \ g ` S) \ sets (lebesgue_on (g ` S))" by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel) then have "({u. h n u = y} \ g ` S) \ sets lebesgue" using gS_in_sets_leb by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def) then have "{x \ S. g x \ ({u. h n u = y} \ g ` S)} \ sets lebesgue" using meas_gim[of "({u. h n u = y} \ g ` S)"] by force moreover have "{t. h n (g t) = y} \ S = {x \ S. g x \ ({u. h n u = y} \ g ` S)}" by blast ultimately show ?thesis by auto qed have hint: "h n integrable_on g ` S \ integral (g ` S) (h n) \ integral S (\x. ?D x * h n (g x))" (is "?INT \ ?lhs \ ?rhs") for n proof - let ?R = "range (h n)" have hn_eq: "h n = (\x. \y\?R. y * indicat_real {x. h n x = y} x)" by (simp add: indicator_def if_distrib fin_R cong: if_cong) have yind: "(\t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \ (integral (g ` S) (\t. y * indicator {x. h n x = y} t)) \ integral S (\t. \det (matrix (g' t))\ * y * indicator {x. h n x = y} (g t))" if y: "y \ ?R" for y::real proof (cases "y=0") case True then show ?thesis using gS_in_sets_leb integrable_0 by force next case False with that have "y > 0" using less_eq_real_def nonneg_h by fastforce have "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. ?D x) \ borel_measurable (lebesgue_on ({t. h n (g t) = y} \ S))" apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) by (meson der_g IntD2 has_derivative_within_subset inf_le2) then have "(\x. if x \ {t. h n (g t) = y} \ S then ?D x else 0) \ borel_measurable lebesgue" by (rule borel_measurable_if_I [OF _ h_lmeas]) then show "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) \ borel_measurable (lebesgue_on S)" by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric]) show "(\x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S" by (rule integrable_cmul) (use det_int_fg in auto) show "norm (if x \ {t. h n (g t) = y} then ?D x else 0) \ ?D x *\<^sub>R f (g x) /\<^sub>R y" if "x \ S" for x using nonneg_h [of n x] \y > 0\ nonneg_fg [of x] h_le_f [of x n] that by (auto simp: divide_simps mult_left_mono) qed (use S in auto) then have int_det: "(\t. \det (matrix (g' t))\) integrable_on ({t. h n (g t) = y} \ S)" using integrable_restrict_Int by force have "(g ` ({t. h n (g t) = y} \ S)) \ lmeasurable" apply (rule measurable_differentiable_image [OF h_lmeas]) apply (blast intro: has_derivative_within_subset [OF der_g]) apply (rule int_det) done moreover have "g ` ({t. h n (g t) = y} \ S) = {x. h n x = y} \ g ` S" by blast moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \ S)) \ integral ({t. h n (g t) = y} \ S) (\t. \det (matrix (g' t))\)" apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) apply (blast intro: has_derivative_within_subset [OF der_g]) done ultimately show ?thesis using \y > 0\ integral_restrict_Int [of S "{t. h n (g t) = y}" "\t. \det (matrix (g' t))\ * y"] apply (simp add: integrable_on_indicator integrable_on_cmult_iff integral_indicator) apply (simp add: indicator_def if_distrib cong: if_cong) done qed have hn_int: "h n integrable_on g ` S" apply (subst hn_eq) using yind by (force intro: integrable_sum [OF fin_R]) then show ?thesis proof have "?lhs = integral (g ` S) (\x. \y\range (h n). y * indicat_real {x. h n x = y} x)" by (metis hn_eq) also have "\ = (\y\range (h n). integral (g ` S) (\x. y * indicat_real {x. h n x = y} x))" by (rule integral_sum [OF fin_R]) (use yind in blast) also have "\ \ (\y\range (h n). integral S (\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)))" using yind by (force intro: sum_mono) also have "\ = integral S (\u. \y\range (h n). \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u))" proof (rule integral_sum [OF fin_R, symmetric]) fix y assume y: "y \ ?R" with nonneg_h have "y \ 0" by auto show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. indicat_real {x. h n x = y} (g x)) \ borel_measurable (lebesgue_on S)" using h_lmeas S by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff) then show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) \ borel_measurable (lebesgue_on S)" by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g]) next fix x assume "x \ S" have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" by (metis (full_types) \x \ S\ h_le_f indicator_def mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg) with \y \ 0\ show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \ ?D x * f(g x)" by (simp add: abs_mult mult.assoc mult_left_mono) qed (use S det_int_fg in auto) qed also have "\ = integral S (\T. \det (matrix (g' T))\ * (\y\range (h n). y * indicat_real {x. h n x = y} (g T)))" by (simp add: sum_distrib_left mult.assoc) also have "\ = ?rhs" by (metis hn_eq) finally show "integral (g ` S) (h n) \ ?rhs" . qed qed have le: "integral S (\T. \det (matrix (g' T))\ * h n (g T)) \ ?b" for n proof (rule integral_le) show "(\T. \det (matrix (g' T))\ * h n (g T)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\T. \det (matrix (g' T))\ *\<^sub>R h n (g T)) \ borel_measurable (lebesgue_on S)" proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \S \ sets lebesgue\) have eq: "{x \ S. f x \ a} = (\b \ (f ` S) \ atMost a. {x. f x = b} \ S)" for f and a::real by auto have "finite ((\x. h n (g x)) ` S \ {..a})" for a by (force intro: finite_subset [OF _ fin_R]) with h_lmeas [of n] show "(\x. h n (g x)) \ borel_measurable (lebesgue_on S)" apply (simp add: borel_measurable_vimage_halfspace_component_le \S \ sets lebesgue\ sets_restrict_space_iff eq) by (metis (mono_tags) SUP_inf sets.finite_UN) qed (use der_g in blast) then show "(\T. \det (matrix (g' T))\ * h n (g T)) \ borel_measurable (lebesgue_on S)" by simp show "norm (?D x * h n (g x)) \ ?D x *\<^sub>R f (g x)" if "x \ S" for x by (simp add: h_le_f mult_left_mono nonneg_h that) qed (use S det_int_fg in auto) show "?D x * h n (g x) \ ?D x * f (g x)" if "x \ S" for x by (simp add: \x \ S\ h_le_f mult_left_mono) show "(\x. ?D x * f (g x)) integrable_on S" using det_int_fg by blast qed have "f integrable_on g ` S \ (\k. integral (g ` S) (h k)) \ integral (g ` S) f" proof (rule monotone_convergence_increasing) have "\integral (g ` S) (h n)\ \ integral S (\x. ?D x * f (g x))" for n proof - have "\integral (g ` S) (h n)\ = integral (g ` S) (h n)" using hint by (simp add: integral_nonneg nonneg_h) also have "\ \ integral S (\x. ?D x * f (g x))" using hint le by (meson order_trans) finally show ?thesis . qed then show "bounded (range (\k. integral (g ` S) (h k)))" by (force simp: bounded_iff) qed (use h_inc lim hint in auto) moreover have "integral (g ` S) (h n) \ integral S (\x. ?D x * f (g x))" for n using hint by (blast intro: le order_trans) ultimately show ?thesis by (auto intro: Lim_bounded) qed lemma integral_on_image_ubound_nonneg: fixes f :: "real^'n::{finite,wellorder} \ real" assumes nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. det (matrix (g' x))" define S' where "S' \ {x \ S. ?D x * f(g x) \ 0}" then have der_gS': "\x. x \ S' \ (g has_derivative g' x) (at x within S')" by (metis (mono_tags, lifting) der_g has_derivative_within_subset mem_Collect_eq subset_iff) have "(\x. if x \ S then \?D x\ * f (g x) else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV intS) then have Df_borel: "(\x. if x \ S then \?D x\ * f (g x) else 0) \ borel_measurable lebesgue" using integrable_imp_measurable lebesgue_on_UNIV_eq by force have S': "S' \ sets lebesgue" proof - from Df_borel borel_measurable_vimage_open [of _ UNIV] have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ T} \ sets lebesgue" if "open T" for T using that unfolding lebesgue_on_UNIV_eq by (fastforce simp add: dest!: spec) then have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ -{0}} \ sets lebesgue" using open_Compl by blast then show ?thesis by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong) qed then have gS': "g ` S' \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "g differentiable_on S'" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_within_subset) qed auto have f: "f \ borel_measurable (lebesgue_on (g ` S'))" proof (clarsimp simp add: borel_measurable_vimage_open) fix T :: "real set" assume "open T" have "{x \ g ` S'. f x \ T} = g ` {x \ S'. f(g x) \ T}" by blast moreover have "g ` {x \ S'. f(g x) \ T} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) let ?h = "\x. \?D x\ * f (g x) /\<^sub>R \?D x\" have "(\x. if x \ S' then \?D x\ * f (g x) else 0) = (\x. if x \ S then \?D x\ * f (g x) else 0)" by (auto simp: S'_def) also have "\ \ borel_measurable lebesgue" by (rule Df_borel) finally have *: "(\x. \?D x\ * f (g x)) \ borel_measurable (lebesgue_on S')" by (simp add: borel_measurable_if_D) have "?h \ borel_measurable (lebesgue_on S')" by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') moreover have "?h x = f(g x)" if "x \ S'" for x using that by (auto simp: S'_def) ultimately have "(\x. f(g x)) \ borel_measurable (lebesgue_on S')" by (metis (no_types, lifting) measurable_lebesgue_cong) then show "{x \ S'. f (g x) \ T} \ sets lebesgue" by (simp add: \S' \ sets lebesgue\ \open T\ borel_measurable_vimage_open sets_restrict_space_iff) show "g differentiable_on {x \ S'. f (g x) \ T}" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_within_subset) qed auto ultimately have "{x \ g ` S'. f x \ T} \ sets lebesgue" by metis then show "{x \ g ` S'. f x \ T} \ sets (lebesgue_on (g ` S'))" by (simp add: \g ` S' \ sets lebesgue\ sets_restrict_space_iff) qed have intS': "(\x. \?D x\ * f (g x)) integrable_on S'" using intS by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible) have lebS': "{x \ S'. g x \ T} \ sets lebesgue" if "T \ g ` S'" "T \ sets lebesgue" for T proof - have "g \ borel_measurable (lebesgue_on S')" using der_gS' has_derivative_continuous_on S' by (blast intro: continuous_imp_measurable_on_sets_lebesgue) moreover have "{x \ S'. g x \ U} \ sets lebesgue" if "negligible U" "U \ g ` S'" for U proof (intro negligible_imp_sets negligible_differentiable_vimage that) fix x assume x: "x \ S'" then have "linear (g' x)" using der_gS' has_derivative_linear by blast with x show "inj (g' x)" by (auto simp: S'_def det_nz_iff_inj) qed (use der_gS' in auto) ultimately show ?thesis using double_lebesgue_sets [OF S' gS' order_refl] that by blast qed have int_gS': "f integrable_on g ` S' \ integral (g ` S') f \ integral S' (\x. \?D x\ * f(g x))" using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast have "negligible (g ` {x \ S. det(matrix(g' x)) = 0})" proof (rule baby_Sard, simp_all) fix x assume x: "x \ S \ det (matrix (g' x)) = 0" then show "(g has_derivative g' x) (at x within {x \ S. det (matrix (g' x)) = 0})" by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI) then show "rank (matrix (g' x)) < CARD('n)" using det_nz_iff_inj matrix_vector_mul_linear x by (fastforce simp add: less_rank_noninjective) qed then have negg: "negligible (g ` S - g ` {x \ S. ?D x \ 0})" by (rule negligible_subset) (auto simp: S'_def) have null: "g ` {x \ S. ?D x \ 0} - g ` S = {}" by (auto simp: S'_def) let ?F = "{x \ S. f (g x) \ 0}" have eq: "g ` S' = g ` ?F \ g ` {x \ S. ?D x \ 0}" by (auto simp: S'_def image_iff) show ?thesis proof have "((\x. if x \ g ` ?F then f x else 0) integrable_on g ` {x \ S. ?D x \ 0})" using int_gS' eq integrable_restrict_Int [where f=f] by simp then have "f integrable_on g ` {x \ S. ?D x \ 0}" by (auto simp: image_iff elim!: integrable_eq) then show "f integrable_on g ` S" apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) using negg null by auto have "integral (g ` S) f = integral (g ` {x \ S. ?D x \ 0}) f" using negg by (auto intro: negligible_subset integral_spike_set) also have "\ = integral (g ` {x \ S. ?D x \ 0}) (\x. if x \ g ` ?F then f x else 0)" by (auto simp: image_iff intro!: integral_cong) also have "\ = integral (g ` S') f" using eq integral_restrict_Int by simp also have "\ \ integral S' (\x. \?D x\ * f(g x))" by (metis int_gS') also have "\ \ ?b" by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto) finally show "integral (g ` S) f \ ?b" . qed qed lemma absolutely_integrable_on_image_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" proof - let ?D = "\x. \det (matrix (g' x))\ * f (g x)" let ?N = "{x \ S. f (g x) < 0}" and ?P = "{x \ S. f (g x) > 0}" have eq: "{x. (if x \ S then ?D x else 0) > 0} = {x \ S. ?D x > 0}" "{x. (if x \ S then ?D x else 0) < 0} = {x \ S. ?D x < 0}" by auto have "?D integrable_on S" using intS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?D x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?D x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have Dlt: "{x \ S. ?D x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have Dgt: "{x \ S. ?D x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have dfgbm: "?D \ borel_measurable (lebesgue_on S)" using intS absolutely_integrable_on_def integrable_imp_measurable by blast have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \ ?N" for x using der_g has_derivative_within_subset that by force have "(\x. - f x) integrable_on g ` ?N \ integral (g ` ?N) (\x. - f x) \ integral ?N (\x. \det (matrix (g' x))\ * - f (g x))" proof (rule integral_on_image_ubound_nonneg [OF _ der_gN]) have 1: "?D integrable_on {x \ S. ?D x < 0}" using Dlt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) have "uminus \ (\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1]) then show "(\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: integrable_neg_iff o_def) qed auto then have "f integrable_on g ` ?N" by (simp add: integrable_neg_iff) moreover have "g ` ?N = {y \ g ` S. f y < 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y < 0}" by simp then have N: "f absolutely_integrable_on {y \ g ` S. f y < 0}" by (rule absolutely_integrable_absolutely_integrable_ubound) auto have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \ ?P" for x using der_g has_derivative_within_subset that by force have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) have "?D integrable_on {x \ S. 0 < ?D x}" using Dgt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) then show "?D integrable_on ?P" apply (rule integrable_spike_set) by (auto simp: zero_less_mult_iff empty_imp_negligible) qed auto then have "f integrable_on g ` ?P" by metis moreover have "g ` ?P = {y \ g ` S. f y > 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y > 0}" by simp then have P: "f absolutely_integrable_on {y \ g ` S. f y > 0}" by (rule absolutely_integrable_absolutely_integrable_lbound) auto have "(\x. if x \ g ` S \ f x < 0 \ x \ g ` S \ 0 < f x then f x else 0) = (\x. if x \ g ` S then f x else 0)" by auto then show ?thesis using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f] by simp qed proposition absolutely_integrable_on_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) using absolutely_integrable_component [OF intS] by auto proposition integral_on_image_ubound: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes"\x. x \ S \ 0 \ f(g x)" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" using integral_on_image_ubound_nonneg [OF assms] by simp subsection\Change-of-variables theorem\ text\The classic change-of-variables theorem. We have two versions with quite general hypotheses, the first that the transforming function has a continuous inverse, the second that the base set is Lebesgue measurable.\ lemma cov_invertible_nonneg_le: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and f0: "\y. y \ T \ 0 \ f y" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "f integrable_on T \ (integral T f) \ b \ (\x. \det (matrix (g' x))\ * f(g x)) integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) \ b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ have gS: "g differentiable_on S" by (meson der_g differentiable_def differentiable_on_def) let ?D = "\x. \det (matrix (g' x))\ * f (g x)" show ?thesis proof assume ?lhs then have fT: "f integrable_on T" and intf: "integral T f \ b" by blast+ show ?rhs proof let ?fgh = "\x. \det (matrix (h' x))\ * (\det (matrix (g' (h x)))\ * f (g (h x)))" have ddf: "?fgh x = f x" if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis by (simp add: gh that) qed have "?D integrable_on (h ` T)" proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) show "(\x. ?fgh x) absolutely_integrable_on T" proof (subst absolutely_integrable_on_iff_nonneg) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (simp add: zero_le_mult_iff f0 gh) qed (use der_h in auto) with Seq show "(\x. ?D x) integrable_on S" by simp have "integral S (\x. ?D x) \ integral T (\x. ?fgh x)" unfolding Seq proof (rule integral_on_image_ubound) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (use f0 gh der_h in auto) also have "\ = integral T f" by (force simp: ddf intro: integral_cong) also have "\ \ b" by (rule intf) finally show "integral S (\x. ?D x) \ b" . qed next assume R: ?rhs then have "f integrable_on g ` S" using der_g f0 hg integral_on_image_ubound_nonneg by blast moreover have "integral (g ` S) f \ integral S (\x. ?D x)" by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto) ultimately show ?lhs using R by (simp add: Teq) qed qed lemma cov_invertible_nonneg_eq: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "\y. y \ T \ (h has_derivative h' y) (at y within T)" and "\y. y \ T \ 0 \ f y" and "\x. x \ S \ g x \ T \ h(g x) = x" and "\y. y \ T \ h y \ S \ g(h y) = y" and "\y. y \ T \ h' y \ g'(h y) = id" shows "((\x. \det (matrix (g' x))\ * f(g x)) has_integral b) S \ (f has_integral b) T" using cov_invertible_nonneg_le [OF assms] by (simp add: has_integral_iff) (meson eq_iff) lemma cov_invertible_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ let ?DP = "\x. \det (matrix (g' x))\ * f(g x)" and ?DN = "\x. \det (matrix (g' x))\ * -f(g x)" have "+": "(?DP has_integral b) {x \ S. f (g x) > 0} \ (f has_integral b) {y \ T. f y > 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. f (g x)) -` Y \ {x \ S. f (g x) > 0} = ((\x. f (g x)) -` Y \ S) \ {x \ S. f (g x) > 0}" for Y by auto show "(g has_derivative g' x) (at x within {x \ S. f (g x) > 0})" if "x \ {x \ S. f (g x) > 0}" for x using that der_g has_derivative_within_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y > 0})" if "y \ {y \ T. f y > 0}" for y using that der_h has_derivative_within_subset by fastforce qed (use gh hg id in auto) have "-": "(?DN has_integral b) {x \ S. f (g x) < 0} \ ((\x. - f x) has_integral b) {y \ T. f y < 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. - f (g x)) -` y \ {x \ S. f (g x) < 0} = ((\x. f (g x)) -` uminus ` y \ S) \ {x \ S. f (g x) < 0}" for y using image_iff by fastforce show "(g has_derivative g' x) (at x within {x \ S. f (g x) < 0})" if "x \ {x \ S. f (g x) < 0}" for x using that der_g has_derivative_within_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y < 0})" if "y \ {y \ T. f y < 0}" for y using that der_h has_derivative_within_subset by fastforce qed (use gh hg id in auto) show ?thesis proof assume LHS: ?lhs have eq: "{x. (if x \ S then ?DP x else 0) > 0} = {x \ S. ?DP x > 0}" "{x. (if x \ S then ?DP x else 0) < 0} = {x \ S. ?DP x < 0}" by auto have "?DP integrable_on S" using LHS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?DP x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?DP x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have SN: "{x \ S. ?DP x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have SP: "{x \ S. ?DP x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have "?DP absolutely_integrable_on {x \ S. ?DP x > 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP) then have aP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible) have "?DP absolutely_integrable_on {x \ S. ?DP x < 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN) then have aN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible) have fN: "f integrable_on {y \ T. f y < 0}" "integral {y \ T. f y < 0} f = integral {x \ S. f (g x) < 0} ?DP" using "-" [of "integral {x \ S. f(g x) < 0} ?DN"] aN by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faN: "f absolutely_integrable_on {y \ T. f y < 0}" apply (rule absolutely_integrable_integrable_bound [where g = "\x. - f x"]) using fN by (auto simp: integrable_neg_iff) have fP: "f integrable_on {y \ T. f y > 0}" "integral {y \ T. f y > 0} f = integral {x \ S. f (g x) > 0} ?DP" using "+" [of "integral {x \ S. f(g x) > 0} ?DP"] aP by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faP: "f absolutely_integrable_on {y \ T. f y > 0}" apply (rule absolutely_integrable_integrable_bound [where g = f]) using fP by auto have fa: "f absolutely_integrable_on ({y \ T. f y < 0} \ {y \ T. f y > 0})" by (rule absolutely_integrable_Un [OF faN faP]) show ?rhs proof have eq: "((if x \ T \ f x < 0 \ x \ T \ 0 < f x then 1 else 0) * f x) = (if x \ T then 1 else 0) * f x" for x by auto show "f absolutely_integrable_on T" using fa by (simp add: indicator_def set_integrable_def eq) have [simp]: "{y \ T. f y < 0} \ {y \ T. 0 < f y} = {}" for T and f :: "(real^'n::_) \ real" by auto have "integral T f = integral ({y \ T. f y < 0} \ {y \ T. f y > 0}) f" by (intro empty_imp_negligible integral_spike_set) (auto simp: eq) also have "\ = integral {y \ T. f y < 0} f + integral {y \ T. f y > 0} f" using fN fP by simp also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" by (simp add: fN fP) also have "\ = integral ({x \ S. f (g x) < 0} \ {x \ S. 0 < f (g x)}) ?DP" using aP aN by (simp add: set_lebesgue_integral_eq_integral) also have "\ = integral S ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using LHS by simp finally show "integral T f = b" . qed next assume RHS: ?rhs have eq: "{x. (if x \ T then f x else 0) > 0} = {x \ T. f x > 0}" "{x. (if x \ T then f x else 0) < 0} = {x \ T. f x < 0}" by auto have "f integrable_on T" using RHS absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have TN: "{x \ T. f x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have TP: "{x \ T. f x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have aint: "f absolutely_integrable_on {y. y \ T \ 0 < (f y)}" "f absolutely_integrable_on {y. y \ T \ (f y) < 0}" and intT: "integral T f = b" using set_integrable_subset [of _ T] TP TN RHS by blast+ show ?lhs proof have fN: "f integrable_on {v \ T. f v < 0}" using absolutely_integrable_on_def aint by blast then have DN: "(?DN has_integral integral {y \ T. f y < 0} (\x. - f x)) {x \ S. f (g x) < 0}" using "-" [of "integral {y \ T. f y < 0} (\x. - f x)"] by (simp add: has_integral_neg_iff integrable_integral) have aDN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DN]) using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+ have fP: "f integrable_on {v \ T. f v > 0}" using absolutely_integrable_on_def aint by blast then have DP: "(?DP has_integral integral {y \ T. f y > 0} f) {x \ S. f (g x) > 0}" using "+" [of "integral {y \ T. f y > 0} f"] by (simp add: has_integral_neg_iff integrable_integral) have aDP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DP]) using DP hg by (fastforce simp: integrable_neg_iff)+ have eq: "(if x \ S then 1 else 0) * ?DP x = (if x \ S \ f (g x) < 0 \ x \ S \ f (g x) > 0 then 1 else 0) * ?DP x" for x by force have "?DP absolutely_integrable_on ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0})" by (rule absolutely_integrable_Un [OF aDN aDP]) then show I: "?DP absolutely_integrable_on S" by (simp add: indicator_def eq set_integrable_def) have [simp]: "{y \ S. f y < 0} \ {y \ S. 0 < f y} = {}" for S and f :: "(real^'n::_) \ real" by auto have "integral S ?DP = integral ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0}) ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" using aDN aDP by (simp add: set_lebesgue_integral_eq_integral) also have "\ = - integral {y \ T. f y < 0} (\x. - f x) + integral {y \ T. f y > 0} f" using DN DP by (auto simp: has_integral_iff) also have "\ = integral ({x \ T. f x < 0} \ {x \ T. 0 < f x}) f" by (simp add: fN fP) also have "\ = integral T f" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using intT by simp finally show "integral S ?DP = b" . qed qed qed lemma cv_inv_version3: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have "((\x. \det (matrix (g' x))\ * f(g x) $ i) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * (f(g x) $ i)) = b $ i) \ ((\x. f x $ i) absolutely_integrable_on T \ integral T (\x. f x $ i) = b $ i)" for i by (rule cov_invertible_real [OF der_g der_h hg gh id]) then have "?D absolutely_integrable_on S \ (?D has_integral b) S \ f absolutely_integrable_on T \ (f has_integral b) T" unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f] absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D] by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric] has_integral_iff set_lebesgue_integral_eq_integral) then show ?thesis using absolutely_integrable_on_def by blast qed lemma cv_inv_version4: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S) \ invertible(matrix(g' x))" and hg: "\x. x \ S \ continuous_on (g ` S) h \ h(g x) = x" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "\x. \h'. x \ S \ (g has_derivative g' x) (at x within S) \ linear h' \ g' x \ h' = id \ h' \ g' x = id" using der_g matrix_invertible has_derivative_linear by blast then obtain h' where h': "\x. x \ S \ (g has_derivative g' x) (at x within S) \ linear (h' x) \ g' x \ (h' x) = id \ (h' x) \ g' x = id" by metis show ?thesis proof (rule cv_inv_version3) show "\y. y \ g ` S \ (h has_derivative h' (h y)) (at y within g ` S)" using h' hg by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within) qed (use h' hg in auto) qed theorem has_absolute_integral_change_of_variables_invertible: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and hg: "\x. x \ S \ h(g x) = x" and conth: "continuous_on (g ` S) h" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" (is "?lhs = ?rhs") proof - let ?S = "{x \ S. invertible (matrix (g' x))}" and ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have *: "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b" proof (rule cv_inv_version4) show "(g has_derivative g' x) (at x within ?S) \ invertible (matrix (g' x))" if "x \ ?S" for x using der_g that has_derivative_within_subset that by fastforce show "continuous_on (g ` ?S) h \ h (g x) = x" if "x \ ?S" for x using that continuous_on_subset [OF conth] by (simp add: hg image_mono) qed have "(g has_derivative g' x) (at x within {x \ S. rank (matrix (g' x)) < CARD('m)})" if "x \ S" for x by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI that) then have "negligible (g ` {x \ S. \ invertible (matrix (g' x))})" by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard) then have neg: "negligible {x \ g ` S. x \ g ` ?S \ f x \ 0}" by (auto intro: negligible_subset) have [simp]: "{x \ g ` ?S. x \ g ` S \ f x \ 0} = {}" by auto have "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" apply (intro conj_cong absolutely_integrable_spike_set_eq) apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg) done moreover have "f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg) ultimately show ?thesis using "*" by blast qed theorem has_absolute_integral_change_of_variables_compact: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "compact S" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b)" proof - obtain h where hg: "\x. x \ S \ h(g x) = x" using inj by (metis the_inv_into_f_f) have conth: "continuous_on (g ` S) h" by (metis \compact S\ continuous_on_inv der_g has_derivative_continuous_on hg) show ?thesis by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth]) qed lemma has_absolute_integral_change_of_variables_compact_family: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes compact: "\n::nat. compact (F n)" and der_g: "\x. x \ (\n. F n) \ (g has_derivative g' x) (at x within (\n. F n))" and inj: "inj_on g (\n. F n)" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on (\n. F n) \ integral (\n. F n) (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` (\n. F n)) \ integral (g ` (\n. F n)) f = b)" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" let ?U = "\n. \m\n. F m" let ?lift = "vec::real\real^1" have F_leb: "F m \ sets lebesgue" for m by (simp add: compact borel_compact) have iff: "(\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \ integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) = b \ f absolutely_integrable_on (g ` (?U n)) \ integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \ real^'k" proof (rule has_absolute_integral_change_of_variables_compact) show "compact (?U n)" by (simp add: compact compact_UN) show "(g has_derivative g' x) (at x within (?U n))" if "x \ ?U n" for x using that by (blast intro!: has_derivative_within_subset [OF der_g]) show "inj_on g (?U n)" using inj by (auto simp: inj_on_def) qed show ?thesis unfolding image_UN proof safe assume DS: "?D absolutely_integrable_on (\n. F n)" and b: "b = integral (\n. F n) ?D" have DU: "\n. ?D absolutely_integrable_on (?U n)" "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" using integral_countable_UN [OF DS F_leb] by auto with iff have fag: "f absolutely_integrable_on g ` (?U n)" and fg_int: "integral (\m\n. g ` F m) f = integral (?U n) ?D" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\m. g ` F m) then norm(f x) else 0" have "(\x. if x \ (\m. g ` F m) then f x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ (\m\k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using fag by (simp add: image_UN) let ?nf = "\n x. if x \ (\m\n. g ` F m) then norm(f x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nf k integrable_on UNIV" for k using fag unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n have "(norm \ ?D) absolutely_integrable_on ?U n" by (intro absolutely_integrable_norm DU) then have "integral (g ` ?U n) (norm \ f) = integral (?U n) (norm \ ?D)" using iff [of n "vec \ norm \ f" "integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R (?lift \ norm \ f) (g x))"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def) } moreover have "bounded (range (\k. integral (?U k) (norm \ ?D)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (?U k) (norm \ ?D)) \ integral (\n. F n) (norm \ ?D)" unfolding integral_restrict_UNIV [of _ "norm \ ?D", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ \ (F ` {..k}) then (norm \ ?D) x else 0) integrable_on UNIV" "(\x. if x \ (\n. F n) then (norm \ ?D) x else 0) integrable_on UNIV" using DU(1) DS unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nf k)))" by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def) next show "(\k. if x \ (\m\k. g ` F m) then norm (f x) else 0) \ (if x \ (\m. g ` F m) then norm (f x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ (\m\k. g ` F m) then f x else 0) \ (if x \ (\m. g ` F m) then f x else 0)" for x proof clarsimp fix m y assume "y \ F m" show "(\k. if \x\{..k}. g y \ g ` F x then f (g y) else 0) \ f (g y)" using \y \ F m\ by (force intro: tendsto_eventually eventually_sequentiallyI [of m]) qed qed auto then show fai: "f absolutely_integrable_on (\m. g ` F m)" using absolutely_integrable_restrict_UNIV by blast show "integral ((\x. g ` F x)) f = integral (\n. F n) ?D" proof (rule LIMSEQ_unique) show "(\n. integral (?U n) ?D) \ integral (\x. g ` F x) f" unfolding fg_int [symmetric] proof (rule integral_countable_UN [OF fai]) show "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) qed auto qed next show "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" by (rule DU) qed next assume fs: "f absolutely_integrable_on (\x. g ` F x)" and b: "b = integral ((\x. g ` F x)) f" have gF_leb: "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" using der_g unfolding differentiable_def differentiable_on_def by (meson Sup_upper UNIV_I UnionI has_derivative_within_subset image_eqI) qed auto have fgU: "\n. f absolutely_integrable_on (\m\n. g ` F m)" "(\n. integral (\m\n. g ` F m) f) \ integral (\m. g ` F m) f" using integral_countable_UN [OF fs gF_leb] by auto with iff have DUn: "?D absolutely_integrable_on ?U n" and D_int: "integral (?U n) ?D = integral (\m\n. g ` F m) f" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\n. F n) then norm(?D x) else 0" have "(\x. if x \ (\n. F n) then ?D x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using DUn by simp let ?nD = "\n x. if x \ ?U n then norm(?D x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nD k integrable_on UNIV" for k using DUn unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n::nat have "(norm \ f) absolutely_integrable_on (\m\n. g ` F m)" apply (rule absolutely_integrable_norm) using fgU by blast then have "integral (?U n) (norm \ ?D) = integral (g ` ?U n) (norm \ f)" using iff [of n "?lift \ norm \ f" "integral (g ` ?U n) (?lift \ norm \ f)"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) } moreover have "bounded (range (\k. integral (g ` ?U k) (norm \ f)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (g ` ?U k) (norm \ f)) \ integral (g ` (\n. F n)) (norm \ f)" unfolding integral_restrict_UNIV [of _ "norm \ f", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ g ` ?U k then (norm \ f) x else 0) integrable_on UNIV" "(\x. if x \ g ` (\n. F n) then (norm \ f) x else 0) integrable_on UNIV" using fgU fs unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by (auto simp: image_UN) qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nD k)))" unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp next show "(\k. if x \ ?U k then norm (?D x) else 0) \ (if x \ (\n. F n) then norm (?D x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ ?U k then ?D x else 0) \ (if x \ (\n. F n) then ?D x else 0)" for x proof clarsimp fix n assume "x \ F n" show "(\m. if \j\{..m}. x \ F j then ?D x else 0) \ ?D x" using \x \ F n\ by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n]) qed qed auto then show Dai: "?D absolutely_integrable_on (\n. F n)" unfolding absolutely_integrable_restrict_UNIV by simp show "integral (\n. F n) ?D = integral ((\x. g ` F x)) f" proof (rule LIMSEQ_unique) show "(\n. integral (\m\n. g ` F m) f) \ integral (\x. g ` F x) f" by (rule fgU) show "(\n. integral (\m\n. g ` F m) f) \ integral (\n. F n) ?D" unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) qed qed qed theorem has_absolute_integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - obtain C N where "fsigma C" and N: "N \ null_sets lebesgue" and CNS: "C \ N = S" and "disjnt C N" using lebesgue_set_almost_fsigma [OF S] . then obtain F :: "nat \ (real^'m::_) set" where F: "range F \ Collect compact" and Ceq: "C = Union(range F)" using fsigma_Union_compact by metis have "negligible N" using N by (simp add: negligible_iff_null_sets) let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" have "?D absolutely_integrable_on C \ integral C ?D = b \ f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b" unfolding Ceq proof (rule has_absolute_integral_change_of_variables_compact_family) fix n x assume "x \ \(F ` UNIV)" then show "(g has_derivative g' x) (at x within \(F ` UNIV))" using Ceq \C \ N = S\ der_g has_derivative_within_subset by blast next have "\(F ` UNIV) \ S" using Ceq \C \ N = S\ by blast then show "inj_on g (\(F ` UNIV))" using inj by (meson inj_on_subset) qed (use F in auto) moreover have "?D absolutely_integrable_on C \ integral C ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" proof (rule conj_cong) have neg: "negligible {x \ C - S. ?D x \ 0}" "negligible {x \ S - C. ?D x \ 0}" using CNS by (blast intro: negligible_subset [OF \negligible N\])+ then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)" by (rule absolutely_integrable_spike_set_eq) show "(integral C ?D = b) \ (integral S ?D = b)" using integral_spike_set [OF neg] by simp qed moreover have "f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (rule conj_cong) have "g differentiable_on N" by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2) with \negligible N\ have neg_gN: "negligible (g ` N)" by (blast intro: negligible_differentiable_image_negligible) have neg: "negligible {x \ g ` C - g ` S. f x \ 0}" "negligible {x \ g ` S - g ` C. f x \ 0}" using CNS by (blast intro: negligible_subset [OF neg_gN])+ then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)" by (rule absolutely_integrable_spike_set_eq) show "(integral (g ` C) f = b) \ (integral (g ` S) f = b)" using integral_spike_set [OF neg] by simp qed ultimately show ?thesis by simp qed corollary absolutely_integrable_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "inj_on g S" shows "f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" using assms has_absolute_integral_change_of_variables by blast corollary integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" and disj: "(f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S)" shows "integral (g ` S) f = integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x))" using has_absolute_integral_change_of_variables [OF S der_g inj] disj by blast lemma has_absolute_integral_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - let ?lift = "vec :: real \ real^1" let ?drop = "(\x::real^1. x $ 1)" have S': "?lift ` S \ sets lebesgue" by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec) have "((\x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)" if "z \ S" for z using der_g [OF that] by (simp add: has_vector_derivative_def has_derivative_vector_1) then have der': "\x. x \ ?lift ` S \ (?lift \ g \ ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)" by (auto simp: o_def) have inj': "inj_on (vec \ g \ ?drop) (vec ` S)" using inj by (simp add: inj_on_def) let ?fg = "\x. \g' x\ *\<^sub>R f(g x)" have "((\x. ?fg x $ i) absolutely_integrable_on S \ ((\x. ?fg x $ i) has_integral b $ i) S \ (\x. f x $ i) absolutely_integrable_on g ` S \ ((\x. f x $ i) has_integral b $ i) (g ` S))" for i using has_absolute_integral_change_of_variables [OF S' der' inj', of "\x. ?lift(f (?drop x) $ i)" "?lift (b$i)"] unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff) then have "?fg absolutely_integrable_on S \ (?fg has_integral b) S \ f absolutely_integrable_on (g ` S) \ (f has_integral b) (g ` S)" unfolding has_integral_componentwise_iff [where y=b] absolutely_integrable_componentwise_iff [where f=f] absolutely_integrable_componentwise_iff [where f = ?fg] by (force simp: Basis_vec_def cart_eq_inner_axis) then show ?thesis using absolutely_integrable_on_def by blast qed corollary absolutely_integrable_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(f absolutely_integrable_on g ` S \ (\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S)" using has_absolute_integral_change_of_variables_1 [OF assms] by auto subsection\Change of variables for integrals: special case of linear function\ lemma has_absolute_integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix g)\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (cases "det(matrix g) = 0") case True then have "negligible(g ` S)" using assms det_nz_iff_inj negligible_linear_singular_image by blast with True show ?thesis by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible) next case False then obtain h where h: "\x. x \ S \ h (g x) = x" "linear h" using assms det_nz_iff_inj linear_injective_isomorphism by blast show ?thesis proof (rule has_absolute_integral_change_of_variables_invertible) show "(g has_derivative g) (at x within S)" for x by (simp add: assms linear_imp_has_derivative) show "continuous_on (g ` S) h" using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast qed (use h in auto) qed lemma absolutely_integrable_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ f absolutely_integrable_on (g ` S)" using assms has_absolute_integral_change_of_variables_linear by blast lemma absolutely_integrable_on_linear_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S \ det(matrix g) = 0" unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff by (auto simp: set_integrable_def) lemma integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" and "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S" shows "integral (g ` S) f = \det (matrix g)\ *\<^sub>R integral S (f \ g)" proof - have "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S) \ (f absolutely_integrable_on g ` S)" using absolutely_integrable_on_linear_image assms by blast moreover have ?thesis if "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)" using has_absolute_integral_change_of_variables_linear [OF \linear g\] that by (auto simp: o_def) ultimately show ?thesis using absolutely_integrable_change_of_variables_linear [OF \linear g\] by blast qed subsection\Change of variable for measure\ lemma has_measure_differentiable_image: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) = m \ ((\x. \det (matrix (f' x))\) has_integral m) S" using has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) lemma measurable_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) integrable_on S" using has_measure_differentiable_image [OF assms] by blast lemma measurable_differentiable_image_alt: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) absolutely_integrable_on S" using measurable_differentiable_image_eq [OF assms] by (simp only: absolutely_integrable_on_iff_nonneg) lemma measure_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and der_f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and inj: "inj_on f S" and intS: "(\x. \det (matrix (f' x))\) integrable_on S" shows "measure lebesgue (f ` S) = integral S (\x. \det (matrix (f' x))\)" using measurable_differentiable_image_eq [OF S der_f inj] assms has_measure_differentiable_image by blast end diff --git a/src/HOL/Analysis/Connected.thy b/src/HOL/Analysis/Connected.thy --- a/src/HOL/Analysis/Connected.thy +++ b/src/HOL/Analysis/Connected.thy @@ -1,741 +1,741 @@ (* Author: L C Paulson, University of Cambridge Material split off from Topology_Euclidean_Space *) section \Connected Components\ theory Connected imports Abstract_Topology_2 begin subsection\<^marker>\tag unimportant\ \Connectedness\ lemma connected_local: "connected S \ \ (\e1 e2. openin (top_of_set S) e1 \ openin (top_of_set S) e2 \ S \ e1 \ e2 \ e1 \ e2 = {} \ e1 \ {} \ e2 \ {})" unfolding connected_def openin_open by safe blast+ lemma exists_diff: fixes P :: "'a set \ bool" shows "(\S. P (- S)) \ (\S. P S)" (is "?lhs \ ?rhs") proof - have ?rhs if ?lhs using that by blast moreover have "P (- (- S))" if "P S" for S proof - have "S = - (- S)" by simp with that show ?thesis by metis qed ultimately show ?thesis by metis qed lemma connected_clopen: "connected S \ (\T. openin (top_of_set S) T \ closedin (top_of_set S) T \ T = {} \ T = S)" (is "?lhs \ ?rhs") proof - have "\ connected S \ (\e1 e2. open e1 \ open (- e2) \ S \ e1 \ (- e2) \ e1 \ (- e2) \ S = {} \ e1 \ S \ {} \ (- e2) \ S \ {})" unfolding connected_def openin_open closedin_closed by (metis double_complement) then have th0: "connected S \ \ (\e2 e1. closed e2 \ open e1 \ S \ e1 \ (- e2) \ e1 \ (- e2) \ S = {} \ e1 \ S \ {} \ (- e2) \ S \ {})" (is " _ \ \ (\e2 e1. ?P e2 e1)") by (simp add: closed_def) metis have th1: "?rhs \ \ (\t' t. closed t'\t = S\t' \ t\{} \ t\S \ (\t'. open t' \ t = S \ t'))" (is "_ \ \ (\t' t. ?Q t' t)") unfolding connected_def openin_open closedin_closed by auto have "(\e1. ?P e2 e1) \ (\t. ?Q e2 t)" for e2 proof - have "?P e2 e1 \ (\t. closed e2 \ t = S\e2 \ open e1 \ t = S\e1 \ t\{} \ t \ S)" for e1 by auto then show ?thesis by metis qed then have "\e2. (\e1. ?P e2 e1) \ (\t. ?Q e2 t)" by blast then show ?thesis by (simp add: th0 th1) qed subsection \Connected components, considered as a connectedness relation or a set\ definition\<^marker>\tag important\ "connected_component s x y \ \t. connected t \ t \ s \ x \ t \ y \ t" abbreviation "connected_component_set s x \ Collect (connected_component s x)" lemma connected_componentI: "connected t \ t \ s \ x \ t \ y \ t \ connected_component s x y" by (auto simp: connected_component_def) lemma connected_component_in: "connected_component s x y \ x \ s \ y \ s" by (auto simp: connected_component_def) lemma connected_component_refl: "x \ s \ connected_component s x x" by (auto simp: connected_component_def) (use connected_sing in blast) lemma connected_component_refl_eq [simp]: "connected_component s x x \ x \ s" by (auto simp: connected_component_refl) (auto simp: connected_component_def) lemma connected_component_sym: "connected_component s x y \ connected_component s y x" by (auto simp: connected_component_def) lemma connected_component_trans: "connected_component s x y \ connected_component s y z \ connected_component s x z" unfolding connected_component_def by (metis Int_iff Un_iff Un_subset_iff equals0D connected_Un) lemma connected_component_of_subset: "connected_component s x y \ s \ t \ connected_component t x y" by (auto simp: connected_component_def) lemma connected_component_Union: "connected_component_set s x = \{t. connected t \ x \ t \ t \ s}" by (auto simp: connected_component_def) lemma connected_connected_component [iff]: "connected (connected_component_set s x)" by (auto simp: connected_component_Union intro: connected_Union) lemma connected_iff_eq_connected_component_set: "connected s \ (\x \ s. connected_component_set s x = s)" proof (cases "s = {}") case True then show ?thesis by simp next case False then obtain x where "x \ s" by auto show ?thesis proof assume "connected s" then show "\x \ s. connected_component_set s x = s" by (force simp: connected_component_def) next assume "\x \ s. connected_component_set s x = s" then show "connected s" by (metis \x \ s\ connected_connected_component) qed qed lemma connected_component_subset: "connected_component_set s x \ s" using connected_component_in by blast lemma connected_component_eq_self: "connected s \ x \ s \ connected_component_set s x = s" by (simp add: connected_iff_eq_connected_component_set) lemma connected_iff_connected_component: "connected s \ (\x \ s. \y \ s. connected_component s x y)" using connected_component_in by (auto simp: connected_iff_eq_connected_component_set) lemma connected_component_maximal: "x \ t \ connected t \ t \ s \ t \ (connected_component_set s x)" using connected_component_eq_self connected_component_of_subset by blast lemma connected_component_mono: "s \ t \ connected_component_set s x \ connected_component_set t x" by (simp add: Collect_mono connected_component_of_subset) lemma connected_component_eq_empty [simp]: "connected_component_set s x = {} \ x \ s" using connected_component_refl by (fastforce simp: connected_component_in) lemma connected_component_set_empty [simp]: "connected_component_set {} x = {}" using connected_component_eq_empty by blast lemma connected_component_eq: "y \ connected_component_set s x \ (connected_component_set s y = connected_component_set s x)" by (metis (no_types, lifting) Collect_cong connected_component_sym connected_component_trans mem_Collect_eq) lemma closed_connected_component: assumes s: "closed s" shows "closed (connected_component_set s x)" proof (cases "x \ s") case False then show ?thesis by (metis connected_component_eq_empty closed_empty) next case True show ?thesis unfolding closure_eq [symmetric] proof show "closure (connected_component_set s x) \ connected_component_set s x" apply (rule connected_component_maximal) apply (simp add: closure_def True) apply (simp add: connected_imp_connected_closure) apply (simp add: s closure_minimal connected_component_subset) done next show "connected_component_set s x \ closure (connected_component_set s x)" by (simp add: closure_subset) qed qed lemma connected_component_disjoint: "connected_component_set s a \ connected_component_set s b = {} \ a \ connected_component_set s b" apply (auto simp: connected_component_eq) using connected_component_eq connected_component_sym apply blast done lemma connected_component_nonoverlap: "connected_component_set s a \ connected_component_set s b = {} \ a \ s \ b \ s \ connected_component_set s a \ connected_component_set s b" apply (auto simp: connected_component_in) using connected_component_refl_eq apply blast apply (metis connected_component_eq mem_Collect_eq) apply (metis connected_component_eq mem_Collect_eq) done lemma connected_component_overlap: "connected_component_set s a \ connected_component_set s b \ {} \ a \ s \ b \ s \ connected_component_set s a = connected_component_set s b" by (auto simp: connected_component_nonoverlap) lemma connected_component_sym_eq: "connected_component s x y \ connected_component s y x" using connected_component_sym by blast lemma connected_component_eq_eq: "connected_component_set s x = connected_component_set s y \ x \ s \ y \ s \ x \ s \ y \ s \ connected_component s x y" apply (cases "y \ s", simp) apply (metis connected_component_eq connected_component_eq_empty connected_component_refl_eq mem_Collect_eq) apply (cases "x \ s", simp) apply (metis connected_component_eq_empty) using connected_component_eq_empty apply blast done lemma connected_iff_connected_component_eq: "connected s \ (\x \ s. \y \ s. connected_component_set s x = connected_component_set s y)" by (simp add: connected_component_eq_eq connected_iff_connected_component) lemma connected_component_idemp: "connected_component_set (connected_component_set s x) x = connected_component_set s x" apply (rule subset_antisym) apply (simp add: connected_component_subset) apply (metis connected_component_eq_empty connected_component_maximal connected_component_refl_eq connected_connected_component mem_Collect_eq set_eq_subset) done lemma connected_component_unique: "\x \ c; c \ s; connected c; \c'. x \ c' \ c' \ s \ connected c' \ c' \ c\ \ connected_component_set s x = c" apply (rule subset_antisym) apply (meson connected_component_maximal connected_component_subset connected_connected_component contra_subsetD) by (simp add: connected_component_maximal) lemma joinable_connected_component_eq: "\connected t; t \ s; connected_component_set s x \ t \ {}; connected_component_set s y \ t \ {}\ \ connected_component_set s x = connected_component_set s y" apply (simp add: ex_in_conv [symmetric]) apply (rule connected_component_eq) by (metis (no_types, hide_lams) connected_component_eq_eq connected_component_in connected_component_maximal subsetD mem_Collect_eq) lemma Union_connected_component: "\(connected_component_set s ` s) = s" apply (rule subset_antisym) apply (simp add: SUP_least connected_component_subset) using connected_component_refl_eq by force lemma complement_connected_component_unions: "s - connected_component_set s x = \(connected_component_set s ` s - {connected_component_set s x})" apply (subst Union_connected_component [symmetric], auto) apply (metis connected_component_eq_eq connected_component_in) by (metis connected_component_eq mem_Collect_eq) lemma connected_component_intermediate_subset: "\connected_component_set u a \ t; t \ u\ \ connected_component_set t a = connected_component_set u a" apply (case_tac "a \ u") apply (simp add: connected_component_maximal connected_component_mono subset_antisym) using connected_component_eq_empty by blast subsection \The set of connected components of a set\ definition\<^marker>\tag important\ components:: "'a::topological_space set \ 'a set set" where "components s \ connected_component_set s ` s" lemma components_iff: "s \ components u \ (\x. x \ u \ s = connected_component_set u x)" by (auto simp: components_def) lemma componentsI: "x \ u \ connected_component_set u x \ components u" by (auto simp: components_def) lemma componentsE: assumes "s \ components u" obtains x where "x \ u" "s = connected_component_set u x" using assms by (auto simp: components_def) lemma Union_components [simp]: "\(components u) = u" apply (rule subset_antisym) using Union_connected_component components_def apply fastforce apply (metis Union_connected_component components_def set_eq_subset) done lemma pairwise_disjoint_components: "pairwise (\X Y. X \ Y = {}) (components u)" apply (simp add: pairwise_def) apply (auto simp: components_iff) apply (metis connected_component_eq_eq connected_component_in)+ done lemma in_components_nonempty: "c \ components s \ c \ {}" by (metis components_iff connected_component_eq_empty) lemma in_components_subset: "c \ components s \ c \ s" using Union_components by blast lemma in_components_connected: "c \ components s \ connected c" by (metis components_iff connected_connected_component) lemma in_components_maximal: "c \ components s \ c \ {} \ c \ s \ connected c \ (\d. d \ {} \ c \ d \ d \ s \ connected d \ d = c)" apply (rule iffI) apply (simp add: in_components_nonempty in_components_connected) apply (metis (full_types) components_iff connected_component_eq_self connected_component_intermediate_subset connected_component_refl in_components_subset mem_Collect_eq rev_subsetD) apply (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI) done lemma joinable_components_eq: "connected t \ t \ s \ c1 \ components s \ c2 \ components s \ c1 \ t \ {} \ c2 \ t \ {} \ c1 = c2" by (metis (full_types) components_iff joinable_connected_component_eq) lemma closed_components: "\closed s; c \ components s\ \ closed c" by (metis closed_connected_component components_iff) lemma components_nonoverlap: "\c \ components s; c' \ components s\ \ (c \ c' = {}) \ (c \ c')" apply (auto simp: in_components_nonempty components_iff) using connected_component_refl apply blast apply (metis connected_component_eq_eq connected_component_in) by (metis connected_component_eq mem_Collect_eq) lemma components_eq: "\c \ components s; c' \ components s\ \ (c = c' \ c \ c' \ {})" by (metis components_nonoverlap) lemma components_eq_empty [simp]: "components s = {} \ s = {}" by (simp add: components_def) lemma components_empty [simp]: "components {} = {}" by simp lemma connected_eq_connected_components_eq: "connected s \ (\c \ components s. \c' \ components s. c = c')" by (metis (no_types, hide_lams) components_iff connected_component_eq_eq connected_iff_connected_component) lemma components_eq_sing_iff: "components s = {s} \ connected s \ s \ {}" apply (rule iffI) using in_components_connected apply fastforce apply safe using Union_components apply fastforce apply (metis components_iff connected_component_eq_self) using in_components_maximal apply auto done lemma components_eq_sing_exists: "(\a. components s = {a}) \ connected s \ s \ {}" apply (rule iffI) using connected_eq_connected_components_eq apply fastforce apply (metis components_eq_sing_iff) done lemma connected_eq_components_subset_sing: "connected s \ components s \ {s}" by (metis Union_components components_empty components_eq_sing_iff connected_empty insert_subset order_refl subset_singletonD) lemma connected_eq_components_subset_sing_exists: "connected s \ (\a. components s \ {a})" by (metis components_eq_sing_exists connected_eq_components_subset_sing empty_iff subset_iff subset_singletonD) lemma in_components_self: "s \ components s \ connected s \ s \ {}" by (metis components_empty components_eq_sing_iff empty_iff in_components_connected insertI1) lemma components_maximal: "\c \ components s; connected t; t \ s; c \ t \ {}\ \ t \ c" apply (simp add: components_def ex_in_conv [symmetric], clarify) by (meson connected_component_def connected_component_trans) lemma exists_component_superset: "\t \ s; s \ {}; connected t\ \ \c. c \ components s \ t \ c" apply (cases "t = {}", force) apply (metis components_def ex_in_conv connected_component_maximal contra_subsetD image_eqI) done lemma components_intermediate_subset: "\s \ components u; s \ t; t \ u\ \ s \ components t" apply (auto simp: components_iff) apply (metis connected_component_eq_empty connected_component_intermediate_subset) done lemma in_components_unions_complement: "c \ components s \ s - c = \(components s - {c})" by (metis complement_connected_component_unions components_def components_iff) lemma connected_intermediate_closure: assumes cs: "connected s" and st: "s \ t" and ts: "t \ closure s" shows "connected t" proof (rule connectedI) fix A B assume A: "open A" and B: "open B" and Alap: "A \ t \ {}" and Blap: "B \ t \ {}" and disj: "A \ B \ t = {}" and cover: "t \ A \ B" have disjs: "A \ B \ s = {}" using disj st by auto have "A \ closure s \ {}" using Alap Int_absorb1 ts by blast then have Alaps: "A \ s \ {}" by (simp add: A open_Int_closure_eq_empty) have "B \ closure s \ {}" using Blap Int_absorb1 ts by blast then have Blaps: "B \ s \ {}" by (simp add: B open_Int_closure_eq_empty) then show False using cs [unfolded connected_def] A B disjs Alaps Blaps cover st by blast qed lemma closedin_connected_component: "closedin (top_of_set s) (connected_component_set s x)" proof (cases "connected_component_set s x = {}") case True then show ?thesis by (metis closedin_empty) next case False then obtain y where y: "connected_component s x y" by blast have *: "connected_component_set s x \ s \ closure (connected_component_set s x)" by (auto simp: closure_def connected_component_in) have "connected_component s x y \ s \ closure (connected_component_set s x) \ connected_component_set s x" apply (rule connected_component_maximal, simp) using closure_subset connected_component_in apply fastforce using * connected_intermediate_closure apply blast+ done with y * show ?thesis by (auto simp: closedin_closed) qed lemma closedin_component: "C \ components s \ closedin (top_of_set s) C" using closedin_connected_component componentsE by blast subsection\<^marker>\tag unimportant\ \Proving a function is constant on a connected set by proving that a level set is open\ lemma continuous_levelset_openin_cases: fixes f :: "_ \ 'b::t1_space" shows "connected s \ continuous_on s f \ openin (top_of_set s) {x \ s. f x = a} \ (\x \ s. f x \ a) \ (\x \ s. f x = a)" unfolding connected_clopen using continuous_closedin_preimage_constant by auto lemma continuous_levelset_openin: fixes f :: "_ \ 'b::t1_space" shows "connected s \ continuous_on s f \ openin (top_of_set s) {x \ s. f x = a} \ (\x \ s. f x = a) \ (\x \ s. f x = a)" using continuous_levelset_openin_cases[of s f ] by meson lemma continuous_levelset_open: fixes f :: "_ \ 'b::t1_space" assumes "connected s" and "continuous_on s f" and "open {x \ s. f x = a}" and "\x \ s. f x = a" shows "\x \ s. f x = a" using continuous_levelset_openin[OF assms(1,2), of a, unfolded openin_open] using assms (3,4) by fast subsection\<^marker>\tag unimportant\ \Preservation of Connectedness\ lemma homeomorphic_connectedness: assumes "s homeomorphic t" shows "connected s \ connected t" using assms unfolding homeomorphic_def homeomorphism_def by (metis connected_continuous_image) lemma connected_monotone_quotient_preimage: assumes "connected T" and contf: "continuous_on S f" and fim: "f ` S = T" and opT: "\U. U \ T \ openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" and connT: "\y. y \ T \ connected (S \ f -` {y})" shows "connected S" proof (rule connectedI) fix U V assume "open U" and "open V" and "U \ S \ {}" and "V \ S \ {}" and "U \ V \ S = {}" and "S \ U \ V" moreover have disjoint: "f ` (S \ U) \ f ` (S \ V) = {}" proof - have False if "y \ f ` (S \ U) \ f ` (S \ V)" for y proof - have "y \ T" using fim that by blast show ?thesis using connectedD [OF connT [OF \y \ T\] \open U\ \open V\] \S \ U \ V\ \U \ V \ S = {}\ that by fastforce qed then show ?thesis by blast qed ultimately have UU: "(S \ f -` f ` (S \ U)) = S \ U" and VV: "(S \ f -` f ` (S \ V)) = S \ V" by auto have opeU: "openin (top_of_set T) (f ` (S \ U))" by (metis UU \open U\ fim image_Int_subset le_inf_iff opT openin_open_Int) have opeV: "openin (top_of_set T) (f ` (S \ V))" by (metis opT fim VV \open V\ openin_open_Int image_Int_subset inf.bounded_iff) have "T \ f ` (S \ U) \ f ` (S \ V)" using \S \ U \ V\ fim by auto then show False using \connected T\ disjoint opeU opeV \U \ S \ {}\ \V \ S \ {}\ by (auto simp: connected_openin) qed lemma connected_open_monotone_preimage: assumes contf: "continuous_on S f" and fim: "f ` S = T" and ST: "\C. openin (top_of_set S) C \ openin (top_of_set T) (f ` C)" and connT: "\y. y \ T \ connected (S \ f -` {y})" and "connected C" "C \ T" shows "connected (S \ f -` C)" proof - have contf': "continuous_on (S \ f -` C) f" by (meson contf continuous_on_subset inf_le1) have eqC: "f ` (S \ f -` C) = C" using \C \ T\ fim by blast show ?thesis proof (rule connected_monotone_quotient_preimage [OF \connected C\ contf' eqC]) show "connected (S \ f -` C \ f -` {y})" if "y \ C" for y proof - have "S \ f -` C \ f -` {y} = S \ f -` {y}" using that by blast moreover have "connected (S \ f -` {y})" using \C \ T\ connT that by blast ultimately show ?thesis by metis qed have "\U. openin (top_of_set (S \ f -` C)) U \ openin (top_of_set C) (f ` U)" using open_map_restrict [OF _ ST \C \ T\] by metis then show "\D. D \ C \ openin (top_of_set (S \ f -` C)) (S \ f -` C \ f -` D) = openin (top_of_set C) D" using open_map_imp_quotient_map [of "(S \ f -` C)" f] contf' by (simp add: eqC) qed qed lemma connected_closed_monotone_preimage: assumes contf: "continuous_on S f" and fim: "f ` S = T" and ST: "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" and connT: "\y. y \ T \ connected (S \ f -` {y})" and "connected C" "C \ T" shows "connected (S \ f -` C)" proof - have contf': "continuous_on (S \ f -` C) f" by (meson contf continuous_on_subset inf_le1) have eqC: "f ` (S \ f -` C) = C" using \C \ T\ fim by blast show ?thesis proof (rule connected_monotone_quotient_preimage [OF \connected C\ contf' eqC]) show "connected (S \ f -` C \ f -` {y})" if "y \ C" for y proof - have "S \ f -` C \ f -` {y} = S \ f -` {y}" using that by blast moreover have "connected (S \ f -` {y})" using \C \ T\ connT that by blast ultimately show ?thesis by metis qed have "\U. closedin (top_of_set (S \ f -` C)) U \ closedin (top_of_set C) (f ` U)" using closed_map_restrict [OF _ ST \C \ T\] by metis then show "\D. D \ C \ openin (top_of_set (S \ f -` C)) (S \ f -` C \ f -` D) = openin (top_of_set C) D" using closed_map_imp_quotient_map [of "(S \ f -` C)" f] contf' by (simp add: eqC) qed qed subsection\Lemmas about components\ text \See Newman IV, 3.3 and 3.4.\ lemma connected_Un_clopen_in_complement: fixes S U :: "'a::metric_space set" assumes "connected S" "connected U" "S \ U" and opeT: "openin (top_of_set (U - S)) T" and cloT: "closedin (top_of_set (U - S)) T" shows "connected (S \ T)" proof - have *: "\\x y. P x y \ P y x; \x y. P x y \ S \ x \ S \ y; \x y. \P x y; S \ x\ \ False\ \ \(\x y. (P x y))" for P by metis show ?thesis unfolding connected_closedin_eq proof (rule *) fix H1 H2 assume H: "closedin (top_of_set (S \ T)) H1 \ closedin (top_of_set (S \ T)) H2 \ H1 \ H2 = S \ T \ H1 \ H2 = {} \ H1 \ {} \ H2 \ {}" then have clo: "closedin (top_of_set S) (S \ H1)" "closedin (top_of_set S) (S \ H2)" by (metis Un_upper1 closedin_closed_subset inf_commute)+ have Seq: "S \ (H1 \ H2) = S" by (simp add: H) have "S \ ((S \ T) \ H1) \ S \ ((S \ T) \ H2) = S" using Seq by auto moreover have "H1 \ (S \ ((S \ T) \ H2)) = {}" using H by blast ultimately have "S \ H1 = {} \ S \ H2 = {}" by (metis (no_types) H Int_assoc \S \ (H1 \ H2) = S\ \connected S\ clo Seq connected_closedin inf_bot_right inf_le1) then show "S \ H1 \ S \ H2" using H \connected S\ unfolding connected_closedin by blast next fix H1 H2 assume H: "closedin (top_of_set (S \ T)) H1 \ closedin (top_of_set (S \ T)) H2 \ H1 \ H2 = S \ T \ H1 \ H2 = {} \ H1 \ {} \ H2 \ {}" and "S \ H1" then have H2T: "H2 \ T" by auto have "T \ U" using Diff_iff opeT openin_imp_subset by auto with \S \ U\ have Ueq: "U = (U - S) \ (S \ T)" by auto have "openin (top_of_set ((U - S) \ (S \ T))) H2" proof (rule openin_subtopology_Un) show "openin (top_of_set (S \ T)) H2" using \H2 \ T\ apply (auto simp: openin_closedin_eq) by (metis Diff_Diff_Int Diff_disjoint Diff_partition Diff_subset H Int_absorb1 Un_Diff) then show "openin (top_of_set (U - S)) H2" by (meson H2T Un_upper2 opeT openin_subset_trans openin_trans) qed moreover have "closedin (top_of_set ((U - S) \ (S \ T))) H2" proof (rule closedin_subtopology_Un) show "closedin (top_of_set (U - S)) H2" using H H2T cloT closedin_subset_trans by (blast intro: closedin_subtopology_Un closedin_trans) qed (simp add: H) ultimately have H2: "H2 = {} \ H2 = U" using Ueq \connected U\ unfolding connected_clopen by metis then have "H2 \ S" by (metis Diff_partition H Un_Diff_cancel Un_subset_iff \H2 \ T\ assms(3) inf.orderE opeT openin_imp_subset) moreover have "T \ H2 - S" by (metis (no_types) H2 H opeT openin_closedin_eq topspace_euclidean_subtopology) ultimately show False using H \S \ H1\ by blast qed blast qed proposition component_diff_connected: fixes S :: "'a::metric_space set" assumes "connected S" "connected U" "S \ U" and C: "C \ components (U - S)" shows "connected(U - C)" using \connected S\ unfolding connected_closedin_eq not_ex de_Morgan_conj proof clarify fix H3 H4 assume clo3: "closedin (top_of_set (U - C)) H3" and clo4: "closedin (top_of_set (U - C)) H4" and "H3 \ H4 = U - C" and "H3 \ H4 = {}" and "H3 \ {}" and "H4 \ {}" and * [rule_format]: "\H1 H2. \ closedin (top_of_set S) H1 \ \ closedin (top_of_set S) H2 \ H1 \ H2 \ S \ H1 \ H2 \ {} \ \ H1 \ {} \ \ H2 \ {}" then have "H3 \ U-C" and ope3: "openin (top_of_set (U - C)) (U - C - H3)" and "H4 \ U-C" and ope4: "openin (top_of_set (U - C)) (U - C - H4)" by (auto simp: closedin_def) have "C \ {}" "C \ U-S" "connected C" using C in_components_nonempty in_components_subset in_components_maximal by blast+ have cCH3: "connected (C \ H3)" proof (rule connected_Un_clopen_in_complement [OF \connected C\ \connected U\ _ _ clo3]) show "openin (top_of_set (U - C)) H3" apply (simp add: openin_closedin_eq \H3 \ U - C\) apply (simp add: closedin_subtopology) by (metis Diff_cancel Diff_triv Un_Diff clo4 \H3 \ H4 = {}\ \H3 \ H4 = U - C\ closedin_closed inf_commute sup_bot.left_neutral) qed (use clo3 \C \ U - S\ in auto) have cCH4: "connected (C \ H4)" proof (rule connected_Un_clopen_in_complement [OF \connected C\ \connected U\ _ _ clo4]) show "openin (top_of_set (U - C)) H4" apply (simp add: openin_closedin_eq \H4 \ U - C\) apply (simp add: closedin_subtopology) by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \H3 \ H4 = {}\ \H3 \ H4 = U - C\ clo3 closedin_closed) qed (use clo4 \C \ U - S\ in auto) have "closedin (top_of_set S) (S \ H3)" "closedin (top_of_set S) (S \ H4)" using clo3 clo4 \S \ U\ \C \ U - S\ by (auto simp: closedin_closed) moreover have "S \ H3 \ {}" using components_maximal [OF C cCH3] \C \ {}\ \C \ U - S\ \H3 \ {}\ \H3 \ U - C\ by auto moreover have "S \ H4 \ {}" using components_maximal [OF C cCH4] \C \ {}\ \C \ U - S\ \H4 \ {}\ \H4 \ U - C\ by auto ultimately show False using * [of "S \ H3" "S \ H4"] \H3 \ H4 = {}\ \C \ U - S\ \H3 \ H4 = U - C\ \S \ U\ by auto qed subsection\<^marker>\tag unimportant\\Constancy of a function from a connected set into a finite, disconnected or discrete set\ text\Still missing: versions for a set that is smaller than R, or countable.\ lemma continuous_disconnected_range_constant: assumes S: "connected S" and conf: "continuous_on S f" and fim: "f ` S \ t" and cct: "\y. y \ t \ connected_component_set t y = {y}" shows "f constant_on S" proof (cases "S = {}") case True then show ?thesis by (simp add: constant_on_def) next case False { fix x assume "x \ S" then have "f ` S \ {f x}" by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI S cct) } with False show ?thesis unfolding constant_on_def by blast qed text\This proof requires the existence of two separate values of the range type.\ lemma finite_range_constant_imp_connected: assumes "\f::'a::topological_space \ 'b::real_normed_algebra_1. \continuous_on S f; finite(f ` S)\ \ f constant_on S" shows "connected S" proof - { fix t u assume clt: "closedin (top_of_set S) t" and clu: "closedin (top_of_set S) u" and tue: "t \ u = {}" and tus: "t \ u = S" have conif: "continuous_on S (\x. if x \ t then 0 else 1)" apply (subst tus [symmetric]) apply (rule continuous_on_cases_local) using clt clu tue - apply (auto simp: tus continuous_on_const) + apply (auto simp: tus) done have fi: "finite ((\x. if x \ t then 0 else 1) ` S)" by (rule finite_subset [of _ "{0,1}"]) auto have "t = {} \ u = {}" using assms [OF conif fi] tus [symmetric] by (auto simp: Ball_def constant_on_def) (metis IntI empty_iff one_neq_zero tue) } then show ?thesis by (simp add: connected_closedin_eq) qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Continuous_Extension.thy b/src/HOL/Analysis/Continuous_Extension.thy --- a/src/HOL/Analysis/Continuous_Extension.thy +++ b/src/HOL/Analysis/Continuous_Extension.thy @@ -1,537 +1,537 @@ (* Title: HOL/Analysis/Continuous_Extension.thy Authors: LC Paulson, based on material from HOL Light *) section \Continuous Extensions of Functions\ theory Continuous_Extension imports Starlike begin subsection\Partitions of unity subordinate to locally finite open coverings\ text\A difference from HOL Light: all summations over infinite sets equal zero, so the "support" must be made explicit in the summation below!\ proposition subordinate_partition_of_unity: fixes S :: "'a::metric_space set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" and fin: "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" obtains F :: "['a set, 'a] \ real" where "\U. U \ \ \ continuous_on S (F U) \ (\x \ S. 0 \ F U x)" and "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" and "\x. x \ S \ supp_sum (\W. F W x) \ = 1" and "\x. x \ S \ \V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" proof (cases "\W. W \ \ \ S \ W") case True then obtain W where "W \ \" "S \ W" by metis then show ?thesis by (rule_tac F = "\V x. if V = W then 1 else 0" in that) (auto simp: supp_sum_def support_on_def) next case False have nonneg: "0 \ supp_sum (\V. setdist {x} (S - V)) \" for x by (simp add: supp_sum_def sum_nonneg) have sd_pos: "0 < setdist {x} (S - V)" if "V \ \" "x \ S" "x \ V" for V x proof - have "closedin (top_of_set S) (S - V)" by (simp add: Diff_Diff_Int closedin_def opC openin_open_Int \V \ \\) with that False setdist_pos_le [of "{x}" "S - V"] show ?thesis using setdist_gt_0_closedin by fastforce qed have ss_pos: "0 < supp_sum (\V. setdist {x} (S - V)) \" if "x \ S" for x proof - obtain U where "U \ \" "x \ U" using \x \ S\ \S \ \\\ by blast obtain V where "open V" "x \ V" "finite {U \ \. U \ V \ {}}" using \x \ S\ fin by blast then have *: "finite {A \ \. \ S \ A \ x \ closure (S - A)}" using closure_def that by (blast intro: rev_finite_subset) have "x \ closure (S - U)" using \U \ \\ \x \ U\ opC open_Int_closure_eq_empty by fastforce then show ?thesis apply (simp add: setdist_eq_0_sing_1 supp_sum_def support_on_def) apply (rule ordered_comm_monoid_add_class.sum_pos2 [OF *, of U]) using \U \ \\ \x \ U\ False apply (auto simp: sd_pos that) done qed define F where "F \ \W x. if x \ S then setdist {x} (S - W) / supp_sum (\V. setdist {x} (S - V)) \ else 0" show ?thesis proof (rule_tac F = F in that) have "continuous_on S (F U)" if "U \ \" for U proof - have *: "continuous_on S (\x. supp_sum (\V. setdist {x} (S - V)) \)" proof (clarsimp simp add: continuous_on_eq_continuous_within) fix x assume "x \ S" then obtain X where "open X" and x: "x \ S \ X" and finX: "finite {U \ \. U \ X \ {}}" using assms by blast then have OSX: "openin (top_of_set S) (S \ X)" by blast have sumeq: "\x. x \ S \ X \ (\V | V \ \ \ V \ X \ {}. setdist {x} (S - V)) = supp_sum (\V. setdist {x} (S - V)) \" apply (simp add: supp_sum_def) apply (rule sum.mono_neutral_right [OF finX]) apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff) apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE) done show "continuous (at x within S) (\x. supp_sum (\V. setdist {x} (S - V)) \)" apply (rule continuous_transform_within_openin [where f = "\x. (sum (\V. setdist {x} (S - V)) {V \ \. V \ X \ {}})" and S ="S \ X"]) apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+ apply (simp add: sumeq) done qed show ?thesis apply (simp add: F_def) apply (rule continuous_intros *)+ using ss_pos apply force done qed moreover have "\U \ \; x \ S\ \ 0 \ F U x" for U x - using nonneg [of x] by (simp add: F_def field_split_simps setdist_pos_le) + using nonneg [of x] by (simp add: F_def field_split_simps) ultimately show "\U. U \ \ \ continuous_on S (F U) \ (\x\S. 0 \ F U x)" by metis next show "\x U. \U \ \; x \ S; x \ U\ \ F U x = 0" by (simp add: setdist_eq_0_sing_1 closure_def F_def) next show "supp_sum (\W. F W x) \ = 1" if "x \ S" for x using that ss_pos [OF that] by (simp add: F_def field_split_simps supp_sum_divide_distrib [symmetric]) next show "\V. open V \ x \ V \ finite {U \ \. \x\V. F U x \ 0}" if "x \ S" for x using fin [OF that] that by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset) qed qed subsection\Urysohn's Lemma for Euclidean Spaces\ text \For Euclidean spaces the proof is easy using distances.\ lemma Urysohn_both_ne: assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" "S \ {}" "T \ {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::real_normed_vector" where "continuous_on U f" "\x. x \ U \ f x \ closed_segment a b" "\x. x \ U \ (f x = a \ x \ S)" "\x. x \ U \ (f x = b \ x \ T)" proof - have S0: "\x. x \ U \ setdist {x} S = 0 \ x \ S" using \S \ {}\ US setdist_eq_0_closedin by auto have T0: "\x. x \ U \ setdist {x} T = 0 \ x \ T" using \T \ {}\ UT setdist_eq_0_closedin by auto have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \ U" for x proof - have "\ (setdist {x} S = 0 \ setdist {x} T = 0)" using assms by (metis IntI empty_iff setdist_eq_0_closedin that) then show ?thesis by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le) qed define f where "f \ \x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)" show ?thesis proof (rule_tac f = f in that) show "continuous_on U f" using sdpos unfolding f_def by (intro continuous_intros | force)+ show "f x \ closed_segment a b" if "x \ U" for x unfolding f_def apply (simp add: closed_segment_def) apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI) using sdpos that apply (simp add: algebra_simps) done show "\x. x \ U \ (f x = a \ x \ S)" using S0 \a \ b\ f_def sdpos by force show "(f x = b \ x \ T)" if "x \ U" for x proof - have "f x = b \ (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1" unfolding f_def apply (rule iffI) apply (metis \a \ b\ add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force) done also have "... \ setdist {x} T = 0 \ setdist {x} S \ 0" using sdpos that by (simp add: field_split_simps) linarith also have "... \ x \ T" using \S \ {}\ \T \ {}\ \S \ T = {}\ that by (force simp: S0 T0) finally show ?thesis . qed qed qed proposition Urysohn_local_strong: assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on U f" "\x. x \ U \ f x \ closed_segment a b" "\x. x \ U \ (f x = a \ x \ S)" "\x. x \ U \ (f x = b \ x \ T)" proof (cases "S = {}") case True show ?thesis proof (cases "T = {}") case True show ?thesis proof (rule_tac f = "\x. midpoint a b" in that) show "continuous_on U (\x. midpoint a b)" by (intro continuous_intros) show "midpoint a b \ closed_segment a b" using csegment_midpoint_subset by blast show "(midpoint a b = a) = (x \ S)" for x using \S = {}\ \a \ b\ by simp show "(midpoint a b = b) = (x \ T)" for x using \T = {}\ \a \ b\ by simp qed next case False show ?thesis proof (cases "T = U") case True with \S = {}\ \a \ b\ show ?thesis - by (rule_tac f = "\x. b" in that) (auto simp: continuous_on_const) + by (rule_tac f = "\x. b" in that) (auto) next case False with UT closedin_subset obtain c where c: "c \ U" "c \ T" by fastforce obtain f where f: "continuous_on U f" "\x. x \ U \ f x \ closed_segment (midpoint a b) b" "\x. x \ U \ (f x = midpoint a b \ x = c)" "\x. x \ U \ (f x = b \ x \ T)" apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"]) using c \T \ {}\ assms apply simp_all done show ?thesis apply (rule_tac f=f in that) using \S = {}\ \T \ {}\ f csegment_midpoint_subset notin_segment_midpoint [OF \a \ b\] apply force+ done qed qed next case False show ?thesis proof (cases "T = {}") case True show ?thesis proof (cases "S = U") case True with \T = {}\ \a \ b\ show ?thesis - by (rule_tac f = "\x. a" in that) (auto simp: continuous_on_const) + by (rule_tac f = "\x. a" in that) (auto) next case False with US closedin_subset obtain c where c: "c \ U" "c \ S" by fastforce obtain f where f: "continuous_on U f" "\x. x \ U \ f x \ closed_segment a (midpoint a b)" "\x. x \ U \ (f x = midpoint a b \ x = c)" "\x. x \ U \ (f x = a \ x \ S)" apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"]) using c \S \ {}\ assms apply simp_all apply (metis midpoint_eq_endpoint) done show ?thesis apply (rule_tac f=f in that) using \S \ {}\ \T = {}\ f \a \ b\ apply simp_all apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff) apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint) done qed next case False show ?thesis using Urysohn_both_ne [OF US UT \S \ T = {}\ \S \ {}\ \T \ {}\ \a \ b\] that by blast qed qed lemma Urysohn_local: assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on U f" "\x. x \ U \ f x \ closed_segment a b" "\x. x \ S \ f x = a" "\x. x \ T \ f x = b" proof (cases "a = b") case True then show ?thesis - by (rule_tac f = "\x. b" in that) (auto simp: continuous_on_const) + by (rule_tac f = "\x. b" in that) (auto) next case False then show ?thesis apply (rule Urysohn_local_strong [OF assms]) apply (erule that, assumption) apply (meson US closedin_singleton closedin_trans) apply (meson UT closedin_singleton closedin_trans) done qed lemma Urysohn_strong: assumes US: "closed S" and UT: "closed T" and "S \ T = {}" "a \ b" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on UNIV f" "\x. f x \ closed_segment a b" "\x. f x = a \ x \ S" "\x. f x = b \ x \ T" using assms by (auto intro: Urysohn_local_strong [of UNIV S T]) proposition Urysohn: assumes US: "closed S" and UT: "closed T" and "S \ T = {}" obtains f :: "'a::euclidean_space \ 'b::euclidean_space" where "continuous_on UNIV f" "\x. f x \ closed_segment a b" "\x. x \ S \ f x = a" "\x. x \ T \ f x = b" using assms by (auto intro: Urysohn_local [of UNIV S T a b]) subsection\Dugundji's Extension Theorem and Tietze Variants\ text \See \cite{dugundji}.\ theorem Dugundji: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "convex C" "C \ {}" and cloin: "closedin (top_of_set U) S" and contf: "continuous_on S f" and "f ` S \ C" obtains g where "continuous_on U g" "g ` U \ C" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True then show thesis apply (rule_tac g="\x. SOME y. y \ C" in that) apply (rule continuous_intros) apply (meson all_not_in_conv \C \ {}\ image_subsetI someI_ex, simp) done next case False then have sd_pos: "\x. \x \ U; x \ S\ \ 0 < setdist {x} S" using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce define \ where "\ = {ball x (setdist {x} S / 2) |x. x \ U - S}" have [simp]: "\T. T \ \ \ open T" by (auto simp: \_def) have USS: "U - S \ \\" by (auto simp: sd_pos \_def) obtain \ where USsub: "U - S \ \\" and nbrhd: "\U. U \ \ \ open U \ (\T. T \ \ \ U \ T)" and fin: "\x. x \ U - S \ \V. open V \ x \ V \ finite {U. U \ \ \ U \ V \ {}}" by (rule paracompact [OF USS]) auto have "\v a. v \ U \ v \ S \ a \ S \ T \ ball v (setdist {v} S / 2) \ dist v a \ 2 * setdist {v} S" if "T \ \" for T proof - obtain v where v: "T \ ball v (setdist {v} S / 2)" "v \ U" "v \ S" using \T \ \\ nbrhd by (force simp: \_def) then obtain a where "a \ S" "dist v a < 2 * setdist {v} S" using setdist_ltE [of "{v}" S "2 * setdist {v} S"] using False sd_pos by force with v show ?thesis apply (rule_tac x=v in exI) apply (rule_tac x=a in exI, auto) done qed then obtain \ \ where VA: "\T. T \ \ \ \ T \ U \ \ T \ S \ \ T \ S \ T \ ball (\ T) (setdist {\ T} S / 2) \ dist (\ T) (\ T) \ 2 * setdist {\ T} S" by metis have sdle: "setdist {\ T} S \ 2 * setdist {v} S" if "T \ \" "v \ T" for T v using setdist_Lipschitz [of "\ T" S v] VA [OF \T \ \\] \v \ T\ by auto have d6: "dist a (\ T) \ 6 * dist a v" if "T \ \" "v \ T" "a \ S" for T v a proof - have "dist (\ T) v < setdist {\ T} S / 2" using that VA mem_ball by blast also have "\ \ setdist {v} S" using sdle [OF \T \ \\ \v \ T\] by simp also have vS: "setdist {v} S \ dist a v" by (simp add: setdist_le_dist setdist_sym \a \ S\) finally have VTV: "dist (\ T) v < dist a v" . have VTS: "setdist {\ T} S \ 2 * dist a v" using sdle that vS by force have "dist a (\ T) \ dist a v + dist v (\ T) + dist (\ T) (\ T)" by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le) also have "\ \ dist a v + dist a v + dist (\ T) (\ T)" using VTV by (simp add: dist_commute) also have "\ \ 2 * dist a v + 2 * setdist {\ T} S" using VA [OF \T \ \\] by auto finally show ?thesis using VTS by linarith qed obtain H :: "['a set, 'a] \ real" where Hcont: "\Z. Z \ \ \ continuous_on (U-S) (H Z)" and Hge0: "\Z x. \Z \ \; x \ U-S\ \ 0 \ H Z x" and Heq0: "\x Z. \Z \ \; x \ U-S; x \ Z\ \ H Z x = 0" and H1: "\x. x \ U-S \ supp_sum (\W. H W x) \ = 1" and Hfin: "\x. x \ U-S \ \V. open V \ x \ V \ finite {U \ \. \x\V. H U x \ 0}" apply (rule subordinate_partition_of_unity [OF USsub _ fin]) using nbrhd by auto define g where "g \ \x. if x \ S then f x else supp_sum (\T. H T x *\<^sub>R f(\ T)) \" show ?thesis proof (rule that) show "continuous_on U g" proof (clarsimp simp: continuous_on_eq_continuous_within) fix a assume "a \ U" show "continuous (at a within U) g" proof (cases "a \ S") case True show ?thesis proof (clarsimp simp add: continuous_within_topological) fix W assume "open W" "g a \ W" then obtain e where "0 < e" and e: "ball (f a) e \ W" using openE True g_def by auto have "continuous (at a within S) f" using True contf continuous_on_eq_continuous_within by blast then obtain d where "0 < d" and d: "\x. \x \ S; dist x a < d\ \ dist (f x) (f a) < e" using continuous_within_eps_delta \0 < e\ by force have "g y \ ball (f a) e" if "y \ U" and y: "y \ ball a (d / 6)" for y proof (cases "y \ S") case True then have "dist (f a) (f y) < e" by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d) then show ?thesis by (simp add: True g_def) next case False have *: "dist (f (\ T)) (f a) < e" if "T \ \" "H T y \ 0" for T proof - have "y \ T" using Heq0 that False \y \ U\ by blast have "dist (\ T) a < d" using d6 [OF \T \ \\ \y \ T\ \a \ S\] y by (simp add: dist_commute mult.commute) then show ?thesis using VA [OF \T \ \\] by (auto simp: d) qed have "supp_sum (\T. H T y *\<^sub>R f (\ T)) \ \ ball (f a) e" apply (rule convex_supp_sum [OF convex_ball]) apply (simp_all add: False H1 Hge0 \y \ U\) by (metis dist_commute *) then show ?thesis by (simp add: False g_def) qed then show "\A. open A \ a \ A \ (\y\U. y \ A \ g y \ W)" apply (rule_tac x = "ball a (d / 6)" in exI) using e \0 < d\ by fastforce qed next case False obtain N where N: "open N" "a \ N" and finN: "finite {U \ \. \a\N. H U a \ 0}" using Hfin False \a \ U\ by auto have oUS: "openin (top_of_set U) (U - S)" using cloin by (simp add: openin_diff) have HcontU: "continuous (at a within U) (H T)" if "T \ \" for T using Hcont [OF \T \ \\] False \a \ U\ \T \ \\ apply (simp add: continuous_on_eq_continuous_within continuous_within) apply (rule Lim_transform_within_set) using oUS apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+ done show ?thesis proof (rule continuous_transform_within_openin [OF _ oUS]) show "continuous (at a within U) (\x. supp_sum (\T. H T x *\<^sub>R f (\ T)) \)" proof (rule continuous_transform_within_openin) show "continuous (at a within U) (\x. \T\{U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T))" by (force intro: continuous_intros HcontU)+ next show "openin (top_of_set U) ((U - S) \ N)" using N oUS openin_trans by blast next show "a \ (U - S) \ N" using False \a \ U\ N by blast next show "\x. x \ (U - S) \ N \ (\T \ {U \ \. \x\N. H U x \ 0}. H T x *\<^sub>R f (\ T)) = supp_sum (\T. H T x *\<^sub>R f (\ T)) \" by (auto simp: supp_sum_def support_on_def intro: sum.mono_neutral_right [OF finN]) qed next show "a \ U - S" using False \a \ U\ by blast next show "\x. x \ U - S \ supp_sum (\T. H T x *\<^sub>R f (\ T)) \ = g x" by (simp add: g_def) qed qed qed show "g ` U \ C" using \f ` S \ C\ VA by (fastforce simp: g_def Hge0 intro!: convex_supp_sum [OF \convex C\] H1) show "\x. x \ S \ g x = f x" by (simp add: g_def) qed qed corollary Tietze: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "0 \ B" and "\x. x \ S \ norm(f x) \ B" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ norm(g x) \ B" using assms by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f]) corollary\<^marker>\tag unimportant\ Tietze_closed_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::euclidean_space" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "cbox a b \ {}" and "\x. x \ S \ f x \ cbox a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ cbox a b" apply (rule Dugundji [of "cbox a b" U S f]) using assms by auto corollary\<^marker>\tag unimportant\ Tietze_closed_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ real" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "a \ b" and "\x. x \ S \ f x \ cbox a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ cbox a b" apply (rule Dugundji [of "cbox a b" U S f]) using assms by (auto simp: image_subset_iff) corollary\<^marker>\tag unimportant\ Tietze_open_interval: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::euclidean_space" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "box a b \ {}" and "\x. x \ S \ f x \ box a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ box a b" apply (rule Dugundji [of "box a b" U S f]) using assms by auto corollary\<^marker>\tag unimportant\ Tietze_open_interval_1: fixes f :: "'a::{metric_space,second_countable_topology} \ real" assumes "continuous_on S f" and "closedin (top_of_set U) S" and "a < b" and no: "\x. x \ S \ f x \ box a b" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" "\x. x \ U \ g x \ box a b" apply (rule Dugundji [of "box a b" U S f]) using assms by (auto simp: image_subset_iff) corollary\<^marker>\tag unimportant\ Tietze_unbounded: fixes f :: "'a::{metric_space,second_countable_topology} \ 'b::real_inner" assumes "continuous_on S f" and "closedin (top_of_set U) S" obtains g where "continuous_on U g" "\x. x \ S \ g x = f x" apply (rule Dugundji [of UNIV U S f]) using assms by auto end diff --git a/src/HOL/Analysis/Convex_Euclidean_Space.thy b/src/HOL/Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Analysis/Convex_Euclidean_Space.thy +++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy @@ -1,2275 +1,2275 @@ (* Title: HOL/Analysis/Convex_Euclidean_Space.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) section \Convex Sets and Functions on (Normed) Euclidean Spaces\ theory Convex_Euclidean_Space imports Convex Topology_Euclidean_Space begin subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets and Functions\ lemma aff_dim_cball: fixes a :: "'n::euclidean_space" assumes "e > 0" shows "aff_dim (cball a e) = int (DIM('n))" proof - have "(\x. a + x) ` (cball 0 e) \ cball a e" unfolding cball_def dist_norm by auto then have "aff_dim (cball (0 :: 'n::euclidean_space) e) \ aff_dim (cball a e)" using aff_dim_translation_eq[of a "cball 0 e"] aff_dim_subset[of "(+) a ` cball 0 e" "cball a e"] by auto moreover have "aff_dim (cball (0 :: 'n::euclidean_space) e) = int (DIM('n))" using hull_inc[of "(0 :: 'n::euclidean_space)" "cball 0 e"] centre_in_cball[of "(0 :: 'n::euclidean_space)"] assms by (simp add: dim_cball[of e] aff_dim_zero[of "cball 0 e"]) ultimately show ?thesis using aff_dim_le_DIM[of "cball a e"] by auto qed lemma aff_dim_open: fixes S :: "'n::euclidean_space set" assumes "open S" and "S \ {}" shows "aff_dim S = int (DIM('n))" proof - obtain x where "x \ S" using assms by auto then obtain e where e: "e > 0" "cball x e \ S" using open_contains_cball[of S] assms by auto then have "aff_dim (cball x e) \ aff_dim S" using aff_dim_subset by auto with e show ?thesis using aff_dim_cball[of e x] aff_dim_le_DIM[of S] by auto qed lemma low_dim_interior: fixes S :: "'n::euclidean_space set" assumes "\ aff_dim S = int (DIM('n))" shows "interior S = {}" proof - have "aff_dim(interior S) \ aff_dim S" using interior_subset aff_dim_subset[of "interior S" S] by auto then show ?thesis using aff_dim_open[of "interior S"] aff_dim_le_DIM[of S] assms by auto qed corollary empty_interior_lowdim: fixes S :: "'n::euclidean_space set" shows "dim S < DIM ('n) \ interior S = {}" by (metis low_dim_interior affine_hull_UNIV dim_affine_hull less_not_refl dim_UNIV) corollary aff_dim_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "interior S \ {} \ aff_dim S = DIM('a)" by (metis low_dim_interior) subsection \Relative interior of a set\ definition\<^marker>\tag important\ "rel_interior S = {x. \T. openin (top_of_set (affine hull S)) T \ x \ T \ T \ S}" lemma rel_interior_mono: "\S \ T; affine hull S = affine hull T\ \ (rel_interior S) \ (rel_interior T)" by (auto simp: rel_interior_def) lemma rel_interior_maximal: "\T \ S; openin(top_of_set (affine hull S)) T\ \ T \ (rel_interior S)" by (auto simp: rel_interior_def) lemma rel_interior: "rel_interior S = {x \ S. \T. open T \ x \ T \ T \ affine hull S \ S}" unfolding rel_interior_def[of S] openin_open[of "affine hull S"] apply auto proof - fix x T assume *: "x \ S" "open T" "x \ T" "T \ affine hull S \ S" then have **: "x \ T \ affine hull S" using hull_inc by auto show "\Tb. (\Ta. open Ta \ Tb = affine hull S \ Ta) \ x \ Tb \ Tb \ S" apply (rule_tac x = "T \ (affine hull S)" in exI) using * ** apply auto done qed lemma mem_rel_interior: "x \ rel_interior S \ (\T. open T \ x \ T \ S \ T \ affine hull S \ S)" by (auto simp: rel_interior) lemma mem_rel_interior_ball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ ball x e \ affine hull S \ S)" apply (simp add: rel_interior, safe) apply (force simp: open_contains_ball) apply (rule_tac x = "ball x e" in exI, simp) done lemma rel_interior_ball: "rel_interior S = {x \ S. \e. e > 0 \ ball x e \ affine hull S \ S}" using mem_rel_interior_ball [of _ S] by auto lemma mem_rel_interior_cball: "x \ rel_interior S \ x \ S \ (\e. e > 0 \ cball x e \ affine hull S \ S)" apply (simp add: rel_interior, safe) apply (force simp: open_contains_cball) apply (rule_tac x = "ball x e" in exI) apply (simp add: subset_trans [OF ball_subset_cball], auto) done lemma rel_interior_cball: "rel_interior S = {x \ S. \e. e > 0 \ cball x e \ affine hull S \ S}" using mem_rel_interior_cball [of _ S] by auto lemma rel_interior_empty [simp]: "rel_interior {} = {}" by (auto simp: rel_interior_def) lemma affine_hull_sing [simp]: "affine hull {a :: 'n::euclidean_space} = {a}" by (metis affine_hull_eq affine_sing) lemma rel_interior_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_interior {a} = {a}" apply (auto simp: rel_interior_ball) apply (rule_tac x=1 in exI, force) done lemma subset_rel_interior: fixes S T :: "'n::euclidean_space set" assumes "S \ T" and "affine hull S = affine hull T" shows "rel_interior S \ rel_interior T" using assms by (auto simp: rel_interior_def) lemma rel_interior_subset: "rel_interior S \ S" by (auto simp: rel_interior_def) lemma rel_interior_subset_closure: "rel_interior S \ closure S" using rel_interior_subset by (auto simp: closure_def) lemma interior_subset_rel_interior: "interior S \ rel_interior S" by (auto simp: rel_interior interior_def) lemma interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "aff_dim S = int(DIM('n))" shows "rel_interior S = interior S" proof - have "affine hull S = UNIV" using assms affine_hull_UNIV[of S] by auto then show ?thesis unfolding rel_interior interior_def by auto qed lemma rel_interior_interior: fixes S :: "'n::euclidean_space set" assumes "affine hull S = UNIV" shows "rel_interior S = interior S" using assms unfolding rel_interior interior_def by auto lemma rel_interior_open: fixes S :: "'n::euclidean_space set" assumes "open S" shows "rel_interior S = S" by (metis assms interior_eq interior_subset_rel_interior rel_interior_subset set_eq_subset) lemma interior_ball [simp]: "interior (ball x e) = ball x e" by (simp add: interior_open) lemma interior_rel_interior_gen: fixes S :: "'n::euclidean_space set" shows "interior S = (if aff_dim S = int(DIM('n)) then rel_interior S else {})" by (metis interior_rel_interior low_dim_interior) lemma rel_interior_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_interior S = interior S" by (metis interior_rel_interior_gen) lemma affine_hull_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ affine hull S = UNIV" by (metis affine_hull_UNIV interior_rel_interior_gen) lemma rel_interior_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "rel_interior (affine hull S) = affine hull S" proof - have *: "rel_interior (affine hull S) \ affine hull S" using rel_interior_subset by auto { fix x assume x: "x \ affine hull S" define e :: real where "e = 1" then have "e > 0" "ball x e \ affine hull (affine hull S) \ affine hull S" using hull_hull[of _ S] by auto then have "x \ rel_interior (affine hull S)" using x rel_interior_ball[of "affine hull S"] by auto } then show ?thesis using * by auto qed lemma rel_interior_UNIV [simp]: "rel_interior (UNIV :: ('n::euclidean_space) set) = UNIV" by (metis open_UNIV rel_interior_open) lemma rel_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where "d > 0" and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto { fix y assume as: "dist (x - e *\<^sub>R (x - c)) y < e * d" "y \ affine hull S" have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "x \ affine hull S" using assms hull_subset[of S] by auto moreover have "1 / e + - ((1 - e) / e) = 1" using \e > 0\ left_diff_distrib[of "1" "(1-e)" "1/e"] by auto ultimately have **: "(1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x \ affine hull S" using as affine_affine_hull[of S] mem_affine[of "affine hull S" y x "(1 / e)" "-((1 - e) / e)"] by (simp add: algebra_simps) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ apply (auto simp: euclidean_eq_iff[where 'a='a] field_simps inner_simps) done also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp:pos_divide_less_eq[OF \e > 0\] mult.commute) finally have "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[THEN subsetD]) unfolding mem_ball using assms(3-5) ** apply auto done } then have "ball (x - e *\<^sub>R (x - c)) (e*d) \ affine hull S \ S" by auto moreover have "e * d > 0" using \e > 0\ \d > 0\ by simp moreover have c: "c \ S" using assms rel_interior_subset by auto moreover from c have "x - e *\<^sub>R (x - c) \ S" using convexD_alt[of S x c e] apply (simp add: algebra_simps) using assms apply auto done ultimately show ?thesis using mem_rel_interior_ball[of "x - e *\<^sub>R (x - c)" S] \e > 0\ by auto qed lemma interior_real_atLeast [simp]: fixes a :: real shows "interior {a..} = {a<..}" proof - { fix y assume "a < y" then have "y \ interior {a..}" apply (simp add: mem_interior) apply (rule_tac x="(y-a)" in exI) apply (auto simp: dist_norm) done } moreover { fix y assume "y \ interior {a..}" then obtain e where e: "e > 0" "cball y e \ {a..}" using mem_interior_cball[of y "{a..}"] by auto moreover from e have "y - e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y - e" by blast then have "a < y" using e by auto } ultimately show ?thesis by auto qed lemma continuous_ge_on_Ioo: assumes "continuous_on {c..d} g" "\x. x \ {c<.. g x \ a" "c < d" "x \ {c..d}" shows "g (x::real) \ (a::real)" proof- from assms(3) have "{c..d} = closure {c<.. (g -` {a..} \ {c..d})" by auto hence "closure {c<.. closure (g -` {a..} \ {c..d})" by (rule closure_mono) also from assms(1) have "closed (g -` {a..} \ {c..d})" by (auto simp: continuous_on_closed_vimage) hence "closure (g -` {a..} \ {c..d}) = g -` {a..} \ {c..d}" by simp finally show ?thesis using \x \ {c..d}\ by auto qed lemma interior_real_atMost [simp]: fixes a :: real shows "interior {..a} = {.. y" then have "y \ interior {..a}" apply (simp add: mem_interior) apply (rule_tac x="(a-y)" in exI) apply (auto simp: dist_norm) done } moreover { fix y assume "y \ interior {..a}" then obtain e where e: "e > 0" "cball y e \ {..a}" using mem_interior_cball[of y "{..a}"] by auto moreover from e have "y + e \ cball y e" by (auto simp: cball_def dist_norm) ultimately have "a \ y + e" by auto then have "a > y" using e by auto } ultimately show ?thesis by auto qed lemma interior_atLeastAtMost_real [simp]: "interior {a..b} = {a<.. {..b}" by auto also have "interior \ = {a<..} \ {.. = {a<.. {}" using assms unfolding set_eq_iff by (auto intro!: exI[of _ "(a + b) / 2"] simp: box_def) then show ?thesis using interior_rel_interior_gen[of "cbox a b", symmetric] by (simp split: if_split_asm del: box_real add: box_real[symmetric] interior_cbox) qed lemma rel_interior_real_semiline [simp]: fixes a :: real shows "rel_interior {a..} = {a<..}" proof - have *: "{a<..} \ {}" unfolding set_eq_iff by (auto intro!: exI[of _ "a + 1"]) then show ?thesis using interior_real_atLeast interior_rel_interior_gen[of "{a..}"] by (auto split: if_split_asm) qed subsubsection \Relative open sets\ definition\<^marker>\tag important\ "rel_open S \ rel_interior S = S" lemma rel_open: "rel_open S \ openin (top_of_set (affine hull S)) S" unfolding rel_open_def rel_interior_def apply auto using openin_subopen[of "top_of_set (affine hull S)" S] apply auto done lemma openin_rel_interior: "openin (top_of_set (affine hull S)) (rel_interior S)" apply (simp add: rel_interior_def) apply (subst openin_subopen, blast) done lemma openin_set_rel_interior: "openin (top_of_set S) (rel_interior S)" by (rule openin_subset_trans [OF openin_rel_interior rel_interior_subset hull_subset]) lemma affine_rel_open: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "rel_open S" unfolding rel_open_def using assms rel_interior_affine_hull[of S] affine_hull_eq[of S] by metis lemma affine_closed: fixes S :: "'n::euclidean_space set" assumes "affine S" shows "closed S" proof - { assume "S \ {}" then obtain L where L: "subspace L" "affine_parallel S L" using assms affine_parallel_subspace[of S] by auto then obtain a where a: "S = ((+) a ` L)" using affine_parallel_def[of L S] affine_parallel_commut by auto from L have "closed L" using closed_subspace by auto then have "closed S" using closed_translation a by auto } then show ?thesis by auto qed lemma closure_affine_hull: fixes S :: "'n::euclidean_space set" shows "closure S \ affine hull S" by (intro closure_minimal hull_subset affine_closed affine_affine_hull) lemma closure_same_affine_hull [simp]: fixes S :: "'n::euclidean_space set" shows "affine hull (closure S) = affine hull S" proof - have "affine hull (closure S) \ affine hull S" using hull_mono[of "closure S" "affine hull S" "affine"] closure_affine_hull[of S] hull_hull[of "affine" S] by auto moreover have "affine hull (closure S) \ affine hull S" using hull_mono[of "S" "closure S" "affine"] closure_subset by auto ultimately show ?thesis by auto qed lemma closure_aff_dim [simp]: fixes S :: "'n::euclidean_space set" shows "aff_dim (closure S) = aff_dim S" proof - have "aff_dim S \ aff_dim (closure S)" using aff_dim_subset closure_subset by auto moreover have "aff_dim (closure S) \ aff_dim (affine hull S)" using aff_dim_subset closure_affine_hull by blast moreover have "aff_dim (affine hull S) = aff_dim S" using aff_dim_affine_hull by auto ultimately show ?thesis by auto qed lemma rel_interior_closure_convex_shrink: fixes S :: "_::euclidean_space set" assumes "convex S" and "c \ rel_interior S" and "x \ closure S" and "e > 0" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ rel_interior S" proof - obtain d where "d > 0" and d: "ball c d \ affine hull S \ S" using assms(2) unfolding mem_rel_interior_ball by auto have "\y \ S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x], auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have zball: "z \ ball c d" using mem_ball z_def dist_norm[of c] using y and assms(4,5) by (simp add: norm_minus_commute) (simp add: field_simps) have "x \ affine hull S" using closure_affine_hull assms by auto moreover have "y \ affine hull S" using \y \ S\ hull_subset[of S] by auto moreover have "c \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto ultimately have "z \ affine hull S" using z_def affine_affine_hull[of S] mem_affine_3_minus [of "affine hull S" c x y "(1 - e) / e"] assms by simp then have "z \ S" using d zball by auto obtain d1 where "d1 > 0" and d1: "ball z d1 \ ball c d" using zball open_ball[of c d] openE[of "ball c d" z] by auto then have "ball z d1 \ affine hull S \ ball c d \ affine hull S" by auto then have "ball z d1 \ affine hull S \ S" using d by auto then have "z \ rel_interior S" using mem_rel_interior_ball using \d1 > 0\ \z \ S\ by auto then have "y - e *\<^sub>R (y - z) \ rel_interior S" using rel_interior_convex_shrink[of S z y e] assms \y \ S\ by auto then show ?thesis using * by auto qed lemma rel_interior_eq: "rel_interior s = s \ openin(top_of_set (affine hull s)) s" using rel_open rel_open_def by blast lemma rel_interior_openin: "openin(top_of_set (affine hull s)) s \ rel_interior s = s" by (simp add: rel_interior_eq) lemma rel_interior_affine: fixes S :: "'n::euclidean_space set" shows "affine S \ rel_interior S = S" using affine_rel_open rel_open_def by auto lemma rel_interior_eq_closure: fixes S :: "'n::euclidean_space set" shows "rel_interior S = closure S \ affine S" proof (cases "S = {}") case True then show ?thesis by auto next case False show ?thesis proof assume eq: "rel_interior S = closure S" have "S = {} \ S = affine hull S" apply (rule connected_clopen [THEN iffD1, rule_format]) apply (simp add: affine_imp_convex convex_connected) apply (rule conjI) apply (metis eq closure_subset openin_rel_interior rel_interior_subset subset_antisym) apply (metis closed_subset closure_subset_eq eq hull_subset rel_interior_subset) done with False have "affine hull S = S" by auto then show "affine S" by (metis affine_hull_eq) next assume "affine S" then show "rel_interior S = closure S" by (simp add: rel_interior_affine affine_closed) qed qed subsubsection\<^marker>\tag unimportant\\Relative interior preserves under linear transformations\ lemma rel_interior_translation_aux: fixes a :: "'n::euclidean_space" shows "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" proof - { fix x assume x: "x \ rel_interior S" then obtain T where "open T" "x \ T \ S" "T \ affine hull S \ S" using mem_rel_interior[of x S] by auto then have "open ((\x. a + x) ` T)" and "a + x \ ((\x. a + x) ` T) \ ((\x. a + x) ` S)" and "((\x. a + x) ` T) \ affine hull ((\x. a + x) ` S) \ (\x. a + x) ` S" using affine_hull_translation[of a S] open_translation[of T a] x by auto then have "a + x \ rel_interior ((\x. a + x) ` S)" using mem_rel_interior[of "a+x" "((\x. a + x) ` S)"] by auto } then show ?thesis by auto qed lemma rel_interior_translation: fixes a :: "'n::euclidean_space" shows "rel_interior ((\x. a + x) ` S) = (\x. a + x) ` rel_interior S" proof - have "(\x. (-a) + x) ` rel_interior ((\x. a + x) ` S) \ rel_interior S" using rel_interior_translation_aux[of "-a" "(\x. a + x) ` S"] translation_assoc[of "-a" "a"] by auto then have "((\x. a + x) ` rel_interior S) \ rel_interior ((\x. a + x) ` S)" using translation_inverse_subset[of a "rel_interior ((+) a ` S)" "rel_interior S"] by auto then show ?thesis using rel_interior_translation_aux[of a S] by auto qed lemma affine_hull_linear_image: assumes "bounded_linear f" shows "f ` (affine hull s) = affine hull f ` s" proof - interpret f: bounded_linear f by fact have "affine {x. f x \ affine hull f ` s}" unfolding affine_def by (auto simp: f.scaleR f.add affine_affine_hull[unfolded affine_def, rule_format]) moreover have "affine {x. x \ f ` (affine hull s)}" using affine_affine_hull[unfolded affine_def, of s] unfolding affine_def by (auto simp: f.scaleR [symmetric] f.add [symmetric]) ultimately show ?thesis by (auto simp: hull_inc elim!: hull_induct) qed lemma rel_interior_injective_on_span_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" and S :: "'m::euclidean_space set" assumes "bounded_linear f" and "inj_on f (span S)" shows "rel_interior (f ` S) = f ` (rel_interior S)" proof - { fix z assume z: "z \ rel_interior (f ` S)" then have "z \ f ` S" using rel_interior_subset[of "f ` S"] by auto then obtain x where x: "x \ S" "f x = z" by auto obtain e2 where e2: "e2 > 0" "cball z e2 \ affine hull (f ` S) \ (f ` S)" using z rel_interior_cball[of "f ` S"] by auto obtain K where K: "K > 0" "\x. norm (f x) \ norm x * K" using assms Real_Vector_Spaces.bounded_linear.pos_bounded[of f] by auto define e1 where "e1 = 1 / K" then have e1: "e1 > 0" "\x. e1 * norm (f x) \ norm x" using K pos_le_divide_eq[of e1] by auto define e where "e = e1 * e2" then have "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball x e \ affine hull S" then have h1: "f y \ affine hull (f ` S)" using affine_hull_linear_image[of f S] assms by auto from y have "norm (x-y) \ e1 * e2" using cball_def[of x e] dist_norm[of x y] e_def by auto moreover have "f x - f y = f (x - y)" using assms linear_diff[of f x y] linear_conv_bounded_linear[of f] by auto moreover have "e1 * norm (f (x-y)) \ norm (x - y)" using e1 by auto ultimately have "e1 * norm ((f x)-(f y)) \ e1 * e2" by auto then have "f y \ cball z e2" using cball_def[of "f x" e2] dist_norm[of "f x" "f y"] e1 x by auto then have "f y \ f ` S" using y e2 h1 by auto then have "y \ S" using assms y hull_subset[of S] affine_hull_subset_span inj_on_image_mem_iff [OF \inj_on f (span S)\] by (metis Int_iff span_superset subsetCE) } then have "z \ f ` (rel_interior S)" using mem_rel_interior_cball[of x S] \e > 0\ x by auto } moreover { fix x assume x: "x \ rel_interior S" then obtain e2 where e2: "e2 > 0" "cball x e2 \ affine hull S \ S" using rel_interior_cball[of S] by auto have "x \ S" using x rel_interior_subset by auto then have *: "f x \ f ` S" by auto have "\x\span S. f x = 0 \ x = 0" using assms subspace_span linear_conv_bounded_linear[of f] linear_injective_on_subspace_0[of f "span S"] by auto then obtain e1 where e1: "e1 > 0" "\x \ span S. e1 * norm x \ norm (f x)" using assms injective_imp_isometric[of "span S" f] subspace_span[of S] closed_subspace[of "span S"] by auto define e where "e = e1 * e2" hence "e > 0" using e1 e2 by auto { fix y assume y: "y \ cball (f x) e \ affine hull (f ` S)" then have "y \ f ` (affine hull S)" using affine_hull_linear_image[of f S] assms by auto then obtain xy where xy: "xy \ affine hull S" "f xy = y" by auto with y have "norm (f x - f xy) \ e1 * e2" using cball_def[of "f x" e] dist_norm[of "f x" y] e_def by auto moreover have "f x - f xy = f (x - xy)" using assms linear_diff[of f x xy] linear_conv_bounded_linear[of f] by auto moreover have *: "x - xy \ span S" using subspace_diff[of "span S" x xy] subspace_span \x \ S\ xy affine_hull_subset_span[of S] span_superset by auto moreover from * have "e1 * norm (x - xy) \ norm (f (x - xy))" using e1 by auto ultimately have "e1 * norm (x - xy) \ e1 * e2" by auto then have "xy \ cball x e2" using cball_def[of x e2] dist_norm[of x xy] e1 by auto then have "y \ f ` S" using xy e2 by auto } then have "f x \ rel_interior (f ` S)" using mem_rel_interior_cball[of "(f x)" "(f ` S)"] * \e > 0\ by auto } ultimately show ?thesis by auto qed lemma rel_interior_injective_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "bounded_linear f" and "inj f" shows "rel_interior (f ` S) = f ` (rel_interior S)" using assms rel_interior_injective_on_span_linear_image[of f S] subset_inj_on[of f "UNIV" "span S"] by auto subsection\<^marker>\tag unimportant\ \Openness and compactness are preserved by convex hull operation\ lemma open_convex_hull[intro]: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (convex hull S)" proof (clarsimp simp: open_contains_cball convex_hull_explicit) fix T and u :: "'a\real" assume obt: "finite T" "T\S" "\x\T. 0 \ u x" "sum u T = 1" from assms[unfolded open_contains_cball] obtain b where b: "\x. x\S \ 0 < b x \ cball x (b x) \ S" by metis have "b ` T \ {}" using obt by auto define i where "i = b ` T" let ?\ = "\y. \F. finite F \ F \ S \ (\u. (\x\F. 0 \ u x) \ sum u F = 1 \ (\v\F. u v *\<^sub>R v) = y)" let ?a = "\v\T. u v *\<^sub>R v" show "\e > 0. cball ?a e \ {y. ?\ y}" proof (intro exI subsetI conjI) show "0 < Min i" unfolding i_def and Min_gr_iff[OF finite_imageI[OF obt(1)] \b ` T\{}\] using b \T\S\ by auto next fix y assume "y \ cball ?a (Min i)" then have y: "norm (?a - y) \ Min i" unfolding dist_norm[symmetric] by auto { fix x assume "x \ T" then have "Min i \ b x" by (simp add: i_def obt(1)) then have "x + (y - ?a) \ cball x (b x)" using y unfolding mem_cball dist_norm by auto moreover have "x \ S" using \x\T\ \T\S\ by auto ultimately have "x + (y - ?a) \ S" using y b by blast } moreover have *: "inj_on (\v. v + (y - ?a)) T" unfolding inj_on_def by auto have "(\v\(\v. v + (y - ?a)) ` T. u (v - (y - ?a)) *\<^sub>R v) = y" unfolding sum.reindex[OF *] o_def using obt(4) by (simp add: sum.distrib sum_subtractf scaleR_left.sum[symmetric] scaleR_right_distrib) ultimately show "y \ {y. ?\ y}" proof (intro CollectI exI conjI) show "finite ((\v. v + (y - ?a)) ` T)" by (simp add: obt(1)) show "sum (\v. u (v - (y - ?a))) ((\v. v + (y - ?a)) ` T) = 1" unfolding sum.reindex[OF *] o_def using obt(4) by auto qed (use obt(1, 3) in auto) qed qed lemma compact_convex_combinations: fixes S T :: "'a::real_normed_vector set" assumes "compact S" "compact T" shows "compact { (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T}" proof - let ?X = "{0..1} \ S \ T" let ?h = "(\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" have *: "{ (1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ T} = ?h ` ?X" by force have "continuous_on ?X (\z. (1 - fst z) *\<^sub>R fst (snd z) + fst z *\<^sub>R snd (snd z))" unfolding continuous_on by (rule ballI) (intro tendsto_intros) with assms show ?thesis by (simp add: * compact_Times compact_continuous_image) qed lemma finite_imp_compact_convex_hull: fixes S :: "'a::real_normed_vector set" assumes "finite S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis by simp next case False with assms show ?thesis proof (induct rule: finite_ne_induct) case (singleton x) show ?case by simp next case (insert x A) let ?f = "\(u, y::'a). u *\<^sub>R x + (1 - u) *\<^sub>R y" let ?T = "{0..1::real} \ (convex hull A)" have "continuous_on ?T ?f" unfolding split_def continuous_on by (intro ballI tendsto_intros) moreover have "compact ?T" by (intro compact_Times compact_Icc insert) ultimately have "compact (?f ` ?T)" by (rule compact_continuous_image) also have "?f ` ?T = convex hull (insert x A)" unfolding convex_hull_insert [OF \A \ {}\] apply safe apply (rule_tac x=a in exI, simp) apply (rule_tac x="1 - a" in exI, simp, fast) apply (rule_tac x="(u, b)" in image_eqI, simp_all) done finally show "compact (convex hull (insert x A))" . qed qed lemma compact_convex_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "compact (convex hull S)" proof (cases "S = {}") case True then show ?thesis using compact_empty by simp next case False then obtain w where "w \ S" by auto show ?thesis unfolding caratheodory[of S] proof (induct ("DIM('a) + 1")) case 0 have *: "{x.\sa. finite sa \ sa \ S \ card sa \ 0 \ x \ convex hull sa} = {}" using compact_empty by auto from 0 show ?case unfolding * by simp next case (Suc n) show ?case proof (cases "n = 0") case True have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = S" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "x \ S" proof (cases "card T = 0") case True then show ?thesis using T(4) unfolding card_0_eq[OF T(1)] by simp next case False then have "card T = Suc 0" using T(3) \n=0\ by auto then obtain a where "T = {a}" unfolding card_Suc_eq by auto then show ?thesis using T(2,4) by simp qed next fix x assume "x\S" then show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="{x}" in exI) unfolding convex_hull_singleton apply auto done qed then show ?thesis using assms by simp next case False have "{x. \T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T} = {(1 - u) *\<^sub>R x + u *\<^sub>R y | x y u. 0 \ u \ u \ 1 \ x \ S \ y \ {x. \T. finite T \ T \ S \ card T \ n \ x \ convex hull T}}" unfolding set_eq_iff and mem_Collect_eq proof (rule, rule) fix x assume "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" then obtain u v c T where obt: "x = (1 - c) *\<^sub>R u + c *\<^sub>R v" "0 \ c \ c \ 1" "u \ S" "finite T" "T \ S" "card T \ n" "v \ convex hull T" by auto moreover have "(1 - c) *\<^sub>R u + c *\<^sub>R v \ convex hull insert u T" apply (rule convexD_alt) using obt(2) and convex_convex_hull and hull_subset[of "insert u T" convex] using obt(7) and hull_mono[of T "insert u T"] apply auto done ultimately show "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" apply (rule_tac x="insert u T" in exI) apply (auto simp: card_insert_if) done next fix x assume "\T. finite T \ T \ S \ card T \ Suc n \ x \ convex hull T" then obtain T where T: "finite T" "T \ S" "card T \ Suc n" "x \ convex hull T" by auto show "\u v c. x = (1 - c) *\<^sub>R u + c *\<^sub>R v \ 0 \ c \ c \ 1 \ u \ S \ (\T. finite T \ T \ S \ card T \ n \ v \ convex hull T)" proof (cases "card T = Suc n") case False then have "card T \ n" using T(3) by auto then show ?thesis apply (rule_tac x=w in exI, rule_tac x=x in exI, rule_tac x=1 in exI) using \w\S\ and T apply (auto intro!: exI[where x=T]) done next case True then obtain a u where au: "T = insert a u" "a\u" apply (drule_tac card_eq_SucD, auto) done show ?thesis proof (cases "u = {}") case True then have "x = a" using T(4)[unfolded au] by auto show ?thesis unfolding \x = a\ apply (rule_tac x=a in exI) apply (rule_tac x=a in exI) apply (rule_tac x=1 in exI) using T and \n \ 0\ unfolding au apply (auto intro!: exI[where x="{a}"]) done next case False obtain ux vx b where obt: "ux\0" "vx\0" "ux + vx = 1" "b \ convex hull u" "x = ux *\<^sub>R a + vx *\<^sub>R b" using T(4)[unfolded au convex_hull_insert[OF False]] by auto have *: "1 - vx = ux" using obt(3) by auto show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=b in exI) apply (rule_tac x=vx in exI) using obt and T(1-3) unfolding au and * using card_insert_disjoint[OF _ au(2)] apply (auto intro!: exI[where x=u]) done qed qed qed then show ?thesis using compact_convex_combinations[OF assms Suc] by simp qed qed qed subsection\<^marker>\tag unimportant\ \Extremal points of a simplex are some vertices\ lemma dist_increases_online: fixes a b d :: "'a::real_inner" assumes "d \ 0" shows "dist a (b + d) > dist a b \ dist a (b - d) > dist a b" proof (cases "inner a d - inner b d > 0") case True then have "0 < inner d d + (inner a d * 2 - inner b d * 2)" apply (rule_tac add_pos_pos) using assms apply auto done then show ?thesis apply (rule_tac disjI2) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff apply (simp add: algebra_simps inner_commute) done next case False then have "0 < inner d d + (inner b d * 2 - inner a d * 2)" apply (rule_tac add_pos_nonneg) using assms apply auto done then show ?thesis apply (rule_tac disjI1) unfolding dist_norm and norm_eq_sqrt_inner and real_sqrt_less_iff apply (simp add: algebra_simps inner_commute) done qed lemma norm_increases_online: fixes d :: "'a::real_inner" shows "d \ 0 \ norm (a + d) > norm a \ norm(a - d) > norm a" using dist_increases_online[of d a 0] unfolding dist_norm by auto lemma simplex_furthest_lt: fixes S :: "'a::real_inner set" assumes "finite S" shows "\x \ convex hull S. x \ S \ (\y \ convex hull S. norm (x - a) < norm(y - a))" using assms proof induct fix x S assume as: "finite S" "x\S" "\x\convex hull S. x \ S \ (\y\convex hull S. norm (x - a) < norm (y - a))" show "\xa\convex hull insert x S. xa \ insert x S \ (\y\convex hull insert x S. norm (xa - a) < norm (y - a))" proof (intro impI ballI, cases "S = {}") case False fix y assume y: "y \ convex hull insert x S" "y \ insert x S" obtain u v b where obt: "u\0" "v\0" "u + v = 1" "b \ convex hull S" "y = u *\<^sub>R x + v *\<^sub>R b" using y(1)[unfolded convex_hull_insert[OF False]] by auto show "\z\convex hull insert x S. norm (y - a) < norm (z - a)" proof (cases "y \ convex hull S") case True then obtain z where "z \ convex hull S" "norm (y - a) < norm (z - a)" using as(3)[THEN bspec[where x=y]] and y(2) by auto then show ?thesis apply (rule_tac x=z in bexI) unfolding convex_hull_insert[OF False] apply auto done next case False show ?thesis using obt(3) proof (cases "u = 0", case_tac[!] "v = 0") assume "u = 0" "v \ 0" then have "y = b" using obt by auto then show ?thesis using False and obt(4) by auto next assume "u \ 0" "v = 0" then have "y = x" using obt by auto then show ?thesis using y(2) by auto next assume "u \ 0" "v \ 0" then obtain w where w: "w>0" "w b" proof assume "x = b" then have "y = b" unfolding obt(5) using obt(3) by (auto simp: scaleR_left_distrib[symmetric]) then show False using obt(4) and False by simp qed then have *: "w *\<^sub>R (x - b) \ 0" using w(1) by auto show ?thesis using dist_increases_online[OF *, of a y] proof (elim disjE) assume "dist a y < dist a (y + w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u + w) *\<^sub>R x + (v - w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u + w) *\<^sub>R x + (v - w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u + w \ 0" "v - w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto next assume "dist a y < dist a (y - w *\<^sub>R (x - b))" then have "norm (y - a) < norm ((u - w) *\<^sub>R x + (v + w) *\<^sub>R b - a)" unfolding dist_commute[of a] unfolding dist_norm obt(5) by (simp add: algebra_simps) moreover have "(u - w) *\<^sub>R x + (v + w) *\<^sub>R b \ convex hull insert x S" unfolding convex_hull_insert[OF \S\{}\] proof (intro CollectI conjI exI) show "u - w \ 0" "v + w \ 0" using obt(1) w by auto qed (use obt in auto) ultimately show ?thesis by auto qed qed auto qed qed auto qed (auto simp: assms) lemma simplex_furthest_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\y\S. \x\ convex hull S. norm (x - a) \ norm (y - a)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain x where x: "x \ convex hull S" "\y\convex hull S. norm (y - a) \ norm (x - a)" using distance_attains_sup[OF finite_imp_compact_convex_hull[OF \finite S\], of a] unfolding dist_commute[of a] unfolding dist_norm by auto show ?thesis proof (cases "x \ S") case False then obtain y where "y \ convex hull S" "norm (x - a) < norm (y - a)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=x]] and x(1) by auto then show ?thesis using x(2)[THEN bspec[where x=y]] by auto next case True with x show ?thesis by auto qed qed lemma simplex_furthest_le_exists: fixes S :: "('a::real_inner) set" shows "finite S \ \x\(convex hull S). \y\S. norm (x - a) \ norm (y - a)" using simplex_furthest_le[of S] by (cases "S = {}") auto lemma simplex_extremal_le: fixes S :: "'a::real_inner set" assumes "finite S" and "S \ {}" shows "\u\S. \v\S. \x\convex hull S. \y \ convex hull S. norm (x - y) \ norm (u - v)" proof - have "convex hull S \ {}" using hull_subset[of S convex] and assms(2) by auto then obtain u v where obt: "u \ convex hull S" "v \ convex hull S" "\x\convex hull S. \y\convex hull S. norm (x - y) \ norm (u - v)" using compact_sup_maxdistance[OF finite_imp_compact_convex_hull[OF assms(1)]] by (auto simp: dist_norm) then show ?thesis proof (cases "u\S \ v\S", elim disjE) assume "u \ S" then obtain y where "y \ convex hull S" "norm (u - v) < norm (y - v)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=u]] and obt(1) by auto then show ?thesis using obt(3)[THEN bspec[where x=y], THEN bspec[where x=v]] and obt(2) by auto next assume "v \ S" then obtain y where "y \ convex hull S" "norm (v - u) < norm (y - u)" using simplex_furthest_lt[OF assms(1), THEN bspec[where x=v]] and obt(2) by auto then show ?thesis using obt(3)[THEN bspec[where x=u], THEN bspec[where x=y]] and obt(1) by (auto simp: norm_minus_commute) qed auto qed lemma simplex_extremal_le_exists: fixes S :: "'a::real_inner set" shows "finite S \ x \ convex hull S \ y \ convex hull S \ \u\S. \v\S. norm (x - y) \ norm (u - v)" using convex_hull_empty simplex_extremal_le[of S] by(cases "S = {}") auto subsection \Closest point of a convex set is unique, with a continuous projection\ definition\<^marker>\tag important\ closest_point :: "'a::{real_inner,heine_borel} set \ 'a \ 'a" where "closest_point S a = (SOME x. x \ S \ (\y\S. dist a x \ dist a y))" lemma closest_point_exists: assumes "closed S" and "S \ {}" shows "closest_point S a \ S" and "\y\S. dist a (closest_point S a) \ dist a y" unfolding closest_point_def apply(rule_tac[!] someI2_ex) apply (auto intro: distance_attains_inf[OF assms(1,2), of a]) done lemma closest_point_in_set: "closed S \ S \ {} \ closest_point S a \ S" by (meson closest_point_exists) lemma closest_point_le: "closed S \ x \ S \ dist a (closest_point S a) \ dist a x" using closest_point_exists[of S] by auto lemma closest_point_self: assumes "x \ S" shows "closest_point S x = x" unfolding closest_point_def apply (rule some1_equality, rule ex1I[of _ x]) using assms apply auto done lemma closest_point_refl: "closed S \ S \ {} \ closest_point S x = x \ x \ S" using closest_point_in_set[of S x] closest_point_self[of x S] by auto lemma closer_points_lemma: assumes "inner y z > 0" shows "\u>0. \v>0. v \ u \ norm(v *\<^sub>R z - y) < norm y" proof - have z: "inner z z > 0" unfolding inner_gt_zero_iff using assms by auto have "norm (v *\<^sub>R z - y) < norm y" if "0 < v" and "v \ inner y z / inner z z" for v unfolding norm_lt using z assms that by (simp add: field_simps inner_diff inner_commute mult_strict_left_mono[OF _ \0]) then show ?thesis using assms z by (rule_tac x = "inner y z / inner z z" in exI) auto qed lemma closer_point_lemma: assumes "inner (y - x) (z - x) > 0" shows "\u>0. u \ 1 \ dist (x + u *\<^sub>R (z - x)) y < dist x y" proof - obtain u where "u > 0" and u: "\v>0. v \ u \ norm (v *\<^sub>R (z - x) - (y - x)) < norm (y - x)" using closer_points_lemma[OF assms] by auto show ?thesis apply (rule_tac x="min u 1" in exI) using u[THEN spec[where x="min u 1"]] and \u > 0\ unfolding dist_norm by (auto simp: norm_minus_commute field_simps) qed lemma any_closest_point_dot: assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" shows "inner (a - x) (y - x) \ 0" proof (rule ccontr) assume "\ ?thesis" then obtain u where u: "u>0" "u\1" "dist (x + u *\<^sub>R (y - x)) a < dist x a" using closer_point_lemma[of a x y] by auto let ?z = "(1 - u) *\<^sub>R x + u *\<^sub>R y" have "?z \ S" using convexD_alt[OF assms(1,3,4), of u] using u by auto then show False using assms(5)[THEN bspec[where x="?z"]] and u(3) by (auto simp: dist_commute algebra_simps) qed lemma any_closest_point_unique: fixes x :: "'a::real_inner" assumes "convex S" "closed S" "x \ S" "y \ S" "\z\S. dist a x \ dist a z" "\z\S. dist a y \ dist a z" shows "x = y" using any_closest_point_dot[OF assms(1-4,5)] and any_closest_point_dot[OF assms(1-2,4,3,6)] unfolding norm_pths(1) and norm_le_square by (auto simp: algebra_simps) lemma closest_point_unique: assumes "convex S" "closed S" "x \ S" "\z\S. dist a x \ dist a z" shows "x = closest_point S a" using any_closest_point_unique[OF assms(1-3) _ assms(4), of "closest_point S a"] using closest_point_exists[OF assms(2)] and assms(3) by auto lemma closest_point_dot: assumes "convex S" "closed S" "x \ S" shows "inner (a - closest_point S a) (x - closest_point S a) \ 0" apply (rule any_closest_point_dot[OF assms(1,2) _ assms(3)]) using closest_point_exists[OF assms(2)] and assms(3) apply auto done lemma closest_point_lt: assumes "convex S" "closed S" "x \ S" "x \ closest_point S a" shows "dist a (closest_point S a) < dist a x" apply (rule ccontr) apply (rule_tac notE[OF assms(4)]) apply (rule closest_point_unique[OF assms(1-3), of a]) using closest_point_le[OF assms(2), of _ a] apply fastforce done lemma setdist_closest_point: "\closed S; S \ {}\ \ setdist {a} S = dist a (closest_point S a)" apply (rule setdist_unique) using closest_point_le apply (auto simp: closest_point_in_set) done lemma closest_point_lipschitz: assumes "convex S" and "closed S" "S \ {}" shows "dist (closest_point S x) (closest_point S y) \ dist x y" proof - have "inner (x - closest_point S x) (closest_point S y - closest_point S x) \ 0" and "inner (y - closest_point S y) (closest_point S x - closest_point S y) \ 0" apply (rule_tac[!] any_closest_point_dot[OF assms(1-2)]) using closest_point_exists[OF assms(2-3)] apply auto done then show ?thesis unfolding dist_norm and norm_le using inner_ge_zero[of "(x - closest_point S x) - (y - closest_point S y)"] by (simp add: inner_add inner_diff inner_commute) qed lemma continuous_at_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous (at x) (closest_point S)" unfolding continuous_at_eps_delta using le_less_trans[OF closest_point_lipschitz[OF assms]] by auto lemma continuous_on_closest_point: assumes "convex S" and "closed S" and "S \ {}" shows "continuous_on t (closest_point S)" by (metis continuous_at_imp_continuous_on continuous_at_closest_point[OF assms]) proposition closest_point_in_rel_interior: assumes "closed S" "S \ {}" and x: "x \ affine hull S" shows "closest_point S x \ rel_interior S \ x \ rel_interior S" proof (cases "x \ S") case True then show ?thesis by (simp add: closest_point_self) next case False then have "False" if asm: "closest_point S x \ rel_interior S" proof - obtain e where "e > 0" and clox: "closest_point S x \ S" and e: "cball (closest_point S x) e \ affine hull S \ S" using asm mem_rel_interior_cball by blast then have clo_notx: "closest_point S x \ x" using \x \ S\ by auto define y where "y \ closest_point S x - (min 1 (e / norm(closest_point S x - x))) *\<^sub>R (closest_point S x - x)" have "x - y = (1 - min 1 (e / norm (closest_point S x - x))) *\<^sub>R (x - closest_point S x)" by (simp add: y_def algebra_simps) then have "norm (x - y) = abs ((1 - min 1 (e / norm (closest_point S x - x)))) * norm(x - closest_point S x)" by simp also have "\ < norm(x - closest_point S x)" using clo_notx \e > 0\ by (auto simp: mult_less_cancel_right2 field_split_simps) finally have no_less: "norm (x - y) < norm (x - closest_point S x)" . have "y \ affine hull S" unfolding y_def by (meson affine_affine_hull clox hull_subset mem_affine_3_minus2 subsetD x) moreover have "dist (closest_point S x) y \ e" using \e > 0\ by (auto simp: y_def min_mult_distrib_right) ultimately have "y \ S" using subsetD [OF e] by simp then have "dist x (closest_point S x) \ dist x y" by (simp add: closest_point_le \closed S\) with no_less show False by (simp add: dist_norm) qed moreover have "x \ rel_interior S" using rel_interior_subset False by blast ultimately show ?thesis by blast qed subsubsection\<^marker>\tag unimportant\ \Various point-to-set separating/supporting hyperplane theorems\ lemma supporting_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "S \ {}" and "z \ S" shows "\a b. \y\S. inner a z < b \ inner a y = b \ (\x\S. inner a x \ b)" proof - obtain y where "y \ S" and y: "\x\S. dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2-3)]) show ?thesis proof (intro exI bexI conjI ballI) show "(y - z) \ z < (y - z) \ y" by (metis \y \ S\ assms(4) diff_gt_0_iff_gt inner_commute inner_diff_left inner_gt_zero_iff right_minus_eq) show "(y - z) \ y \ (y - z) \ x" if "x \ S" for x proof (rule ccontr) have *: "\u. 0 \ u \ u \ 1 \ dist z y \ dist z ((1 - u) *\<^sub>R y + u *\<^sub>R x)" using assms(1)[unfolded convex_alt] and y and \x\S\ and \y\S\ by auto assume "\ (y - z) \ y \ (y - z) \ x" then obtain v where "v > 0" "v \ 1" "dist (y + v *\<^sub>R (x - y)) z < dist y z" using closer_point_lemma[of z y x] by (auto simp: inner_diff) then show False using *[of v] by (auto simp: dist_commute algebra_simps) qed qed (use \y \ S\ in auto) qed lemma separating_hyperplane_closed_point: fixes z :: "'a::{real_inner,heine_borel}" assumes "convex S" and "closed S" and "z \ S" shows "\a b. inner a z < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True then show ?thesis by (simp add: gt_ex) next case False obtain y where "y \ S" and y: "\x. x \ S \ dist z y \ dist z x" by (metis distance_attains_inf[OF assms(2) False]) show ?thesis proof (intro exI conjI ballI) show "(y - z) \ z < inner (y - z) z + (norm (y - z))\<^sup>2 / 2" using \y\S\ \z\S\ by auto next fix x assume "x \ S" have "False" if *: "0 < inner (z - y) (x - y)" proof - obtain u where "u > 0" "u \ 1" "dist (y + u *\<^sub>R (x - y)) z < dist y z" using * closer_point_lemma by blast then show False using y[of "y + u *\<^sub>R (x - y)"] convexD_alt [OF \convex S\] using \x\S\ \y\S\ by (auto simp: dist_commute algebra_simps) qed moreover have "0 < (norm (y - z))\<^sup>2" using \y\S\ \z\S\ by auto then have "0 < inner (y - z) (y - z)" unfolding power2_norm_eq_inner by simp ultimately show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (force simp: field_simps power2_norm_eq_inner inner_commute inner_diff) qed qed lemma separating_hyperplane_closed_0: assumes "convex (S::('a::euclidean_space) set)" and "closed S" and "0 \ S" shows "\a b. a \ 0 \ 0 < b \ (\x\S. inner a x > b)" proof (cases "S = {}") case True have "(SOME i. i\Basis) \ (0::'a)" by (metis Basis_zero SOME_Basis) then show ?thesis using True zero_less_one by blast next case False then show ?thesis using False using separating_hyperplane_closed_point[OF assms] by (metis all_not_in_conv inner_zero_left inner_zero_right less_eq_real_def not_le) qed subsubsection\<^marker>\tag unimportant\ \Now set-to-set for closed/compact sets\ lemma separating_hyperplane_closed_compact: fixes S :: "'a::euclidean_space set" assumes "convex S" and "closed S" and "convex T" and "compact T" and "T \ {}" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof (cases "S = {}") case True obtain b where b: "b > 0" "\x\T. norm x \ b" using compact_imp_bounded[OF assms(4)] unfolding bounded_pos by auto obtain z :: 'a where z: "norm z = b + 1" using vector_choose_size[of "b + 1"] and b(1) by auto then have "z \ T" using b(2)[THEN bspec[where x=z]] by auto then obtain a b where ab: "inner a z < b" "\x\T. b < inner a x" using separating_hyperplane_closed_point[OF assms(3) compact_imp_closed[OF assms(4)], of z] by auto then show ?thesis using True by auto next case False then obtain y where "y \ S" by auto obtain a b where "0 < b" "\x \ (\x\ S. \y \ T. {x - y}). b < inner a x" using separating_hyperplane_closed_point[OF convex_differences[OF assms(1,3)], of 0] using closed_compact_differences[OF assms(2,4)] using assms(6) by auto then have ab: "\x\S. \y\T. b + inner a y < inner a x" apply - apply rule apply rule apply (erule_tac x="x - y" in ballE) apply (auto simp: inner_diff) done define k where "k = (SUP x\T. a \ x)" show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-(k + b / 2)" in exI) apply (intro conjI ballI) unfolding inner_minus_left and neg_less_iff_less proof - fix x assume "x \ T" then have "inner a x - b / 2 < k" unfolding k_def proof (subst less_cSUP_iff) show "T \ {}" by fact show "bdd_above ((\) a ` T)" using ab[rule_format, of y] \y \ S\ by (intro bdd_aboveI2[where M="inner a y - b"]) (auto simp: field_simps intro: less_imp_le) qed (auto intro!: bexI[of _ x] \0) then show "inner a x < k + b / 2" by auto next fix x assume "x \ S" then have "k \ inner a x - b" unfolding k_def apply (rule_tac cSUP_least) using assms(5) using ab[THEN bspec[where x=x]] apply auto done then show "k + b / 2 < inner a x" using \0 < b\ by auto qed qed lemma separating_hyperplane_compact_closed: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "S \ {}" and "convex T" and "closed T" and "S \ T = {}" shows "\a b. (\x\S. inner a x < b) \ (\x\T. inner a x > b)" proof - obtain a b where "(\x\T. inner a x < b) \ (\x\S. b < inner a x)" using separating_hyperplane_closed_compact[OF assms(4-5,1-2,3)] and assms(6) by auto then show ?thesis apply (rule_tac x="-a" in exI) apply (rule_tac x="-b" in exI, auto) done qed subsubsection\<^marker>\tag unimportant\ \General case without assuming closure and getting non-strict separation\ lemma separating_hyperplane_set_0: assumes "convex S" "(0::'a::euclidean_space) \ S" shows "\a. a \ 0 \ (\x\S. 0 \ inner a x)" proof - let ?k = "\c. {x::'a. 0 \ inner c x}" have *: "frontier (cball 0 1) \ \f \ {}" if as: "f \ ?k ` S" "finite f" for f proof - obtain c where c: "f = ?k ` c" "c \ S" "finite c" using finite_subset_image[OF as(2,1)] by auto then obtain a b where ab: "a \ 0" "0 < b" "\x\convex hull c. b < inner a x" using separating_hyperplane_closed_0[OF convex_convex_hull, of c] using finite_imp_compact_convex_hull[OF c(3), THEN compact_imp_closed] and assms(2) using subset_hull[of convex, OF assms(1), symmetric, of c] by force then have "\x. norm x = 1 \ (\y\c. 0 \ inner y x)" apply (rule_tac x = "inverse(norm a) *\<^sub>R a" in exI) using hull_subset[of c convex] unfolding subset_eq and inner_scaleR by (auto simp: inner_commute del: ballE elim!: ballE) then show "frontier (cball 0 1) \ \f \ {}" unfolding c(1) frontier_cball sphere_def dist_norm by auto qed have "frontier (cball 0 1) \ (\(?k ` S)) \ {}" apply (rule compact_imp_fip) apply (rule compact_frontier[OF compact_cball]) using * closed_halfspace_ge by auto then obtain x where "norm x = 1" "\y\S. x\?k y" unfolding frontier_cball dist_norm sphere_def by auto then show ?thesis by (metis inner_commute mem_Collect_eq norm_eq_zero zero_neq_one) qed lemma separating_hyperplane_sets: fixes S T :: "'a::euclidean_space set" assumes "convex S" and "convex T" and "S \ {}" and "T \ {}" and "S \ T = {}" shows "\a b. a \ 0 \ (\x\S. inner a x \ b) \ (\x\T. inner a x \ b)" proof - from separating_hyperplane_set_0[OF convex_differences[OF assms(2,1)]] obtain a where "a \ 0" "\x\{x - y |x y. x \ T \ y \ S}. 0 \ inner a x" using assms(3-5) by force then have *: "\x y. x \ T \ y \ S \ inner a y \ inner a x" by (force simp: inner_diff) then have bdd: "bdd_above (((\) a)`S)" using \T \ {}\ by (auto intro: bdd_aboveI2[OF *]) show ?thesis using \a\0\ by (intro exI[of _ a] exI[of _ "SUP x\S. a \ x"]) (auto intro!: cSUP_upper bdd cSUP_least \a \ 0\ \S \ {}\ *) qed subsection\<^marker>\tag unimportant\ \More convexity generalities\ lemma convex_closure [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (closure S)" apply (rule convexI) apply (unfold closure_sequential, elim exE) apply (rule_tac x="\n. u *\<^sub>R xa n + v *\<^sub>R xb n" in exI) apply (rule,rule) apply (rule convexD [OF assms]) apply (auto del: tendsto_const intro!: tendsto_intros) done lemma convex_interior [intro,simp]: fixes S :: "'a::real_normed_vector set" assumes "convex S" shows "convex (interior S)" unfolding convex_alt Ball_def mem_interior proof clarify fix x y u assume u: "0 \ u" "u \ (1::real)" fix e d assume ed: "ball x e \ S" "ball y d \ S" "0e>0. ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) e \ S" proof (intro exI conjI subsetI) fix z assume "z \ ball ((1 - u) *\<^sub>R x + u *\<^sub>R y) (min d e)" then have "(1- u) *\<^sub>R (z - u *\<^sub>R (y - x)) + u *\<^sub>R (z + (1 - u) *\<^sub>R (y - x)) \ S" apply (rule_tac assms[unfolded convex_alt, rule_format]) using ed(1,2) and u unfolding subset_eq mem_ball Ball_def dist_norm apply (auto simp: algebra_simps) done then show "z \ S" using u by (auto simp: algebra_simps) qed(insert u ed(3-4), auto) qed lemma convex_hull_eq_empty[simp]: "convex hull S = {} \ S = {}" using hull_subset[of S convex] convex_hull_empty by auto subsection\<^marker>\tag unimportant\ \Convex set as intersection of halfspaces\ lemma convex_halfspace_intersection: fixes s :: "('a::euclidean_space) set" assumes "closed s" "convex s" shows "s = \{h. s \ h \ (\a b. h = {x. inner a x \ b})}" apply (rule set_eqI, rule) unfolding Inter_iff Ball_def mem_Collect_eq apply (rule,rule,erule conjE) proof - fix x assume "\xa. s \ xa \ (\a b. xa = {x. inner a x \ b}) \ x \ xa" then have "\a b. s \ {x. inner a x \ b} \ x \ {x. inner a x \ b}" by blast then show "x \ s" apply (rule_tac ccontr) apply (drule separating_hyperplane_closed_point[OF assms(2,1)]) apply (erule exE)+ apply (erule_tac x="-a" in allE) apply (erule_tac x="-b" in allE, auto) done qed auto subsection\<^marker>\tag unimportant\ \Convexity of general and special intervals\ lemma is_interval_convex: fixes S :: "'a::euclidean_space set" assumes "is_interval S" shows "convex S" proof (rule convexI) fix x y and u v :: real assume as: "x \ S" "y \ S" "0 \ u" "0 \ v" "u + v = 1" then have *: "u = 1 - v" "1 - v \ 0" and **: "v = 1 - u" "1 - u \ 0" by auto { fix a b assume "\ b \ u * a + v * b" then have "u * a < (1 - v) * b" unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) *(2) apply (rule_tac mult_left_less_imp_less[of "1 - v"]) apply (auto simp: field_simps) done then have "a \ u * a + v * b" unfolding * using as(4) by (auto simp: field_simps intro!:mult_right_mono) } moreover { fix a b assume "\ u * a + v * b \ a" then have "v * b > (1 - u) * a" unfolding not_le using as(4) by (auto simp: field_simps) then have "a < b" unfolding * using as(4) apply (rule_tac mult_left_less_imp_less) apply (auto simp: field_simps) done then have "u * a + v * b \ b" unfolding ** using **(2) as(3) by (auto simp: field_simps intro!:mult_right_mono) } ultimately show "u *\<^sub>R x + v *\<^sub>R y \ S" apply - apply (rule assms[unfolded is_interval_def, rule_format, OF as(1,2)]) using as(3-) DIM_positive[where 'a='a] apply (auto simp: inner_simps) done qed lemma is_interval_connected: fixes S :: "'a::euclidean_space set" shows "is_interval S \ connected S" using is_interval_convex convex_connected by auto lemma convex_box [simp]: "convex (cbox a b)" "convex (box a (b::'a::euclidean_space))" apply (rule_tac[!] is_interval_convex)+ using is_interval_box is_interval_cbox apply auto done text\A non-singleton connected set is perfect (i.e. has no isolated points). \ lemma connected_imp_perfect: fixes a :: "'a::metric_space" assumes "connected S" "a \ S" and S: "\x. S \ {x}" shows "a islimpt S" proof - have False if "a \ T" "open T" "\y. \y \ S; y \ T\ \ y = a" for T proof - obtain e where "e > 0" and e: "cball a e \ T" using \open T\ \a \ T\ by (auto simp: open_contains_cball) have "openin (top_of_set S) {a}" unfolding openin_open using that \a \ S\ by blast moreover have "closedin (top_of_set S) {a}" by (simp add: assms) ultimately show "False" using \connected S\ connected_clopen S by blast qed then show ?thesis unfolding islimpt_def by blast qed lemma connected_imp_perfect_aff_dim: "\connected S; aff_dim S \ 0; a \ S\ \ a islimpt S" using aff_dim_sing connected_imp_perfect by blast subsection\<^marker>\tag unimportant\ \On \real\, \is_interval\, \convex\ and \connected\ are all equivalent\ lemma mem_is_interval_1_I: fixes a b c::real assumes "is_interval S" assumes "a \ S" "c \ S" assumes "a \ b" "b \ c" shows "b \ S" using assms is_interval_1 by blast lemma is_interval_connected_1: fixes s :: "real set" shows "is_interval s \ connected s" apply rule apply (rule is_interval_connected, assumption) unfolding is_interval_1 apply rule apply rule apply rule apply rule apply (erule conjE) apply (rule ccontr) proof - fix a b x assume as: "connected s" "a \ s" "b \ s" "a \ x" "x \ b" "x \ s" then have *: "a < x" "x < b" unfolding not_le [symmetric] by auto let ?halfl = "{.. s" with \x \ s\ have "x \ y" by auto then have "y \ ?halfr \ ?halfl" by auto } moreover have "a \ ?halfl" "b \ ?halfr" using * by auto then have "?halfl \ s \ {}" "?halfr \ s \ {}" using as(2-3) by auto ultimately show False apply (rule_tac notE[OF as(1)[unfolded connected_def]]) apply (rule_tac x = ?halfl in exI) apply (rule_tac x = ?halfr in exI, rule) apply (rule open_lessThan, rule) apply (rule open_greaterThan, auto) done qed lemma is_interval_convex_1: fixes s :: "real set" shows "is_interval s \ convex s" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_compact_interval_1: "connected S \ compact S \ (\a b. S = {a..b::real})" by (auto simp: is_interval_connected_1 [symmetric] is_interval_compact) lemma connected_convex_1: fixes s :: "real set" shows "connected s \ convex s" by (metis is_interval_convex convex_connected is_interval_connected_1) lemma connected_convex_1_gen: fixes s :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected s \ convex s" proof - obtain f:: "'a \ real" where linf: "linear f" and "inj f" using subspace_isomorphism[OF subspace_UNIV subspace_UNIV, where 'a='a and 'b=real] unfolding Euclidean_Space.dim_UNIV by (auto simp: assms) then have "f -` (f ` s) = s" by (simp add: inj_vimage_image_eq) then show ?thesis by (metis connected_convex_1 convex_linear_vimage linf convex_connected connected_linear_image) qed lemma [simp]: fixes r s::real shows is_interval_io: "is_interval {..\tag unimportant\ \Another intermediate value theorem formulation\ lemma ivt_increasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f a)\k \ y" "y \ (f b)\k" shows "\x\{a..b}. (f x)\k = y" proof - have "f a \ f ` cbox a b" "f b \ f ` cbox a b" apply (rule_tac[!] imageI) using assms(1) apply auto done then show ?thesis using connected_ivt_component[of "f ` cbox a b" "f a" "f b" k y] by (simp add: connected_continuous_image assms) qed lemma ivt_increasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f a\k \ y \ y \ f b\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_increasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) lemma ivt_decreasing_component_on_1: fixes f :: "real \ 'a::euclidean_space" assumes "a \ b" and "continuous_on {a..b} f" and "(f b)\k \ y" and "y \ (f a)\k" shows "\x\{a..b}. (f x)\k = y" apply (subst neg_equal_iff_equal[symmetric]) using ivt_increasing_component_on_1[of a b "\x. - f x" k "- y"] using assms using continuous_on_minus apply auto done lemma ivt_decreasing_component_1: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ \x\{a..b}. continuous (at x) f \ f b\k \ y \ y \ f a\k \ \x\{a..b}. (f x)\k = y" by (rule ivt_decreasing_component_on_1) (auto simp: continuous_at_imp_continuous_on) subsection\<^marker>\tag unimportant\ \A bound within an interval\ lemma convex_hull_eq_real_cbox: fixes x y :: real assumes "x \ y" shows "convex hull {x, y} = cbox x y" proof (rule hull_unique) show "{x, y} \ cbox x y" using \x \ y\ by auto show "convex (cbox x y)" by (rule convex_box) next fix S assume "{x, y} \ S" and "convex S" then show "cbox x y \ S" unfolding is_interval_convex_1 [symmetric] is_interval_def Basis_real_def by - (clarify, simp (no_asm_use), fast) qed lemma unit_interval_convex_hull: "cbox (0::'a::euclidean_space) One = convex hull {x. \i\Basis. (x\i = 0) \ (x\i = 1)}" (is "?int = convex hull ?points") proof - have One[simp]: "\i. i \ Basis \ One \ i = 1" by (simp add: inner_sum_left sum.If_cases inner_Basis) have "?int = {x. \i\Basis. x \ i \ cbox 0 1}" by (auto simp: cbox_def) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` cbox 0 1)" by (simp only: box_eq_set_sum_Basis) also have "\ = (\i\Basis. (\x. x *\<^sub>R i) ` (convex hull {0, 1}))" by (simp only: convex_hull_eq_real_cbox zero_le_one) also have "\ = (\i\Basis. convex hull ((\x. x *\<^sub>R i) ` {0, 1}))" by (simp add: convex_hull_linear_image) also have "\ = convex hull (\i\Basis. (\x. x *\<^sub>R i) ` {0, 1})" by (simp only: convex_hull_set_sum) also have "\ = convex hull {x. \i\Basis. x\i \ {0, 1}}" by (simp only: box_eq_set_sum_Basis) also have "convex hull {x. \i\Basis. x\i \ {0, 1}} = convex hull ?points" by simp finally show ?thesis . qed text \And this is a finite set of vertices.\ lemma unit_cube_convex_hull: obtains S :: "'a::euclidean_space set" where "finite S" and "cbox 0 (\Basis) = convex hull S" proof show "finite {x::'a. \i\Basis. x \ i = 0 \ x \ i = 1}" proof (rule finite_subset, clarify) show "finite ((\S. \i\Basis. (if i \ S then 1 else 0) *\<^sub>R i) ` Pow Basis)" using finite_Basis by blast fix x :: 'a assume as: "\i\Basis. x \ i = 0 \ x \ i = 1" show "x \ (\S. \i\Basis. (if i\S then 1 else 0) *\<^sub>R i) ` Pow Basis" apply (rule image_eqI[where x="{i. i\Basis \ x\i = 1}"]) using as apply (subst euclidean_eq_iff, auto) done qed show "cbox 0 One = convex hull {x. \i\Basis. x \ i = 0 \ x \ i = 1}" using unit_interval_convex_hull by blast qed text \Hence any cube (could do any nonempty interval).\ lemma cube_convex_hull: assumes "d > 0" obtains S :: "'a::euclidean_space set" where "finite S" and "cbox (x - (\i\Basis. d*\<^sub>Ri)) (x + (\i\Basis. d*\<^sub>Ri)) = convex hull S" proof - let ?d = "(\i\Basis. d *\<^sub>R i)::'a" have *: "cbox (x - ?d) (x + ?d) = (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 (\Basis)" proof (intro set_eqI iffI) fix y assume "y \ cbox (x - ?d) (x + ?d)" then have "inverse (2 * d) *\<^sub>R (y - (x - ?d)) \ cbox 0 (\Basis)" using assms by (simp add: mem_box inner_simps) (simp add: field_simps) with \0 < d\ show "y \ (\y. x - sum ((*\<^sub>R) d) Basis + (2 * d) *\<^sub>R y) ` cbox 0 One" by (auto intro: image_eqI[where x= "inverse (2 * d) *\<^sub>R (y - (x - ?d))"]) next fix y assume "y \ (\y. x - ?d + (2 * d) *\<^sub>R y) ` cbox 0 One" then obtain z where z: "z \ cbox 0 One" "y = x - ?d + (2*d) *\<^sub>R z" by auto then show "y \ cbox (x - ?d) (x + ?d)" using z assms by (auto simp: mem_box inner_simps) qed obtain S where "finite S" "cbox 0 (\Basis::'a) = convex hull S" using unit_cube_convex_hull by auto then show ?thesis by (rule_tac that[of "(\y. x - ?d + (2 * d) *\<^sub>R y)` S"]) (auto simp: convex_hull_affinity *) qed subsection\<^marker>\tag unimportant\\Representation of any interval as a finite convex hull\ lemma image_stretch_interval: "(\x. \k\Basis. (m k * (x\k)) *\<^sub>R k) ` cbox a (b::'a::euclidean_space) = (if (cbox a b) = {} then {} else cbox (\k\Basis. (min (m k * (a\k)) (m k * (b\k))) *\<^sub>R k::'a) (\k\Basis. (max (m k * (a\k)) (m k * (b\k))) *\<^sub>R k))" proof cases assume *: "cbox a b \ {}" show ?thesis unfolding box_ne_empty if_not_P[OF *] apply (simp add: cbox_def image_Collect set_eq_iff euclidean_eq_iff[where 'a='a] ball_conj_distrib[symmetric]) apply (subst choice_Basis_iff[symmetric]) proof (intro allI ball_cong refl) fix x i :: 'a assume "i \ Basis" with * have a_le_b: "a \ i \ b \ i" unfolding box_ne_empty by auto show "(\xa. x \ i = m i * xa \ a \ i \ xa \ xa \ b \ i) \ min (m i * (a \ i)) (m i * (b \ i)) \ x \ i \ x \ i \ max (m i * (a \ i)) (m i * (b \ i))" proof (cases "m i = 0") case True with a_le_b show ?thesis by auto next case False then have *: "\a b. a = m i * b \ b = a / m i" by (auto simp: field_simps) from False have "min (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (a \ i) else m i * (b \ i))" "max (m i * (a \ i)) (m i * (b \ i)) = (if 0 < m i then m i * (b \ i) else m i * (a \ i))" using a_le_b by (auto simp: min_def max_def mult_le_cancel_left) with False show ?thesis using a_le_b unfolding * by (auto simp: le_divide_eq divide_le_eq ac_simps) qed qed qed simp lemma interval_image_stretch_interval: "\u v. (\x. \k\Basis. (m k * (x\k))*\<^sub>R k) ` cbox a (b::'a::euclidean_space) = cbox u (v::'a::euclidean_space)" unfolding image_stretch_interval by auto lemma cbox_translation: "cbox (c + a) (c + b) = image (\x. c + x) (cbox a b)" using image_affinity_cbox [of 1 c a b] using box_ne_empty [of "a+c" "b+c"] box_ne_empty [of a b] by (auto simp: inner_left_distrib add.commute) lemma cbox_image_unit_interval: fixes a :: "'a::euclidean_space" assumes "cbox a b \ {}" shows "cbox a b = (+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` cbox 0 One" using assms apply (simp add: box_ne_empty image_stretch_interval cbox_translation [symmetric]) apply (simp add: min_def max_def algebra_simps sum_subtractf euclidean_representation) done lemma closed_interval_as_convex_hull: fixes a :: "'a::euclidean_space" obtains S where "finite S" "cbox a b = convex hull S" proof (cases "cbox a b = {}") case True with convex_hull_empty that show ?thesis by blast next case False obtain S::"'a set" where "finite S" and eq: "cbox 0 One = convex hull S" by (blast intro: unit_cube_convex_hull) have lin: "linear (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k)" by (rule linear_compose_sum) (auto simp: algebra_simps linearI) have "finite ((+) a ` (\x. \k\Basis. ((b \ k - a \ k) * (x \ k)) *\<^sub>R k) ` S)" by (rule finite_imageI \finite S\)+ then show ?thesis apply (rule that) apply (simp add: convex_hull_translation convex_hull_linear_image [OF lin, symmetric]) apply (simp add: eq [symmetric] cbox_image_unit_interval [OF False]) done qed subsection\<^marker>\tag unimportant\ \Bounded convex function on open set is continuous\ lemma convex_on_bounded_continuous: fixes S :: "('a::real_normed_vector) set" assumes "open S" and "convex_on S f" and "\x\S. \f x\ \ b" shows "continuous_on S f" apply (rule continuous_at_imp_continuous_on) unfolding continuous_at_real_range proof (rule,rule,rule) fix x and e :: real assume "x \ S" "e > 0" define B where "B = \b\ + 1" then have B: "0 < B""\x. x\S \ \f x\ \ B" using assms(3) by auto obtain k where "k > 0" and k: "cball x k \ S" using \x \ S\ assms(1) open_contains_cball_eq by blast show "\d>0. \x'. norm (x' - x) < d \ \f x' - f x\ < e" proof (intro exI conjI allI impI) fix y assume as: "norm (y - x) < min (k / 2) (e / (2 * B) * k)" show "\f y - f x\ < e" proof (cases "y = x") case False define t where "t = k / norm (y - x)" have "2 < t" "0k>0\ by (auto simp:field_simps) have "y \ S" apply (rule k[THEN subsetD]) unfolding mem_cball dist_norm apply (rule order_trans[of _ "2 * norm (x - y)"]) using as by (auto simp: field_simps norm_minus_commute) { define w where "w = x + t *\<^sub>R (y - x)" have "w \ S" using \k>0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / t) *\<^sub>R x + - x + ((t - 1) / t) *\<^sub>R x = (1 / t - 1 + (t - 1) / t) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = 0" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / t) *\<^sub>R w + ((t - 1) / t) *\<^sub>R x = y" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have 2: "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) have "f y - f x \ (f w - f x) / t" using assms(2)[unfolded convex_on_def,rule_format,of w x "1/t" "(t - 1)/t", unfolded w] using \0 < t\ \2 < t\ and \x \ S\ \w \ S\ by (auto simp:field_simps) also have "... < e" using B(2)[OF \w\S\] and B(2)[OF \x\S\] 2 \t > 0\ by (auto simp: field_simps) finally have th1: "f y - f x < e" . } moreover { define w where "w = x - t *\<^sub>R (y - x)" have "w \ S" using \k > 0\ by (auto simp: dist_norm t_def w_def k[THEN subsetD]) have "(1 / (1 + t)) *\<^sub>R x + (t / (1 + t)) *\<^sub>R x = (1 / (1 + t) + t / (1 + t)) *\<^sub>R x" by (auto simp: algebra_simps) also have "\ = x" using \t > 0\ by (auto simp:field_simps) finally have w: "(1 / (1+t)) *\<^sub>R w + (t / (1 + t)) *\<^sub>R y = x" unfolding w_def using False and \t > 0\ by (auto simp: algebra_simps) have "2 * B < e * t" unfolding t_def using \0 < e\ \0 < k\ \B > 0\ and as and False by (auto simp:field_simps) then have *: "(f w - f y) / t < e" using B(2)[OF \w\S\] and B(2)[OF \y\S\] using \t > 0\ by (auto simp:field_simps) have "f x \ 1 / (1 + t) * f w + (t / (1 + t)) * f y" using assms(2)[unfolded convex_on_def,rule_format,of w y "1/(1+t)" "t / (1+t)",unfolded w] using \0 < t\ \2 < t\ and \y \ S\ \w \ S\ by (auto simp:field_simps) also have "\ = (f w + t * f y) / (1 + t)" using \t > 0\ by (simp add: add_divide_distrib) also have "\ < e + f y" using \t > 0\ * \e > 0\ by (auto simp: field_simps) finally have "f x - f y < e" by auto } ultimately show ?thesis by auto qed (insert \0, auto) qed (insert \0 \0 \0, auto simp: field_simps) qed subsection\<^marker>\tag unimportant\ \Upper bound on a ball implies upper and lower bounds\ lemma convex_bounds_lemma: fixes x :: "'a::real_normed_vector" assumes "convex_on (cball x e) f" and "\y \ cball x e. f y \ b" shows "\y \ cball x e. \f y\ \ b + 2 * \f x\" apply rule proof (cases "0 \ e") case True fix y assume y: "y \ cball x e" define z where "z = 2 *\<^sub>R x - y" have *: "x - (2 *\<^sub>R x - y) = y - x" by (simp add: scaleR_2) have z: "z \ cball x e" using y unfolding z_def mem_cball dist_norm * by (auto simp: norm_minus_commute) have "(1 / 2) *\<^sub>R y + (1 / 2) *\<^sub>R z = x" unfolding z_def by (auto simp: algebra_simps) then show "\f y\ \ b + 2 * \f x\" using assms(1)[unfolded convex_on_def,rule_format, OF y z, of "1/2" "1/2"] using assms(2)[rule_format,OF y] assms(2)[rule_format,OF z] by (auto simp:field_simps) next case False fix y assume "y \ cball x e" then have "dist x y < 0" using False unfolding mem_cball not_le by (auto simp del: dist_not_less_zero) then show "\f y\ \ b + 2 * \f x\" using zero_le_dist[of x y] by auto qed subsubsection\<^marker>\tag unimportant\ \Hence a convex function on an open set is continuous\ lemma real_of_nat_ge_one_iff: "1 \ real (n::nat) \ 1 \ n" by auto lemma convex_on_continuous: assumes "open (s::('a::euclidean_space) set)" "convex_on s f" shows "continuous_on s f" unfolding continuous_on_eq_continuous_at[OF assms(1)] proof note dimge1 = DIM_positive[where 'a='a] fix x assume "x \ s" then obtain e where e: "cball x e \ s" "e > 0" using assms(1) unfolding open_contains_cball by auto define d where "d = e / real DIM('a)" have "0 < d" unfolding d_def using \e > 0\ dimge1 by auto let ?d = "(\i\Basis. d *\<^sub>R i)::'a" obtain c where c: "finite c" and c1: "convex hull c \ cball x e" and c2: "cball x d \ convex hull c" proof define c where "c = (\i\Basis. (\a. a *\<^sub>R i) ` {x\i - d, x\i + d})" show "finite c" unfolding c_def by (simp add: finite_set_sum) have 1: "convex hull c = {a. \i\Basis. a \ i \ cbox (x \ i - d) (x \ i + d)}" unfolding box_eq_set_sum_Basis unfolding c_def convex_hull_set_sum apply (subst convex_hull_linear_image [symmetric]) apply (simp add: linear_iff scaleR_add_left) apply (rule sum.cong [OF refl]) apply (rule image_cong [OF _ refl]) apply (rule convex_hull_eq_real_cbox) apply (cut_tac \0 < d\, simp) done then have 2: "convex hull c = {a. \i\Basis. a \ i \ cball (x \ i) d}" by (simp add: dist_norm abs_le_iff algebra_simps) show "cball x d \ convex hull c" unfolding 2 by (clarsimp simp: dist_norm) (metis inner_commute inner_diff_right norm_bound_Basis_le) have e': "e = (\(i::'a)\Basis. d)" - by (simp add: d_def DIM_positive) + by (simp add: d_def) show "convex hull c \ cball x e" unfolding 2 apply clarsimp apply (subst euclidean_dist_l2) apply (rule order_trans [OF L2_set_le_sum]) apply (rule zero_le_dist) unfolding e' apply (rule sum_mono, simp) done qed define k where "k = Max (f ` c)" have "convex_on (convex hull c) f" apply(rule convex_on_subset[OF assms(2)]) apply(rule subset_trans[OF c1 e(1)]) done then have k: "\y\convex hull c. f y \ k" apply (rule_tac convex_on_convex_hull_bound, assumption) by (simp add: k_def c) have "e \ e * real DIM('a)" using e(2) real_of_nat_ge_one_iff by auto then have "d \ e" by (simp add: d_def field_split_simps) then have dsube: "cball x d \ cball x e" by (rule subset_cball) have conv: "convex_on (cball x d) f" using \convex_on (convex hull c) f\ c2 convex_on_subset by blast then have "\y\cball x d. \f y\ \ k + 2 * \f x\" by (rule convex_bounds_lemma) (use c2 k in blast) then have "continuous_on (ball x d) f" apply (rule_tac convex_on_bounded_continuous) apply (rule open_ball, rule convex_on_subset[OF conv]) apply (rule ball_subset_cball, force) done then show "continuous (at x) f" unfolding continuous_on_eq_continuous_at[OF open_ball] using \d > 0\ by auto qed end diff --git a/src/HOL/Analysis/Extended_Real_Limits.thy b/src/HOL/Analysis/Extended_Real_Limits.thy --- a/src/HOL/Analysis/Extended_Real_Limits.thy +++ b/src/HOL/Analysis/Extended_Real_Limits.thy @@ -1,1782 +1,1782 @@ (* Title: HOL/Analysis/Extended_Real_Limits.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Author: Armin Heller, TU München Author: Bogdan Grechuk, University of Edinburgh *) section \Limits on the Extended Real Number Line\ (* TO FIX: perhaps put all Nonstandard Analysis related topics together? *) theory Extended_Real_Limits imports Topology_Euclidean_Space "HOL-Library.Extended_Real" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Indicator_Function" begin lemma compact_UNIV: "compact (UNIV :: 'a::{complete_linorder,linorder_topology,second_countable_topology} set)" using compact_complete_linorder by (auto simp: seq_compact_eq_compact[symmetric] seq_compact_def) lemma compact_eq_closed: fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set" shows "compact S \ closed S" using closed_Int_compact[of S, OF _ compact_UNIV] compact_imp_closed by auto lemma closed_contains_Sup_cl: fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set" assumes "closed S" and "S \ {}" shows "Sup S \ S" proof - from compact_eq_closed[of S] compact_attains_sup[of S] assms obtain s where S: "s \ S" "\t\S. t \ s" by auto then have "Sup S = s" by (auto intro!: Sup_eqI) with S show ?thesis by simp qed lemma closed_contains_Inf_cl: fixes S :: "'a::{complete_linorder,linorder_topology,second_countable_topology} set" assumes "closed S" and "S \ {}" shows "Inf S \ S" proof - from compact_eq_closed[of S] compact_attains_inf[of S] assms obtain s where S: "s \ S" "\t\S. s \ t" by auto then have "Inf S = s" by (auto intro!: Inf_eqI) with S show ?thesis by simp qed instance\<^marker>\tag unimportant\ enat :: second_countable_topology proof show "\B::enat set set. countable B \ open = generate_topology B" proof (intro exI conjI) show "countable (range lessThan \ range greaterThan::enat set set)" by auto qed (simp add: open_enat_def) qed instance\<^marker>\tag unimportant\ ereal :: second_countable_topology proof (standard, intro exI conjI) let ?B = "(\r\\. {{..< r}, {r <..}} :: ereal set set)" show "countable ?B" by (auto intro: countable_rat) show "open = generate_topology ?B" proof (intro ext iffI) fix S :: "ereal set" assume "open S" then show "generate_topology ?B S" unfolding open_generated_order proof induct case (Basis b) then obtain e where "b = {.. b = {e<..}" by auto moreover have "{..{{.. \ \ x < e}" "{e<..} = \{{x<..}|x. x \ \ \ e < x}" by (auto dest: ereal_dense3 simp del: ex_simps simp add: ex_simps[symmetric] conj_commute Rats_def image_iff) ultimately show ?case by (auto intro: generate_topology.intros) qed (auto intro: generate_topology.intros) next fix S assume "generate_topology ?B S" then show "open S" by induct auto qed qed text \This is a copy from \ereal :: second_countable_topology\. Maybe find a common super class of topological spaces where the rational numbers are densely embedded ?\ instance ennreal :: second_countable_topology proof (standard, intro exI conjI) let ?B = "(\r\\. {{..< r}, {r <..}} :: ennreal set set)" show "countable ?B" by (auto intro: countable_rat) show "open = generate_topology ?B" proof (intro ext iffI) fix S :: "ennreal set" assume "open S" then show "generate_topology ?B S" unfolding open_generated_order proof induct case (Basis b) then obtain e where "b = {.. b = {e<..}" by auto moreover have "{..{{.. \ \ x < e}" "{e<..} = \{{x<..}|x. x \ \ \ e < x}" by (auto dest: ennreal_rat_dense simp del: ex_simps simp add: ex_simps[symmetric] conj_commute Rats_def image_iff) ultimately show ?case by (auto intro: generate_topology.intros) qed (auto intro: generate_topology.intros) next fix S assume "generate_topology ?B S" then show "open S" by induct auto qed qed lemma ereal_open_closed_aux: fixes S :: "ereal set" assumes "open S" and "closed S" and S: "(-\) \ S" shows "S = {}" proof (rule ccontr) assume "\ ?thesis" then have *: "Inf S \ S" by (metis assms(2) closed_contains_Inf_cl) { assume "Inf S = -\" then have False using * assms(3) by auto } moreover { assume "Inf S = \" then have "S = {\}" by (metis Inf_eq_PInfty \S \ {}\) then have False by (metis assms(1) not_open_singleton) } moreover { assume fin: "\Inf S\ \ \" from ereal_open_cont_interval[OF assms(1) * fin] obtain e where e: "e > 0" "{Inf S - e<.. S" . then obtain b where b: "Inf S - e < b" "b < Inf S" using fin ereal_between[of "Inf S" e] dense[of "Inf S - e"] by auto then have "b \ {Inf S - e <..< Inf S + e}" using e fin ereal_between[of "Inf S" e] by auto then have "b \ S" using e by auto then have False using b by (metis complete_lattice_class.Inf_lower leD) } ultimately show False by auto qed lemma ereal_open_closed: fixes S :: "ereal set" shows "open S \ closed S \ S = {} \ S = UNIV" proof - { assume lhs: "open S \ closed S" { assume "-\ \ S" then have "S = {}" using lhs ereal_open_closed_aux by auto } moreover { assume "-\ \ S" then have "- S = {}" using lhs ereal_open_closed_aux[of "-S"] by auto } ultimately have "S = {} \ S = UNIV" by auto } then show ?thesis by auto qed lemma ereal_open_atLeast: fixes x :: ereal shows "open {x..} \ x = -\" proof assume "x = -\" then have "{x..} = UNIV" by auto then show "open {x..}" by auto next assume "open {x..}" then have "open {x..} \ closed {x..}" by auto then have "{x..} = UNIV" unfolding ereal_open_closed by auto then show "x = -\" by (simp add: bot_ereal_def atLeast_eq_UNIV_iff) qed lemma mono_closed_real: fixes S :: "real set" assumes mono: "\y z. y \ S \ y \ z \ z \ S" and "closed S" shows "S = {} \ S = UNIV \ (\a. S = {a..})" proof - { assume "S \ {}" { assume ex: "\B. \x\S. B \ x" then have *: "\x\S. Inf S \ x" using cInf_lower[of _ S] ex by (metis bdd_below_def) then have "Inf S \ S" apply (subst closed_contains_Inf) using ex \S \ {}\ \closed S\ apply auto done then have "\x. Inf S \ x \ x \ S" using mono[rule_format, of "Inf S"] * by auto then have "S = {Inf S ..}" by auto then have "\a. S = {a ..}" by auto } moreover { assume "\ (\B. \x\S. B \ x)" then have nex: "\B. \x\S. x < B" by (simp add: not_le) { fix y obtain x where "x\S" and "x < y" using nex by auto then have "y \ S" using mono[rule_format, of x y] by auto } then have "S = UNIV" by auto } ultimately have "S = UNIV \ (\a. S = {a ..})" by blast } then show ?thesis by blast qed lemma mono_closed_ereal: fixes S :: "real set" assumes mono: "\y z. y \ S \ y \ z \ z \ S" and "closed S" shows "\a. S = {x. a \ ereal x}" proof - { assume "S = {}" then have ?thesis apply (rule_tac x=PInfty in exI) apply auto done } moreover { assume "S = UNIV" then have ?thesis apply (rule_tac x="-\" in exI) apply auto done } moreover { assume "\a. S = {a ..}" then obtain a where "S = {a ..}" by auto then have ?thesis apply (rule_tac x="ereal a" in exI) apply auto done } ultimately show ?thesis using mono_closed_real[of S] assms by auto qed lemma Liminf_within: fixes f :: "'a::metric_space \ 'b::complete_lattice" shows "Liminf (at x within S) f = (SUP e\{0<..}. INF y\(S \ ball x e - {x}). f y)" unfolding Liminf_def eventually_at proof (rule SUP_eq, simp_all add: Ball_def Bex_def, safe) fix P d assume "0 < d" and "\y. y \ S \ y \ x \ dist y x < d \ P y" then have "S \ ball x d - {x} \ {x. P x}" by (auto simp: zero_less_dist_iff dist_commute) then show "\r>0. Inf (f ` (Collect P)) \ Inf (f ` (S \ ball x r - {x}))" by (intro exI[of _ d] INF_mono conjI \0 < d\) auto next fix d :: real assume "0 < d" then show "\P. (\d>0. \xa. xa \ S \ xa \ x \ dist xa x < d \ P xa) \ Inf (f ` (S \ ball x d - {x})) \ Inf (f ` (Collect P))" by (intro exI[of _ "\y. y \ S \ ball x d - {x}"]) (auto intro!: INF_mono exI[of _ d] simp: dist_commute) qed lemma Limsup_within: fixes f :: "'a::metric_space \ 'b::complete_lattice" shows "Limsup (at x within S) f = (INF e\{0<..}. SUP y\(S \ ball x e - {x}). f y)" unfolding Limsup_def eventually_at proof (rule INF_eq, simp_all add: Ball_def Bex_def, safe) fix P d assume "0 < d" and "\y. y \ S \ y \ x \ dist y x < d \ P y" then have "S \ ball x d - {x} \ {x. P x}" by (auto simp: zero_less_dist_iff dist_commute) then show "\r>0. Sup (f ` (S \ ball x r - {x})) \ Sup (f ` (Collect P))" by (intro exI[of _ d] SUP_mono conjI \0 < d\) auto next fix d :: real assume "0 < d" then show "\P. (\d>0. \xa. xa \ S \ xa \ x \ dist xa x < d \ P xa) \ Sup (f ` (Collect P)) \ Sup (f ` (S \ ball x d - {x}))" by (intro exI[of _ "\y. y \ S \ ball x d - {x}"]) (auto intro!: SUP_mono exI[of _ d] simp: dist_commute) qed lemma Liminf_at: fixes f :: "'a::metric_space \ 'b::complete_lattice" shows "Liminf (at x) f = (SUP e\{0<..}. INF y\(ball x e - {x}). f y)" using Liminf_within[of x UNIV f] by simp lemma Limsup_at: fixes f :: "'a::metric_space \ 'b::complete_lattice" shows "Limsup (at x) f = (INF e\{0<..}. SUP y\(ball x e - {x}). f y)" using Limsup_within[of x UNIV f] by simp lemma min_Liminf_at: fixes f :: "'a::metric_space \ 'b::complete_linorder" shows "min (f x) (Liminf (at x) f) = (SUP e\{0<..}. INF y\ball x e. f y)" apply (simp add: inf_min [symmetric] Liminf_at) apply (subst inf_commute) apply (subst SUP_inf) apply auto apply (metis (no_types, lifting) INF_insert centre_in_ball greaterThan_iff image_cong inf_commute insert_Diff) done subsection \Extended-Real.thy\ (*FIX ME change title *) lemma sum_constant_ereal: fixes a::ereal shows "(\i\I. a) = a * card I" apply (cases "finite I", induct set: finite, simp_all) apply (cases a, auto, metis (no_types, hide_lams) add.commute mult.commute semiring_normalization_rules(3)) done lemma real_lim_then_eventually_real: assumes "(u \ ereal l) F" shows "eventually (\n. u n = ereal(real_of_ereal(u n))) F" proof - have "ereal l \ {-\<..<(\::ereal)}" by simp moreover have "open {-\<..<(\::ereal)}" by simp ultimately have "eventually (\n. u n \ {-\<..<(\::ereal)}) F" using assms tendsto_def by blast moreover have "\x. x \ {-\<..<(\::ereal)} \ x = ereal(real_of_ereal x)" using ereal_real by auto ultimately show ?thesis by (metis (mono_tags, lifting) eventually_mono) qed lemma ereal_Inf_cmult: assumes "c>(0::real)" shows "Inf {ereal c * x |x. P x} = ereal c * Inf {x. P x}" proof - have "(\x::ereal. c * x) (Inf {x::ereal. P x}) = Inf ((\x::ereal. c * x)`{x::ereal. P x})" apply (rule mono_bij_Inf) apply (simp add: assms ereal_mult_left_mono less_imp_le mono_def) apply (rule bij_betw_byWitness[of _ "\x. (x::ereal) / c"], auto simp add: assms ereal_mult_divide) using assms ereal_divide_eq apply auto done then show ?thesis by (simp only: setcompr_eq_image[symmetric]) qed subsubsection\<^marker>\tag important\ \Continuity of addition\ text \The next few lemmas remove an unnecessary assumption in \tendsto_add_ereal\, culminating in \tendsto_add_ereal_general\ which essentially says that the addition is continuous on ereal times ereal, except at \(-\, \)\ and \(\, -\)\. It is much more convenient in many situations, see for instance the proof of \tendsto_sum_ereal\ below.\ lemma tendsto_add_ereal_PInf: fixes y :: ereal assumes y: "y \ -\" assumes f: "(f \ \) F" and g: "(g \ y) F" shows "((\x. f x + g x) \ \) F" proof - have "\C. eventually (\x. g x > ereal C) F" proof (cases y) case (real r) have "y > y-1" using y real by (simp add: ereal_between(1)) then have "eventually (\x. g x > y - 1) F" using g y order_tendsto_iff by auto moreover have "y-1 = ereal(real_of_ereal(y-1))" by (metis real ereal_eq_1(1) ereal_minus(1) real_of_ereal.simps(1)) ultimately have "eventually (\x. g x > ereal(real_of_ereal(y - 1))) F" by simp then show ?thesis by auto next case (PInf) have "eventually (\x. g x > ereal 0) F" using g PInf by (simp add: tendsto_PInfty) then show ?thesis by auto qed (simp add: y) then obtain C::real where ge: "eventually (\x. g x > ereal C) F" by auto { fix M::real have "eventually (\x. f x > ereal(M - C)) F" using f by (simp add: tendsto_PInfty) then have "eventually (\x. (f x > ereal (M-C)) \ (g x > ereal C)) F" by (auto simp add: ge eventually_conj_iff) moreover have "\x. ((f x > ereal (M-C)) \ (g x > ereal C)) \ (f x + g x > ereal M)" using ereal_add_strict_mono2 by fastforce ultimately have "eventually (\x. f x + g x > ereal M) F" using eventually_mono by force } then show ?thesis by (simp add: tendsto_PInfty) qed text\One would like to deduce the next lemma from the previous one, but the fact that \- (x + y)\ is in general different from \(- x) + (- y)\ in ereal creates difficulties, so it is more efficient to copy the previous proof.\ lemma tendsto_add_ereal_MInf: fixes y :: ereal assumes y: "y \ \" assumes f: "(f \ -\) F" and g: "(g \ y) F" shows "((\x. f x + g x) \ -\) F" proof - have "\C. eventually (\x. g x < ereal C) F" proof (cases y) case (real r) have "y < y+1" using y real by (simp add: ereal_between(1)) then have "eventually (\x. g x < y + 1) F" using g y order_tendsto_iff by force moreover have "y+1 = ereal(real_of_ereal (y+1))" by (simp add: real) ultimately have "eventually (\x. g x < ereal(real_of_ereal(y + 1))) F" by simp then show ?thesis by auto next case (MInf) have "eventually (\x. g x < ereal 0) F" using g MInf by (simp add: tendsto_MInfty) then show ?thesis by auto qed (simp add: y) then obtain C::real where ge: "eventually (\x. g x < ereal C) F" by auto { fix M::real have "eventually (\x. f x < ereal(M - C)) F" using f by (simp add: tendsto_MInfty) then have "eventually (\x. (f x < ereal (M- C)) \ (g x < ereal C)) F" by (auto simp add: ge eventually_conj_iff) moreover have "\x. ((f x < ereal (M-C)) \ (g x < ereal C)) \ (f x + g x < ereal M)" using ereal_add_strict_mono2 by fastforce ultimately have "eventually (\x. f x + g x < ereal M) F" using eventually_mono by force } then show ?thesis by (simp add: tendsto_MInfty) qed lemma tendsto_add_ereal_general1: fixes x y :: ereal assumes y: "\y\ \ \" assumes f: "(f \ x) F" and g: "(g \ y) F" shows "((\x. f x + g x) \ x + y) F" proof (cases x) case (real r) have a: "\x\ \ \" by (simp add: real) show ?thesis by (rule tendsto_add_ereal[OF a, OF y, OF f, OF g]) next case PInf then show ?thesis using tendsto_add_ereal_PInf assms by force next case MInf then show ?thesis using tendsto_add_ereal_MInf assms by (metis abs_ereal.simps(3) ereal_MInfty_eq_plus) qed lemma tendsto_add_ereal_general2: fixes x y :: ereal assumes x: "\x\ \ \" and f: "(f \ x) F" and g: "(g \ y) F" shows "((\x. f x + g x) \ x + y) F" proof - have "((\x. g x + f x) \ x + y) F" using tendsto_add_ereal_general1[OF x, OF g, OF f] add.commute[of "y", of "x"] by simp moreover have "\x. g x + f x = f x + g x" using add.commute by auto ultimately show ?thesis by simp qed text \The next lemma says that the addition is continuous on \ereal\, except at the pairs \(-\, \)\ and \(\, -\)\.\ lemma tendsto_add_ereal_general [tendsto_intros]: fixes x y :: ereal assumes "\((x=\ \ y=-\) \ (x=-\ \ y=\))" and f: "(f \ x) F" and g: "(g \ y) F" shows "((\x. f x + g x) \ x + y) F" proof (cases x) case (real r) show ?thesis apply (rule tendsto_add_ereal_general2) using real assms by auto next case (PInf) then have "y \ -\" using assms by simp then show ?thesis using tendsto_add_ereal_PInf PInf assms by auto next case (MInf) then have "y \ \" using assms by simp then show ?thesis using tendsto_add_ereal_MInf MInf f g by (metis ereal_MInfty_eq_plus) qed subsubsection\<^marker>\tag important\ \Continuity of multiplication\ text \In the same way as for addition, we prove that the multiplication is continuous on ereal times ereal, except at \(\, 0)\ and \(-\, 0)\ and \(0, \)\ and \(0, -\)\, starting with specific situations.\ lemma tendsto_mult_real_ereal: assumes "(u \ ereal l) F" "(v \ ereal m) F" shows "((\n. u n * v n) \ ereal l * ereal m) F" proof - have ureal: "eventually (\n. u n = ereal(real_of_ereal(u n))) F" by (rule real_lim_then_eventually_real[OF assms(1)]) then have "((\n. ereal(real_of_ereal(u n))) \ ereal l) F" using assms by auto then have limu: "((\n. real_of_ereal(u n)) \ l) F" by auto have vreal: "eventually (\n. v n = ereal(real_of_ereal(v n))) F" by (rule real_lim_then_eventually_real[OF assms(2)]) then have "((\n. ereal(real_of_ereal(v n))) \ ereal m) F" using assms by auto then have limv: "((\n. real_of_ereal(v n)) \ m) F" by auto { fix n assume "u n = ereal(real_of_ereal(u n))" "v n = ereal(real_of_ereal(v n))" then have "ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n" by (metis times_ereal.simps(1)) } then have *: "eventually (\n. ereal(real_of_ereal(u n) * real_of_ereal(v n)) = u n * v n) F" using eventually_elim2[OF ureal vreal] by auto have "((\n. real_of_ereal(u n) * real_of_ereal(v n)) \ l * m) F" using tendsto_mult[OF limu limv] by auto then have "((\n. ereal(real_of_ereal(u n)) * real_of_ereal(v n)) \ ereal(l * m)) F" by auto then show ?thesis using * filterlim_cong by fastforce qed lemma tendsto_mult_ereal_PInf: fixes f g::"_ \ ereal" assumes "(f \ l) F" "l>0" "(g \ \) F" shows "((\x. f x * g x) \ \) F" proof - obtain a::real where "0 < ereal a" "a < l" using assms(2) using ereal_dense2 by blast have *: "eventually (\x. f x > a) F" using \a < l\ assms(1) by (simp add: order_tendsto_iff) { fix K::real define M where "M = max K 1" then have "M > 0" by simp then have "ereal(M/a) > 0" using \ereal a > 0\ by simp then have "\x. ((f x > a) \ (g x > M/a)) \ (f x * g x > ereal a * ereal(M/a))" using ereal_mult_mono_strict'[where ?c = "M/a", OF \0 < ereal a\] by auto moreover have "ereal a * ereal(M/a) = M" using \ereal a > 0\ by simp ultimately have "\x. ((f x > a) \ (g x > M/a)) \ (f x * g x > M)" by simp moreover have "M \ K" unfolding M_def by simp ultimately have Imp: "\x. ((f x > a) \ (g x > M/a)) \ (f x * g x > K)" using ereal_less_eq(3) le_less_trans by blast have "eventually (\x. g x > M/a) F" using assms(3) by (simp add: tendsto_PInfty) then have "eventually (\x. (f x > a) \ (g x > M/a)) F" using * by (auto simp add: eventually_conj_iff) then have "eventually (\x. f x * g x > K) F" using eventually_mono Imp by force } then show ?thesis by (auto simp add: tendsto_PInfty) qed lemma tendsto_mult_ereal_pos: fixes f g::"_ \ ereal" assumes "(f \ l) F" "(g \ m) F" "l>0" "m>0" shows "((\x. f x * g x) \ l * m) F" proof (cases) assume *: "l = \ \ m = \" then show ?thesis proof (cases) assume "m = \" then show ?thesis using tendsto_mult_ereal_PInf assms by auto next assume "\(m = \)" then have "l = \" using * by simp then have "((\x. g x * f x) \ l * m) F" using tendsto_mult_ereal_PInf assms by auto moreover have "\x. g x * f x = f x * g x" using mult.commute by auto ultimately show ?thesis by simp qed next assume "\(l = \ \ m = \)" then have "l < \" "m < \" by auto then obtain lr mr where "l = ereal lr" "m = ereal mr" using \l>0\ \m>0\ by (metis ereal_cases ereal_less(6) not_less_iff_gr_or_eq) then show ?thesis using tendsto_mult_real_ereal assms by auto qed text \We reduce the general situation to the positive case by multiplying by suitable signs. Unfortunately, as ereal is not a ring, all the neat sign lemmas are not available there. We give the bare minimum we need.\ lemma ereal_sgn_abs: fixes l::ereal shows "sgn(l) * l = abs(l)" apply (cases l) by (auto simp add: sgn_if ereal_less_uminus_reorder) lemma sgn_squared_ereal: assumes "l \ (0::ereal)" shows "sgn(l) * sgn(l) = 1" apply (cases l) using assms by (auto simp add: one_ereal_def sgn_if) lemma tendsto_mult_ereal [tendsto_intros]: fixes f g::"_ \ ereal" assumes "(f \ l) F" "(g \ m) F" "\((l=0 \ abs(m) = \) \ (m=0 \ abs(l) = \))" shows "((\x. f x * g x) \ l * m) F" proof (cases) assume "l=0 \ m=0" then have "abs(l) \ \" "abs(m) \ \" using assms(3) by auto then obtain lr mr where "l = ereal lr" "m = ereal mr" by auto then show ?thesis using tendsto_mult_real_ereal assms by auto next have sgn_finite: "\a::ereal. abs(sgn a) \ \" by (metis MInfty_neq_ereal(2) PInfty_neq_ereal(2) abs_eq_infinity_cases ereal_times(1) ereal_times(3) ereal_uminus_eq_reorder sgn_ereal.elims) then have sgn_finite2: "\a b::ereal. abs(sgn a * sgn b) \ \" by (metis abs_eq_infinity_cases abs_ereal.simps(2) abs_ereal.simps(3) ereal_mult_eq_MInfty ereal_mult_eq_PInfty) assume "\(l=0 \ m=0)" then have "l \ 0" "m \ 0" by auto then have "abs(l) > 0" "abs(m) > 0" by (metis abs_ereal_ge0 abs_ereal_less0 abs_ereal_pos ereal_uminus_uminus ereal_uminus_zero less_le not_less)+ then have "sgn(l) * l > 0" "sgn(m) * m > 0" using ereal_sgn_abs by auto moreover have "((\x. sgn(l) * f x) \ (sgn(l) * l)) F" by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(1)) moreover have "((\x. sgn(m) * g x) \ (sgn(m) * m)) F" by (rule tendsto_cmult_ereal, auto simp add: sgn_finite assms(2)) ultimately have *: "((\x. (sgn(l) * f x) * (sgn(m) * g x)) \ (sgn(l) * l) * (sgn(m) * m)) F" using tendsto_mult_ereal_pos by force have "((\x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x))) \ (sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m))) F" by (rule tendsto_cmult_ereal, auto simp add: sgn_finite2 *) moreover have "\x. (sgn(l) * sgn(m)) * ((sgn(l) * f x) * (sgn(m) * g x)) = f x * g x" by (metis mult.left_neutral sgn_squared_ereal[OF \l \ 0\] sgn_squared_ereal[OF \m \ 0\] mult.assoc mult.commute) moreover have "(sgn(l) * sgn(m)) * ((sgn(l) * l) * (sgn(m) * m)) = l * m" by (metis mult.left_neutral sgn_squared_ereal[OF \l \ 0\] sgn_squared_ereal[OF \m \ 0\] mult.assoc mult.commute) ultimately show ?thesis by auto qed lemma tendsto_cmult_ereal_general [tendsto_intros]: fixes f::"_ \ ereal" and c::ereal assumes "(f \ l) F" "\ (l=0 \ abs(c) = \)" shows "((\x. c * f x) \ c * l) F" by (cases "c = 0", auto simp add: assms tendsto_mult_ereal) subsubsection\<^marker>\tag important\ \Continuity of division\ lemma tendsto_inverse_ereal_PInf: fixes u::"_ \ ereal" assumes "(u \ \) F" shows "((\x. 1/ u x) \ 0) F" proof - { fix e::real assume "e>0" have "1/e < \" by auto then have "eventually (\n. u n > 1/e) F" using assms(1) by (simp add: tendsto_PInfty) moreover { fix z::ereal assume "z>1/e" then have "z>0" using \e>0\ using less_le_trans not_le by fastforce then have "1/z \ 0" by auto moreover have "1/z < e" using \e>0\ \z>1/e\ apply (cases z) apply auto by (metis (mono_tags, hide_lams) less_ereal.simps(2) less_ereal.simps(4) divide_less_eq ereal_divide_less_pos ereal_less(4) ereal_less_eq(4) less_le_trans mult_eq_0_iff not_le not_one_less_zero times_ereal.simps(1)) ultimately have "1/z \ 0" "1/z < e" by auto } ultimately have "eventually (\n. 1/u nn. 1/u n\0) F" by (auto simp add: eventually_mono) } note * = this show ?thesis proof (subst order_tendsto_iff, auto) fix a::ereal assume "a<0" then show "eventually (\n. 1/u n > a) F" using *(2) eventually_mono less_le_trans linordered_field_no_ub by fastforce next fix a::ereal assume "a>0" then obtain e::real where "e>0" "a>e" using ereal_dense2 ereal_less(2) by blast then have "eventually (\n. 1/u n < e) F" using *(1) by auto then show "eventually (\n. 1/u n < a) F" using \a>e\ by (metis (mono_tags, lifting) eventually_mono less_trans) qed qed text \The next lemma deserves to exist by itself, as it is so common and useful.\ lemma tendsto_inverse_real [tendsto_intros]: fixes u::"_ \ real" shows "(u \ l) F \ l \ 0 \ ((\x. 1/ u x) \ 1/l) F" using tendsto_inverse unfolding inverse_eq_divide . lemma tendsto_inverse_ereal [tendsto_intros]: fixes u::"_ \ ereal" assumes "(u \ l) F" "l \ 0" shows "((\x. 1/ u x) \ 1/l) F" proof (cases l) case (real r) then have "r \ 0" using assms(2) by auto then have "1/l = ereal(1/r)" using real by (simp add: one_ereal_def) define v where "v = (\n. real_of_ereal(u n))" have ureal: "eventually (\n. u n = ereal(v n)) F" unfolding v_def using real_lim_then_eventually_real assms(1) real by auto then have "((\n. ereal(v n)) \ ereal r) F" using assms real v_def by auto then have *: "((\n. v n) \ r) F" by auto then have "((\n. 1/v n) \ 1/r) F" using \r \ 0\ tendsto_inverse_real by auto then have lim: "((\n. ereal(1/v n)) \ 1/l) F" using \1/l = ereal(1/r)\ by auto have "r \ -{0}" "open (-{(0::real)})" using \r \ 0\ by auto then have "eventually (\n. v n \ -{0}) F" using * using topological_tendstoD by blast then have "eventually (\n. v n \ 0) F" by auto moreover { fix n assume H: "v n \ 0" "u n = ereal(v n)" then have "ereal(1/v n) = 1/ereal(v n)" by (simp add: one_ereal_def) then have "ereal(1/v n) = 1/u n" using H(2) by simp } ultimately have "eventually (\n. ereal(1/v n) = 1/u n) F" using ureal eventually_elim2 by force with Lim_transform_eventually[OF lim this] show ?thesis by simp next case (PInf) then have "1/l = 0" by auto then show ?thesis using tendsto_inverse_ereal_PInf assms PInf by auto next case (MInf) then have "1/l = 0" by auto have "1/z = -1/ -z" if "z < 0" for z::ereal apply (cases z) using divide_ereal_def \ z < 0 \ by auto moreover have "eventually (\n. u n < 0) F" by (metis (no_types) MInf assms(1) tendsto_MInfty zero_ereal_def) ultimately have *: "eventually (\n. -1/-u n = 1/u n) F" by (simp add: eventually_mono) define v where "v = (\n. - u n)" have "(v \ \) F" unfolding v_def using MInf assms(1) tendsto_uminus_ereal by fastforce then have "((\n. 1/v n) \ 0) F" using tendsto_inverse_ereal_PInf by auto then have "((\n. -1/v n) \ 0) F" using tendsto_uminus_ereal by fastforce then show ?thesis unfolding v_def using Lim_transform_eventually[OF _ *] \ 1/l = 0 \ by auto qed lemma tendsto_divide_ereal [tendsto_intros]: fixes f g::"_ \ ereal" assumes "(f \ l) F" "(g \ m) F" "m \ 0" "\(abs(l) = \ \ abs(m) = \)" shows "((\x. f x / g x) \ l / m) F" proof - define h where "h = (\x. 1/ g x)" have *: "(h \ 1/m) F" unfolding h_def using assms(2) assms(3) tendsto_inverse_ereal by auto have "((\x. f x * h x) \ l * (1/m)) F" apply (rule tendsto_mult_ereal[OF assms(1) *]) using assms(3) assms(4) by (auto simp add: divide_ereal_def) moreover have "f x * h x = f x / g x" for x unfolding h_def by (simp add: divide_ereal_def) moreover have "l * (1/m) = l/m" by (simp add: divide_ereal_def) ultimately show ?thesis unfolding h_def using Lim_transform_eventually by auto qed subsubsection \Further limits\ text \The assumptions of @{thm tendsto_diff_ereal} are too strong, we weaken them here.\ lemma tendsto_diff_ereal_general [tendsto_intros]: fixes u v::"'a \ ereal" assumes "(u \ l) F" "(v \ m) F" "\((l = \ \ m = \) \ (l = -\ \ m = -\))" shows "((\n. u n - v n) \ l - m) F" proof - have "((\n. u n + (-v n)) \ l + (-m)) F" apply (intro tendsto_intros assms) using assms by (auto simp add: ereal_uminus_eq_reorder) then show ?thesis by (simp add: minus_ereal_def) qed lemma id_nat_ereal_tendsto_PInf [tendsto_intros]: "(\ n::nat. real n) \ \" by (simp add: filterlim_real_sequentially tendsto_PInfty_eq_at_top) lemma tendsto_at_top_pseudo_inverse [tendsto_intros]: fixes u::"nat \ nat" assumes "LIM n sequentially. u n :> at_top" shows "LIM n sequentially. Inf {N. u N \ n} :> at_top" proof - { fix C::nat define M where "M = Max {u n| n. n \ C}+1" { fix n assume "n \ M" have "eventually (\N. u N \ n) sequentially" using assms by (simp add: filterlim_at_top) then have *: "{N. u N \ n} \ {}" by force have "N > C" if "u N \ n" for N proof (rule ccontr) assume "\(N > C)" have "u N \ Max {u n| n. n \ C}" apply (rule Max_ge) using \\(N > C)\ by auto then show False using \u N \ n\ \n \ M\ unfolding M_def by auto qed then have **: "{N. u N \ n} \ {C..}" by fastforce have "Inf {N. u N \ n} \ C" by (metis "*" "**" Inf_nat_def1 atLeast_iff subset_eq) } then have "eventually (\n. Inf {N. u N \ n} \ C) sequentially" using eventually_sequentially by auto } then show ?thesis using filterlim_at_top by auto qed lemma pseudo_inverse_finite_set: fixes u::"nat \ nat" assumes "LIM n sequentially. u n :> at_top" shows "finite {N. u N \ n}" proof - fix n have "eventually (\N. u N \ n+1) sequentially" using assms by (simp add: filterlim_at_top) then obtain N1 where N1: "\N. N \ N1 \ u N \ n + 1" using eventually_sequentially by auto have "{N. u N \ n} \ {.. n}" by (simp add: finite_subset) qed lemma tendsto_at_top_pseudo_inverse2 [tendsto_intros]: fixes u::"nat \ nat" assumes "LIM n sequentially. u n :> at_top" shows "LIM n sequentially. Max {N. u N \ n} :> at_top" proof - { fix N0::nat have "N0 \ Max {N. u N \ n}" if "n \ u N0" for n apply (rule Max.coboundedI) using pseudo_inverse_finite_set[OF assms] that by auto then have "eventually (\n. N0 \ Max {N. u N \ n}) sequentially" using eventually_sequentially by blast } then show ?thesis using filterlim_at_top by auto qed lemma ereal_truncation_top [tendsto_intros]: fixes x::ereal shows "(\n::nat. min x n) \ x" proof (cases x) case (real r) then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto then have "min x n = x" if "n \ K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce then have "eventually (\n. min x n = x) sequentially" using eventually_at_top_linorder by blast then show ?thesis by (simp add: tendsto_eventually) next case (PInf) then have "min x n = n" for n::nat by (auto simp add: min_def) then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto next case (MInf) then have "min x n = x" for n::nat by (auto simp add: min_def) then show ?thesis by auto qed lemma ereal_truncation_real_top [tendsto_intros]: fixes x::ereal assumes "x \ - \" shows "(\n::nat. real_of_ereal(min x n)) \ x" proof (cases x) case (real r) then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto then have "min x n = x" if "n \ K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce then have "real_of_ereal(min x n) = r" if "n \ K" for n using real that by auto then have "eventually (\n. real_of_ereal(min x n) = r) sequentially" using eventually_at_top_linorder by blast then have "(\n. real_of_ereal(min x n)) \ r" by (simp add: tendsto_eventually) then show ?thesis using real by auto next case (PInf) then have "real_of_ereal(min x n) = n" for n::nat by (auto simp add: min_def) then show ?thesis using id_nat_ereal_tendsto_PInf PInf by auto qed (simp add: assms) lemma ereal_truncation_bottom [tendsto_intros]: fixes x::ereal shows "(\n::nat. max x (- real n)) \ x" proof (cases x) case (real r) then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto then have "max x (-real n) = x" if "n \ K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce then have "eventually (\n. max x (-real n) = x) sequentially" using eventually_at_top_linorder by blast then show ?thesis by (simp add: tendsto_eventually) next case (MInf) then have "max x (-real n) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def) moreover have "(\n. (-1)* ereal(real n)) \ -\" using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def) ultimately show ?thesis using MInf by auto next case (PInf) then have "max x (-real n) = x" for n::nat by (auto simp add: max_def) then show ?thesis by auto qed lemma ereal_truncation_real_bottom [tendsto_intros]: fixes x::ereal assumes "x \ \" shows "(\n::nat. real_of_ereal(max x (- real n))) \ x" proof (cases x) case (real r) then obtain K::nat where "K>0" "K > abs(r)" using reals_Archimedean2 gr0I by auto then have "max x (-real n) = x" if "n \ K" for n apply (subst real, subst real, auto) using that eq_iff by fastforce then have "real_of_ereal(max x (-real n)) = r" if "n \ K" for n using real that by auto then have "eventually (\n. real_of_ereal(max x (-real n)) = r) sequentially" using eventually_at_top_linorder by blast then have "(\n. real_of_ereal(max x (-real n))) \ r" by (simp add: tendsto_eventually) then show ?thesis using real by auto next case (MInf) then have "real_of_ereal(max x (-real n)) = (-1)* ereal(real n)" for n::nat by (auto simp add: max_def) moreover have "(\n. (-1)* ereal(real n)) \ -\" using tendsto_cmult_ereal[of "-1", OF _ id_nat_ereal_tendsto_PInf] by (simp add: one_ereal_def) ultimately show ?thesis using MInf by auto qed (simp add: assms) text \the next one is copied from \tendsto_sum\.\ lemma tendsto_sum_ereal [tendsto_intros]: fixes f :: "'a \ 'b \ ereal" assumes "\i. i \ S \ (f i \ a i) F" "\i. abs(a i) \ \" shows "((\x. \i\S. f i x) \ (\i\S. a i)) F" proof (cases "finite S") assume "finite S" then show ?thesis using assms by (induct, simp, simp add: tendsto_add_ereal_general2 assms) qed(simp) lemma continuous_ereal_abs: "continuous_on (UNIV::ereal set) abs" proof - have "continuous_on ({..0} \ {(0::ereal)..}) abs" apply (rule continuous_on_closed_Un, auto) apply (rule iffD1[OF continuous_on_cong, of "{..0}" _ "\x. -x"]) using less_eq_ereal_def apply (auto simp add: continuous_uminus_ereal) apply (rule iffD1[OF continuous_on_cong, of "{0..}" _ "\x. x"]) - apply (auto simp add: continuous_on_id) + apply (auto) done moreover have "(UNIV::ereal set) = {..0} \ {(0::ereal)..}" by auto ultimately show ?thesis by auto qed lemmas continuous_on_compose_ereal_abs[continuous_intros] = continuous_on_compose2[OF continuous_ereal_abs _ subset_UNIV] lemma tendsto_abs_ereal [tendsto_intros]: assumes "(u \ (l::ereal)) F" shows "((\n. abs(u n)) \ abs l) F" using continuous_ereal_abs assms by (metis UNIV_I continuous_on tendsto_compose) lemma ereal_minus_real_tendsto_MInf [tendsto_intros]: "(\x. ereal (- real x)) \ - \" by (subst uminus_ereal.simps(1)[symmetric], intro tendsto_intros) subsection \Extended-Nonnegative-Real.thy\ (*FIX title *) lemma tendsto_diff_ennreal_general [tendsto_intros]: fixes u v::"'a \ ennreal" assumes "(u \ l) F" "(v \ m) F" "\(l = \ \ m = \)" shows "((\n. u n - v n) \ l - m) F" proof - have "((\n. e2ennreal(enn2ereal(u n) - enn2ereal(v n))) \ e2ennreal(enn2ereal l - enn2ereal m)) F" apply (intro tendsto_intros) using assms by auto then show ?thesis by auto qed lemma tendsto_mult_ennreal [tendsto_intros]: fixes l m::ennreal assumes "(u \ l) F" "(v \ m) F" "\((l = 0 \ m = \) \ (l = \ \ m = 0))" shows "((\n. u n * v n) \ l * m) F" proof - have "((\n. e2ennreal(enn2ereal (u n) * enn2ereal (v n))) \ e2ennreal(enn2ereal l * enn2ereal m)) F" apply (intro tendsto_intros) using assms apply auto using enn2ereal_inject zero_ennreal.rep_eq by fastforce+ moreover have "e2ennreal(enn2ereal (u n) * enn2ereal (v n)) = u n * v n" for n by (subst times_ennreal.abs_eq[symmetric], auto simp add: eq_onp_same_args) moreover have "e2ennreal(enn2ereal l * enn2ereal m) = l * m" by (subst times_ennreal.abs_eq[symmetric], auto simp add: eq_onp_same_args) ultimately show ?thesis by auto qed subsection \monoset\ (*FIX ME title *) definition (in order) mono_set: "mono_set S \ (\x y. x \ y \ x \ S \ y \ S)" lemma (in order) mono_greaterThan [intro, simp]: "mono_set {B<..}" unfolding mono_set by auto lemma (in order) mono_atLeast [intro, simp]: "mono_set {B..}" unfolding mono_set by auto lemma (in order) mono_UNIV [intro, simp]: "mono_set UNIV" unfolding mono_set by auto lemma (in order) mono_empty [intro, simp]: "mono_set {}" unfolding mono_set by auto lemma (in complete_linorder) mono_set_iff: fixes S :: "'a set" defines "a \ Inf S" shows "mono_set S \ S = {a <..} \ S = {a..}" (is "_ = ?c") proof assume "mono_set S" then have mono: "\x y. x \ y \ x \ S \ y \ S" by (auto simp: mono_set) show ?c proof cases assume "a \ S" show ?c using mono[OF _ \a \ S\] by (auto intro: Inf_lower simp: a_def) next assume "a \ S" have "S = {a <..}" proof safe fix x assume "x \ S" then have "a \ x" unfolding a_def by (rule Inf_lower) then show "a < x" using \x \ S\ \a \ S\ by (cases "a = x") auto next fix x assume "a < x" then obtain y where "y < x" "y \ S" unfolding a_def Inf_less_iff .. with mono[of y x] show "x \ S" by auto qed then show ?c .. qed qed auto lemma ereal_open_mono_set: fixes S :: "ereal set" shows "open S \ mono_set S \ S = UNIV \ S = {Inf S <..}" by (metis Inf_UNIV atLeast_eq_UNIV_iff ereal_open_atLeast ereal_open_closed mono_set_iff open_ereal_greaterThan) lemma ereal_closed_mono_set: fixes S :: "ereal set" shows "closed S \ mono_set S \ S = {} \ S = {Inf S ..}" by (metis Inf_UNIV atLeast_eq_UNIV_iff closed_ereal_atLeast ereal_open_closed mono_empty mono_set_iff open_ereal_greaterThan) lemma ereal_Liminf_Sup_monoset: fixes f :: "'a \ ereal" shows "Liminf net f = Sup {l. \S. open S \ mono_set S \ l \ S \ eventually (\x. f x \ S) net}" (is "_ = Sup ?A") proof (safe intro!: Liminf_eqI complete_lattice_class.Sup_upper complete_lattice_class.Sup_least) fix P assume P: "eventually P net" fix S assume S: "mono_set S" "Inf (f ` (Collect P)) \ S" { fix x assume "P x" then have "Inf (f ` (Collect P)) \ f x" by (intro complete_lattice_class.INF_lower) simp with S have "f x \ S" by (simp add: mono_set) } with P show "eventually (\x. f x \ S) net" by (auto elim: eventually_mono) next fix y l assume S: "\S. open S \ mono_set S \ l \ S \ eventually (\x. f x \ S) net" assume P: "\P. eventually P net \ Inf (f ` (Collect P)) \ y" show "l \ y" proof (rule dense_le) fix B assume "B < l" then have "eventually (\x. f x \ {B <..}) net" by (intro S[rule_format]) auto then have "Inf (f ` {x. B < f x}) \ y" using P by auto moreover have "B \ Inf (f ` {x. B < f x})" by (intro INF_greatest) auto ultimately show "B \ y" by simp qed qed lemma ereal_Limsup_Inf_monoset: fixes f :: "'a \ ereal" shows "Limsup net f = Inf {l. \S. open S \ mono_set (uminus ` S) \ l \ S \ eventually (\x. f x \ S) net}" (is "_ = Inf ?A") proof (safe intro!: Limsup_eqI complete_lattice_class.Inf_lower complete_lattice_class.Inf_greatest) fix P assume P: "eventually P net" fix S assume S: "mono_set (uminus`S)" "Sup (f ` (Collect P)) \ S" { fix x assume "P x" then have "f x \ Sup (f ` (Collect P))" by (intro complete_lattice_class.SUP_upper) simp with S(1)[unfolded mono_set, rule_format, of "- Sup (f ` (Collect P))" "- f x"] S(2) have "f x \ S" by (simp add: inj_image_mem_iff) } with P show "eventually (\x. f x \ S) net" by (auto elim: eventually_mono) next fix y l assume S: "\S. open S \ mono_set (uminus ` S) \ l \ S \ eventually (\x. f x \ S) net" assume P: "\P. eventually P net \ y \ Sup (f ` (Collect P))" show "y \ l" proof (rule dense_ge) fix B assume "l < B" then have "eventually (\x. f x \ {..< B}) net" by (intro S[rule_format]) auto then have "y \ Sup (f ` {x. f x < B})" using P by auto moreover have "Sup (f ` {x. f x < B}) \ B" by (intro SUP_least) auto ultimately show "y \ B" by simp qed qed lemma liminf_bounded_open: fixes x :: "nat \ ereal" shows "x0 \ liminf x \ (\S. open S \ mono_set S \ x0 \ S \ (\N. \n\N. x n \ S))" (is "_ \ ?P x0") proof assume "?P x0" then show "x0 \ liminf x" unfolding ereal_Liminf_Sup_monoset eventually_sequentially by (intro complete_lattice_class.Sup_upper) auto next assume "x0 \ liminf x" { fix S :: "ereal set" assume om: "open S" "mono_set S" "x0 \ S" { assume "S = UNIV" then have "\N. \n\N. x n \ S" by auto } moreover { assume "S \ UNIV" then obtain B where B: "S = {B<..}" using om ereal_open_mono_set by auto then have "B < x0" using om by auto then have "\N. \n\N. x n \ S" unfolding B using \x0 \ liminf x\ liminf_bounded_iff by auto } ultimately have "\N. \n\N. x n \ S" by auto } then show "?P x0" by auto qed lemma limsup_finite_then_bounded: fixes u::"nat \ real" assumes "limsup u < \" shows "\C. \n. u n \ C" proof - obtain C where C: "limsup u < C" "C < \" using assms ereal_dense2 by blast then have "C = ereal(real_of_ereal C)" using ereal_real by force have "eventually (\n. u n < C) sequentially" using C(1) unfolding Limsup_def apply (auto simp add: INF_less_iff) using SUP_lessD eventually_mono by fastforce then obtain N where N: "\n. n \ N \ u n < C" using eventually_sequentially by auto define D where "D = max (real_of_ereal C) (Max {u n |n. n \ N})" have "\n. u n \ D" proof - fix n show "u n \ D" proof (cases) assume *: "n \ N" have "u n \ Max {u n |n. n \ N}" by (rule Max_ge, auto simp add: *) then show "u n \ D" unfolding D_def by linarith next assume "\(n \ N)" then have "n \ N" by simp then have "u n < C" using N by auto then have "u n < real_of_ereal C" using \C = ereal(real_of_ereal C)\ less_ereal.simps(1) by fastforce then show "u n \ D" unfolding D_def by linarith qed qed then show ?thesis by blast qed lemma liminf_finite_then_bounded_below: fixes u::"nat \ real" assumes "liminf u > -\" shows "\C. \n. u n \ C" proof - obtain C where C: "liminf u > C" "C > -\" using assms using ereal_dense2 by blast then have "C = ereal(real_of_ereal C)" using ereal_real by force have "eventually (\n. u n > C) sequentially" using C(1) unfolding Liminf_def apply (auto simp add: less_SUP_iff) using eventually_elim2 less_INF_D by fastforce then obtain N where N: "\n. n \ N \ u n > C" using eventually_sequentially by auto define D where "D = min (real_of_ereal C) (Min {u n |n. n \ N})" have "\n. u n \ D" proof - fix n show "u n \ D" proof (cases) assume *: "n \ N" have "u n \ Min {u n |n. n \ N}" by (rule Min_le, auto simp add: *) then show "u n \ D" unfolding D_def by linarith next assume "\(n \ N)" then have "n \ N" by simp then have "u n > C" using N by auto then have "u n > real_of_ereal C" using \C = ereal(real_of_ereal C)\ less_ereal.simps(1) by fastforce then show "u n \ D" unfolding D_def by linarith qed qed then show ?thesis by blast qed lemma liminf_upper_bound: fixes u:: "nat \ ereal" assumes "liminf u < l" shows "\N>k. u N < l" by (metis assms gt_ex less_le_trans liminf_bounded_iff not_less) lemma limsup_shift: "limsup (\n. u (n+1)) = limsup u" proof - have "(SUP m\{n+1..}. u m) = (SUP m\{n..}. u (m + 1))" for n apply (rule SUP_eq) using Suc_le_D by auto then have a: "(INF n. SUP m\{n..}. u (m + 1)) = (INF n. (SUP m\{n+1..}. u m))" by auto have b: "(INF n. (SUP m\{n+1..}. u m)) = (INF n\{1..}. (SUP m\{n..}. u m))" apply (rule INF_eq) using Suc_le_D by auto have "(INF n\{1..}. v n) = (INF n. v n)" if "decseq v" for v::"nat \ 'a" apply (rule INF_eq) using \decseq v\ decseq_Suc_iff by auto moreover have "decseq (\n. (SUP m\{n..}. u m))" by (simp add: SUP_subset_mono decseq_def) ultimately have c: "(INF n\{1..}. (SUP m\{n..}. u m)) = (INF n. (SUP m\{n..}. u m))" by simp have "(INF n. Sup (u ` {n..})) = (INF n. SUP m\{n..}. u (m + 1))" using a b c by simp then show ?thesis by (auto cong: limsup_INF_SUP) qed lemma limsup_shift_k: "limsup (\n. u (n+k)) = limsup u" proof (induction k) case (Suc k) have "limsup (\n. u (n+k+1)) = limsup (\n. u (n+k))" using limsup_shift[where ?u="\n. u(n+k)"] by simp then show ?case using Suc.IH by simp qed (auto) lemma liminf_shift: "liminf (\n. u (n+1)) = liminf u" proof - have "(INF m\{n+1..}. u m) = (INF m\{n..}. u (m + 1))" for n apply (rule INF_eq) using Suc_le_D by (auto) then have a: "(SUP n. INF m\{n..}. u (m + 1)) = (SUP n. (INF m\{n+1..}. u m))" by auto have b: "(SUP n. (INF m\{n+1..}. u m)) = (SUP n\{1..}. (INF m\{n..}. u m))" apply (rule SUP_eq) using Suc_le_D by (auto) have "(SUP n\{1..}. v n) = (SUP n. v n)" if "incseq v" for v::"nat \ 'a" apply (rule SUP_eq) using \incseq v\ incseq_Suc_iff by auto moreover have "incseq (\n. (INF m\{n..}. u m))" by (simp add: INF_superset_mono mono_def) ultimately have c: "(SUP n\{1..}. (INF m\{n..}. u m)) = (SUP n. (INF m\{n..}. u m))" by simp have "(SUP n. Inf (u ` {n..})) = (SUP n. INF m\{n..}. u (m + 1))" using a b c by simp then show ?thesis by (auto cong: liminf_SUP_INF) qed lemma liminf_shift_k: "liminf (\n. u (n+k)) = liminf u" proof (induction k) case (Suc k) have "liminf (\n. u (n+k+1)) = liminf (\n. u (n+k))" using liminf_shift[where ?u="\n. u(n+k)"] by simp then show ?case using Suc.IH by simp qed (auto) lemma Limsup_obtain: fixes u::"_ \ 'a :: complete_linorder" assumes "Limsup F u > c" shows "\i. u i > c" proof - have "(INF P\{P. eventually P F}. SUP x\{x. P x}. u x) > c" using assms by (simp add: Limsup_def) then show ?thesis by (metis eventually_True mem_Collect_eq less_INF_D less_SUP_iff) qed text \The next lemma is extremely useful, as it often makes it possible to reduce statements about limsups to statements about limits.\ lemma limsup_subseq_lim: fixes u::"nat \ 'a :: {complete_linorder, linorder_topology}" shows "\r::nat\nat. strict_mono r \ (u o r) \ limsup u" proof (cases) assume "\n. \p>n. \m\p. u m \ u p" then have "\r. \n. (\m\r n. u m \ u (r n)) \ r n < r (Suc n)" by (intro dependent_nat_choice) (auto simp: conj_commute) then obtain r :: "nat \ nat" where "strict_mono r" and mono: "\n m. r n \ m \ u m \ u (r n)" by (auto simp: strict_mono_Suc_iff) define umax where "umax = (\n. (SUP m\{n..}. u m))" have "decseq umax" unfolding umax_def by (simp add: SUP_subset_mono antimono_def) then have "umax \ limsup u" unfolding umax_def by (metis LIMSEQ_INF limsup_INF_SUP) then have *: "(umax o r) \ limsup u" by (simp add: LIMSEQ_subseq_LIMSEQ \strict_mono r\) have "\n. umax(r n) = u(r n)" unfolding umax_def using mono by (metis SUP_le_iff antisym atLeast_def mem_Collect_eq order_refl) then have "umax o r = u o r" unfolding o_def by simp then have "(u o r) \ limsup u" using * by simp then show ?thesis using \strict_mono r\ by blast next assume "\ (\n. \p>n. (\m\p. u m \ u p))" then obtain N where N: "\p. p > N \ \m>p. u p < u m" by (force simp: not_le le_less) have "\r. \n. N < r n \ r n < r (Suc n) \ (\i\ {N<..r (Suc n)}. u i \ u (r (Suc n)))" proof (rule dependent_nat_choice) fix x assume "N < x" then have a: "finite {N<..x}" "{N<..x} \ {}" by simp_all have "Max {u i |i. i \ {N<..x}} \ {u i |i. i \ {N<..x}}" apply (rule Max_in) using a by (auto) then obtain p where "p \ {N<..x}" and upmax: "u p = Max{u i |i. i \ {N<..x}}" by auto define U where "U = {m. m > p \ u p < u m}" have "U \ {}" unfolding U_def using N[of p] \p \ {N<..x}\ by auto define y where "y = Inf U" then have "y \ U" using \U \ {}\ by (simp add: Inf_nat_def1) have a: "\i. i \ {N<..x} \ u i \ u p" proof - fix i assume "i \ {N<..x}" then have "u i \ {u i |i. i \ {N<..x}}" by blast then show "u i \ u p" using upmax by simp qed moreover have "u p < u y" using \y \ U\ U_def by auto ultimately have "y \ {N<..x}" using not_le by blast moreover have "y > N" using \y \ U\ U_def \p \ {N<..x}\ by auto ultimately have "y > x" by auto have "\i. i \ {N<..y} \ u i \ u y" proof - fix i assume "i \ {N<..y}" show "u i \ u y" proof (cases) assume "i = y" then show ?thesis by simp next assume "\(i=y)" then have i:"i \ {N<..i \ {N<..y}\ by simp have "u i \ u p" proof (cases) assume "i \ x" then have "i \ {N<..x}" using i by simp then show ?thesis using a by simp next assume "\(i \ x)" then have "i > x" by simp then have *: "i > p" using \p \ {N<..x}\ by simp have "i < Inf U" using i y_def by simp then have "i \ U" using Inf_nat_def not_less_Least by auto then show ?thesis using U_def * by auto qed then show "u i \ u y" using \u p < u y\ by auto qed qed then have "N < y \ x < y \ (\i\{N<..y}. u i \ u y)" using \y > x\ \y > N\ by auto then show "\y>N. x < y \ (\i\{N<..y}. u i \ u y)" by auto qed (auto) then obtain r where r: "\n. N < r n \ r n < r (Suc n) \ (\i\ {N<..r (Suc n)}. u i \ u (r (Suc n)))" by auto have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff) have "incseq (u o r)" unfolding o_def using r by (simp add: incseq_SucI order.strict_implies_order) then have "(u o r) \ (SUP n. (u o r) n)" using LIMSEQ_SUP by blast then have "limsup (u o r) = (SUP n. (u o r) n)" by (simp add: lim_imp_Limsup) moreover have "limsup (u o r) \ limsup u" using \strict_mono r\ by (simp add: limsup_subseq_mono) ultimately have "(SUP n. (u o r) n) \ limsup u" by simp { fix i assume i: "i \ {N<..}" obtain n where "i < r (Suc n)" using \strict_mono r\ using Suc_le_eq seq_suble by blast then have "i \ {N<..r(Suc n)}" using i by simp then have "u i \ u (r(Suc n))" using r by simp then have "u i \ (SUP n. (u o r) n)" unfolding o_def by (meson SUP_upper2 UNIV_I) } then have "(SUP i\{N<..}. u i) \ (SUP n. (u o r) n)" using SUP_least by blast then have "limsup u \ (SUP n. (u o r) n)" unfolding Limsup_def by (metis (mono_tags, lifting) INF_lower2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq) then have "limsup u = (SUP n. (u o r) n)" using \(SUP n. (u o r) n) \ limsup u\ by simp then have "(u o r) \ limsup u" using \(u o r) \ (SUP n. (u o r) n)\ by simp then show ?thesis using \strict_mono r\ by auto qed lemma liminf_subseq_lim: fixes u::"nat \ 'a :: {complete_linorder, linorder_topology}" shows "\r::nat\nat. strict_mono r \ (u o r) \ liminf u" proof (cases) assume "\n. \p>n. \m\p. u m \ u p" then have "\r. \n. (\m\r n. u m \ u (r n)) \ r n < r (Suc n)" by (intro dependent_nat_choice) (auto simp: conj_commute) then obtain r :: "nat \ nat" where "strict_mono r" and mono: "\n m. r n \ m \ u m \ u (r n)" by (auto simp: strict_mono_Suc_iff) define umin where "umin = (\n. (INF m\{n..}. u m))" have "incseq umin" unfolding umin_def by (simp add: INF_superset_mono incseq_def) then have "umin \ liminf u" unfolding umin_def by (metis LIMSEQ_SUP liminf_SUP_INF) then have *: "(umin o r) \ liminf u" by (simp add: LIMSEQ_subseq_LIMSEQ \strict_mono r\) have "\n. umin(r n) = u(r n)" unfolding umin_def using mono by (metis le_INF_iff antisym atLeast_def mem_Collect_eq order_refl) then have "umin o r = u o r" unfolding o_def by simp then have "(u o r) \ liminf u" using * by simp then show ?thesis using \strict_mono r\ by blast next assume "\ (\n. \p>n. (\m\p. u m \ u p))" then obtain N where N: "\p. p > N \ \m>p. u p > u m" by (force simp: not_le le_less) have "\r. \n. N < r n \ r n < r (Suc n) \ (\i\ {N<..r (Suc n)}. u i \ u (r (Suc n)))" proof (rule dependent_nat_choice) fix x assume "N < x" then have a: "finite {N<..x}" "{N<..x} \ {}" by simp_all have "Min {u i |i. i \ {N<..x}} \ {u i |i. i \ {N<..x}}" apply (rule Min_in) using a by (auto) then obtain p where "p \ {N<..x}" and upmin: "u p = Min{u i |i. i \ {N<..x}}" by auto define U where "U = {m. m > p \ u p > u m}" have "U \ {}" unfolding U_def using N[of p] \p \ {N<..x}\ by auto define y where "y = Inf U" then have "y \ U" using \U \ {}\ by (simp add: Inf_nat_def1) have a: "\i. i \ {N<..x} \ u i \ u p" proof - fix i assume "i \ {N<..x}" then have "u i \ {u i |i. i \ {N<..x}}" by blast then show "u i \ u p" using upmin by simp qed moreover have "u p > u y" using \y \ U\ U_def by auto ultimately have "y \ {N<..x}" using not_le by blast moreover have "y > N" using \y \ U\ U_def \p \ {N<..x}\ by auto ultimately have "y > x" by auto have "\i. i \ {N<..y} \ u i \ u y" proof - fix i assume "i \ {N<..y}" show "u i \ u y" proof (cases) assume "i = y" then show ?thesis by simp next assume "\(i=y)" then have i:"i \ {N<..i \ {N<..y}\ by simp have "u i \ u p" proof (cases) assume "i \ x" then have "i \ {N<..x}" using i by simp then show ?thesis using a by simp next assume "\(i \ x)" then have "i > x" by simp then have *: "i > p" using \p \ {N<..x}\ by simp have "i < Inf U" using i y_def by simp then have "i \ U" using Inf_nat_def not_less_Least by auto then show ?thesis using U_def * by auto qed then show "u i \ u y" using \u p > u y\ by auto qed qed then have "N < y \ x < y \ (\i\{N<..y}. u i \ u y)" using \y > x\ \y > N\ by auto then show "\y>N. x < y \ (\i\{N<..y}. u i \ u y)" by auto qed (auto) then obtain r :: "nat \ nat" where r: "\n. N < r n \ r n < r (Suc n) \ (\i\ {N<..r (Suc n)}. u i \ u (r (Suc n)))" by auto have "strict_mono r" using r by (auto simp: strict_mono_Suc_iff) have "decseq (u o r)" unfolding o_def using r by (simp add: decseq_SucI order.strict_implies_order) then have "(u o r) \ (INF n. (u o r) n)" using LIMSEQ_INF by blast then have "liminf (u o r) = (INF n. (u o r) n)" by (simp add: lim_imp_Liminf) moreover have "liminf (u o r) \ liminf u" using \strict_mono r\ by (simp add: liminf_subseq_mono) ultimately have "(INF n. (u o r) n) \ liminf u" by simp { fix i assume i: "i \ {N<..}" obtain n where "i < r (Suc n)" using \strict_mono r\ using Suc_le_eq seq_suble by blast then have "i \ {N<..r(Suc n)}" using i by simp then have "u i \ u (r(Suc n))" using r by simp then have "u i \ (INF n. (u o r) n)" unfolding o_def by (meson INF_lower2 UNIV_I) } then have "(INF i\{N<..}. u i) \ (INF n. (u o r) n)" using INF_greatest by blast then have "liminf u \ (INF n. (u o r) n)" unfolding Liminf_def by (metis (mono_tags, lifting) SUP_upper2 atLeast_Suc_greaterThan atLeast_def eventually_ge_at_top mem_Collect_eq) then have "liminf u = (INF n. (u o r) n)" using \(INF n. (u o r) n) \ liminf u\ by simp then have "(u o r) \ liminf u" using \(u o r) \ (INF n. (u o r) n)\ by simp then show ?thesis using \strict_mono r\ by auto qed text \The following statement about limsups is reduced to a statement about limits using subsequences thanks to \limsup_subseq_lim\. The statement for limits follows for instance from \tendsto_add_ereal_general\.\ lemma ereal_limsup_add_mono: fixes u v::"nat \ ereal" shows "limsup (\n. u n + v n) \ limsup u + limsup v" proof (cases) assume "(limsup u = \) \ (limsup v = \)" then have "limsup u + limsup v = \" by simp then show ?thesis by auto next assume "\((limsup u = \) \ (limsup v = \))" then have "limsup u < \" "limsup v < \" by auto define w where "w = (\n. u n + v n)" obtain r where r: "strict_mono r" "(w o r) \ limsup w" using limsup_subseq_lim by auto obtain s where s: "strict_mono s" "(u o r o s) \ limsup (u o r)" using limsup_subseq_lim by auto obtain t where t: "strict_mono t" "(v o r o s o t) \ limsup (v o r o s)" using limsup_subseq_lim by auto define a where "a = r o s o t" have "strict_mono a" using r s t by (simp add: a_def strict_mono_o) have l:"(w o a) \ limsup w" "(u o a) \ limsup (u o r)" "(v o a) \ limsup (v o r o s)" apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc) apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc) apply (metis (no_types, lifting) t(2) a_def comp_assoc) done have "limsup (u o r) \ limsup u" by (simp add: limsup_subseq_mono r(1)) then have a: "limsup (u o r) \ \" using \limsup u < \\ by auto have "limsup (v o r o s) \ limsup v" by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o) then have b: "limsup (v o r o s) \ \" using \limsup v < \\ by auto have "(\n. (u o a) n + (v o a) n) \ limsup (u o r) + limsup (v o r o s)" using l tendsto_add_ereal_general a b by fastforce moreover have "(\n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto ultimately have "(w o a) \ limsup (u o r) + limsup (v o r o s)" by simp then have "limsup w = limsup (u o r) + limsup (v o r o s)" using l(1) LIMSEQ_unique by blast then have "limsup w \ limsup u + limsup v" using \limsup (u o r) \ limsup u\ \limsup (v o r o s) \ limsup v\ add_mono by simp then show ?thesis unfolding w_def by simp qed text \There is an asymmetry between liminfs and limsups in \ereal\, as \\ + (-\) = \\. This explains why there are more assumptions in the next lemma dealing with liminfs that in the previous one about limsups.\ lemma ereal_liminf_add_mono: fixes u v::"nat \ ereal" assumes "\((liminf u = \ \ liminf v = -\) \ (liminf u = -\ \ liminf v = \))" shows "liminf (\n. u n + v n) \ liminf u + liminf v" proof (cases) assume "(liminf u = -\) \ (liminf v = -\)" then have *: "liminf u + liminf v = -\" using assms by auto show ?thesis by (simp add: *) next assume "\((liminf u = -\) \ (liminf v = -\))" then have "liminf u > -\" "liminf v > -\" by auto define w where "w = (\n. u n + v n)" obtain r where r: "strict_mono r" "(w o r) \ liminf w" using liminf_subseq_lim by auto obtain s where s: "strict_mono s" "(u o r o s) \ liminf (u o r)" using liminf_subseq_lim by auto obtain t where t: "strict_mono t" "(v o r o s o t) \ liminf (v o r o s)" using liminf_subseq_lim by auto define a where "a = r o s o t" have "strict_mono a" using r s t by (simp add: a_def strict_mono_o) have l:"(w o a) \ liminf w" "(u o a) \ liminf (u o r)" "(v o a) \ liminf (v o r o s)" apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc) apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc) apply (metis (no_types, lifting) t(2) a_def comp_assoc) done have "liminf (u o r) \ liminf u" by (simp add: liminf_subseq_mono r(1)) then have a: "liminf (u o r) \ -\" using \liminf u > -\\ by auto have "liminf (v o r o s) \ liminf v" by (simp add: comp_assoc liminf_subseq_mono r(1) s(1) strict_mono_o) then have b: "liminf (v o r o s) \ -\" using \liminf v > -\\ by auto have "(\n. (u o a) n + (v o a) n) \ liminf (u o r) + liminf (v o r o s)" using l tendsto_add_ereal_general a b by fastforce moreover have "(\n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto ultimately have "(w o a) \ liminf (u o r) + liminf (v o r o s)" by simp then have "liminf w = liminf (u o r) + liminf (v o r o s)" using l(1) LIMSEQ_unique by blast then have "liminf w \ liminf u + liminf v" using \liminf (u o r) \ liminf u\ \liminf (v o r o s) \ liminf v\ add_mono by simp then show ?thesis unfolding w_def by simp qed lemma ereal_limsup_lim_add: fixes u v::"nat \ ereal" assumes "u \ a" "abs(a) \ \" shows "limsup (\n. u n + v n) = a + limsup v" proof - have "limsup u = a" using assms(1) using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast have "(\n. -u n) \ -a" using assms(1) by auto then have "limsup (\n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast have "limsup (\n. u n + v n) \ limsup u + limsup v" by (rule ereal_limsup_add_mono) then have up: "limsup (\n. u n + v n) \ a + limsup v" using \limsup u = a\ by simp have a: "limsup (\n. (u n + v n) + (-u n)) \ limsup (\n. u n + v n) + limsup (\n. -u n)" by (rule ereal_limsup_add_mono) have "eventually (\n. u n = ereal(real_of_ereal(u n))) sequentially" using assms real_lim_then_eventually_real by auto moreover have "\x. x = ereal(real_of_ereal(x)) \ x + (-x) = 0" by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def) ultimately have "eventually (\n. u n + (-u n) = 0) sequentially" by (metis (mono_tags, lifting) eventually_mono) moreover have "\n. u n + (-u n) = 0 \ u n + v n + (-u n) = v n" by (metis add.commute add.left_commute add.left_neutral) ultimately have "eventually (\n. u n + v n + (-u n) = v n) sequentially" using eventually_mono by force then have "limsup v = limsup (\n. u n + v n + (-u n))" using Limsup_eq by force then have "limsup v \ limsup (\n. u n + v n) -a" using a \limsup (\n. -u n) = -a\ by (simp add: minus_ereal_def) then have "limsup (\n. u n + v n) \ a + limsup v" using assms(2) by (metis add.commute ereal_le_minus) then show ?thesis using up by simp qed lemma ereal_limsup_lim_mult: fixes u v::"nat \ ereal" assumes "u \ a" "a>0" "a \ \" shows "limsup (\n. u n * v n) = a * limsup v" proof - define w where "w = (\n. u n * v n)" obtain r where r: "strict_mono r" "(v o r) \ limsup v" using limsup_subseq_lim by auto have "(u o r) \ a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto with tendsto_mult_ereal[OF this r(2)] have "(\n. (u o r) n * (v o r) n) \ a * limsup v" using assms(2) assms(3) by auto moreover have "\n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto ultimately have "(w o r) \ a * limsup v" unfolding w_def by presburger then have "limsup (w o r) = a * limsup v" by (simp add: tendsto_iff_Liminf_eq_Limsup) then have I: "limsup w \ a * limsup v" by (metis limsup_subseq_mono r(1)) obtain s where s: "strict_mono s" "(w o s) \ limsup w" using limsup_subseq_lim by auto have *: "(u o s) \ a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto have "eventually (\n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast moreover have "eventually (\n. (u o s) n < \) sequentially" using assms(3) * order_tendsto_iff by blast moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \" for n unfolding w_def using that by (auto simp add: ereal_divide_eq) ultimately have "eventually (\n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force moreover have "(\n. (w o s) n / (u o s) n) \ (limsup w) / a" apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto ultimately have "(v o s) \ (limsup w) / a" using Lim_transform_eventually by fastforce then have "limsup (v o s) = (limsup w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup) then have "limsup v \ (limsup w) / a" by (metis limsup_subseq_mono s(1)) then have "a * limsup v \ limsup w" using assms(2) assms(3) by (simp add: ereal_divide_le_pos) then show ?thesis using I unfolding w_def by auto qed lemma ereal_liminf_lim_mult: fixes u v::"nat \ ereal" assumes "u \ a" "a>0" "a \ \" shows "liminf (\n. u n * v n) = a * liminf v" proof - define w where "w = (\n. u n * v n)" obtain r where r: "strict_mono r" "(v o r) \ liminf v" using liminf_subseq_lim by auto have "(u o r) \ a" using assms(1) LIMSEQ_subseq_LIMSEQ r by auto with tendsto_mult_ereal[OF this r(2)] have "(\n. (u o r) n * (v o r) n) \ a * liminf v" using assms(2) assms(3) by auto moreover have "\n. (w o r) n = (u o r) n * (v o r) n" unfolding w_def by auto ultimately have "(w o r) \ a * liminf v" unfolding w_def by presburger then have "liminf (w o r) = a * liminf v" by (simp add: tendsto_iff_Liminf_eq_Limsup) then have I: "liminf w \ a * liminf v" by (metis liminf_subseq_mono r(1)) obtain s where s: "strict_mono s" "(w o s) \ liminf w" using liminf_subseq_lim by auto have *: "(u o s) \ a" using assms(1) LIMSEQ_subseq_LIMSEQ s by auto have "eventually (\n. (u o s) n > 0) sequentially" using assms(2) * order_tendsto_iff by blast moreover have "eventually (\n. (u o s) n < \) sequentially" using assms(3) * order_tendsto_iff by blast moreover have "(w o s) n / (u o s) n = (v o s) n" if "(u o s) n > 0" "(u o s) n < \" for n unfolding w_def using that by (auto simp add: ereal_divide_eq) ultimately have "eventually (\n. (w o s) n / (u o s) n = (v o s) n) sequentially" using eventually_elim2 by force moreover have "(\n. (w o s) n / (u o s) n) \ (liminf w) / a" apply (rule tendsto_divide_ereal[OF s(2) *]) using assms(2) assms(3) by auto ultimately have "(v o s) \ (liminf w) / a" using Lim_transform_eventually by fastforce then have "liminf (v o s) = (liminf w) / a" by (simp add: tendsto_iff_Liminf_eq_Limsup) then have "liminf v \ (liminf w) / a" by (metis liminf_subseq_mono s(1)) then have "a * liminf v \ liminf w" using assms(2) assms(3) by (simp add: ereal_le_divide_pos) then show ?thesis using I unfolding w_def by auto qed lemma ereal_liminf_lim_add: fixes u v::"nat \ ereal" assumes "u \ a" "abs(a) \ \" shows "liminf (\n. u n + v n) = a + liminf v" proof - have "liminf u = a" using assms(1) tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast then have *: "abs(liminf u) \ \" using assms(2) by auto have "(\n. -u n) \ -a" using assms(1) by auto then have "liminf (\n. -u n) = -a" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast then have **: "abs(liminf (\n. -u n)) \ \" using assms(2) by auto have "liminf (\n. u n + v n) \ liminf u + liminf v" apply (rule ereal_liminf_add_mono) using * by auto then have up: "liminf (\n. u n + v n) \ a + liminf v" using \liminf u = a\ by simp have a: "liminf (\n. (u n + v n) + (-u n)) \ liminf (\n. u n + v n) + liminf (\n. -u n)" apply (rule ereal_liminf_add_mono) using ** by auto have "eventually (\n. u n = ereal(real_of_ereal(u n))) sequentially" using assms real_lim_then_eventually_real by auto moreover have "\x. x = ereal(real_of_ereal(x)) \ x + (-x) = 0" by (metis plus_ereal.simps(1) right_minus uminus_ereal.simps(1) zero_ereal_def) ultimately have "eventually (\n. u n + (-u n) = 0) sequentially" by (metis (mono_tags, lifting) eventually_mono) moreover have "\n. u n + (-u n) = 0 \ u n + v n + (-u n) = v n" by (metis add.commute add.left_commute add.left_neutral) ultimately have "eventually (\n. u n + v n + (-u n) = v n) sequentially" using eventually_mono by force then have "liminf v = liminf (\n. u n + v n + (-u n))" using Liminf_eq by force then have "liminf v \ liminf (\n. u n + v n) -a" using a \liminf (\n. -u n) = -a\ by (simp add: minus_ereal_def) then have "liminf (\n. u n + v n) \ a + liminf v" using assms(2) by (metis add.commute ereal_minus_le) then show ?thesis using up by simp qed lemma ereal_liminf_limsup_add: fixes u v::"nat \ ereal" shows "liminf (\n. u n + v n) \ liminf u + limsup v" proof (cases) assume "limsup v = \ \ liminf u = \" then show ?thesis by auto next assume "\(limsup v = \ \ liminf u = \)" then have "limsup v < \" "liminf u < \" by auto define w where "w = (\n. u n + v n)" obtain r where r: "strict_mono r" "(u o r) \ liminf u" using liminf_subseq_lim by auto obtain s where s: "strict_mono s" "(w o r o s) \ liminf (w o r)" using liminf_subseq_lim by auto obtain t where t: "strict_mono t" "(v o r o s o t) \ limsup (v o r o s)" using limsup_subseq_lim by auto define a where "a = r o s o t" have "strict_mono a" using r s t by (simp add: a_def strict_mono_o) have l:"(u o a) \ liminf u" "(w o a) \ liminf (w o r)" "(v o a) \ limsup (v o r o s)" apply (metis (no_types, lifting) r(2) s(1) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc) apply (metis (no_types, lifting) s(2) t(1) LIMSEQ_subseq_LIMSEQ a_def comp_assoc) apply (metis (no_types, lifting) t(2) a_def comp_assoc) done have "liminf (w o r) \ liminf w" by (simp add: liminf_subseq_mono r(1)) have "limsup (v o r o s) \ limsup v" by (simp add: comp_assoc limsup_subseq_mono r(1) s(1) strict_mono_o) then have b: "limsup (v o r o s) < \" using \limsup v < \\ by auto have "(\n. (u o a) n + (v o a) n) \ liminf u + limsup (v o r o s)" apply (rule tendsto_add_ereal_general) using b \liminf u < \\ l(1) l(3) by force+ moreover have "(\n. (u o a) n + (v o a) n) = (w o a)" unfolding w_def by auto ultimately have "(w o a) \ liminf u + limsup (v o r o s)" by simp then have "liminf (w o r) = liminf u + limsup (v o r o s)" using l(2) using LIMSEQ_unique by blast then have "liminf w \ liminf u + limsup v" using \liminf (w o r) \ liminf w\ \limsup (v o r o s) \ limsup v\ by (metis add_mono_thms_linordered_semiring(2) le_less_trans not_less) then show ?thesis unfolding w_def by simp qed lemma ereal_liminf_limsup_minus: fixes u v::"nat \ ereal" shows "liminf (\n. u n - v n) \ limsup u - limsup v" unfolding minus_ereal_def apply (subst add.commute) apply (rule order_trans[OF ereal_liminf_limsup_add]) using ereal_Limsup_uminus[of sequentially "\n. - v n"] apply (simp add: add.commute) done lemma liminf_minus_ennreal: fixes u v::"nat \ ennreal" shows "(\n. v n \ u n) \ liminf (\n. u n - v n) \ limsup u - limsup v" unfolding liminf_SUP_INF limsup_INF_SUP including ennreal.lifting proof (transfer, clarsimp) fix v u :: "nat \ ereal" assume *: "\x. 0 \ v x" "\x. 0 \ u x" "\n. v n \ u n" moreover have "0 \ limsup u - limsup v" using * by (intro ereal_diff_positive Limsup_mono always_eventually) simp moreover have "0 \ Sup (u ` {x..})" for x using * by (intro SUP_upper2[of x]) auto moreover have "0 \ Sup (v ` {x..})" for x using * by (intro SUP_upper2[of x]) auto ultimately show "(SUP n. INF n\{n..}. max 0 (u n - v n)) \ max 0 ((INF x. max 0 (Sup (u ` {x..}))) - (INF x. max 0 (Sup (v ` {x..}))))" by (auto simp: * ereal_diff_positive max.absorb2 liminf_SUP_INF[symmetric] limsup_INF_SUP[symmetric] ereal_liminf_limsup_minus) qed subsection "Relate extended reals and the indicator function" lemma ereal_indicator_le_0: "(indicator S x::ereal) \ 0 \ x \ S" by (auto split: split_indicator simp: one_ereal_def) lemma ereal_indicator: "ereal (indicator A x) = indicator A x" by (auto simp: indicator_def one_ereal_def) lemma ereal_mult_indicator: "ereal (x * indicator A y) = ereal x * indicator A y" by (simp split: split_indicator) lemma ereal_indicator_mult: "ereal (indicator A y * x) = indicator A y * ereal x" by (simp split: split_indicator) lemma ereal_indicator_nonneg[simp, intro]: "0 \ (indicator A x ::ereal)" unfolding indicator_def by auto lemma indicator_inter_arith_ereal: "indicator A x * indicator B x = (indicator (A \ B) x :: ereal)" by (simp split: split_indicator) end diff --git a/src/HOL/Analysis/Function_Topology.thy b/src/HOL/Analysis/Function_Topology.thy --- a/src/HOL/Analysis/Function_Topology.thy +++ b/src/HOL/Analysis/Function_Topology.thy @@ -1,2062 +1,2061 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP License: BSD *) theory Function_Topology imports Topology_Euclidean_Space Abstract_Limits begin section \Function Topology\ text \We want to define the general product topology. The product topology on a product of topological spaces is generated by the sets which are products of open sets along finitely many coordinates, and the whole space along the other coordinates. This is the coarsest topology for which the projection to each factor is continuous. To form a product of objects in Isabelle/HOL, all these objects should be subsets of a common type 'a. The product is then \<^term>\Pi\<^sub>E I X\, the set of elements from \'i\ to \'a\ such that the \i\-th coordinate belongs to \X i\ for all \i \ I\. Hence, to form a product of topological spaces, all these spaces should be subsets of a common type. This means that type classes can not be used to define such a product if one wants to take the product of different topological spaces (as the type 'a can only be given one structure of topological space using type classes). On the other hand, one can define different topologies (as introduced in \thy\) on one type, and these topologies do not need to share the same maximal open set. Hence, one can form a product of topologies in this sense, and this works well. The big caveat is that it does not interact well with the main body of topology in Isabelle/HOL defined in terms of type classes... For instance, continuity of maps is not defined in this setting. As the product of different topological spaces is very important in several areas of mathematics (for instance adeles), I introduce below the product topology in terms of topologies, and reformulate afterwards the consequences in terms of type classes (which are of course very handy for applications). Given this limitation, it looks to me that it would be very beneficial to revamp the theory of topological spaces in Isabelle/HOL in terms of topologies, and keep the statements involving type classes as consequences of more general statements in terms of topologies (but I am probably too naive here). Here is an example of a reformulation using topologies. Let @{text [display] \continuous_map T1 T2 f = ((\ U. openin T2 U \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (topspace T2)))\} be the natural continuity definition of a map from the topology \T1\ to the topology \T2\. Then the current \continuous_on\ (with type classes) can be redefined as @{text [display] \continuous_on s f = continuous_map (top_of_set s) (topology euclidean) f\} In fact, I need \continuous_map\ to express the continuity of the projection on subfactors for the product topology, in Lemma~\continuous_on_restrict_product_topology\, and I show the above equivalence in Lemma~\continuous_map_iff_continuous\. I only develop the basics of the product topology in this theory. The most important missing piece is Tychonov theorem, stating that a product of compact spaces is always compact for the product topology, even when the product is not finite (or even countable). I realized afterwards that this theory has a lot in common with \<^file>\~~/src/HOL/Library/Finite_Map.thy\. \ subsection \The product topology\ text \We can now define the product topology, as generated by the sets which are products of open sets along finitely many coordinates, and the whole space along the other coordinates. Equivalently, it is generated by sets which are one open set along one single coordinate, and the whole space along other coordinates. In fact, this is only equivalent for nonempty products, but for the empty product the first formulation is better (the second one gives an empty product space, while an empty product should have exactly one point, equal to \undefined\ along all coordinates. So, we use the first formulation, which moreover seems to give rise to more straightforward proofs. \ definition\<^marker>\tag important\ product_topology::"('i \ ('a topology)) \ ('i set) \ (('i \ 'a) topology)" where "product_topology T I = topology_generated_by {(\\<^sub>E i\I. X i) |X. (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}}" abbreviation powertop_real :: "'a set \ ('a \ real) topology" where "powertop_real \ product_topology (\i. euclideanreal)" text \The total set of the product topology is the product of the total sets along each coordinate.\ proposition product_topology: "product_topology X I = topology (arbitrary union_of ((finite intersection_of (\F. \i U. F = {f. f i \ U} \ i \ I \ openin (X i) U)) relative_to (\\<^sub>E i\I. topspace (X i))))" (is "_ = topology (_ union_of ((_ intersection_of ?\) relative_to ?TOP))") proof - let ?\ = "(\F. \Y. F = Pi\<^sub>E I Y \ (\i. openin (X i) (Y i)) \ finite {i. Y i \ topspace (X i)})" have *: "(finite' intersection_of ?\) A = (finite intersection_of ?\ relative_to ?TOP) A" for A proof - have 1: "\U. (\\. finite \ \ \ \ Collect ?\ \ \\ = U) \ ?TOP \ U = \\" if \: "\ \ Collect ?\" and "finite' \" "A = \\" "\ \ {}" for \ proof - have "\U \ \. \Y. U = Pi\<^sub>E I Y \ (\i. openin (X i) (Y i)) \ finite {i. Y i \ topspace (X i)}" using \ by auto then obtain Y where Y: "\U. U \ \ \ U = Pi\<^sub>E I (Y U) \ (\i. openin (X i) (Y U i)) \ finite {i. (Y U) i \ topspace (X i)}" by metis obtain U where "U \ \" using \\ \ {}\ by blast let ?F = "\U. (\i. {f. f i \ Y U i}) ` {i \ I. Y U i \ topspace (X i)}" show ?thesis proof (intro conjI exI) show "finite (\U\\. ?F U)" using Y \finite' \\ by auto show "?TOP \ \(\U\\. ?F U) = \\" proof have *: "f \ U" if "U \ \" and "\V\\. \i. i \ I \ Y V i \ topspace (X i) \ f i \ Y V i" and "\i\I. f i \ topspace (X i)" and "f \ extensional I" for f U proof - have "Pi\<^sub>E I (Y U) = U" using Y \U \ \\ by blast then show "f \ U" using that unfolding PiE_def Pi_def by blast qed show "?TOP \ \(\U\\. ?F U) \ \\" by (auto simp: PiE_iff *) show "\\ \ ?TOP \ \(\U\\. ?F U)" using Y openin_subset \finite' \\ by fastforce qed qed (use Y openin_subset in \blast+\) qed have 2: "\\'. finite' \' \ \' \ Collect ?\ \ \\' = ?TOP \ \\" if \: "\ \ Collect ?\" and "finite \" for \ proof (cases "\={}") case True then show ?thesis apply (rule_tac x="{?TOP}" in exI, simp) apply (rule_tac x="\i. topspace (X i)" in exI) apply force done next case False then obtain U where "U \ \" by blast have "\U \ \. \i Y. U = {f. f i \ Y} \ i \ I \ openin (X i) Y" using \ by auto then obtain J Y where Y: "\U. U \ \ \ U = {f. f (J U) \ Y U} \ J U \ I \ openin (X (J U)) (Y U)" by metis let ?G = "\U. \\<^sub>E i\I. if i = J U then Y U else topspace (X i)" show ?thesis proof (intro conjI exI) show "finite (?G ` \)" "?G ` \ \ {}" using \finite \\ \U \ \\ by blast+ have *: "\U. U \ \ \ openin (X (J U)) (Y U)" using Y by force show "?G ` \ \ {Pi\<^sub>E I Y |Y. (\i. openin (X i) (Y i)) \ finite {i. Y i \ topspace (X i)}}" apply clarsimp apply (rule_tac x=" (\i. if i = J U then Y U else topspace (X i))" in exI) apply (auto simp: *) done next show "(\U\\. ?G U) = ?TOP \ \\" proof have "(\\<^sub>E i\I. if i = J U then Y U else topspace (X i)) \ (\\<^sub>E i\I. topspace (X i))" apply (clarsimp simp: PiE_iff Ball_def all_conj_distrib split: if_split_asm) using Y \U \ \\ openin_subset by blast+ then have "(\U\\. ?G U) \ ?TOP" using \U \ \\ by fastforce moreover have "(\U\\. ?G U) \ \\" using PiE_mem Y by fastforce ultimately show "(\U\\. ?G U) \ ?TOP \ \\" by auto qed (use Y in fastforce) qed qed show ?thesis unfolding relative_to_def intersection_of_def by (safe; blast dest!: 1 2) qed show ?thesis unfolding product_topology_def generate_topology_on_eq apply (rule arg_cong [where f = topology]) apply (rule arg_cong [where f = "(union_of)arbitrary"]) apply (force simp: *) done qed lemma topspace_product_topology [simp]: "topspace (product_topology T I) = (\\<^sub>E i\I. topspace(T i))" proof show "topspace (product_topology T I) \ (\\<^sub>E i\I. topspace (T i))" unfolding product_topology_def topology_generated_by_topspace unfolding topspace_def by auto have "(\\<^sub>E i\I. topspace (T i)) \ {(\\<^sub>E i\I. X i) |X. (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}}" using openin_topspace not_finite_existsD by auto then show "(\\<^sub>E i\I. topspace (T i)) \ topspace (product_topology T I)" unfolding product_topology_def using PiE_def by (auto simp add: topology_generated_by_topspace) qed lemma product_topology_basis: assumes "\i. openin (T i) (X i)" "finite {i. X i \ topspace (T i)}" shows "openin (product_topology T I) (\\<^sub>E i\I. X i)" unfolding product_topology_def by (rule topology_generated_by_Basis) (use assms in auto) proposition product_topology_open_contains_basis: assumes "openin (product_topology T I) U" "x \ U" shows "\X. x \ (\\<^sub>E i\I. X i) \ (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)} \ (\\<^sub>E i\I. X i) \ U" proof - have "generate_topology_on {(\\<^sub>E i\I. X i) |X. (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}} U" using assms unfolding product_topology_def by (intro openin_topology_generated_by) auto then have "\x. x\U \ \X. x \ (\\<^sub>E i\I. X i) \ (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)} \ (\\<^sub>E i\I. X i) \ U" proof induction case (Int U V x) then obtain XU XV where H: "x \ Pi\<^sub>E I XU" "(\i. openin (T i) (XU i))" "finite {i. XU i \ topspace (T i)}" "Pi\<^sub>E I XU \ U" "x \ Pi\<^sub>E I XV" "(\i. openin (T i) (XV i))" "finite {i. XV i \ topspace (T i)}" "Pi\<^sub>E I XV \ V" by auto meson define X where "X = (\i. XU i \ XV i)" have "Pi\<^sub>E I X \ Pi\<^sub>E I XU \ Pi\<^sub>E I XV" unfolding X_def by (auto simp add: PiE_iff) then have "Pi\<^sub>E I X \ U \ V" using H by auto moreover have "\i. openin (T i) (X i)" unfolding X_def using H by auto moreover have "finite {i. X i \ topspace (T i)}" apply (rule rev_finite_subset[of "{i. XU i \ topspace (T i)} \ {i. XV i \ topspace (T i)}"]) unfolding X_def using H by auto moreover have "x \ Pi\<^sub>E I X" unfolding X_def using H by auto ultimately show ?case by auto next case (UN K x) then obtain k where "k \ K" "x \ k" by auto with UN have "\X. x \ Pi\<^sub>E I X \ (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)} \ Pi\<^sub>E I X \ k" by simp then obtain X where "x \ Pi\<^sub>E I X \ (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)} \ Pi\<^sub>E I X \ k" by blast then have "x \ Pi\<^sub>E I X \ (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)} \ Pi\<^sub>E I X \ (\K)" using \k \ K\ by auto then show ?case by auto qed auto then show ?thesis using \x \ U\ by auto qed lemma product_topology_empty_discrete: "product_topology T {} = discrete_topology {(\x. undefined)}" by (simp add: subtopology_eq_discrete_topology_sing) lemma openin_product_topology: "openin (product_topology X I) = arbitrary union_of ((finite intersection_of (\F. (\i U. F = {f. f i \ U} \ i \ I \ openin (X i) U))) relative_to topspace (product_topology X I))" apply (subst product_topology) apply (simp add: topology_inverse' [OF istopology_subbase]) done lemma subtopology_PiE: "subtopology (product_topology X I) (\\<^sub>E i\I. (S i)) = product_topology (\i. subtopology (X i) (S i)) I" proof - let ?P = "\F. \i U. F = {f. f i \ U} \ i \ I \ openin (X i) U" let ?X = "\\<^sub>E i\I. topspace (X i)" have "finite intersection_of ?P relative_to ?X \ Pi\<^sub>E I S = finite intersection_of (?P relative_to ?X \ Pi\<^sub>E I S) relative_to ?X \ Pi\<^sub>E I S" by (rule finite_intersection_of_relative_to) also have "\ = finite intersection_of ((\F. \i U. F = {f. f i \ U} \ i \ I \ (openin (X i) relative_to S i) U) relative_to ?X \ Pi\<^sub>E I S) relative_to ?X \ Pi\<^sub>E I S" apply (rule arg_cong2 [where f = "(relative_to)"]) apply (rule arg_cong [where f = "(intersection_of)finite"]) apply (rule ext) apply (auto simp: relative_to_def intersection_of_def) done finally have "finite intersection_of ?P relative_to ?X \ Pi\<^sub>E I S = finite intersection_of (\F. \i U. F = {f. f i \ U} \ i \ I \ (openin (X i) relative_to S i) U) relative_to ?X \ Pi\<^sub>E I S" by (metis finite_intersection_of_relative_to) then show ?thesis unfolding topology_eq apply clarify apply (simp add: openin_product_topology flip: openin_relative_to) - apply (simp add: arbitrary_union_of_relative_to topspace_product_topology topspace_subtopology flip: PiE_Int) + apply (simp add: arbitrary_union_of_relative_to flip: PiE_Int) done qed lemma product_topology_base_alt: "finite intersection_of (\F. (\i U. F = {f. f i \ U} \ i \ I \ openin (X i) U)) relative_to (\\<^sub>E i\I. topspace (X i)) = (\F. (\U. F = Pi\<^sub>E I U \ finite {i \ I. U i \ topspace(X i)} \ (\i \ I. openin (X i) (U i))))" (is "?lhs = ?rhs") proof - have "(\F. ?lhs F \ ?rhs F)" unfolding all_relative_to all_intersection_of topspace_product_topology proof clarify fix \ assume "finite \" and "\ \ {{f. f i \ U} |i U. i \ I \ openin (X i) U}" then show "\U. (\\<^sub>E i\I. topspace (X i)) \ \\ = Pi\<^sub>E I U \ finite {i \ I. U i \ topspace (X i)} \ (\i\I. openin (X i) (U i))" proof (induction) case (insert F \) then obtain U where eq: "(\\<^sub>E i\I. topspace (X i)) \ \\ = Pi\<^sub>E I U" and fin: "finite {i \ I. U i \ topspace (X i)}" and ope: "\i. i \ I \ openin (X i) (U i)" by auto obtain i V where "F = {f. f i \ V}" "i \ I" "openin (X i) V" using insert by auto let ?U = "\j. U j \ (if j = i then V else topspace(X j))" show ?case proof (intro exI conjI) show "(\\<^sub>E i\I. topspace (X i)) \ \(insert F \) = Pi\<^sub>E I ?U" using eq PiE_mem \i \ I\ by (auto simp: \F = {f. f i \ V}\) fastforce next show "finite {i \ I. ?U i \ topspace (X i)}" by (rule rev_finite_subset [OF finite.insertI [OF fin]]) auto next show "\i\I. openin (X i) (?U i)" by (simp add: \openin (X i) V\ ope openin_Int) qed qed (auto intro: dest: not_finite_existsD) qed moreover have "(\F. ?rhs F \ ?lhs F)" proof clarify fix U :: "'a \ 'b set" assume fin: "finite {i \ I. U i \ topspace (X i)}" and ope: "\i\I. openin (X i) (U i)" let ?U = "\i\{i \ I. U i \ topspace (X i)}. {x. x i \ U i}" show "?lhs (Pi\<^sub>E I U)" unfolding relative_to_def topspace_product_topology proof (intro exI conjI) show "(finite intersection_of (\F. \i U. F = {f. f i \ U} \ i \ I \ openin (X i) U)) ?U" using fin ope by (intro finite_intersection_of_Inter finite_intersection_of_inc) auto show "(\\<^sub>E i\I. topspace (X i)) \ ?U = Pi\<^sub>E I U" using ope openin_subset by fastforce qed qed ultimately show ?thesis by meson qed corollary openin_product_topology_alt: "openin (product_topology X I) S \ (\x \ S. \U. finite {i \ I. U i \ topspace(X i)} \ (\i \ I. openin (X i) (U i)) \ x \ Pi\<^sub>E I U \ Pi\<^sub>E I U \ S)" unfolding openin_product_topology arbitrary_union_of_alt product_topology_base_alt topspace_product_topology apply safe apply (drule bspec; blast)+ done lemma closure_of_product_topology: "(product_topology X I) closure_of (PiE I S) = PiE I (\i. (X i) closure_of (S i))" proof - have *: "(\T. f \ T \ openin (product_topology X I) T \ (\y\Pi\<^sub>E I S. y \ T)) \ (\i \ I. \T. f i \ T \ openin (X i) T \ S i \ T \ {})" (is "?lhs = ?rhs") if top: "\i. i \ I \ f i \ topspace (X i)" and ext: "f \ extensional I" for f proof assume L[rule_format]: ?lhs show ?rhs proof clarify fix i T assume "i \ I" "f i \ T" "openin (X i) T" "S i \ T = {}" then have "openin (product_topology X I) ((\\<^sub>E i\I. topspace (X i)) \ {x. x i \ T})" by (force simp: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc) then show "False" using L [of "topspace (product_topology X I) \ {f. f i \ T}"] \S i \ T = {}\ \f i \ T\ \i \ I\ by (auto simp: top ext PiE_iff) qed next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp: openin_product_topology union_of_def arbitrary_def) fix \ U assume \: "\ \ Collect (finite intersection_of (\F. \i U. F = {x. x i \ U} \ i \ I \ openin (X i) U) relative_to (\\<^sub>E i\I. topspace (X i)))" and "f \ U" "U \ \" then have "(finite intersection_of (\F. \i U. F = {x. x i \ U} \ i \ I \ openin (X i) U) relative_to (\\<^sub>E i\I. topspace (X i))) U" by blast with \f \ U\ \U \ \\ obtain \ where "finite \" and \: "\C. C \ \ \ \i \ I. \V. openin (X i) V \ C = {x. x i \ V}" and "topspace (product_topology X I) \ \ \ \ U" "f \ topspace (product_topology X I) \ \ \" apply (clarsimp simp add: relative_to_def intersection_of_def) apply (rule that, auto dest!: subsetD) done then have "f \ PiE I (topspace \ X)" "f \ \\" and subU: "PiE I (topspace \ X) \ \\ \ U" by (auto simp: PiE_iff) have *: "f i \ topspace (X i) \ \{U. openin (X i) U \ {x. x i \ U} \ \} \ openin (X i) (topspace (X i) \ \{U. openin (X i) U \ {x. x i \ U} \ \})" if "i \ I" for i proof - have "finite ((\U. {x. x i \ U}) -` \)" proof (rule finite_vimageI [OF \finite \\]) show "inj (\U. {x. x i \ U})" by (auto simp: inj_on_def) qed then have fin: "finite {U. openin (X i) U \ {x. x i \ U} \ \}" by (rule rev_finite_subset) auto have "openin (X i) (\ (insert (topspace (X i)) {U. openin (X i) U \ {x. x i \ U} \ \}))" by (rule openin_Inter) (auto simp: fin) then show ?thesis using \f \ \ \\ by (fastforce simp: that top) qed define \ where "\ \ \i. topspace (X i) \ \{U. openin (X i) U \ {f. f i \ U} \ \}" have "\i \ I. \x. x \ S i \ \ i" using R [OF _ *] unfolding \_def by blast then obtain \ where \ [rule_format]: "\i \ I. \ i \ S i \ \ i" by metis show "\y\Pi\<^sub>E I S. \x\\. y \ x" proof show "\U \ \. (\i \ I. \ i) \ U" proof have "restrict \ I \ PiE I (topspace \ X) \ \\" using \ by (fastforce simp: \_def PiE_def dest: \) then show "restrict \ I \ U" using subU by blast qed (rule \U \ \\) next show "(\i \ I. \ i) \ Pi\<^sub>E I S" using \ by simp qed qed qed show ?thesis apply (simp add: * closure_of_def PiE_iff set_eq_iff cong: conj_cong) by metis qed corollary closedin_product_topology: "closedin (product_topology X I) (PiE I S) \ PiE I S = {} \ (\i \ I. closedin (X i) (S i))" apply (simp add: PiE_eq PiE_eq_empty_iff closure_of_product_topology flip: closure_of_eq) apply (metis closure_of_empty) done corollary closedin_product_topology_singleton: "f \ extensional I \ closedin (product_topology X I) {f} \ (\i \ I. closedin (X i) {f i})" using PiE_singleton closedin_product_topology [of X I] by (metis (no_types, lifting) all_not_in_conv insertI1) lemma product_topology_empty: "product_topology X {} = topology (\S. S \ {{},{\k. undefined}})" unfolding product_topology union_of_def intersection_of_def arbitrary_def relative_to_def by (auto intro: arg_cong [where f=topology]) lemma openin_product_topology_empty: "openin (product_topology X {}) S \ S \ {{},{\k. undefined}}" unfolding union_of_def intersection_of_def arbitrary_def relative_to_def openin_product_topology by auto subsubsection \The basic property of the product topology is the continuity of projections:\ lemma continuous_map_product_coordinates [simp]: assumes "i \ I" shows "continuous_map (product_topology T I) (T i) (\x. x i)" proof - { fix U assume "openin (T i) U" define X where "X = (\j. if j = i then U else topspace (T j))" then have *: "(\x. x i) -` U \ (\\<^sub>E i\I. topspace (T i)) = (\\<^sub>E j\I. X j)" unfolding X_def using assms openin_subset[OF \openin (T i) U\] by (auto simp add: PiE_iff, auto, metis subsetCE) have **: "(\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}" unfolding X_def using \openin (T i) U\ by auto have "openin (product_topology T I) ((\x. x i) -` U \ (\\<^sub>E i\I. topspace (T i)))" unfolding product_topology_def apply (rule topology_generated_by_Basis) apply (subst *) using ** by auto } then show ?thesis unfolding continuous_map_alt by (auto simp add: assms PiE_iff) qed lemma continuous_map_coordinatewise_then_product [intro]: assumes "\i. i \ I \ continuous_map T1 (T i) (\x. f x i)" "\i x. i \ I \ x \ topspace T1 \ f x i = undefined" shows "continuous_map T1 (product_topology T I) f" unfolding product_topology_def proof (rule continuous_on_generated_topo) fix U assume "U \ {Pi\<^sub>E I X |X. (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}}" then obtain X where H: "U = Pi\<^sub>E I X" "\i. openin (T i) (X i)" "finite {i. X i \ topspace (T i)}" by blast define J where "J = {i \ I. X i \ topspace (T i)}" have "finite J" "J \ I" unfolding J_def using H(3) by auto have "(\x. f x i)-`(topspace(T i)) \ topspace T1 = topspace T1" if "i \ I" for i using that assms(1) by (simp add: continuous_map_preimage_topspace) then have *: "(\x. f x i)-`(X i) \ topspace T1 = topspace T1" if "i \ I-J" for i using that unfolding J_def by auto have "f-`U \ topspace T1 = (\i\I. (\x. f x i)-`(X i) \ topspace T1) \ (topspace T1)" by (subst H(1), auto simp add: PiE_iff assms) also have "... = (\i\J. (\x. f x i)-`(X i) \ topspace T1) \ (topspace T1)" using * \J \ I\ by auto also have "openin T1 (...)" apply (rule openin_INT) apply (simp add: \finite J\) using H(2) assms(1) \J \ I\ by auto ultimately show "openin T1 (f-`U \ topspace T1)" by simp next show "f `topspace T1 \ \{Pi\<^sub>E I X |X. (\i. openin (T i) (X i)) \ finite {i. X i \ topspace (T i)}}" apply (subst topology_generated_by_topspace[symmetric]) apply (subst product_topology_def[symmetric]) apply (simp only: topspace_product_topology) apply (auto simp add: PiE_iff) using assms unfolding continuous_map_def by auto qed lemma continuous_map_product_then_coordinatewise [intro]: assumes "continuous_map T1 (product_topology T I) f" shows "\i. i \ I \ continuous_map T1 (T i) (\x. f x i)" "\i x. i \ I \ x \ topspace T1 \ f x i = undefined" proof - fix i assume "i \ I" have "(\x. f x i) = (\y. y i) o f" by auto also have "continuous_map T1 (T i) (...)" apply (rule continuous_map_compose[of _ "product_topology T I"]) using assms \i \ I\ by auto ultimately show "continuous_map T1 (T i) (\x. f x i)" by simp next fix i x assume "i \ I" "x \ topspace T1" have "f x \ topspace (product_topology T I)" using assms \x \ topspace T1\ unfolding continuous_map_def by auto then have "f x \ (\\<^sub>E i\I. topspace (T i))" using topspace_product_topology by metis then show "f x i = undefined" using \i \ I\ by (auto simp add: PiE_iff extensional_def) qed lemma continuous_on_restrict: assumes "J \ I" shows "continuous_map (product_topology T I) (product_topology T J) (\x. restrict x J)" proof (rule continuous_map_coordinatewise_then_product) fix i assume "i \ J" then have "(\x. restrict x J i) = (\x. x i)" unfolding restrict_def by auto then show "continuous_map (product_topology T I) (T i) (\x. restrict x J i)" using \i \ J\ \J \ I\ by auto next fix i assume "i \ J" then show "restrict x J i = undefined" for x::"'a \ 'b" unfolding restrict_def by auto qed subsubsection\<^marker>\tag important\ \Powers of a single topological space as a topological space, using type classes\ instantiation "fun" :: (type, topological_space) topological_space begin definition\<^marker>\tag important\ open_fun_def: "open U = openin (product_topology (\i. euclidean) UNIV) U" instance proof have "topspace (product_topology (\(i::'a). euclidean::('b topology)) UNIV) = UNIV" unfolding topspace_product_topology topspace_euclidean by auto then show "open (UNIV::('a \ 'b) set)" unfolding open_fun_def by (metis openin_topspace) qed (auto simp add: open_fun_def) end lemma open_PiE [intro?]: fixes X::"'i \ ('b::topological_space) set" assumes "\i. open (X i)" "finite {i. X i \ UNIV}" shows "open (Pi\<^sub>E UNIV X)" by (simp add: assms open_fun_def product_topology_basis) lemma euclidean_product_topology: "product_topology (\i. euclidean::('b::topological_space) topology) UNIV = euclidean" by (metis open_openin topology_eq open_fun_def) proposition product_topology_basis': fixes x::"'i \ 'a" and U::"'i \ ('b::topological_space) set" assumes "finite I" "\i. i \ I \ open (U i)" shows "open {f. \i\I. f (x i) \ U i}" proof - define J where "J = x`I" define V where "V = (\y. if y \ J then \{U i|i. i\I \ x i = y} else UNIV)" define X where "X = (\y. if y \ J then V y else UNIV)" have *: "open (X i)" for i unfolding X_def V_def using assms by auto have **: "finite {i. X i \ UNIV}" unfolding X_def V_def J_def using assms(1) by auto have "open (Pi\<^sub>E UNIV X)" by (simp add: "*" "**" open_PiE) moreover have "Pi\<^sub>E UNIV X = {f. \i\I. f (x i) \ U i}" apply (auto simp add: PiE_iff) unfolding X_def V_def J_def proof (auto) fix f :: "'a \ 'b" and i :: 'i assume a1: "i \ I" assume a2: "\i. f i \ (if i \ x`I then if i \ x`I then \{U ia |ia. ia \ I \ x ia = i} else UNIV else UNIV)" have f3: "x i \ x`I" using a1 by blast have "U i \ {U ia |ia. ia \ I \ x ia = x i}" using a1 by blast then show "f (x i) \ U i" using f3 a2 by (meson Inter_iff) qed ultimately show ?thesis by simp qed text \The results proved in the general situation of products of possibly different spaces have their counterparts in this simpler setting.\ lemma continuous_on_product_coordinates [simp]: "continuous_on UNIV (\x. x i::('b::topological_space))" using continuous_map_product_coordinates [of _ UNIV "\i. euclidean"] by (metis (no_types) continuous_map_iff_continuous euclidean_product_topology iso_tuple_UNIV_I subtopology_UNIV) lemma continuous_on_coordinatewise_then_product [continuous_intros]: fixes f :: "'a::topological_space \ 'b \ 'c::topological_space" assumes "\i. continuous_on S (\x. f x i)" shows "continuous_on S f" using continuous_map_coordinatewise_then_product [of UNIV, where T = "\i. euclidean"] by (metis UNIV_I assms continuous_map_iff_continuous euclidean_product_topology) lemma continuous_on_product_then_coordinatewise_UNIV: assumes "continuous_on UNIV f" shows "continuous_on UNIV (\x. f x i)" unfolding continuous_map_iff_continuous2 [symmetric] by (rule continuous_map_product_then_coordinatewise [where I=UNIV]) (use assms in \auto simp: euclidean_product_topology\) lemma continuous_on_product_then_coordinatewise: assumes "continuous_on S f" shows "continuous_on S (\x. f x i)" proof - have "continuous_on S ((\q. q i) \ f)" by (metis assms continuous_on_compose continuous_on_id continuous_on_product_then_coordinatewise_UNIV continuous_on_subset subset_UNIV) then show ?thesis by auto qed lemma continuous_on_coordinatewise_iff: fixes f :: "('a \ real) \ 'b \ real" shows "continuous_on (A \ S) f \ (\i. continuous_on (A \ S) (\x. f x i))" by (auto simp: continuous_on_product_then_coordinatewise continuous_on_coordinatewise_then_product) lemma continuous_map_span_sum: fixes B :: "'a::real_inner set" assumes biB: "\i. i \ I \ b i \ B" shows "continuous_map euclidean (top_of_set (span B)) (\x. \i\I. x i *\<^sub>R b i)" proof (rule continuous_map_euclidean_top_of_set) show "(\x. \i\I. x i *\<^sub>R b i) -` span B = UNIV" by auto (meson biB lessThan_iff span_base span_scale span_sum) show "continuous_on UNIV (\x. \i\ I. x i *\<^sub>R b i)" by (intro continuous_intros) auto qed subsubsection\<^marker>\tag important\ \Topological countability for product spaces\ text \The next two lemmas are useful to prove first or second countability of product spaces, but they have more to do with countability and could be put in the corresponding theory.\ lemma countable_nat_product_event_const: fixes F::"'a set" and a::'a assumes "a \ F" "countable F" shows "countable {x::(nat \ 'a). (\i. x i \ F) \ finite {i. x i \ a}}" proof - have *: "{x::(nat \ 'a). (\i. x i \ F) \ finite {i. x i \ a}} \ (\N. {x. (\i. x i \ F) \ (\i\N. x i = a)})" using infinite_nat_iff_unbounded_le by fastforce have "countable {x. (\i. x i \ F) \ (\i\N. x i = a)}" for N::nat proof (induction N) case 0 have "{x. (\i. x i \ F) \ (\i\(0::nat). x i = a)} = {(\i. a)}" using \a \ F\ by auto then show ?case by auto next case (Suc N) define f::"((nat \ 'a) \ 'a) \ (nat \ 'a)" where "f = (\(x, b). (\i. if i = N then b else x i))" have "{x. (\i. x i \ F) \ (\i\Suc N. x i = a)} \ f`({x. (\i. x i \ F) \ (\i\N. x i = a)} \ F)" proof (auto) fix x assume H: "\i::nat. x i \ F" "\i\Suc N. x i = a" define y where "y = (\i. if i = N then a else x i)" have "f (y, x N) = x" unfolding f_def y_def by auto moreover have "(y, x N) \ {x. (\i. x i \ F) \ (\i\N. x i = a)} \ F" unfolding y_def using H \a \ F\ by auto ultimately show "x \ f`({x. (\i. x i \ F) \ (\i\N. x i = a)} \ F)" by (metis (no_types, lifting) image_eqI) qed moreover have "countable ({x. (\i. x i \ F) \ (\i\N. x i = a)} \ F)" using Suc.IH assms(2) by auto ultimately show ?case by (meson countable_image countable_subset) qed then show ?thesis using countable_subset[OF *] by auto qed lemma countable_product_event_const: fixes F::"('a::countable) \ 'b set" and b::'b assumes "\i. countable (F i)" shows "countable {f::('a \ 'b). (\i. f i \ F i) \ (finite {i. f i \ b})}" proof - define G where "G = (\i. F i) \ {b}" have "countable G" unfolding G_def using assms by auto have "b \ G" unfolding G_def by auto define pi where "pi = (\(x::(nat \ 'b)). (\ i::'a. x ((to_nat::('a \ nat)) i)))" have "{f::('a \ 'b). (\i. f i \ F i) \ (finite {i. f i \ b})} \ pi`{g::(nat \ 'b). (\j. g j \ G) \ (finite {j. g j \ b})}" proof (auto) fix f assume H: "\i. f i \ F i" "finite {i. f i \ b}" define I where "I = {i. f i \ b}" define g where "g = (\j. if j \ to_nat`I then f (from_nat j) else b)" have "{j. g j \ b} \ to_nat`I" unfolding g_def by auto then have "finite {j. g j \ b}" unfolding I_def using H(2) using finite_surj by blast moreover have "g j \ G" for j unfolding g_def G_def using H by auto ultimately have "g \ {g::(nat \ 'b). (\j. g j \ G) \ (finite {j. g j \ b})}" by auto moreover have "f = pi g" unfolding pi_def g_def I_def using H by fastforce ultimately show "f \ pi`{g. (\j. g j \ G) \ finite {j. g j \ b}}" by auto qed then show ?thesis using countable_nat_product_event_const[OF \b \ G\ \countable G\] by (meson countable_image countable_subset) qed instance "fun" :: (countable, first_countable_topology) first_countable_topology proof fix x::"'a \ 'b" have "\A::('b \ nat \ 'b set). \x. (\i. x \ A x i \ open (A x i)) \ (\S. open S \ x \ S \ (\i. A x i \ S))" apply (rule choice) using first_countable_basis by auto then obtain A::"('b \ nat \ 'b set)" where A: "\x i. x \ A x i" "\x i. open (A x i)" "\x S. open S \ x \ S \ (\i. A x i \ S)" by metis text \\B i\ is a countable basis of neighborhoods of \x\<^sub>i\.\ define B where "B = (\i. (A (x i))`UNIV \ {UNIV})" have "countable (B i)" for i unfolding B_def by auto have open_B: "\X i. X \ B i \ open X" by (auto simp: B_def A) define K where "K = {Pi\<^sub>E UNIV X | X. (\i. X i \ B i) \ finite {i. X i \ UNIV}}" have "Pi\<^sub>E UNIV (\i. UNIV) \ K" unfolding K_def B_def by auto then have "K \ {}" by auto have "countable {X. (\i. X i \ B i) \ finite {i. X i \ UNIV}}" apply (rule countable_product_event_const) using \\i. countable (B i)\ by auto moreover have "K = (\X. Pi\<^sub>E UNIV X)`{X. (\i. X i \ B i) \ finite {i. X i \ UNIV}}" unfolding K_def by auto ultimately have "countable K" by auto have "x \ k" if "k \ K" for k using that unfolding K_def B_def apply auto using A(1) by auto have "open k" if "k \ K" for k using that unfolding K_def by (blast intro: open_B open_PiE elim: ) have Inc: "\k\K. k \ U" if "open U \ x \ U" for U proof - have "openin (product_topology (\i. euclidean) UNIV) U" "x \ U" using \open U \ x \ U\ unfolding open_fun_def by auto with product_topology_open_contains_basis[OF this] have "\X. x \ (\\<^sub>E i\UNIV. X i) \ (\i. open (X i)) \ finite {i. X i \ UNIV} \ (\\<^sub>E i\UNIV. X i) \ U" by simp then obtain X where H: "x \ (\\<^sub>E i\UNIV. X i)" "\i. open (X i)" "finite {i. X i \ UNIV}" "(\\<^sub>E i\UNIV. X i) \ U" by auto define I where "I = {i. X i \ UNIV}" define Y where "Y = (\i. if i \ I then (SOME y. y \ B i \ y \ X i) else UNIV)" have *: "\y. y \ B i \ y \ X i" for i unfolding B_def using A(3)[OF H(2)] H(1) by (metis PiE_E UNIV_I UnCI image_iff) have **: "Y i \ B i \ Y i \ X i" for i apply (cases "i \ I") unfolding Y_def using * that apply (auto) apply (metis (no_types, lifting) someI, metis (no_types, lifting) someI_ex subset_iff) unfolding B_def apply simp unfolding I_def apply auto done have "{i. Y i \ UNIV} \ I" unfolding Y_def by auto then have ***: "finite {i. Y i \ UNIV}" unfolding I_def using H(3) rev_finite_subset by blast have "(\i. Y i \ B i) \ finite {i. Y i \ UNIV}" using ** *** by auto then have "Pi\<^sub>E UNIV Y \ K" unfolding K_def by auto have "Y i \ X i" for i apply (cases "i \ I") using ** apply simp unfolding Y_def I_def by auto then have "Pi\<^sub>E UNIV Y \ Pi\<^sub>E UNIV X" by auto then have "Pi\<^sub>E UNIV Y \ U" using H(4) by auto then show ?thesis using \Pi\<^sub>E UNIV Y \ K\ by auto qed show "\L. (\(i::nat). x \ L i \ open (L i)) \ (\U. open U \ x \ U \ (\i. L i \ U))" apply (rule first_countableI[of K]) using \countable K\ \\k. k \ K \ x \ k\ \\k. k \ K \ open k\ Inc by auto qed proposition product_topology_countable_basis: shows "\K::(('a::countable \ 'b::second_countable_topology) set set). topological_basis K \ countable K \ (\k\K. \X. (k = Pi\<^sub>E UNIV X) \ (\i. open (X i)) \ finite {i. X i \ UNIV})" proof - obtain B::"'b set set" where B: "countable B \ topological_basis B" using ex_countable_basis by auto then have "B \ {}" by (meson UNIV_I empty_iff open_UNIV topological_basisE) define B2 where "B2 = B \ {UNIV}" have "countable B2" unfolding B2_def using B by auto have "open U" if "U \ B2" for U using that unfolding B2_def using B topological_basis_open by auto define K where "K = {Pi\<^sub>E UNIV X | X. (\i::'a. X i \ B2) \ finite {i. X i \ UNIV}}" have i: "\k\K. \X. (k = Pi\<^sub>E UNIV X) \ (\i. open (X i)) \ finite {i. X i \ UNIV}" unfolding K_def using \\U. U \ B2 \ open U\ by auto have "countable {X. (\(i::'a). X i \ B2) \ finite {i. X i \ UNIV}}" apply (rule countable_product_event_const) using \countable B2\ by auto moreover have "K = (\X. Pi\<^sub>E UNIV X)`{X. (\i. X i \ B2) \ finite {i. X i \ UNIV}}" unfolding K_def by auto ultimately have ii: "countable K" by auto have iii: "topological_basis K" proof (rule topological_basisI) fix U and x::"'a\'b" assume "open U" "x \ U" then have "openin (product_topology (\i. euclidean) UNIV) U" unfolding open_fun_def by auto with product_topology_open_contains_basis[OF this \x \ U\] have "\X. x \ (\\<^sub>E i\UNIV. X i) \ (\i. open (X i)) \ finite {i. X i \ UNIV} \ (\\<^sub>E i\UNIV. X i) \ U" by simp then obtain X where H: "x \ (\\<^sub>E i\UNIV. X i)" "\i. open (X i)" "finite {i. X i \ UNIV}" "(\\<^sub>E i\UNIV. X i) \ U" by auto then have "x i \ X i" for i by auto define I where "I = {i. X i \ UNIV}" define Y where "Y = (\i. if i \ I then (SOME y. y \ B2 \ y \ X i \ x i \ y) else UNIV)" have *: "\y. y \ B2 \ y \ X i \ x i \ y" for i unfolding B2_def using B \open (X i)\ \x i \ X i\ by (meson UnCI topological_basisE) have **: "Y i \ B2 \ Y i \ X i \ x i \ Y i" for i using someI_ex[OF *] apply (cases "i \ I") unfolding Y_def using * apply (auto) unfolding B2_def I_def by auto have "{i. Y i \ UNIV} \ I" unfolding Y_def by auto then have ***: "finite {i. Y i \ UNIV}" unfolding I_def using H(3) rev_finite_subset by blast have "(\i. Y i \ B2) \ finite {i. Y i \ UNIV}" using ** *** by auto then have "Pi\<^sub>E UNIV Y \ K" unfolding K_def by auto have "Y i \ X i" for i apply (cases "i \ I") using ** apply simp unfolding Y_def I_def by auto then have "Pi\<^sub>E UNIV Y \ Pi\<^sub>E UNIV X" by auto then have "Pi\<^sub>E UNIV Y \ U" using H(4) by auto have "x \ Pi\<^sub>E UNIV Y" using ** by auto show "\V\K. x \ V \ V \ U" using \Pi\<^sub>E UNIV Y \ K\ \Pi\<^sub>E UNIV Y \ U\ \x \ Pi\<^sub>E UNIV Y\ by auto next fix U assume "U \ K" show "open U" using \U \ K\ unfolding open_fun_def K_def by clarify (metis \U \ K\ i open_PiE open_fun_def) qed show ?thesis using i ii iii by auto qed instance "fun" :: (countable, second_countable_topology) second_countable_topology apply standard using product_topology_countable_basis topological_basis_imp_subbasis by auto subsection \Metrics on product spaces\ text \In general, the product topology is not metrizable, unless the index set is countable. When the index set is countable, essentially any (convergent) combination of the metrics on the factors will do. We use below the simplest one, based on \L\<^sup>1\, but \L\<^sup>2\ would also work, for instance. What is not completely trivial is that the distance thus defined induces the same topology as the product topology. This is what we have to prove to show that we have an instance of \<^class>\metric_space\. The proofs below would work verbatim for general countable products of metric spaces. However, since distances are only implemented in terms of type classes, we only develop the theory for countable products of the same space.\ instantiation "fun" :: (countable, metric_space) metric_space begin definition\<^marker>\tag important\ dist_fun_def: "dist x y = (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" definition\<^marker>\tag important\ uniformity_fun_def: "(uniformity::(('a \ 'b) \ ('a \ 'b)) filter) = (INF e\{0<..}. principal {(x, y). dist (x::('a\'b)) y < e})" text \Except for the first one, the auxiliary lemmas below are only useful when proving the instance: once it is proved, they become trivial consequences of the general theory of metric spaces. It would thus be desirable to hide them once the instance is proved, but I do not know how to do this.\ lemma dist_fun_le_dist_first_terms: "dist x y \ 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N} + (1/2)^N" proof - have "(\n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) = (\n. (1 / 2) ^ (Suc N) * ((1/2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1))" by (rule suminf_cong, simp add: power_add) also have "... = (1/2)^(Suc N) * (\n. (1 / 2) ^ n * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1)" apply (rule suminf_mult) by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... \ (1/2)^(Suc N) * (\n. (1 / 2) ^ n)" apply (simp, rule suminf_le, simp) by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... = (1/2)^(Suc N) * 2" using suminf_geometric[of "1/2"] by auto also have "... = (1/2)^N" by auto finally have *: "(\n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) \ (1/2)^N" by simp define M where "M = Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N}" have "dist (x (from_nat 0)) (y (from_nat 0)) \ M" unfolding M_def by (rule Max_ge, auto) then have [simp]: "M \ 0" by (meson dual_order.trans zero_le_dist) have "dist (x (from_nat n)) (y (from_nat n)) \ M" if "n\N" for n unfolding M_def apply (rule Max_ge) using that by auto then have i: "min (dist (x (from_nat n)) (y (from_nat n))) 1 \ M" if "n\N" for n using that by force have "(\n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \ (\n< Suc N. M * (1 / 2) ^ n)" by (rule sum_mono, simp add: i) also have "... = M * (\n M * (\n. (1 / 2) ^ n)" by (rule mult_left_mono, rule sum_le_suminf, auto simp add: summable_geometric_iff) also have "... = M * 2" using suminf_geometric[of "1/2"] by auto finally have **: "(\n< Suc N. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1) \ 2 * M" by simp have "dist x y = (\n. (1 / 2) ^ n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" unfolding dist_fun_def by simp also have "... = (\n. (1 / 2) ^ (n+Suc N) * min (dist (x (from_nat (n+Suc N))) (y (from_nat (n+Suc N)))) 1) + (\nn. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... \ 2 * M + (1/2)^N" using * ** by auto finally show ?thesis unfolding M_def by simp qed lemma open_fun_contains_ball_aux: assumes "open (U::(('a \ 'b) set))" "x \ U" shows "\e>0. {y. dist x y < e} \ U" proof - have *: "openin (product_topology (\i. euclidean) UNIV) U" using open_fun_def assms by auto obtain X where H: "Pi\<^sub>E UNIV X \ U" "\i. openin euclidean (X i)" "finite {i. X i \ topspace euclidean}" "x \ Pi\<^sub>E UNIV X" using product_topology_open_contains_basis[OF * \x \ U\] by auto define I where "I = {i. X i \ topspace euclidean}" have "finite I" unfolding I_def using H(3) by auto { fix i have "x i \ X i" using \x \ U\ H by auto then have "\e. e>0 \ ball (x i) e \ X i" using \openin euclidean (X i)\ open_openin open_contains_ball by blast then obtain e where "e>0" "ball (x i) e \ X i" by blast define f where "f = min e (1/2)" have "f>0" "f<1" unfolding f_def using \e>0\ by auto moreover have "ball (x i) f \ X i" unfolding f_def using \ball (x i) e \ X i\ by auto ultimately have "\f. f > 0 \ f < 1 \ ball (x i) f \ X i" by auto } note * = this have "\e. \i. e i > 0 \ e i < 1 \ ball (x i) (e i) \ X i" by (rule choice, auto simp add: *) then obtain e where "\i. e i > 0" "\i. e i < 1" "\i. ball (x i) (e i) \ X i" by blast define m where "m = Min {(1/2)^(to_nat i) * e i|i. i \ I}" have "m > 0" if "I\{}" unfolding m_def Min_gr_iff using \finite I\ \I \ {}\ \\i. e i > 0\ by auto moreover have "{y. dist x y < m} \ U" proof (auto) fix y assume "dist x y < m" have "y i \ X i" if "i \ I" for i proof - have *: "summable (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) define n where "n = to_nat i" have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 < m" using \dist x y < m\ unfolding dist_fun_def using sum_le_suminf[OF *, of "{n}"] by auto then have "(1/2)^(to_nat i) * min (dist (x i) (y i)) 1 < m" using \n = to_nat i\ by auto also have "... \ (1/2)^(to_nat i) * e i" unfolding m_def apply (rule Min_le) using \finite I\ \i \ I\ by auto ultimately have "min (dist (x i) (y i)) 1 < e i" by (auto simp add: field_split_simps) then have "dist (x i) (y i) < e i" using \e i < 1\ by auto then show "y i \ X i" using \ball (x i) (e i) \ X i\ by auto qed then have "y \ Pi\<^sub>E UNIV X" using H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff) then show "y \ U" using \Pi\<^sub>E UNIV X \ U\ by auto qed ultimately have *: "\m>0. {y. dist x y < m} \ U" if "I \ {}" using that by auto have "U = UNIV" if "I = {}" using that H(1) unfolding I_def topspace_euclidean by (auto simp add: PiE_iff) then have "\m>0. {y. dist x y < m} \ U" if "I = {}" using that zero_less_one by blast then show "\m>0. {y. dist x y < m} \ U" using * by blast qed lemma ball_fun_contains_open_aux: fixes x::"('a \ 'b)" and e::real assumes "e>0" shows "\U. open U \ x \ U \ U \ {y. dist x y < e}" proof - have "\N::nat. 2^N > 8/e" by (simp add: real_arch_pow) then obtain N::nat where "2^N > 8/e" by auto define f where "f = e/4" have [simp]: "e>0" "f > 0" unfolding f_def using assms by auto define X::"('a \ 'b set)" where "X = (\i. if (\n\N. i = from_nat n) then {z. dist (x i) z < f} else UNIV)" have "{i. X i \ UNIV} \ from_nat`{0..N}" unfolding X_def by auto then have "finite {i. X i \ topspace euclidean}" unfolding topspace_euclidean using finite_surj by blast have "\i. open (X i)" unfolding X_def using metric_space_class.open_ball open_UNIV by auto then have "\i. openin euclidean (X i)" using open_openin by auto define U where "U = Pi\<^sub>E UNIV X" have "open U" unfolding open_fun_def product_topology_def apply (rule topology_generated_by_Basis) unfolding U_def using \\i. openin euclidean (X i)\ \finite {i. X i \ topspace euclidean}\ by auto moreover have "x \ U" unfolding U_def X_def by (auto simp add: PiE_iff) moreover have "dist x y < e" if "y \ U" for y proof - have *: "dist (x (from_nat n)) (y (from_nat n)) \ f" if "n \ N" for n using \y \ U\ unfolding U_def apply (auto simp add: PiE_iff) unfolding X_def using that by (metis less_imp_le mem_Collect_eq) have **: "Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N} \ f" apply (rule Max.boundedI) using * by auto have "dist x y \ 2 * Max {dist (x (from_nat n)) (y (from_nat n))|n. n \ N} + (1/2)^N" by (rule dist_fun_le_dist_first_terms) also have "... \ 2 * f + e / 8" apply (rule add_mono) using ** \2^N > 8/e\ by(auto simp add: algebra_simps field_split_simps) also have "... \ e/2 + e/8" unfolding f_def by auto also have "... < e" by auto finally show "dist x y < e" by simp qed ultimately show ?thesis by auto qed lemma fun_open_ball_aux: fixes U::"('a \ 'b) set" shows "open U \ (\x\U. \e>0. \y. dist x y < e \ y \ U)" proof (auto) assume "open U" fix x assume "x \ U" then show "\e>0. \y. dist x y < e \ y \ U" using open_fun_contains_ball_aux[OF \open U\ \x \ U\] by auto next assume H: "\x\U. \e>0. \y. dist x y < e \ y \ U" define K where "K = {V. open V \ V \ U}" { fix x assume "x \ U" then obtain e where e: "e>0" "{y. dist x y < e} \ U" using H by blast then obtain V where V: "open V" "x \ V" "V \ {y. dist x y < e}" using ball_fun_contains_open_aux[OF \e>0\, of x] by auto have "V \ K" unfolding K_def using e(2) V(1) V(3) by auto then have "x \ (\K)" using \x \ V\ by auto } then have "(\K) = U" unfolding K_def by auto moreover have "open (\K)" unfolding K_def by auto ultimately show "open U" by simp qed instance proof fix x y::"'a \ 'b" show "(dist x y = 0) = (x = y)" proof assume "x = y" then show "dist x y = 0" unfolding dist_fun_def using \x = y\ by auto next assume "dist x y = 0" have *: "summable (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) have "(1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1 = 0" for n using \dist x y = 0\ unfolding dist_fun_def by (simp add: "*" suminf_eq_zero_iff) then have "dist (x (from_nat n)) (y (from_nat n)) = 0" for n by (metis div_0 min_def nonzero_mult_div_cancel_left power_eq_0_iff zero_eq_1_divide_iff zero_neq_numeral) then have "x (from_nat n) = y (from_nat n)" for n by auto then have "x i = y i" for i by (metis from_nat_to_nat) then show "x = y" by auto qed next text\The proof of the triangular inequality is trivial, modulo the fact that we are dealing with infinite series, hence we should justify the convergence at each step. In this respect, the following simplification rule is extremely handy.\ have [simp]: "summable (\n. (1/2)^n * min (dist (u (from_nat n)) (v (from_nat n))) 1)" for u v::"'a \ 'b" by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) fix x y z::"'a \ 'b" { fix n have *: "dist (x (from_nat n)) (y (from_nat n)) \ dist (x (from_nat n)) (z (from_nat n)) + dist (y (from_nat n)) (z (from_nat n))" by (simp add: dist_triangle2) have "0 \ dist (y (from_nat n)) (z (from_nat n))" using zero_le_dist by blast moreover { assume "min (dist (y (from_nat n)) (z (from_nat n))) 1 \ dist (y (from_nat n)) (z (from_nat n))" then have "1 \ min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1" by (metis (no_types) diff_le_eq diff_self min_def zero_le_dist zero_le_one) } ultimately have "min (dist (x (from_nat n)) (y (from_nat n))) 1 \ min (dist (x (from_nat n)) (z (from_nat n))) 1 + min (dist (y (from_nat n)) (z (from_nat n))) 1" using * by linarith } note ineq = this have "dist x y = (\n. (1/2)^n * min (dist (x (from_nat n)) (y (from_nat n))) 1)" unfolding dist_fun_def by simp also have "... \ (\n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1 + (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)" apply (rule suminf_le) using ineq apply (metis (no_types, hide_lams) add.right_neutral distrib_left le_divide_eq_numeral1(1) mult_2_right mult_left_mono zero_le_one zero_le_power) by (auto simp add: summable_add) also have "... = (\n. (1/2)^n * min (dist (x (from_nat n)) (z (from_nat n))) 1) + (\n. (1/2)^n * min (dist (y (from_nat n)) (z (from_nat n))) 1)" by (rule suminf_add[symmetric], auto) also have "... = dist x z + dist y z" unfolding dist_fun_def by simp finally show "dist x y \ dist x z + dist y z" by simp next text\Finally, we show that the topology generated by the distance and the product topology coincide. This is essentially contained in Lemma \fun_open_ball_aux\, except that the condition to prove is expressed with filters. To deal with this, we copy some mumbo jumbo from Lemma \eventually_uniformity_metric\ in \<^file>\~~/src/HOL/Real_Vector_Spaces.thy\\ fix U::"('a \ 'b) set" have "eventually P uniformity \ (\e>0. \x (y::('a \ 'b)). dist x y < e \ P (x, y))" for P unfolding uniformity_fun_def apply (subst eventually_INF_base) by (auto simp: eventually_principal subset_eq intro: bexI[of _ "min _ _"]) then show "open U = (\x\U. \\<^sub>F (x', y) in uniformity. x' = x \ y \ U)" unfolding fun_open_ball_aux by simp qed (simp add: uniformity_fun_def) end text \Nice properties of spaces are preserved under countable products. In addition to first countability, second countability and metrizability, as we have seen above, completeness is also preserved, and therefore being Polish.\ instance "fun" :: (countable, complete_space) complete_space proof fix u::"nat \ ('a \ 'b)" assume "Cauchy u" have "Cauchy (\n. u n i)" for i unfolding cauchy_def proof (intro impI allI) fix e assume "e>(0::real)" obtain k where "i = from_nat k" using from_nat_to_nat[of i] by metis have "(1/2)^k * min e 1 > 0" using \e>0\ by auto then have "\N. \m n. N \ m \ N \ n \ dist (u m) (u n) < (1/2)^k * min e 1" using \Cauchy u\ unfolding cauchy_def by blast then obtain N::nat where C: "\m n. N \ m \ N \ n \ dist (u m) (u n) < (1/2)^k * min e 1" by blast have "\m n. N \ m \ N \ n \ dist (u m i) (u n i) < e" proof (auto) fix m n::nat assume "m \ N" "n \ N" have "(1/2)^k * min (dist (u m i) (u n i)) 1 = sum (\p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1) {k}" using \i = from_nat k\ by auto also have "... \ (\p. (1/2)^p * min (dist (u m (from_nat p)) (u n (from_nat p))) 1)" apply (rule sum_le_suminf) by (rule summable_comparison_test'[of "\n. (1/2)^n"], auto simp add: summable_geometric_iff) also have "... = dist (u m) (u n)" unfolding dist_fun_def by auto also have "... < (1/2)^k * min e 1" using C \m \ N\ \n \ N\ by auto finally have "min (dist (u m i) (u n i)) 1 < min e 1" by (auto simp add: algebra_simps field_split_simps) then show "dist (u m i) (u n i) < e" by auto qed then show "\N. \m n. N \ m \ N \ n \ dist (u m i) (u n i) < e" by blast qed then have "\x. (\n. u n i) \ x" for i using Cauchy_convergent convergent_def by auto then have "\x. \i. (\n. u n i) \ x i" using choice by force then obtain x where *: "\i. (\n. u n i) \ x i" by blast have "u \ x" proof (rule metric_LIMSEQ_I) fix e assume [simp]: "e>(0::real)" have i: "\K. \n\K. dist (u n i) (x i) < e/4" for i by (rule metric_LIMSEQ_D, auto simp add: *) have "\K. \i. \n\K i. dist (u n i) (x i) < e/4" apply (rule choice) using i by auto then obtain K where K: "\i n. n \ K i \ dist (u n i) (x i) < e/4" by blast have "\N::nat. 2^N > 4/e" by (simp add: real_arch_pow) then obtain N::nat where "2^N > 4/e" by auto define L where "L = Max {K (from_nat n)|n. n \ N}" have "dist (u k) x < e" if "k \ L" for k proof - have *: "dist (u k (from_nat n)) (x (from_nat n)) \ e / 4" if "n \ N" for n proof - have "K (from_nat n) \ L" unfolding L_def apply (rule Max_ge) using \n \ N\ by auto then have "k \ K (from_nat n)" using \k \ L\ by auto then show ?thesis using K less_imp_le by auto qed have **: "Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \ N} \ e/4" apply (rule Max.boundedI) using * by auto have "dist (u k) x \ 2 * Max {dist (u k (from_nat n)) (x (from_nat n)) |n. n \ N} + (1/2)^N" using dist_fun_le_dist_first_terms by auto also have "... \ 2 * e/4 + e/4" apply (rule add_mono) using ** \2^N > 4/e\ less_imp_le by (auto simp add: algebra_simps field_split_simps) also have "... < e" by auto finally show ?thesis by simp qed then show "\L. \k\L. dist (u k) x < e" by blast qed then show "convergent u" using convergent_def by blast qed instance "fun" :: (countable, polish_space) polish_space by standard subsection\The Alexander subbase theorem\ theorem Alexander_subbase: assumes X: "topology (arbitrary union_of (finite intersection_of (\x. x \ \) relative_to \\)) = X" and fin: "\C. \C \ \; \C = topspace X\ \ \C'. finite C' \ C' \ C \ \C' = topspace X" shows "compact_space X" proof - have U\: "\\ = topspace X" by (simp flip: X) have False if \: "\U\\. openin X U" and sub: "topspace X \ \\" and neg: "\\. \\ \ \; finite \\ \ \ topspace X \ \\" for \ proof - define \ where "\ \ {\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ ~(topspace X \ \\))}" have 1: "\ \ {}" unfolding \_def using sub \ neg by force have 2: "\\ \ \" if "\\{}" and \: "subset.chain \ \" for \ unfolding \_def proof (intro CollectI conjI ballI allI impI notI) show "openin X U" if U: "U \ \\" for U using U \ unfolding \_def subset_chain_def by force have "\ \ \" using subset_chain_def \ by blast with that \_def show UUC: "topspace X \ \(\\)" by blast show "False" if "finite \" and "\ \ \\" and "topspace X \ \\" for \ proof - obtain \ where "\ \ \" "\ \ \" by (metis Sup_empty \ \\ \ \\\ \finite \\ UUC empty_subsetI finite.emptyI finite_subset_Union_chain neg) then show False using \_def \\ \ \\ \finite \\ \topspace X \ \\\ by blast qed qed obtain \ where "\ \ \" and "\X. \X\\; \ \ X\ \ X = \" using subset_Zorn_nonempty [OF 1 2] by metis then have *: "\\. \\W. W\\ \ openin X W; topspace X \ \\; \ \ \; \\. \finite \; \ \ \; topspace X \ \\\ \ False\ \ \ = \" and ope: "\U\\. openin X U" and top: "topspace X \ \\" and non: "\\. \finite \; \ \ \; topspace X \ \\\ \ False" unfolding \_def by simp_all metis+ then obtain x where "x \ topspace X" "x \ \(\ \ \)" proof - have "\(\ \ \) \ \\" by (metis \\\ = topspace X\ fin inf.bounded_iff non order_refl) then have "\a. a \ \(\ \ \) \ a \ \\" by blast then show ?thesis using that by (metis U\) qed obtain C where C: "openin X C" "C \ \" "x \ C" using \x \ topspace X\ ope top by auto then have "C \ topspace X" by (metis openin_subset) then have "(arbitrary union_of (finite intersection_of (\x. x \ \) relative_to \\)) C" using openin_subbase C unfolding X [symmetric] by blast moreover have "C \ topspace X" using \\ \ \\ \C \ \\ unfolding \_def by blast ultimately obtain \ W where W: "(finite intersection_of (\x. x \ \) relative_to topspace X) W" and "x \ W" "W \ \" "\\ \ topspace X" "C = \\" using C by (auto simp: union_of_def U\) then have "\\ \ topspace X" by (metis \C \ topspace X\) then have "topspace X \ \" using \\\ \ topspace X\ by blast then obtain \' where \': "finite \'" "\' \ \" "x \ \\'" "W = topspace X \ \\'" using W \x \ W\ unfolding relative_to_def intersection_of_def by auto then have "\\' \ \\" using \W \ \\ \\\ \ topspace X\ \\\ \ topspace X\ by blast then have "\\' \ C" using U\ \C = \\\ \W = topspace X \ \\'\ \W \ \\ by auto have "\b \ \'. \C'. finite C' \ C' \ \ \ topspace X \ \(insert b C')" proof fix b assume "b \ \'" have "insert b \ = \" if neg: "\ (\C'. finite C' \ C' \ \ \ topspace X \ \(insert b C'))" proof (rule *) show "openin X W" if "W \ insert b \" for W using that proof have "b \ \" using \b \ \'\ \\' \ \\ by blast then have "\\. finite \ \ \ \ \ \ \\ = b" by (rule_tac x="{b}" in exI) auto moreover have "\\ \ b = b" using \'(2) \b \ \'\ by auto ultimately show "openin X W" if "W = b" using that \b \ \'\ apply (simp add: openin_subbase flip: X) apply (auto simp: arbitrary_def intersection_of_def relative_to_def intro!: union_of_inc) done show "openin X W" if "W \ \" by (simp add: \W \ \\ ope) qed next show "topspace X \ \ (insert b \)" using top by auto next show False if "finite \" and "\ \ insert b \" "topspace X \ \\" for \ proof - have "insert b (\ \ \) = \" using non that by blast then show False by (metis Int_lower2 finite_insert neg that(1) that(3)) qed qed auto then show "\C'. finite C' \ C' \ \ \ topspace X \ \(insert b C')" using \b \ \'\ \x \ \(\ \ \)\ \' by (metis IntI InterE Union_iff subsetD insertI1) qed then obtain F where F: "\b \ \'. finite (F b) \ F b \ \ \ topspace X \ \(insert b (F b))" by metis let ?\ = "insert C (\(F ` \'))" show False proof (rule non) have "topspace X \ (\b \ \'. \(insert b (F b)))" using F by (simp add: INT_greatest) also have "\ \ \?\" using \\\' \ C\ by force finally show "topspace X \ \?\" . show "?\ \ \" using \C \ \\ F by auto show "finite ?\" using \finite \'\ F by auto qed qed then show ?thesis by (force simp: compact_space_def compactin_def) qed corollary Alexander_subbase_alt: assumes "U \ \\" and fin: "\C. \C \ \; U \ \C\ \ \C'. finite C' \ C' \ C \ U \ \C'" and X: "topology (arbitrary union_of (finite intersection_of (\x. x \ \) relative_to U)) = X" shows "compact_space X" proof - have "topspace X = U" using X topspace_subbase by fastforce have eq: "\ (Collect ((\x. x \ \) relative_to U)) = U" unfolding relative_to_def using \U \ \\\ by blast have *: "\\. finite \ \ \ \ \ \ \\ = topspace X" if "\ \ Collect ((\x. x \ \) relative_to topspace X)" and UC: "\\ = topspace X" for \ proof - have "\ \ (\U. topspace X \ U) ` \" using that by (auto simp: relative_to_def) then obtain \' where "\' \ \" and \': "\ = (\) (topspace X) ` \'" by (auto simp: subset_image_iff) moreover have "U \ \\'" using \' \topspace X = U\ UC by auto ultimately obtain \' where "finite \'" "\' \ \'" "U \ \\'" using fin [of \'] \topspace X = U\ \U \ \\\ by blast then show ?thesis unfolding \' ex_finite_subset_image \topspace X = U\ by auto qed show ?thesis apply (rule Alexander_subbase [where \ = "Collect ((\x. x \ \) relative_to (topspace X))"]) apply (simp flip: X) apply (metis finite_intersection_of_relative_to eq) apply (blast intro: *) done qed proposition continuous_map_componentwise: "continuous_map X (product_topology Y I) f \ f ` (topspace X) \ extensional I \ (\k \ I. continuous_map X (Y k) (\x. f x k))" (is "?lhs \ _ \ ?rhs") proof (cases "\x \ topspace X. f x \ extensional I") case True then have "f ` (topspace X) \ extensional I" by force moreover have ?rhs if L: ?lhs proof - have "openin X {x \ topspace X. f x k \ U}" if "k \ I" and "openin (Y k) U" for k U proof - have "openin (product_topology Y I) ({Y. Y k \ U} \ (\\<^sub>E i\I. topspace (Y i)))" apply (simp add: openin_product_topology flip: arbitrary_union_of_relative_to) apply (simp add: relative_to_def) using that apply (blast intro: arbitrary_union_of_inc finite_intersection_of_inc) done with that have "openin X {x \ topspace X. f x \ ({Y. Y k \ U} \ (\\<^sub>E i\I. topspace (Y i)))}" using L unfolding continuous_map_def by blast moreover have "{x \ topspace X. f x \ ({Y. Y k \ U} \ (\\<^sub>E i\I. topspace (Y i)))} = {x \ topspace X. f x k \ U}" using L by (auto simp: continuous_map_def) ultimately show ?thesis by metis qed with that show ?thesis by (auto simp: continuous_map_def) qed moreover have ?lhs if ?rhs proof - have 1: "\x. x \ topspace X \ f x \ (\\<^sub>E i\I. topspace (Y i))" using that True by (auto simp: continuous_map_def PiE_iff) have 2: "{x \ S. \T\\. f x \ T} = (\T\\. {x \ S. f x \ T})" for S \ by blast have 3: "{x \ S. \U\\. f x \ U} = (\ (insert S ((\U. {x \ S. f x \ U}) ` \)))" for S \ by blast show ?thesis unfolding continuous_map_def openin_product_topology arbitrary_def proof (clarsimp simp: all_union_of 1 2) fix \ assume \: "\ \ Collect (finite intersection_of (\F. \i U. F = {f. f i \ U} \ i \ I \ openin (Y i) U) relative_to (\\<^sub>E i\I. topspace (Y i)))" show "openin X (\T\\. {x \ topspace X. f x \ T})" proof (rule openin_Union; clarify) fix S T assume "T \ \" obtain \ where "T = (\\<^sub>E i\I. topspace (Y i)) \ \\" and "finite \" "\ \ {{f. f i \ U} |i U. i \ I \ openin (Y i) U}" using subsetD [OF \ \T \ \\] by (auto simp: intersection_of_def relative_to_def) with that show "openin X {x \ topspace X. f x \ T}" apply (simp add: continuous_map_def 1 cong: conj_cong) unfolding 3 apply (rule openin_Inter; auto) done qed qed qed ultimately show ?thesis by metis next case False then show ?thesis by (auto simp: continuous_map_def PiE_def) qed lemma continuous_map_componentwise_UNIV: "continuous_map X (product_topology Y UNIV) f \ (\k. continuous_map X (Y k) (\x. f x k))" by (simp add: continuous_map_componentwise) lemma continuous_map_product_projection [continuous_intros]: "k \ I \ continuous_map (product_topology X I) (X k) (\x. x k)" using continuous_map_componentwise [of "product_topology X I" X I id] by simp declare continuous_map_from_subtopology [OF continuous_map_product_projection, continuous_intros] proposition open_map_product_projection: assumes "i \ I" shows "open_map (product_topology Y I) (Y i) (\f. f i)" unfolding openin_product_topology all_union_of arbitrary_def open_map_def image_Union proof clarify fix \ assume \: "\ \ Collect (finite intersection_of (\F. \i U. F = {f. f i \ U} \ i \ I \ openin (Y i) U) relative_to topspace (product_topology Y I))" show "openin (Y i) (\x\\. (\f. f i) ` x)" proof (rule openin_Union, clarify) fix S V assume "V \ \" obtain \ where "finite \" and V: "V = (\\<^sub>E i\I. topspace (Y i)) \ \\" and \: "\ \ {{f. f i \ U} |i U. i \ I \ openin (Y i) U}" using subsetD [OF \ \V \ \\] by (auto simp: intersection_of_def relative_to_def) show "openin (Y i) ((\f. f i) ` V)" proof (subst openin_subopen; clarify) fix x f assume "f \ V" let ?T = "{a \ topspace(Y i). (\j. if j = i then a else if j \ I then f j else undefined) \ (\\<^sub>E i\I. topspace (Y i)) \ \\}" show "\T. openin (Y i) T \ f i \ T \ T \ (\f. f i) ` V" proof (intro exI conjI) show "openin (Y i) ?T" proof (rule openin_continuous_map_preimage) have "continuous_map (Y i) (Y k) (\x. if k = i then x else f k)" if "k \ I" for k proof (cases "k=i") case True then show ?thesis by (metis (mono_tags) continuous_map_id eq_id_iff) next case False then show ?thesis by simp (metis IntD1 PiE_iff V \f \ V\ that) qed then show "continuous_map (Y i) (product_topology Y I) (\x j. if j = i then x else if j \ I then f j else undefined)" by (auto simp: continuous_map_componentwise assms extensional_def) next have "openin (product_topology Y I) (\\<^sub>E i\I. topspace (Y i))" by (metis openin_topspace topspace_product_topology) moreover have "openin (product_topology Y I) (\B\\. (\\<^sub>E i\I. topspace (Y i)) \ B)" if "\ \ {}" proof - show ?thesis proof (rule openin_Inter) show "\X. X \ (\) (\\<^sub>E i\I. topspace (Y i)) ` \ \ openin (product_topology Y I) X" unfolding openin_product_topology relative_to_def apply (clarify intro!: arbitrary_union_of_inc) apply (rename_tac F) apply (rule_tac x=F in exI) using subsetD [OF \] apply (force intro: finite_intersection_of_inc) done qed (use \finite \\ \\ \ {}\ in auto) qed ultimately show "openin (product_topology Y I) ((\\<^sub>E i\I. topspace (Y i)) \ \\)" by (auto simp only: Int_Inter_eq split: if_split) qed next have eqf: "(\j. if j = i then f i else if j \ I then f j else undefined) = f" using PiE_arb V \f \ V\ by force show "f i \ ?T" using V assms \f \ V\ by (auto simp: PiE_iff eqf) next show "?T \ (\f. f i) ` V" unfolding V by (auto simp: intro!: rev_image_eqI) qed qed qed qed lemma retraction_map_product_projection: assumes "i \ I" shows "(retraction_map (product_topology X I) (X i) (\x. x i) \ (topspace (product_topology X I) = {}) \ topspace (X i) = {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using retraction_imp_surjective_map by force next assume R: ?rhs show ?lhs proof (cases "topspace (product_topology X I) = {}") case True then show ?thesis using R by (auto simp: retraction_map_def retraction_maps_def continuous_map_on_empty) next case False have *: "\g. continuous_map (X i) (product_topology X I) g \ (\x\topspace (X i). g x i = x)" if z: "z \ (\\<^sub>E i\I. topspace (X i))" for z proof - have cm: "continuous_map (X i) (X j) (\x. if j = i then x else z j)" if "j \ I" for j using \j \ I\ z by (case_tac "j = i") auto show ?thesis using \i \ I\ that by (rule_tac x="\x j. if j = i then x else z j" in exI) (auto simp: continuous_map_componentwise PiE_iff extensional_def cm) qed show ?thesis using \i \ I\ False by (auto simp: retraction_map_def retraction_maps_def assms continuous_map_product_projection *) qed qed subsection \Open Pi-sets in the product topology\ proposition openin_PiE_gen: "openin (product_topology X I) (PiE I S) \ PiE I S = {} \ finite {i \ I. ~(S i = topspace(X i))} \ (\i \ I. openin (X i) (S i))" (is "?lhs \ _ \ ?rhs") proof (cases "PiE I S = {}") case False moreover have "?lhs = ?rhs" proof assume L: ?lhs moreover obtain z where z: "z \ PiE I S" using False by blast ultimately obtain U where fin: "finite {i \ I. U i \ topspace (X i)}" and "Pi\<^sub>E I U \ {}" and sub: "Pi\<^sub>E I U \ Pi\<^sub>E I S" by (fastforce simp add: openin_product_topology_alt) then have *: "\i. i \ I \ U i \ S i" by (simp add: subset_PiE) show ?rhs proof (intro conjI ballI) show "finite {i \ I. S i \ topspace (X i)}" apply (rule finite_subset [OF _ fin], clarify) using * by (metis False L openin_subset topspace_product_topology subset_PiE subset_antisym) next fix i :: "'a" assume "i \ I" then show "openin (X i) (S i)" using open_map_product_projection [of i I X] L apply (simp add: open_map_def) apply (drule_tac x="PiE I S" in spec) apply (simp add: False image_projection_PiE split: if_split_asm) done qed next assume ?rhs then show ?lhs apply (simp only: openin_product_topology) apply (rule arbitrary_union_of_inc) apply (auto simp: product_topology_base_alt) done qed ultimately show ?thesis by simp qed simp corollary openin_PiE: "finite I \ openin (product_topology X I) (PiE I S) \ PiE I S = {} \ (\i \ I. openin (X i) (S i))" by (simp add: openin_PiE_gen) proposition compact_space_product_topology: "compact_space(product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. compact_space(X i))" (is "?lhs = ?rhs") proof (cases "topspace(product_topology X I) = {}") case False then obtain z where z: "z \ (\\<^sub>E i\I. topspace(X i))" by auto show ?thesis proof assume L: ?lhs show ?rhs proof (clarsimp simp add: False compact_space_def) fix i assume "i \ I" with L have "continuous_map (product_topology X I) (X i) (\f. f i)" by (simp add: continuous_map_product_projection) moreover have "\x. x \ topspace (X i) \ x \ (\f. f i) ` (\\<^sub>E i\I. topspace (X i))" using \i \ I\ z apply (rule_tac x="\j. if j = i then x else if j \ I then z j else undefined" in image_eqI, auto) done then have "(\f. f i) ` (\\<^sub>E i\I. topspace (X i)) = topspace (X i)" using \i \ I\ z by auto ultimately show "compactin (X i) (topspace (X i))" by (metis L compact_space_def image_compactin topspace_product_topology) qed next assume R: ?rhs show ?lhs proof (cases "I = {}") case True with R show ?thesis by (simp add: compact_space_def) next case False then obtain i where "i \ I" by blast show ?thesis using R proof assume com [rule_format]: "\i\I. compact_space (X i)" let ?\ = "{{f. f i \ U} |i U. i \ I \ openin (X i) U}" show "compact_space (product_topology X I)" proof (rule Alexander_subbase_alt) show "topspace (product_topology X I) \ \?\" unfolding topspace_product_topology using \i \ I\ by blast next fix C assume Csub: "C \ ?\" and UC: "topspace (product_topology X I) \ \C" define \ where "\ \ \i. {U. openin (X i) U \ {f. f i \ U} \ C}" show "\C'. finite C' \ C' \ C \ topspace (product_topology X I) \ \C'" proof (cases "\i. i \ I \ topspace (X i) \ \(\ i)") case True then obtain i where "i \ I" and i: "topspace (X i) \ \(\ i)" unfolding \_def by blast then have *: "\\. \Ball \ (openin (X i)); topspace (X i) \ \\\ \ \\. finite \ \ \ \ \ \ topspace (X i) \ \\" using com [OF \i \ I\] by (auto simp: compact_space_def compactin_def) have "topspace (X i) \ \(\ i)" using i by auto with * obtain \ where "finite \ \ \ \ (\ i) \ topspace (X i) \ \\" unfolding \_def by fastforce with \i \ I\ show ?thesis unfolding \_def by (rule_tac x="(\U. {x. x i \ U}) ` \" in exI) auto next case False then have "\i \ I. \y. y \ topspace (X i) \ y \ \(\ i)" by force then obtain g where g: "\i. i \ I \ g i \ topspace (X i) \ g i \ \(\ i)" by metis then have "(\i. if i \ I then g i else undefined) \ topspace (product_topology X I)" by (simp add: PiE_I) moreover have "(\i. if i \ I then g i else undefined) \ \C" using Csub g unfolding \_def by force ultimately show ?thesis using UC by blast qed qed (simp add: product_topology) qed (simp add: compact_space_topspace_empty) qed qed qed (simp add: compact_space_topspace_empty) corollary compactin_PiE: "compactin (product_topology X I) (PiE I S) \ PiE I S = {} \ (\i \ I. compactin (X i) (S i))" by (auto simp: compactin_subspace subtopology_PiE subset_PiE compact_space_product_topology PiE_eq_empty_iff) lemma in_product_topology_closure_of: "z \ (product_topology X I) closure_of S \ i \ I \ z i \ ((X i) closure_of ((\x. x i) ` S))" using continuous_map_product_projection by (force simp: continuous_map_eq_image_closure_subset image_subset_iff) lemma homeomorphic_space_singleton_product: "product_topology X {k} homeomorphic_space (X k)" unfolding homeomorphic_space apply (rule_tac x="\x. x k" in exI) apply (rule bijective_open_imp_homeomorphic_map) apply (simp_all add: continuous_map_product_projection open_map_product_projection) unfolding PiE_over_singleton_iff apply (auto simp: image_iff inj_on_def) done subsection\Relationship with connected spaces, paths, etc.\ proposition connected_space_product_topology: "connected_space(product_topology X I) \ (\\<^sub>E i\I. topspace (X i)) = {} \ (\i \ I. connected_space(X i))" (is "?lhs \ ?eq \ ?rhs") proof (cases ?eq) case False moreover have "?lhs = ?rhs" proof assume ?lhs moreover have "connectedin(X i) (topspace(X i))" if "i \ I" and ci: "connectedin(product_topology X I) (topspace(product_topology X I))" for i proof - have cm: "continuous_map (product_topology X I) (X i) (\f. f i)" by (simp add: \i \ I\ continuous_map_product_projection) show ?thesis using connectedin_continuous_map_image [OF cm ci] \i \ I\ by (simp add: False image_projection_PiE) qed ultimately show ?rhs by (meson connectedin_topspace) next assume cs [rule_format]: ?rhs have False if disj: "U \ V = {}" and subUV: "(\\<^sub>E i\I. topspace (X i)) \ U \ V" and U: "openin (product_topology X I) U" and V: "openin (product_topology X I) V" and "U \ {}" "V \ {}" for U V proof - obtain f where "f \ U" using \U \ {}\ by blast then have f: "f \ (\\<^sub>E i\I. topspace (X i))" using U openin_subset by fastforce have "U \ topspace(product_topology X I)" "V \ topspace(product_topology X I)" using U V openin_subset by blast+ moreover have "(\\<^sub>E i\I. topspace (X i)) \ U" proof - obtain C where "(finite intersection_of (\F. \i U. F = {x. x i \ U} \ i \ I \ openin (X i) U) relative_to (\\<^sub>E i\I. topspace (X i))) C" "C \ U" "f \ C" using U \f \ U\ unfolding openin_product_topology union_of_def by auto then obtain \ where "finite \" and t: "\C. C \ \ \ \i u. (i \ I \ openin (X i) u) \ C = {x. x i \ u}" and subU: "topspace (product_topology X I) \ \\ \ U" and ftop: "f \ topspace (product_topology X I)" and fint: "f \ \ \" by (fastforce simp: relative_to_def intersection_of_def subset_iff) let ?L = "\C\\. {i. (\x. x i) ` C \ topspace (X i)}" obtain L where "finite L" and L: "\i U. \i \ I; openin (X i) U; U \ topspace(X i); {x. x i \ U} \ \\ \ i \ L" proof show "finite ?L" proof (rule finite_Union) fix M assume "M \ (\C. {i. (\x. x i) ` C \ topspace (X i)}) ` \" then obtain C where "C \ \" and C: "M = {i. (\x. x i) ` C \ topspace (X i)}" by blast then obtain j V where "j \ I" and ope: "openin (X j) V" and Ceq: "C = {x. x j \ V}" using t by meson then have "f j \ V" using \C \ \\ fint by force then have "(\x. x k) ` {x. x j \ V} = UNIV" if "k \ j" for k using that apply (clarsimp simp add: set_eq_iff) apply (rule_tac x="\m. if m = k then x else f m" in image_eqI, auto) done then have "{i. (\x. x i) ` C \ topspace (X i)} \ {j}" using Ceq by auto then show "finite M" using C finite_subset by fastforce qed (use \finite \\ in blast) next fix i U assume "i \ I" and ope: "openin (X i) U" and psub: "U \ topspace (X i)" and int: "{x. x i \ U} \ \" then show "i \ ?L" by (rule_tac a="{x. x i \ U}" in UN_I) (force+) qed show ?thesis proof fix h assume h: "h \ (\\<^sub>E i\I. topspace (X i))" define g where "g \ \i. if i \ L then f i else h i" have gin: "g \ (\\<^sub>E i\I. topspace (X i))" unfolding g_def using f h by auto moreover have "g \ X" if "X \ \" for X using fint openin_subset t [OF that] L g_def h that by fastforce ultimately have "g \ U" using subU by auto have "h \ U" if "finite M" "h \ PiE I (topspace \ X)" "{i \ I. h i \ g i} \ M" for M h using that proof (induction arbitrary: h) case empty then show ?case using PiE_ext \g \ U\ gin by force next case (insert i M) define f where "f \ \j. if j = i then g i else h j" have fin: "f \ PiE I (topspace \ X)" unfolding f_def using gin insert.prems(1) by auto have subM: "{j \ I. f j \ g j} \ M" unfolding f_def using insert.prems(2) by auto have "f \ U" using insert.IH [OF fin subM] . show ?case proof (cases "h \ V") case True show ?thesis proof (cases "i \ I") case True let ?U = "{x \ topspace(X i). (\j. if j = i then x else h j) \ U}" let ?V = "{x \ topspace(X i). (\j. if j = i then x else h j) \ V}" have False proof (rule connected_spaceD [OF cs [OF \i \ I\]]) have "\k. k \ I \ continuous_map (X i) (X k) (\x. if k = i then x else h k)" using continuous_map_eq_topcontinuous_at insert.prems(1) topcontinuous_at_def by fastforce then have cm: "continuous_map (X i) (product_topology X I) (\x j. if j = i then x else h j)" using \i \ I\ insert.prems(1) by (auto simp: continuous_map_componentwise extensional_def) show "openin (X i) ?U" by (rule openin_continuous_map_preimage [OF cm U]) show "openin (X i) ?V" by (rule openin_continuous_map_preimage [OF cm V]) show "topspace (X i) \ ?U \ ?V" proof clarsimp fix x assume "x \ topspace (X i)" and "(\j. if j = i then x else h j) \ V" with True subUV \h \ Pi\<^sub>E I (topspace \ X)\ show "(\j. if j = i then x else h j) \ U" by (drule_tac c="(\j. if j = i then x else h j)" in subsetD) auto qed show "?U \ ?V = {}" using disj by blast show "?U \ {}" using \U \ {}\ f_def proof - have "(\j. if j = i then g i else h j) \ U" using \f \ U\ f_def by blast moreover have "f i \ topspace (X i)" by (metis PiE_iff True comp_apply fin) ultimately have "\b. b \ topspace (X i) \ (\a. if a = i then b else h a) \ U" using f_def by auto then show ?thesis by blast qed have "(\j. if j = i then h i else h j) = h" by force moreover have "h i \ topspace (X i)" using True insert.prems(1) by auto ultimately show "?V \ {}" using \h \ V\ by force qed then show ?thesis .. next case False show ?thesis proof (cases "h = f") case True show ?thesis by (rule insert.IH [OF insert.prems(1)]) (simp add: True subM) next case False then show ?thesis using gin insert.prems(1) \i \ I\ unfolding f_def by fastforce qed qed next case False then show ?thesis using subUV insert.prems(1) by auto qed qed then show "h \ U" unfolding g_def using PiE_iff \finite L\ h by fastforce qed qed ultimately show ?thesis using disj inf_absorb2 \V \ {}\ by fastforce qed then show ?lhs unfolding connected_space_def by auto qed ultimately show ?thesis by simp qed (simp add: connected_space_topspace_empty) lemma connectedin_PiE: "connectedin (product_topology X I) (PiE I S) \ PiE I S = {} \ (\i \ I. connectedin (X i) (S i))" - by (fastforce simp add: connectedin_def subtopology_PiE connected_space_product_topology subset_PiE PiE_eq_empty_iff - topspace_subtopology_subset) + by (fastforce simp add: connectedin_def subtopology_PiE connected_space_product_topology subset_PiE PiE_eq_empty_iff) lemma path_connected_space_product_topology: "path_connected_space(product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. path_connected_space(X i))" (is "?lhs \ ?eq \ ?rhs") proof (cases ?eq) case False moreover have "?lhs = ?rhs" proof assume L: ?lhs show ?rhs proof (clarsimp simp flip: path_connectedin_topspace) fix i :: "'a" assume "i \ I" have cm: "continuous_map (product_topology X I) (X i) (\f. f i)" by (simp add: \i \ I\ continuous_map_product_projection) show "path_connectedin (X i) (topspace (X i))" using path_connectedin_continuous_map_image [OF cm L [unfolded path_connectedin_topspace [symmetric]]] by (metis \i \ I\ False retraction_imp_surjective_map retraction_map_product_projection) qed next assume R [rule_format]: ?rhs show ?lhs unfolding path_connected_space_def topspace_product_topology proof clarify fix x y assume x: "x \ (\\<^sub>E i\I. topspace (X i))" and y: "y \ (\\<^sub>E i\I. topspace (X i))" have "\i. \g. i \ I \ pathin (X i) g \ g 0 = x i \ g 1 = y i" using PiE_mem R path_connected_space_def x y by force then obtain g where g: "\i. i \ I \ pathin (X i) (g i) \ g i 0 = x i \ g i 1 = y i" by metis with x y show "\g. pathin (product_topology X I) g \ g 0 = x \ g 1 = y" apply (rule_tac x="\a. \i \ I. g i a" in exI) apply (force simp: pathin_def continuous_map_componentwise) done qed qed ultimately show ?thesis by simp qed (simp add: path_connected_space_topspace_empty) lemma path_connectedin_PiE: "path_connectedin (product_topology X I) (PiE I S) \ PiE I S = {} \ (\i \ I. path_connectedin (X i) (S i))" by (fastforce simp add: path_connectedin_def subtopology_PiE path_connected_space_product_topology subset_PiE PiE_eq_empty_iff topspace_subtopology_subset) subsection \Projections from a function topology to a component\ lemma quotient_map_product_projection: assumes "i \ I" shows "quotient_map(product_topology X I) (X i) (\x. x i) \ (topspace(product_topology X I) = {} \ topspace(X i) = {})" (is "?lhs = ?rhs") proof assume ?lhs with assms show ?rhs by (auto simp: continuous_open_quotient_map open_map_product_projection) next assume ?rhs with assms show ?lhs by (auto simp: Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection) qed lemma product_topology_homeomorphic_component: assumes "i \ I" "\j. \j \ I; j \ i\ \ \a. topspace(X j) = {a}" shows "product_topology X I homeomorphic_space (X i)" proof - have "quotient_map (product_topology X I) (X i) (\x. x i)" using assms by (force simp add: quotient_map_product_projection PiE_eq_empty_iff) moreover have "inj_on (\x. x i) (\\<^sub>E i\I. topspace (X i))" using assms by (auto simp: inj_on_def PiE_iff) (metis extensionalityI singletonD) ultimately show ?thesis unfolding homeomorphic_space_def by (rule_tac x="\x. x i" in exI) (simp add: homeomorphic_map_def flip: homeomorphic_map_maps) qed lemma topological_property_of_product_component: assumes major: "P (product_topology X I)" and minor: "\z i. \z \ (\\<^sub>E i\I. topspace(X i)); P(product_topology X I); i \ I\ \ P(subtopology (product_topology X I) (PiE I (\j. if j = i then topspace(X i) else {z j})))" (is "\z i. \_; _; _\ \ P (?SX z i)") and PQ: "\X X'. X homeomorphic_space X' \ (P X \ Q X')" shows "(\\<^sub>E i\I. topspace(X i)) = {} \ (\i \ I. Q(X i))" proof - have "Q(X i)" if "(\\<^sub>E i\I. topspace(X i)) \ {}" "i \ I" for i proof - from that obtain f where f: "f \ (\\<^sub>E i\I. topspace (X i))" by force have "?SX f i homeomorphic_space X i" apply (simp add: subtopology_PiE ) using product_topology_homeomorphic_component [OF \i \ I\, of "\j. subtopology (X j) (if j = i then topspace (X i) else {f j})"] using f by fastforce then show ?thesis using minor [OF f major \i \ I\] PQ by auto qed then show ?thesis by metis qed end diff --git a/src/HOL/Analysis/Further_Topology.thy b/src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy +++ b/src/HOL/Analysis/Further_Topology.thy @@ -1,5714 +1,5712 @@ section \Extending Continous Maps, Invariance of Domain, etc\ (*FIX rename? *) text\Ported from HOL Light (moretop.ml) by L C Paulson\ theory Further_Topology imports Weierstrass_Theorems Polytope Complex_Transcendental Equivalence_Lebesgue_Henstock_Integration Retracts begin subsection\A map from a sphere to a higher dimensional sphere is nullhomotopic\ lemma spheremap_lemma1: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "subspace S" "subspace T" and dimST: "dim S < dim T" and "S \ T" and diff_f: "f differentiable_on sphere 0 1 \ S" shows "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" proof assume fim: "f ` (sphere 0 1 \ S) = sphere 0 1 \ T" have inS: "\x. \x \ S; x \ 0\ \ (x /\<^sub>R norm x) \ S" using subspace_mul \subspace S\ by blast have subS01: "(\x. x /\<^sub>R norm x) ` (S - {0}) \ sphere 0 1 \ S" using \subspace S\ subspace_mul by fastforce then have diff_f': "f differentiable_on (\x. x /\<^sub>R norm x) ` (S - {0})" by (rule differentiable_on_subset [OF diff_f]) define g where "g \ \x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)" have gdiff: "g differentiable_on S - {0}" unfolding g_def by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+ have geq: "g ` (S - {0}) = T - {0}" proof have "g ` (S - {0}) \ T" apply (auto simp: g_def subspace_mul [OF \subspace T\]) apply (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \subspace T\] fim image_subset_iff inf_le2 singletonD) done moreover have "g ` (S - {0}) \ UNIV - {0}" proof (clarsimp simp: g_def) fix y assume "y \ S" and f0: "f (y /\<^sub>R norm y) = 0" then have "y \ 0 \ y /\<^sub>R norm y \ sphere 0 1 \ S" by (auto simp: subspace_mul [OF \subspace S\]) then show "y = 0" by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one) qed ultimately show "g ` (S - {0}) \ T - {0}" by auto next have *: "sphere 0 1 \ T \ f ` (sphere 0 1 \ S)" using fim by (simp add: image_subset_iff) have "x \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" if "x \ T" "x \ 0" for x proof - have "x /\<^sub>R norm x \ T" using \subspace T\ subspace_mul that by blast then show ?thesis using * [THEN subsetD, of "x /\<^sub>R norm x"] that apply clarsimp apply (rule_tac x="norm x *\<^sub>R xa" in image_eqI, simp) apply (metis norm_eq_zero right_inverse scaleR_one scaleR_scaleR) using \subspace S\ subspace_mul apply force done qed then have "T - {0} \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" by force then show "T - {0} \ g ` (S - {0})" by (simp add: g_def) qed define T' where "T' \ {y. \x \ T. orthogonal x y}" have "subspace T'" by (simp add: subspace_orthogonal_to_vectors T'_def) have dim_eq: "dim T' + dim T = DIM('a)" using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ by (simp add: T'_def) have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x by (force intro: orthogonal_subspace_decomp_exists [of T x]) then obtain p1 p2 where p1span: "p1 x \ span T" and "\w. w \ span T \ orthogonal (p2 x) w" and eq: "p1 x + p2 x = x" for x by metis then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x using span_eq_iff \subspace T\ by blast+ then have p2: "\z. p2 z \ T'" by (simp add: T'_def orthogonal_commute) have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" using span_eq_iff p2 \subspace T'\ by blast show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: span_base) then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" proof - fix c :: real and x :: 'a have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x" by (metis eq pth_6) have f2: "c *\<^sub>R p2 x \ T'" by (simp add: \subspace T'\ p2 subspace_scale) have "c *\<^sub>R p1 x \ T" by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale) then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" using f2 f1 p12_eq by presburger qed moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" by (simp add: add.assoc add.left_commute eq) show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: p1span p2 span_base span_add) ultimately have "linear p1" "linear p2" by unfold_locales auto have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" apply (rule differentiable_on_compose [where f=g]) apply (rule linear_imp_differentiable_on [OF \linear p1\]) apply (rule differentiable_on_subset [OF gdiff]) using p12_eq \S \ T\ apply auto done then have diff: "(\x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (intro derivative_intros linear_imp_differentiable_on [OF \linear p2\]) have "dim {x + y |x y. x \ S - {0} \ y \ T'} \ dim {x + y |x y. x \ S \ y \ T'}" by (blast intro: dim_subset) also have "... = dim S + dim T' - dim (S \ T')" using dim_sums_Int [OF \subspace S\ \subspace T'\] by (simp add: algebra_simps) also have "... < DIM('a)" using dimST dim_eq by auto finally have neg: "negligible {x + y |x y. x \ S - {0} \ y \ T'}" by (rule negligible_lowdim) have "negligible ((\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'})" by (rule negligible_differentiable_image_negligible [OF order_refl neg diff]) then have "negligible {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof (rule negligible_subset) have "\t' \ T'; s \ S; s \ 0\ \ g s + t' \ (\x. g (p1 x) + p2 x) ` {x + t' |x t'. x \ S \ x \ 0 \ t' \ T'}" for t' s apply (rule_tac x="s + t'" in image_eqI) using \S \ T\ p12_eq by auto then show "{x + y |x y. x \ g ` (S - {0}) \ y \ T'} \ (\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'}" by auto qed moreover have "- T' \ {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof clarsimp fix z assume "z \ T'" show "\x y. z = x + y \ x \ g ` (S - {0}) \ y \ T'" apply (rule_tac x="p1 z" in exI) apply (rule_tac x="p2 z" in exI) apply (simp add: p1 eq p2 geq) by (metis \z \ T'\ add.left_neutral eq p2) qed ultimately have "negligible (-T')" using negligible_subset by blast moreover have "negligible T'" using negligible_lowdim by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0) ultimately have "negligible (-T' \ T')" by (metis negligible_Un_eq) then show False using negligible_Un_eq non_negligible_UNIV by simp qed lemma spheremap_lemma2: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes ST: "subspace S" "subspace T" "dim S < dim T" and "S \ T" and contf: "continuous_on (sphere 0 1 \ S) f" and fim: "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" shows "\c. homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f (\x. c)" proof - have [simp]: "\x. \norm x = 1; x \ S\ \ norm (f x) = 1" using fim by (simp add: image_subset_iff) have "compact (sphere 0 1 \ S)" by (simp add: \subspace S\ closed_subspace compact_Int_closed) then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"]) using fim apply auto done have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x proof - have "norm (f x) = 1" using fim that by (simp add: image_subset_iff) then show ?thesis using g12 [OF that] by auto qed have diffg: "g differentiable_on sphere 0 1 \ S" by (metis pfg differentiable_on_polynomial_function) define h where "h \ \x. inverse(norm(g x)) *\<^sub>R g x" have h: "x \ sphere 0 1 \ S \ h x \ sphere 0 1 \ T" for x unfolding h_def using gnz [of x] by (auto simp: subspace_mul [OF \subspace T\] subsetD [OF gim]) have diffh: "h differentiable_on sphere 0 1 \ S" unfolding h_def apply (intro derivative_intros diffg differentiable_on_compose [OF diffg]) using gnz apply auto done have homfg: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) f g" proof (rule homotopic_with_linear [OF contf]) show "continuous_on (sphere 0 1 \ S) g" using pfg by (simp add: differentiable_imp_continuous_on diffg) next have non0fg: "0 \ closed_segment (f x) (g x)" if "norm x = 1" "x \ S" for x proof - have "f x \ sphere 0 1" using fim that by (simp add: image_subset_iff) moreover have "norm(f x - g x) < 1/2" apply (rule g12) using that by force ultimately show ?thesis by (auto simp: norm_minus_commute dest: segment_bound) qed show "\x. x \ sphere 0 1 \ S \ closed_segment (f x) (g x) \ T - {0}" apply (simp add: subset_Diff_insert non0fg) apply (simp add: segment_convex_hull) apply (rule hull_minimal) using fim image_eqI gim apply force apply (rule subspace_imp_convex [OF \subspace T\]) done qed obtain d where d: "d \ (sphere 0 1 \ T) - h ` (sphere 0 1 \ S)" using h spheremap_lemma1 [OF ST \S \ T\ diffh] by force then have non0hd: "0 \ closed_segment (h x) (- d)" if "norm x = 1" "x \ S" for x using midpoint_between [of 0 "h x" "-d"] that h [of x] by (auto simp: between_mem_segment midpoint_def) have conth: "continuous_on (sphere 0 1 \ S) h" using differentiable_imp_continuous_on diffh by blast have hom_hd: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) h (\x. -d)" apply (rule homotopic_with_linear [OF conth continuous_on_const]) apply (simp add: subset_Diff_insert non0hd) apply (simp add: segment_convex_hull) apply (rule hull_minimal) using h d apply (force simp: subspace_neg [OF \subspace T\]) apply (rule subspace_imp_convex [OF \subspace T\]) done have conT0: "continuous_on (T - {0}) (\y. inverse(norm y) *\<^sub>R y)" by (intro continuous_intros) auto have sub0T: "(\y. y /\<^sub>R norm y) ` (T - {0}) \ sphere 0 1 \ T" by (fastforce simp: assms(2) subspace_mul) obtain c where homhc: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. c)" apply (rule_tac c="-d" in that) apply (rule homotopic_with_eq) apply (rule homotopic_compose_continuous_left [OF hom_hd conT0 sub0T]) using d apply (auto simp: h_def) done show ?thesis apply (rule_tac x=c in exI) apply (rule homotopic_with_trans [OF _ homhc]) apply (rule homotopic_with_eq) apply (rule homotopic_compose_continuous_left [OF homfg conT0 sub0T]) apply (auto simp: h_def) done qed lemma spheremap_lemma3: assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \ dim U" obtains T where "subspace T" "T \ U" "S \ {} \ aff_dim T = aff_dim S" "(rel_frontier S) homeomorphic (sphere 0 1 \ T)" proof (cases "S = {}") case True with \subspace U\ subspace_0 show ?thesis by (rule_tac T = "{0}" in that) auto next case False then obtain a where "a \ S" by auto then have affS: "aff_dim S = int (dim ((\x. -a+x) ` S))" by (metis hull_inc aff_dim_eq_dim) with affSU have "dim ((\x. -a+x) ` S) \ dim U" by linarith with choose_subspace_of_subspace obtain T where "subspace T" "T \ span U" and dimT: "dim T = dim ((\x. -a+x) ` S)" . show ?thesis proof (rule that [OF \subspace T\]) show "T \ U" using span_eq_iff \subspace U\ \T \ span U\ by blast show "aff_dim T = aff_dim S" using dimT \subspace T\ affS aff_dim_subspace by fastforce show "rel_frontier S homeomorphic sphere 0 1 \ T" proof - have "aff_dim (ball 0 1 \ T) = aff_dim (T)" by (metis IntI interior_ball \subspace T\ aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one) then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \ T)" using \aff_dim T = aff_dim S\ by simp have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \ T)" apply (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \convex S\ \bounded S\]) apply (simp add: \subspace T\ convex_Int subspace_imp_convex) apply (simp add: bounded_Int) apply (rule affS_eq) done also have "... = frontier (ball 0 1) \ T" apply (rule convex_affine_rel_frontier_Int [OF convex_ball]) apply (simp add: \subspace T\ subspace_imp_affine) using \subspace T\ subspace_0 by force also have "... = sphere 0 1 \ T" by auto finally show ?thesis . qed qed qed proposition inessential_spheremap_lowdim_gen: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "convex S" "bounded S" "convex T" "bounded T" and affST: "aff_dim S < aff_dim T" and contf: "continuous_on (rel_frontier S) f" and fim: "f ` (rel_frontier S) \ rel_frontier T" obtains c where "homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" proof (cases "S = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis proof (cases "T = {}") case True then show ?thesis using fim that by auto next case False obtain T':: "'a set" where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) apply (simp add: aff_dim_le_DIM) using \T \ {}\ by blast with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" using homotopy_equivalent_space_sym by blast have "aff_dim S \ int (dim T')" using affT' \subspace T'\ affST aff_dim_subspace by force with spheremap_lemma3 [OF \bounded S\ \convex S\ \subspace T'\] \S \ {}\ obtain S':: "'a set" where "subspace S'" "S' \ T'" and affS': "aff_dim S' = aff_dim S" and homT: "rel_frontier S homeomorphic sphere 0 1 \ S'" by metis with homeomorphic_imp_homotopy_eqv have relS: "sphere 0 1 \ S' homotopy_eqv rel_frontier S" using homotopy_equivalent_space_sym by blast have dimST': "dim S' < dim T'" by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) have "\c. homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim]) apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format]) apply (metis dimST' \subspace S'\ \subspace T'\ \S' \ T'\ spheremap_lemma2, blast) done with that show ?thesis by blast qed qed lemma inessential_spheremap_lowdim: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis by (meson nullhomotopic_into_contractible f contractible_sphere that) next case False show ?thesis proof (cases "r \ 0") case True then show ?thesis by (meson f nullhomotopic_from_contractible contractible_sphere that) next case False with \\ s \ 0\ have "r > 0" "s > 0" by auto show ?thesis apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) using \0 < r\ \0 < s\ assms(1) apply (simp_all add: f aff_dim_cball) using that by blast qed qed subsection\ Some technical lemmas about extending maps from cell complexes\ lemma extending_maps_Union_aux: assumes fin: "finite \" and "\S. S \ \ \ closed S" and "\S T. \S \ \; T \ \; S \ T\ \ S \ T \ K" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" using assms proof (induction \) case empty show ?case by simp next case (insert S \) then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \ T" and feq: "\x \ S \ K. f x = h x" by (meson insertI1) obtain g where contg: "continuous_on (\\) g" and gim: "g ` \\ \ T" and geq: "\x \ \\ \ K. g x = h x" using insert by auto have fg: "f x = g x" if "x \ T" "T \ \" "x \ S" for x T proof - have "T \ S \ K \ S = T" using that by (metis (no_types) insert.prems(2) insertCI) then show ?thesis using UnionI feq geq \S \ \\ subsetD that by fastforce qed show ?case apply (rule_tac x="\x. if x \ S then f x else g x" in exI, simp) apply (intro conjI continuous_on_cases) apply (simp_all add: insert closed_Union contf contg) using fim gim feq geq apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+ done qed lemma extending_maps_Union: assumes fin: "finite \" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" apply (simp add: Union_maximal_sets [OF fin, symmetric]) apply (rule extending_maps_Union_aux) apply (simp_all add: Union_maximal_sets [OF fin] assms) by (metis K psubsetI) lemma extend_map_lemma: assumes "finite \" "\ \ \" "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ - \ \ aff_dim X < aff_dim T" and face: "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof (cases "\ - \ = {}") case True then have "\\ \ \\" by (simp add: Union_mono) then show ?thesis apply (rule_tac g=f in that) using contf continuous_on_subset apply blast using fim apply blast by simp next case False then have "0 \ aff_dim T" by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less) then obtain i::nat where i: "int i = aff_dim T" by (metis nonneg_eq_int) have Union_empty_eq: "\{D. D = {} \ P D} = {}" for P :: "'a set \ bool" by auto have extendf: "\g. continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) g \ g ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) \ rel_frontier T \ (\x \ \\. g x = f x)" if "i \ aff_dim T" for i::nat using that proof (induction i) case 0 then show ?case apply (simp add: Union_empty_eq) apply (rule_tac x=f in exI) apply (intro conjI) using contf continuous_on_subset apply blast using fim apply blast by simp next case (Suc p) with \bounded T\ have "rel_frontier T \ {}" by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T]) then obtain t where t: "t \ rel_frontier T" by auto have ple: "int p \ aff_dim T" using Suc.prems by force obtain h where conth: "continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) h" and him: "h ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) \ rel_frontier T" and heq: "\x. x \ \\ \ h x = f x" using Suc.IH [OF ple] by auto let ?Faces = "{D. \C \ \. D face_of C \ aff_dim D \ p}" have extendh: "\g. continuous_on D g \ g ` D \ rel_frontier T \ (\x \ D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p}). g x = h x)" if D: "D \ \ \ ?Faces" for D proof (cases "D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})") case True then show ?thesis apply (rule_tac x=h in exI) apply (intro conjI) apply (blast intro: continuous_on_subset [OF conth]) using him apply blast by simp next case False note notDsub = False show ?thesis proof (cases "\a. D = {a}") case True then obtain a where "D = {a}" by auto with notDsub t show ?thesis by (rule_tac x="\x. t" in exI) simp next case False have "D \ {}" using notDsub by auto have Dnotin: "D \ \ \ {D. \C \ \. D face_of C \ aff_dim D < p}" using notDsub by auto then have "D \ \" by simp have "D \ ?Faces - {D. \C \ \. D face_of C \ aff_dim D < p}" using Dnotin that by auto then obtain C where "C \ \" "D face_of C" and affD: "aff_dim D = int p" by auto then have "bounded D" using face_of_polytope_polytope poly polytope_imp_bounded by blast then have [simp]: "\ affine D" using affine_bounded_eq_trivial False \D \ {}\ \bounded D\ by blast have "{F. F facet_of D} \ {E. E face_of C \ aff_dim E < int p}" apply clarify apply (metis \D face_of C\ affD eq_iff face_of_trans facet_of_def zle_diff1_eq) done moreover have "polyhedron D" using \C \ \\ \D face_of C\ face_of_polytope_polytope poly polytope_imp_polyhedron by auto ultimately have relf_sub: "rel_frontier D \ \ {E. E face_of C \ aff_dim E < p}" by (simp add: rel_frontier_of_polyhedron Union_mono) then have him_relf: "h ` rel_frontier D \ rel_frontier T" using \C \ \\ him by blast have "convex D" by (simp add: \polyhedron D\ polyhedron_imp_convex) have affD_lessT: "aff_dim D < aff_dim T" using Suc.prems affD by linarith have contDh: "continuous_on (rel_frontier D) h" using \C \ \\ relf_sub by (blast intro: continuous_on_subset [OF conth]) then have *: "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c)) = (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x\rel_frontier D. g x = h x))" apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) apply (simp_all add: assms rel_frontier_eq_empty him_relf) done have "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c))" by (metis inessential_spheremap_lowdim_gen [OF \convex D\ \bounded D\ \convex T\ \bounded T\ affD_lessT contDh him_relf]) then obtain g where contg: "continuous_on UNIV g" and gim: "range g \ rel_frontier T" and gh: "\x. x \ rel_frontier D \ g x = h x" by (metis *) have "D \ E \ rel_frontier D" if "E \ \ \ {D. Bex \ ((face_of) D) \ aff_dim D < int p}" for E proof (rule face_of_subset_rel_frontier) show "D \ E face_of D" using that \C \ \\ \D face_of C\ face apply auto apply (meson face_of_Int_subface \\ \ \\ face_of_refl_eq poly polytope_imp_convex subsetD) using face_of_Int_subface apply blast done show "D \ E \ D" using that notDsub by auto qed then show ?thesis apply (rule_tac x=g in exI) apply (intro conjI ballI) using continuous_on_subset contg apply blast using gim apply blast using gh by fastforce qed qed have intle: "i < 1 + int j \ i \ int j" for i j by auto have "finite \" using \finite \\ \\ \ \\ rev_finite_subset by blast then have fin: "finite (\ \ ?Faces)" apply simp apply (rule_tac B = "\{{D. D face_of C}| C. C \ \}" in finite_subset) by (auto simp: \finite \\ finite_polytope_faces poly) have clo: "closed S" if "S \ \ \ ?Faces" for S using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "\ Y \ X" for X Y proof - have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E apply (rule face_of_Int_subface [OF _ _ XY]) apply (auto simp: face DE) done show ?thesis using that apply auto apply (drule_tac x="X \ Y" in spec, safe) using ff face_of_imp_convex [of X] face_of_imp_convex [of Y] apply (fastforce dest: face_of_aff_dim_lt) by (meson face_of_trans ff) qed obtain g where "continuous_on (\(\ \ ?Faces)) g" "g ` \(\ \ ?Faces) \ rel_frontier T" "(\x \ \(\ \ ?Faces) \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < p}). g x = h x)" apply (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+) done then show ?case apply (simp add: intle local.heq [symmetric], blast) done qed have eq: "\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i}) = \\" proof show "\(\ \ {D. \C\\. D face_of C \ aff_dim D < int i}) \ \\" apply (rule Union_subsetI) using \\ \ \\ face_of_imp_subset apply force done show "\\ \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < i})" apply (rule Union_mono) using face apply (fastforce simp: aff i) done qed have "int i \ aff_dim T" by (simp add: i) then show ?thesis using extendf [of i] unfolding eq by (metis that) qed lemma extend_map_lemma_cofinite0: assumes "finite \" and "pairwise (\S T. S \ T \ K) \" and "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" shows "\C g. finite C \ disjnt C U \ card C \ card \ \ continuous_on (\\ - C) g \ g ` (\\ - C) \ T \ (\x \ (\\ - C) \ K. g x = h x)" using assms proof induction case empty then show ?case by force next case (insert X \) then have "closed X" and clo: "\X. X \ \ \ closed X" and \: "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and pwX: "\Y. Y \ \ \ Y \ X \ X \ Y \ K \ Y \ X \ K" and pwF: "pairwise (\ S T. S \ T \ K) \" by (simp_all add: pairwise_insert) obtain C g where C: "finite C" "disjnt C U" "card C \ card \" and contg: "continuous_on (\\ - C) g" and gim: "g ` (\\ - C) \ T" and gh: "\x. x \ (\\ - C) \ K \ g x = h x" using insert.IH [OF pwF \ clo] by auto obtain a f where "a \ U" and contf: "continuous_on (X - {a}) f" and fim: "f ` (X - {a}) \ T" and fh: "(\x \ X \ K. f x = h x)" using insert.prems by (meson insertI1) show ?case proof (intro exI conjI) show "finite (insert a C)" by (simp add: C) show "disjnt (insert a C) U" using C \a \ U\ by simp show "card (insert a C) \ card (insert X \)" by (simp add: C card_insert_if insert.hyps le_SucI) have "closed (\\)" using clo insert.hyps by blast have "continuous_on (X - insert a C \ (\\ - insert a C)) (\x. if x \ X then f x else g x)" apply (rule continuous_on_cases_local) apply (simp_all add: closedin_closed) using \closed X\ apply blast using \closed (\\)\ apply blast using contf apply (force simp: elim: continuous_on_subset) using contg apply (force simp: elim: continuous_on_subset) using fh gh insert.hyps pwX by fastforce then show "continuous_on (\(insert X \) - insert a C) (\a. if a \ X then f a else g a)" by (blast intro: continuous_on_subset) show "\x\(\(insert X \) - insert a C) \ K. (if x \ X then f x else g x) = h x" using gh by (auto simp: fh) show "(\a. if a \ X then f a else g a) ` (\(insert X \) - insert a C) \ T" using fim gim by auto force qed qed lemma extend_map_lemma_cofinite1: assumes "finite \" and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" and clo: "\X. X \ \ \ closed X" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ T" "\x. x \ (\\ - C) \ K \ g x = h x" proof - let ?\ = "{X \ \. \Y\\. \ X \ Y}" have [simp]: "\?\ = \\" by (simp add: Union_maximal_sets assms) have fin: "finite ?\" by (force intro: finite_subset [OF _ \finite \\]) have pw: "pairwise (\ S T. S \ T \ K) ?\" by (simp add: pairwise_def) (metis K psubsetI) have "card {X \ \. \Y\\. \ X \ Y} \ card \" by (simp add: \finite \\ card_mono) moreover obtain C g where "finite C \ disjnt C U \ card C \ card ?\ \ continuous_on (\?\ - C) g \ g ` (\?\ - C) \ T \ (\x \ (\?\ - C) \ K. g x = h x)" apply (rule exE [OF extend_map_lemma_cofinite0 [OF fin pw, of U T h]]) apply (fastforce intro!: clo \)+ done ultimately show ?thesis by (rule_tac C=C and g=g in that) auto qed lemma extend_map_lemma_cofinite: assumes "finite \" "\ \ \" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and aff: "\X. X \ \ - \ \ aff_dim X \ aff_dim T" obtains C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\ \ - C) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof - define \ where "\ \ \ \ {D. \C \ \ - \. D face_of C \ aff_dim D < aff_dim T}" have "finite \" using assms finite_subset by blast moreover have "finite (\{{D. D face_of C} |C. C \ \})" apply (rule finite_Union) apply (simp add: \finite \\) using finite_polytope_faces poly by auto ultimately have "finite \" apply (simp add: \_def) apply (rule finite_subset [of _ "\ {{D. D face_of C} | C. C \ \}"], auto) done have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" unfolding \_def apply (elim UnE bexE CollectE DiffE) using subsetD [OF \\ \ \\] apply (simp_all add: face) apply (meson subsetD [OF \\ \ \\] face face_of_Int_subface face_of_imp_subset face_of_refl poly polytope_imp_convex)+ done obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" and hf: "\x. x \ \\ \ h x = f x" using \finite \\ unfolding \_def apply (rule extend_map_lemma [OF _ Un_upper1 T _ _ _ contf fim]) using \\ \ \\ face_of_polytope_polytope poly apply fastforce using * apply (auto simp: \_def) done have "bounded (\\)" using \finite \\ \\ \ \\ poly polytope_imp_bounded by blast then have "\\ \ UNIV" by auto then obtain a where a: "a \ \\" by blast have \: "\a g. a \ \\ \ continuous_on (D - {a}) g \ g ` (D - {a}) \ rel_frontier T \ (\x \ D \ \\. g x = h x)" if "D \ \" for D proof (cases "D \ \\") case True then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=h in exI) using him apply (blast intro!: \a \ \\\ continuous_on_subset [OF conth]) + done next case False note D_not_subset = False show ?thesis proof (cases "D \ \") case True with D_not_subset show ?thesis by (auto simp: \_def) next case False then have affD: "aff_dim D \ aff_dim T" by (simp add: \D \ \\ aff) show ?thesis proof (cases "rel_interior D = {}") case True with \D \ \\ poly a show ?thesis by (force simp: rel_interior_eq_empty polytope_imp_convex) next case False then obtain b where brelD: "b \ rel_interior D" by blast have "polyhedron D" by (simp add: poly polytope_imp_polyhedron that) have "rel_frontier D retract_of affine hull D - {b}" by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD) then obtain r where relfD: "rel_frontier D \ affine hull D - {b}" and contr: "continuous_on (affine hull D - {b}) r" and rim: "r ` (affine hull D - {b}) \ rel_frontier D" and rid: "\x. x \ rel_frontier D \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof (intro exI conjI ballI) show "b \ \\" proof clarify fix E assume "b \ E" "E \ \" then have "E \ D face_of E \ E \ D face_of D" using \\ \ \\ face that by auto with face_of_subset_rel_frontier \E \ \\ \b \ E\ brelD rel_interior_subset [of D] D_not_subset rel_frontier_def \_def show False by blast qed have "r ` (D - {b}) \ r ` (affine hull D - {b})" by (simp add: Diff_mono hull_subset image_mono) also have "... \ rel_frontier D" by (rule rim) also have "... \ \{E. E face_of D \ aff_dim E < aff_dim T}" using affD by (force simp: rel_frontier_of_polyhedron [OF \polyhedron D\] facet_of_def) also have "... \ \(\)" using D_not_subset \_def that by fastforce finally have rsub: "r ` (D - {b}) \ \(\)" . show "continuous_on (D - {b}) (h \ r)" apply (intro conjI \b \ \\\ continuous_on_compose) apply (rule continuous_on_subset [OF contr]) apply (simp add: Diff_mono hull_subset) apply (rule continuous_on_subset [OF conth rsub]) done show "(h \ r) ` (D - {b}) \ rel_frontier T" using brelD him rsub by fastforce show "(h \ r) x = h x" if x: "x \ D \ \\" for x proof - consider A where "x \ D" "A \ \" "x \ A" | A B where "x \ D" "A face_of B" "B \ \" "B \ \" "aff_dim A < aff_dim T" "x \ A" using x by (auto simp: \_def) then have xrel: "x \ rel_frontier D" proof cases case 1 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" using \A \ \\ \\ \ \\ face \D \ \\ by blast show "D \ A \ D" using \A \ \\ D_not_subset \_def by blast qed (auto simp: 1) next case 2 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" apply (rule face_of_Int_subface [of D B _ A, THEN conjunct1]) apply (simp_all add: 2 \D \ \\ face) apply (simp add: \polyhedron D\ polyhedron_imp_convex face_of_refl) done show "D \ A \ D" using "2" D_not_subset \_def by blast qed (auto simp: 2) qed show ?thesis by (simp add: rid xrel) qed qed qed qed qed have clo: "\S. S \ \ \ closed S" by (simp add: poly polytope_imp_closed) obtain C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" and gh: "\x. x \ (\\ - C) \ \\ \ g x = h x" proof (rule extend_map_lemma_cofinite1 [OF \finite \\ \ clo]) show "X \ Y \ \\" if XY: "X \ \" "Y \ \" and "\ X \ Y" "\ Y \ X" for X Y proof (cases "X \ \") case True then show ?thesis by (auto simp: \_def) next case False have "X \ Y \ X" using \\ X \ Y\ by blast with XY show ?thesis by (clarsimp simp: \_def) (metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl not_le poly polytope_imp_convex) qed qed (blast)+ with \\ \ \\ show ?thesis apply (rule_tac C=C and g=g in that) apply (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) done qed text\The next two proofs are similar\ theorem extend_map_cell_complex_to_sphere: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X < aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T - 1" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" proof (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly _ face]) show "\X. X \ \ \ aff_dim X \ aff_dim T - 1" by (simp add: aff) qed auto obtain h where conth: "continuous_on (\\) h" and him: "h ` \\ \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) qed (use \finite \\ T polyG affG faceG gim in fastforce)+ show ?thesis proof show "continuous_on (\\) h" using \\\ = \\\ conth by auto show "h ` \\ \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "x \ \\" using \\\ = \\\ \S \ \\\ that by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce have "h x = g x" apply (rule hg) using \X \ \\ \X \ V\ \x \ X\ by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed qed qed theorem extend_map_cell_complex_to_sphere_cofinite: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains C g where "finite C" "disjnt C S" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" by (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly aff face]) auto obtain C h where "finite C" and dis: "disjnt C (\(\ \ Pow V))" and card: "card C \ card \" and conth: "continuous_on (\\ - C) h" and him: "h ` (\\ - C) \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma_cofinite [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) show "g ` \(\ \ Pow V) \ rel_frontier T" using gim by force qed (auto intro: \finite \\ T polyG affG dest: faceG) have Ssub: "S \ \(\ \ Pow V)" proof fix x assume "x \ S" then have "x \ \\" using \\\ = \\\ \S \ \\\ by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce then show "x \ \(\ \ Pow V)" using \X \ \\ \x \ X\ by blast qed show ?thesis proof show "continuous_on (\\-C) h" using \\\ = \\\ conth by auto show "h ` (\\ - C) \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "h x = g x" apply (rule hg) using Ssub that by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed show "disjnt C S" using dis Ssub by (meson disjnt_iff subset_eq) qed (intro \finite C\) qed subsection\ Special cases and corollaries involving spheres\ lemma disjnt_Diff1: "X \ Y' \ disjnt (X - Y) (X' - Y')" by (auto simp: disjnt_def) proposition extend_map_affine_to_sphere_cofinite_simple: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" and aff: "aff_dim T \ aff_dim U" and "S \ T" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" obtains K g where "finite K" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - have "\K g. finite K \ disjnt K S \ continuous_on (T - K) g \ g ` (T - K) \ rel_frontier U \ (\x \ S. g x = f x)" if "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" for T proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with \bounded U\ have "aff_dim U \ 0" using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto with aff have "aff_dim T \ 0" by auto then obtain a where "T \ {a}" using \affine T\ affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto then show ?thesis using \S = {}\ fim by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset) next case False then obtain a where "a \ rel_frontier U" by auto then show ?thesis using continuous_on_const [of _ a] \S = {}\ by force qed next case False have "bounded S" by (simp add: \compact S\ compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define bbox where "bbox \ cbox (-(b+One)) (b+One)" have "cbox (-b) b \ bbox" by (auto simp: bbox_def algebra_simps intro!: subset_box_imp) with b \S \ T\ have "S \ bbox \ T" by auto then have Ssub: "S \ \{bbox \ T}" by auto then have "aff_dim (bbox \ T) \ aff_dim U" by (metis aff aff_dim_subset inf_commute inf_le1 order_trans) obtain K g where K: "finite K" "disjnt K S" and contg: "continuous_on (\{bbox \ T} - K) g" and gim: "g ` (\{bbox \ T} - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_cell_complex_to_sphere_cofinite [OF _ Ssub _ \convex U\ \bounded U\ _ _ _ contf fim]) show "closed S" using \compact S\ compact_eq_bounded_closed by auto show poly: "\X. X \ {bbox \ T} \ polytope X" by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \affine T\) show "\X Y. \X \ {bbox \ T}; Y \ {bbox \ T}\ \ X \ Y face_of X \ X \ Y face_of Y" by (simp add:poly face_of_refl polytope_imp_convex) show "\X. X \ {bbox \ T} \ aff_dim X \ aff_dim U" by (simp add: \aff_dim (bbox \ T) \ aff_dim U\) qed auto define fro where "fro \ \d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))" obtain d where d12: "1/2 \ d" "d \ 1" and dd: "disjnt K (fro d)" proof (rule disjoint_family_elem_disjnt [OF _ \finite K\]) show "infinite {1/2..1::real}" by (simp add: infinite_Icc) have dis1: "disjnt (fro x) (fro y)" if "x b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def) have clo_cbT: "closed (cbox (- c) c \ T)" by (simp add: affine_closed closed_Int closed_cbox \affine T\) have cpT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ b cbsub(2) \S \ T\ by fastforce have "closest_point (cbox (- c) c \ T) x \ K" if "x \ T" "x \ K" for x proof (cases "x \ cbox (-c) c") case True with that show ?thesis by (simp add: closest_point_self) next case False have int_ne: "interior (cbox (-c) c) \ T \ {}" using \S \ {}\ \S \ T\ b \cbox (- b) b \ box (- c) c\ by force have "convex T" by (meson \affine T\ affine_imp_convex) then have "x \ affine hull (cbox (- c) c \ T)" by (metis Int_commute Int_iff \S \ {}\ \S \ T\ cbsub(1) \x \ T\ affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox) then have "x \ affine hull (cbox (- c) c \ T) - rel_interior (cbox (- c) c \ T)" by (meson DiffI False Int_iff rel_interior_subset subsetCE) then have "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne]) moreover have "(rel_frontier (cbox (- c) c \ T)) \ fro d" apply (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) apply (auto simp: fro_def c_def) done ultimately show ?thesis using dd by (force simp: disjnt_def) qed then have cpt_subset: "closest_point (cbox (- c) c \ T) ` (T - K) \ \{bbox \ T} - K" using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force show ?thesis proof (intro conjI ballI exI) have "continuous_on (T - K) (closest_point (cbox (- c) c \ T))" apply (rule continuous_on_closest_point) using \S \ {}\ cbsub(2) b that by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \affine T\) then show "continuous_on (T - K) (g \ closest_point (cbox (- c) c \ T))" by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset]) have "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ g ` (\{bbox \ T} - K)" by (metis image_comp image_mono cpt_subset) also have "... \ rel_frontier U" by (rule gim) finally show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ rel_frontier U" . show "(g \ closest_point (cbox (- c) c \ T)) x = f x" if "x \ S" for x proof - have "(g \ closest_point (cbox (- c) c \ T)) x = g x" unfolding o_def by (metis IntI \S \ T\ b cbsub(2) closest_point_self subset_eq that) also have "... = f x" by (simp add: that gf) finally show ?thesis . qed qed (auto simp: K) qed then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (affine hull T - K) g" and gim: "g ` (affine hull T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (metis aff affine_affine_hull aff_dim_affine_hull order_trans [OF \S \ T\ hull_subset [of T affine]]) then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset) then show ?thesis by (rule_tac K="K \ T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg) qed subsection\Extending maps to spheres\ (*Up to extend_map_affine_to_sphere_cofinite_gen*) lemma extend_map_affine_to_sphere1: fixes f :: "'a::euclidean_space \ 'b::topological_space" assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" and fim: "f ` (U - K) \ T" and comps: "\C. \C \ components(U - S); C \ K \ {}\ \ C \ L \ {}" and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \ U" obtains g where "continuous_on (U - L) g" "g ` (U - L) \ T" "\x. x \ S \ g x = f x" proof (cases "K = {}") case True then show ?thesis by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that) next case False have "S \ U" using clo closedin_limpt by blast then have "(U - S) \ K \ {}" by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute) then have "\(components (U - S)) \ K \ {}" using Union_components by simp then obtain C0 where C0: "C0 \ components (U - S)" "C0 \ K \ {}" by blast have "convex U" by (simp add: affine_imp_convex \affine U\) then have "locally connected U" by (rule convex_imp_locally_connected) have "\a g. a \ C \ a \ L \ continuous_on (S \ (C - {a})) g \ g ` (S \ (C - {a})) \ T \ (\x \ S. g x = f x)" if C: "C \ components (U - S)" and CK: "C \ K \ {}" for C proof - have "C \ U-S" "C \ L \ {}" by (simp_all add: in_components_subset comps that) then obtain a where a: "a \ C" "a \ L" by auto have opeUC: "openin (top_of_set U) C" proof (rule openin_trans) show "openin (top_of_set (U-S)) C" by (simp add: \locally connected U\ clo locally_diff_closed openin_components_locally_connected [OF _ C]) show "openin (top_of_set U) (U - S)" by (simp add: clo openin_diff) qed then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" using openin_contains_cball by (metis \a \ C\) then have "ball a d \ U \ C" by auto obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" and subC: "{x. (\ (h x = x \ k x = x))} \ C" and bou: "bounded {x. (\ (h x = x \ k x = x))}" and hin: "\x. x \ C \ K \ h x \ ball a d \ U" proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) show "openin (top_of_set C) (ball a d \ U)" by (metis open_ball \C \ U\ \ball a d \ U \ C\ inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology) show "openin (top_of_set (affine hull C)) C" by (metis \a \ C\ \openin (top_of_set U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) show "ball a d \ U \ {}" using \0 < d\ \C \ U\ \a \ C\ by force show "finite (C \ K)" by (simp add: \finite K\) show "S \ C \ affine hull C" by (metis \C \ U\ \S \ U\ \a \ C\ opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff) show "connected C" by (metis C in_components_connected) qed auto have a_BU: "a \ ball a d \ U" using \0 < d\ \C \ U\ \a \ C\ by auto have "rel_frontier (cball a d \ U) retract_of (affine hull (cball a d \ U) - {a})" apply (rule rel_frontier_retract_of_punctured_affine_hull) apply (auto simp: \convex U\ convex_Int) by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) moreover have "rel_frontier (cball a d \ U) = frontier (cball a d) \ U" apply (rule convex_affine_rel_frontier_Int) using a_BU by (force simp: \affine U\)+ moreover have "affine hull (cball a d \ U) = U" by (metis \convex U\ a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \affine U\ equals0D inf.commute interior_cball) ultimately have "frontier (cball a d) \ U retract_of (U - {a})" by metis then obtain r where contr: "continuous_on (U - {a}) r" and rim: "r ` (U - {a}) \ sphere a d" "r ` (U - {a}) \ U" and req: "\x. x \ sphere a d \ U \ r x = x" using \affine U\ by (auto simp: retract_of_def retraction_def hull_same) define j where "j \ \x. if x \ ball a d then r x else x" have kj: "\x. x \ S \ k (j x) = x" using \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def subC by auto have Uaeq: "U - {a} = (cball a d - {a}) \ U \ (U - ball a d)" using \0 < d\ by auto have jim: "j ` (S \ (C - {a})) \ (S \ C) - ball a d" proof clarify fix y assume "y \ S \ (C - {a})" then have "y \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by auto then have "r y \ sphere a d" using rim by auto then show "j y \ S \ C - ball a d" apply (simp add: j_def) using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim by fastforce qed have contj: "continuous_on (U - {a}) j" unfolding j_def Uaeq proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric]) show "\T. closed T \ (cball a d - {a}) \ U = (U - {a}) \ T" apply (rule_tac x="(cball a d) \ U" in exI) using affine_closed \affine U\ by blast show "\T. closed T \ U - ball a d = (U - {a}) \ T" apply (rule_tac x="U - ball a d" in exI) using \0 < d\ by (force simp: affine_closed \affine U\ closed_Diff) show "continuous_on ((cball a d - {a}) \ U) r" by (force intro: continuous_on_subset [OF contr]) qed have fT: "x \ U - K \ f x \ T" for x using fim by blast show ?thesis proof (intro conjI exI) show "continuous_on (S \ (C - {a})) (f \ k \ j)" proof (intro continuous_on_compose) show "continuous_on (S \ (C - {a})) j" apply (rule continuous_on_subset [OF contj]) using \C \ U - S\ \S \ U\ \a \ C\ by force show "continuous_on (j ` (S \ (C - {a}))) k" apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by fastforce show "continuous_on (k ` j ` (S \ (C - {a}))) f" proof (clarify intro!: continuous_on_subset [OF contf]) fix y assume "y \ S \ (C - {a})" have ky: "k y \ S \ C" using homeomorphism_image2 [OF homhk] \y \ S \ (C - {a})\ by blast have jy: "j y \ S \ C - ball a d" using Un_iff \y \ S \ (C - {a})\ jim by auto show "k (j y) \ U - K" apply safe using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy apply blast by (metis DiffD1 DiffD2 Int_iff Un_iff \disjnt K S\ disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy) qed qed have ST: "\x. x \ S \ (f \ k \ j) x \ T" apply (simp add: kj) apply (metis DiffI \S \ U\ \disjnt K S\ subsetD disjnt_iff fim image_subset_iff) done moreover have "(f \ k \ j) x \ T" if "x \ C" "x \ a" "x \ S" for x proof - have rx: "r x \ sphere a d" using \C \ U\ rim that by fastforce have jj: "j x \ S \ C - ball a d" using jim that by blast have "k (j x) = j x \ k (j x) \ C \ j x \ C" by (metis Diff_iff Int_iff Un_iff \S \ U\ subsetD d j_def jj rx sphere_cball that(1)) then have "k (j x) \ C" using homeomorphism_apply2 [OF homhk, of "j x"] \C \ U\ \S \ U\ a rx by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC) with jj \C \ U\ show ?thesis apply safe using ST j_def apply fastforce apply (auto simp: not_less intro!: fT) by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj) qed ultimately show "(f \ k \ j) ` (S \ (C - {a})) \ T" by force show "\x\S. (f \ k \ j) x = f x" using kj by simp qed (auto simp: a) qed then obtain a h where ah: "\C. \C \ components (U - S); C \ K \ {}\ \ a C \ C \ a C \ L \ continuous_on (S \ (C - {a C})) (h C) \ h C ` (S \ (C - {a C})) \ T \ (\x \ S. h C x = f x)" using that by metis define F where "F \ {C \ components (U - S). C \ K \ {}}" define G where "G \ {C \ components (U - S). C \ K = {}}" define UF where "UF \ (\C\F. C - {a C})" have "C0 \ F" by (auto simp: F_def C0) have "finite F" proof (subst finite_image_iff [of "\C. C \ K" F, symmetric]) show "inj_on (\C. C \ K) F" unfolding F_def inj_on_def using components_nonoverlap by blast show "finite ((\C. C \ K) ` F)" unfolding F_def by (rule finite_subset [of _ "Pow K"]) (auto simp: \finite K\) qed obtain g where contg: "continuous_on (S \ UF) g" and gh: "\x i. \i \ F; x \ (S \ UF) \ (S \ (i - {a i}))\ \ g x = h i x" proof (rule pasting_lemma_exists_closed [OF \finite F\]) let ?X = "top_of_set (S \ UF)" show "topspace ?X \ (\C\F. S \ (C - {a C}))" using \C0 \ F\ by (force simp: UF_def) show "closedin (top_of_set (S \ UF)) (S \ (C - {a C}))" if "C \ F" for C proof (rule closedin_closed_subset [of U "S \ C"]) show "closedin (top_of_set U) (S \ C)" apply (rule closedin_Un_complement_component [OF \locally connected U\ clo]) using F_def that by blast next have "x = a C'" if "C' \ F" "x \ C'" "x \ U" for x C' proof - have "\A. x \ \A \ C' \ A" using \x \ C'\ by blast with that show "x = a C'" by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq) qed then show "S \ UF \ U" using \S \ U\ by (force simp: UF_def) next show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" using F_def UF_def components_nonoverlap that by auto qed show "continuous_map (subtopology ?X (S \ (C' - {a C'}))) euclidean (h C')" if "C' \ F" for C' proof - have C': "C' \ components (U - S)" "C' \ K \ {}" using F_def that by blast+ show ?thesis using ah [OF C'] by (auto simp: F_def subtopology_subtopology intro: continuous_on_subset) qed show "\i j x. \i \ F; j \ F; x \ topspace ?X \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ \ h i x = h j x" using components_eq by (fastforce simp: components_eq F_def ah) qed auto have SU': "S \ \G \ (S \ UF) \ U" using \S \ U\ in_components_subset by (auto simp: F_def G_def UF_def) have clo1: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ \G)" proof (rule closedin_closed_subset [OF _ SU']) have *: "\C. C \ F \ openin (top_of_set U) C" unfolding F_def by clarify (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology) show "closedin (top_of_set U) (U - UF)" unfolding UF_def by (force intro: openin_delete *) show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) apply (metis Diff_iff UnionI Union_components) apply (metis DiffD1 UnionI Union_components) by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) qed have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) show "closedin (top_of_set U) (\C\F. S \ C)" apply (rule closedin_Union) apply (simp add: \finite F\) using F_def \locally connected U\ clo closedin_Un_complement_component by blast show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) using C0 apply blast by (metis components_nonoverlap disjnt_def disjnt_iff) qed have SUG: "S \ \G \ U - K" using \S \ U\ K apply (auto simp: G_def disjnt_iff) by (meson Diff_iff subsetD in_components_subset) then have contf': "continuous_on (S \ \G) f" by (rule continuous_on_subset [OF contf]) have contg': "continuous_on (S \ UF) g" apply (rule continuous_on_subset [OF contg]) using \S \ U\ by (auto simp: F_def G_def) have "\x. \S \ U; x \ S\ \ f x = g x" by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" using \S \ U\ apply (auto simp: F_def G_def UF_def dest: in_components_subset) using components_eq by blast have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) show ?thesis proof have UF: "\F - L \ UF" unfolding F_def UF_def using ah by blast have "U - S - L = \(components (U - S)) - L" by simp also have "... = \F \ \G - L" unfolding F_def G_def by blast also have "... \ UF \ \G" using UF by blast finally have "U - L \ S \ \G \ (S \ UF)" by blast then show "continuous_on (U - L) (\x. if x \ S \ \G then f x else g x)" by (rule continuous_on_subset [OF cont]) have "((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ ((U - L) \ (-S \ UF))" using \U - L \ S \ \G \ (S \ UF)\ by auto moreover have "g ` ((U - L) \ (-S \ UF)) \ T" proof - have "g x \ T" if "x \ U" "x \ L" "x \ S" "C \ F" "x \ C" "x \ a C" for x C proof (subst gh) show "x \ (S \ UF) \ (S \ (C - {a C}))" using that by (auto simp: UF_def) show "h C x \ T" using ah that by (fastforce simp add: F_def) qed (rule that) then show ?thesis by (force simp: UF_def) qed ultimately have "g ` ((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ T" using image_mono order_trans by blast moreover have "f ` ((U - L) \ (S \ \G)) \ T" using fim SUG by blast ultimately show "(\x. if x \ S \ \G then f x else g x) ` (U - L) \ T" by force show "\x. x \ S \ (if x \ S \ \G then f x else g x) = f x" by (simp add: F_def G_def) qed qed lemma extend_map_affine_to_sphere2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" "affine T" "S \ T" and affTU: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and ovlap: "\C. C \ components(T - S) \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - obtain K g where K: "finite K" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" using assms extend_map_affine_to_sphere_cofinite_simple by metis have "(\y C. C \ components (T - S) \ x \ C \ y \ C \ y \ L)" if "x \ K" for x proof - have "x \ T-S" using \K \ T\ \disjnt K S\ disjnt_def that by fastforce then obtain C where "C \ components(T - S)" "x \ C" by (metis UnionE Union_components) with ovlap [of C] show ?thesis by blast qed then obtain \ where \: "\x. x \ K \ \C. C \ components (T - S) \ x \ C \ \ x \ C \ \ x \ L" by metis obtain h where conth: "continuous_on (T - \ ` K) h" and him: "h ` (T - \ ` K) \ rel_frontier U" and hg: "\x. x \ S \ h x = g x" proof (rule extend_map_affine_to_sphere1 [OF \finite K\ \affine T\ contg gim, of S "\ ` K"]) show cloTS: "closedin (top_of_set T) S" by (simp add: \compact S\ \S \ T\ closed_subset compact_imp_closed) show "\C. \C \ components (T - S); C \ K \ {}\ \ C \ \ ` K \ {}" using \ components_eq by blast qed (use K in auto) show ?thesis proof show *: "\ ` K \ L" using \ by blast show "finite (\ ` K)" by (simp add: K) show "\ ` K \ T" by clarify (meson \ Diff_iff contra_subsetD in_components_subset) show "continuous_on (T - \ ` K) h" by (rule conth) show "disjnt (\ ` K) S" using K apply (auto simp: disjnt_def) by (metis \ DiffD2 UnionI Union_components) qed (simp_all add: him hg gf) qed proposition extend_map_affine_to_sphere_cofinite_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with aff have "aff_dim T \ 0" apply (simp add: rel_frontier_eq_empty) using affine_bounded_eq_lowdim \bounded U\ order_trans by auto with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0" by linarith then show ?thesis proof cases assume "aff_dim T = -1" then have "T = {}" by (simp add: aff_dim_empty) then show ?thesis by (rule_tac K="{}" in that) auto next assume "aff_dim T = 0" then obtain a where "T = {a}" using aff_dim_eq_0 by blast then have "a \ L" using dis [of "{a}"] \S = {}\ by (auto simp: in_components_self) with \S = {}\ \T = {a}\ show ?thesis by (rule_tac K="{a}" and g=f in that) auto qed next case False then obtain y where "y \ rel_frontier U" by auto with \S = {}\ show ?thesis - by (rule_tac K="{}" and g="\x. y" in that) (auto simp: continuous_on_const) + by (rule_tac K="{}" and g="\x. y" in that) (auto) qed next case False have "bounded S" by (simp add: assms compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define LU where "LU \ L \ (\ {C \ components (T - S). \bounded C} - cbox (-(b+One)) (b+One))" obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim]) show "C \ LU \ {}" if "C \ components (T - S)" for C proof (cases "bounded C") case True with dis that show ?thesis unfolding LU_def by fastforce next case False then have "\ bounded (\{C \ components (T - S). \ bounded C})" by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) then show ?thesis apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib) by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that) qed qed blast have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" "x \ box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)" "0 \ m" "m < n" "n \ 1" for m n x using that by (auto simp: mem_box algebra_simps) have "disjoint_family_on (\d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}" by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *) then obtain d where d12: "1/2 \ d" "d \ 1" and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))" using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"] by (auto simp: \finite K\) define c where "c \ b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ cbox (-(b+One)) (b+One)" using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib) have clo_cT: "closed (cbox (- c) c \ T)" using affine_closed \affine T\ by blast have cT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ \S \ T\ b cbsub by fastforce have S_sub_cc: "S \ cbox (- c) c" using \cbox (- b) b \ cbox (- c) c\ b by auto show ?thesis proof show "finite (K \ cbox (-(b+One)) (b+One))" using \finite K\ by blast show "K \ cbox (- (b + One)) (b + One) \ L" using \K \ LU\ by (auto simp: LU_def) show "K \ cbox (- (b + One)) (b + One) \ T" using \K \ T\ by auto show "disjnt (K \ cbox (- (b + One)) (b + One)) S" using \disjnt K S\ by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1) have cloTK: "closest_point (cbox (- c) c \ T) x \ T - K" if "x \ T" and Knot: "x \ K \ x \ cbox (- b - One) (b + One)" for x proof (cases "x \ cbox (- c) c") case True with \x \ T\ show ?thesis using cbsub(3) Knot by (force simp: closest_point_self) next case False have clo_in_rf: "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI) have "T \ interior (cbox (- c) c) \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then show "x \ affine hull (cbox (- c) c \ T)" by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \affine T\ hull_inc that(1)) next show "False" if "x \ rel_interior (cbox (- c) c \ T)" proof - have "interior (cbox (- c) c) \ T \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then have "affine hull (T \ cbox (- c) c) = T" using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"] by (simp add: affine_imp_convex \affine T\ inf_commute) then show ?thesis by (meson subsetD le_inf_iff rel_interior_subset that False) qed qed have "closest_point (cbox (- c) c \ T) x \ K" proof assume inK: "closest_point (cbox (- c) c \ T) x \ K" have "\x. x \ K \ x \ frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))" by (metis ddis disjnt_iff) then show False by (metis DiffI Int_iff \affine T\ cT_ne c_def clo_cT clo_in_rf closest_point_in_set convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox) qed then show ?thesis using cT_ne clo_cT closest_point_in_set by blast qed show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]) apply (simp_all add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) using cloTK by blast have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x apply (rule gim [THEN subsetD]) using that cloTK by blast then show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K \ cbox (- (b + One)) (b + One)) \ rel_frontier U" by force show "\x. x \ S \ (g \ closest_point (cbox (- c) c \ T)) x = f x" by simp (metis (mono_tags, lifting) IntI \S \ T\ cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc) qed qed corollary extend_map_affine_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "affine T" "S \ T" and aff: "aff_dim T \ DIM('b)" and "0 \ r" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ sphere a r" "\x. x \ S \ g x = f x" proof (cases "r = 0") case True with fim show ?thesis - by (rule_tac K="{}" and g = "\x. a" in that) (auto simp: continuous_on_const) + by (rule_tac K="{}" and g = "\x. a" in that) (auto) next case False with assms have "0 < r" by auto then have "aff_dim T \ aff_dim (cball a r)" by (simp add: aff aff_dim_cball) then show ?thesis apply (rule extend_map_affine_to_sphere_cofinite_gen [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf]) using fim apply (auto simp: assms False that dest: dis) done qed corollary extend_map_UNIV_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "disjnt K S" "continuous_on (- K) g" "g ` (- K) \ sphere a r" "\x. x \ S \ g x = f x" apply (rule extend_map_affine_to_sphere_cofinite [OF \compact S\ affine_UNIV subset_UNIV _ \0 \ r\ contf fim dis]) apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric]) done corollary extend_map_UNIV_to_sphere_no_bounded_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. C \ components(- S) \ \ bounded C" obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) apply (auto simp: that dest: dis) done theorem Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "(\c \ components(- S). \bounded c) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof clarify fix f :: "'a \ 'a" assume contf: "continuous_on S f" and fim: "f ` S \ sphere 0 1" obtain g where contg: "continuous_on UNIV g" and gim: "range g \ sphere 0 1" and gf: "\x. x \ S \ g x = f x" by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \compact S\ contf fim L]) auto then obtain c where c: "homotopic_with_canon (\h. True) UNIV (sphere 0 1) g (\x. c)" using contractible_UNIV nullhomotopic_from_contractible by blast then show "\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)" by (metis assms compact_imp_closed contf contg contractible_empty fim gf gim nullhomotopic_from_contractible nullhomotopic_into_sphere_extension) qed next assume R [rule_format]: ?rhs show ?lhs unfolding components_def proof clarify fix a assume "a \ S" and a: "bounded (connected_component_set (- S) a)" have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" apply (intro continuous_intros) using \a \ S\ by auto have im: "(\x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \ sphere 0 1" by clarsimp (metis \a \ S\ eq_iff_diff_eq_0 left_inverse norm_eq_zero) show False using R cont im Borsuk_map_essential_bounded_component [OF \compact S\ \a \ S\] a by blast qed qed corollary Borsuk_separation_theorem: fixes S :: "'a::euclidean_space set" assumes "compact S" and 2: "2 \ DIM('a)" shows "connected(- S) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "S = {}") case True then show ?thesis by auto next case False then have "(\c\components (- S). \ bounded c)" by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl) then show ?thesis by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed next assume R: ?rhs then show ?lhs apply (simp add: Borsuk_separation_theorem_gen [OF \compact S\, symmetric]) apply (auto simp: components_def connected_iff_eq_connected_component_set) using connected_component_in apply fastforce using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \compact S\ compact_eq_bounded_closed by fastforce qed lemma homotopy_eqv_separation: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "S homotopy_eqv T" and "compact S" and "compact T" shows "connected(- S) \ connected(- T)" proof - consider "DIM('a) = 1" | "2 \ DIM('a)" by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq) then show ?thesis proof cases case 1 then show ?thesis using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis next case 2 with assms show ?thesis by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null) qed qed proposition Jordan_Brouwer_separation: fixes S :: "'a::euclidean_space set" and a::'a assumes hom: "S homeomorphic sphere a r" and "0 < r" shows "\ connected(- S)" proof - have "- sphere a r \ ball a r \ {}" using \0 < r\ by (simp add: Int_absorb1 subset_eq) moreover have eq: "- sphere a r - ball a r = - cball a r" by auto have "- cball a r \ {}" proof - have "frontier (cball a r) \ {}" using \0 < r\ by auto then show ?thesis by (metis frontier_complement frontier_empty) qed with eq have "- sphere a r - ball a r \ {}" by auto moreover have "connected (- S) = connected (- sphere a r)" proof (rule homotopy_eqv_separation) show "S homotopy_eqv sphere a r" using hom homeomorphic_imp_homotopy_eqv by blast show "compact (sphere a r)" by simp then show " compact S" using hom homeomorphic_compactness by blast qed ultimately show ?thesis using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \0 < r\) qed proposition Jordan_Brouwer_frontier: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and T: "T \ components(- S)" and 2: "2 \ DIM('a)" shows "frontier T = S" proof (cases r rule: linorder_cases) assume "r < 0" with S T show ?thesis by auto next assume "r = 0" with S T card_eq_SucD obtain b where "S = {b}" by (auto simp: homeomorphic_finite [of "{a}" S]) have "components (- {b}) = { -{b}}" using T \S = {b}\ by (auto simp: components_eq_sing_iff connected_punctured_universe 2) with T show ?thesis by (metis \S = {b}\ cball_trivial frontier_cball frontier_complement singletonD sphere_trivial) next assume "r > 0" have "compact S" using homeomorphic_compactness compact_sphere S by blast show ?thesis proof (rule frontier_minimal_separating_closed) show "closed S" using \compact S\ compact_eq_bounded_closed by blast show "\ connected (- S)" using Jordan_Brouwer_separation S \0 < r\ by blast obtain f g where hom: "homeomorphism S (sphere a r) f g" using S by (auto simp: homeomorphic_def) show "connected (- T)" if "closed T" "T \ S" for T proof - have "f ` T \ sphere a r" using \T \ S\ hom homeomorphism_image1 by blast moreover have "f ` T \ sphere a r" using \T \ S\ hom by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE) ultimately have "f ` T \ sphere a r" by blast then have "connected (- f ` T)" by (rule psubset_sphere_Compl_connected [OF _ \0 < r\ 2]) moreover have "compact T" using \compact S\ bounded_subset compact_eq_bounded_closed that by blast moreover then have "compact (f ` T)" by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \T \ S\) moreover have "T homotopy_eqv f ` T" by (meson \f ` T \ sphere a r\ dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \T \ S\) ultimately show ?thesis using homotopy_eqv_separation [of T "f`T"] by blast qed qed (rule T) qed proposition Jordan_Brouwer_nonseparation: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and "T \ S" and 2: "2 \ DIM('a)" shows "connected(- T)" proof - have *: "connected(C \ (S - T))" if "C \ components(- S)" for C proof (rule connected_intermediate_closure) show "connected C" using in_components_connected that by auto have "S = frontier C" using "2" Jordan_Brouwer_frontier S that by blast with closure_subset show "C \ (S - T) \ closure C" by (auto simp: frontier_def) qed auto have "components(- S) \ {}" by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere components_eq_empty homeomorphic_compactness) then have "- T = (\C \ components(- S). C \ (S - T))" using Union_components [of "-S"] \T \ S\ by auto then show ?thesis apply (rule ssubst) apply (rule connected_Union) using \T \ S\ apply (auto simp: *) done qed subsection\ Invariance of domain and corollaries\ lemma invariance_of_domain_ball: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on (cball a r) f" and "0 < r" and inj: "inj_on f (cball a r)" shows "open(f ` ball a r)" proof (cases "DIM('a) = 1") case True obtain h::"'a\real" and k where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" "\x. norm(h x) = norm x" "\x. norm(k x) = norm x" "\x. k(h x) = x" "\x. h(k x) = x" apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real]) using True apply force by (metis UNIV_I UNIV_eq_I imageI) have cont: "continuous_on S h" "continuous_on T k" for S T by (simp_all add: \linear h\ \linear k\ linear_continuous_on linear_linear) have "continuous_on (h ` cball a r) (h \ f \ k)" apply (intro continuous_on_compose cont continuous_on_subset [OF contf]) apply (auto simp: \\x. k (h x) = x\) done moreover have "is_interval (h ` cball a r)" by (simp add: is_interval_connected_1 \linear h\ linear_continuous_on linear_linear connected_continuous_image) moreover have "inj_on (h \ f \ k) (h ` cball a r)" using inj by (simp add: inj_on_def) (metis \\x. k (h x) = x\) ultimately have *: "\T. \open T; T \ h ` cball a r\ \ open ((h \ f \ k) ` T)" using injective_eq_1d_open_map_UNIV by blast have "open ((h \ f \ k) ` (h ` ball a r))" by (rule *) (auto simp: \linear h\ \range h = UNIV\ open_surjective_linear_image) then have "open ((h \ f) ` ball a r)" by (simp add: image_comp \\x. k (h x) = x\ cong: image_cong) then show ?thesis apply (simp only: image_comp [symmetric]) apply (metis open_bijective_linear_image_eq \linear h\ \\x. k (h x) = x\ \range h = UNIV\ bijI inj_on_def) done next case False then have 2: "DIM('a) \ 2" by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq) have fimsub: "f ` ball a r \ - f ` sphere a r" using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl) have hom: "f ` sphere a r homeomorphic sphere a r" by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball) then have nconn: "\ connected (- f ` sphere a r)" by (rule Jordan_Brouwer_separation) (auto simp: \0 < r\) obtain C where C: "C \ components (- f ` sphere a r)" and "bounded C" apply (rule cobounded_has_bounded_component [OF _ nconn]) apply (simp_all add: 2) by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball) moreover have "f ` (ball a r) = C" proof have "C \ {}" by (rule in_components_nonempty [OF C]) show "C \ f ` ball a r" proof (rule ccontr) assume nonsub: "\ C \ f ` ball a r" have "- f ` cball a r \ C" proof (rule components_maximal [OF C]) have "f ` cball a r homeomorphic cball a r" using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast then show "connected (- f ` cball a r)" by (auto intro: connected_complement_homeomorphic_convex_compact 2) show "- f ` cball a r \ - f ` sphere a r" by auto then show "C \ - f ` cball a r \ {}" using \C \ {}\ in_components_subset [OF C] nonsub using image_iff by fastforce qed then have "bounded (- f ` cball a r)" using bounded_subset \bounded C\ by auto then have "\ bounded (f ` cball a r)" using cobounded_imp_unbounded by blast then show "False" using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast qed with \C \ {}\ have "C \ f ` ball a r \ {}" by (simp add: inf.absorb_iff1) then show "f ` ball a r \ C" by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset) qed moreover have "open (- f ` sphere a r)" using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast ultimately show ?thesis using open_components by blast qed text\Proved by L. E. J. Brouwer (1912)\ theorem invariance_of_domain: fixes f :: "'a \ 'a::euclidean_space" assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)" unfolding open_subopen [of "f`S"] proof clarify fix a assume "a \ S" obtain \ where "\ > 0" and \: "cball a \ \ S" using \open S\ \a \ S\ open_contains_cball_eq by blast show "\T. open T \ f a \ T \ T \ f ` S" proof (intro exI conjI) show "open (f ` (ball a \))" by (meson \ \0 < \\ assms continuous_on_subset inj_on_subset invariance_of_domain_ball) show "f a \ f ` ball a \" by (simp add: \0 < \\) show "f ` ball a \ \ f ` S" using \ ball_subset_cball by blast qed qed lemma inv_of_domain_ss0: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - have "U \ S" using ope openin_imp_subset by blast have "(UNIV::'b set) homeomorphic S" by (simp add: \subspace S\ dimS homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" using homeomorphic_def by blast have homkh: "homeomorphism S (k ` S) k h" using homhk homeomorphism_image2 homeomorphism_sym by fastforce have "open ((k \ f \ h) ` k ` U)" proof (rule invariance_of_domain) show "continuous_on (k ` U) (k \ f \ h)" proof (intro continuous_intros) show "continuous_on (k ` U) h" by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest) show "continuous_on (h ` k ` U) f" apply (rule continuous_on_subset [OF contf], clarify) apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD) done show "continuous_on (f ` h ` k ` U) k" apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) using fim homhk homeomorphism_apply2 ope openin_subset by fastforce qed have ope_iff: "\T. open T \ openin (top_of_set (k ` S)) T" using homhk homeomorphism_image2 open_openin by fastforce show "open (k ` U)" by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) show "inj_on (k \ f \ h) (k ` U)" apply (clarsimp simp: inj_on_def) by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \U \ S\) qed moreover have eq: "f ` U = h ` (k \ f \ h \ k) ` U" unfolding image_comp [symmetric] using \U \ S\ fim by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff) ultimately show ?thesis by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV) qed lemma inv_of_domain_ss1: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - define S' where "S' \ {y. \x \ S. orthogonal x y}" have "subspace S'" by (simp add: S'_def subspace_orthogonal_to_vectors) define g where "g \ \z::'a*'a. ((f \ fst)z, snd z)" have "openin (top_of_set (S \ S')) (g ` (U \ S'))" proof (rule inv_of_domain_ss0) show "continuous_on (U \ S') g" apply (simp add: g_def) apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto) done show "g ` (U \ S') \ S \ S'" using fim by (auto simp: g_def) show "inj_on g (U \ S')" using injf by (auto simp: g_def inj_on_def) show "subspace (S \ S')" by (simp add: \subspace S'\ \subspace S\ subspace_Times) show "openin (top_of_set (S \ S')) (U \ S')" by (simp add: openin_Times [OF ope]) have "dim (S \ S') = dim S + dim S'" by (simp add: \subspace S'\ \subspace S\ dim_Times) also have "... = DIM('a)" using dim_subspace_orthogonal_to_vectors [OF \subspace S\ subspace_UNIV] by (simp add: add.commute S'_def) finally show "dim (S \ S') = DIM('a)" . qed moreover have "g ` (U \ S') = f ` U \ S'" by (auto simp: g_def image_iff) moreover have "0 \ S'" using \subspace S'\ subspace_affine by blast ultimately show ?thesis by (auto simp: openin_Times_eq) qed corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and VU: "dim V \ dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof - obtain V' where "subspace V'" "V' \ U" "dim V' = dim V" using choose_subspace_of_subspace [OF VU] by (metis span_eq_iff \subspace U\) then have "V homeomorphic V'" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V V' h k" using homeomorphic_def by blast have eq: "f ` S = k ` (h \ f) ` S" proof - have "k ` h ` f ` S = f ` S" by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl) then show ?thesis by (simp add: image_comp) qed show ?thesis unfolding eq proof (rule homeomorphism_imp_open_map) show homkh: "homeomorphism V' V k h" by (simp add: homeomorphism_symD homhk) have hfV': "(h \ f) ` S \ V'" using fim homeomorphism_image1 homhk by fastforce moreover have "openin (top_of_set U) ((h \ f) ` S)" proof (rule inv_of_domain_ss1) show "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) show "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf) show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) ultimately show "openin (top_of_set V') ((h \ f) ` S)" using openin_subset_trans \V' \ U\ by force qed qed corollary invariance_of_dimension_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "dim U \ dim V" proof - have "False" if "dim V < dim U" proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" using homeomorphic_def by blast have "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" using fim homeomorphism_image1 homhk by fastforce then have "dim ((h \ f) ` S) \ dim T" by (rule dim_subset) also have "dim ((h \ f) ` S) = dim U" using \S \ {}\ \subspace U\ by (blast intro: dim_openin ope_hf) finally show False using \dim V < dim U\ \dim T = dim V\ by simp qed then show ?thesis using not_less by blast qed corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" "aff_dim V \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case True then show ?thesis by auto next case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "dim ((+) (- b) ` V) \ dim ((+) (- a) ` U)" by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed then show ?thesis by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) qed corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "aff_dim U \ aff_dim V" proof - obtain a b where "a \ S" "a \ U" "b \ V" using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed (use \S \ {}\ in auto) then show ?thesis by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) qed corollary invariance_of_dimension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and "open S" and injf: "inj_on f S" and "S \ {}" shows "DIM('a) \ DIM('b)" using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms by auto corollary continuous_injective_image_subspace_dim_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace S" "subspace T" and contf: "continuous_on S f" and fim: "f ` S \ T" and injf: "inj_on f S" shows "dim S \ dim T" apply (rule invariance_of_dimension_subspaces [of S S _ f]) using assms by (auto simp: subspace_affine) lemma invariance_of_dimension_convex_domain: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "convex S" and contf: "continuous_on S f" and fim: "f ` S \ affine hull T" and injf: "inj_on f S" shows "aff_dim S \ aff_dim T" proof (cases "S = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" using fim rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast show "rel_interior S \ {}" by (simp add: False \convex S\ rel_interior_eq_empty) qed auto then show ?thesis by simp qed lemma homeomorphic_convex_sets_le: assumes "convex S" "S homeomorphic T" shows "aff_dim S \ aff_dim T" proof - obtain h k where homhk: "homeomorphism S T h k" using homeomorphic_def assms by blast show ?thesis proof (rule invariance_of_dimension_convex_domain [OF \convex S\]) show "continuous_on S h" using homeomorphism_def homhk by blast show "h ` S \ affine hull T" by (metis homeomorphism_def homhk hull_subset) show "inj_on h S" by (meson homeomorphism_apply1 homhk inj_on_inverseI) qed qed lemma homeomorphic_convex_sets: assumes "convex S" "convex T" "S homeomorphic T" shows "aff_dim S = aff_dim T" by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) lemma homeomorphic_convex_compact_sets_eq: assumes "convex S" "compact S" "convex T" "compact T" shows "S homeomorphic T \ aff_dim S = aff_dim T" by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) lemma invariance_of_domain_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" shows "open (f ` T)" apply (rule invariance_of_domain_gen [OF \open T\]) using assms apply (auto simp: elim: continuous_on_subset subset_inj_on) done lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" and gf: "\x. x \ S \ g(f x) = x" shows "continuous_on (f ` S) g" proof (clarsimp simp add: continuous_openin_preimage_eq) fix T :: "'a set" assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" apply (subst eq) apply (rule open_openin_trans) apply (rule invariance_of_domain_gen) using assms apply auto using inj_on_inverseI apply auto[1] by (metis \open T\ continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq) qed lemma invariance_of_domain_homeomorphism: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" obtains g where "homeomorphism S (f ` S) f g" proof show "homeomorphism S (f ` S) f (inv_into S f)" by (simp add: assms continuous_on_inverse_open homeomorphism_def) qed corollary invariance_of_domain_homeomorphic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" shows "S homeomorphic (f ` S)" using invariance_of_domain_homeomorphism [OF assms] by (meson homeomorphic_def) lemma continuous_image_subset_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "f ` (interior S) \ interior(f ` S)" apply (rule interior_maximal) apply (simp add: image_mono interior_subset) apply (rule invariance_of_domain_gen) using assms apply (auto simp: subset_inj_on interior_subset continuous_on_subset) done lemma homeomorphic_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)" shows "(interior S) homeomorphic (interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp have gim: "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp show "homeomorphism (interior S) (interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show "\x. x \ interior S \ g (f x) = x" by (meson \\x\S. f x \ T \ g (f x) = x\ subsetD interior_subset) have "interior T \ f ` interior S" proof fix x assume "x \ interior T" then have "g x \ interior S" using gim by blast then show "x \ f ` interior S" by (metis T \x \ interior T\ image_iff interior_subset subsetCE) qed then show "f ` interior S = interior T" using fim by blast show "continuous_on (interior S) f" by (metis interior_subset continuous_on_subset contf) show "\y. y \ interior T \ f (g y) = y" by (meson T subsetD interior_subset) have "interior S \ g ` interior T" proof fix x assume "x \ interior S" then have "f x \ interior T" using fim by blast then show "x \ g ` interior T" by (metis S \x \ interior S\ image_iff interior_subset subsetCE) qed then show "g ` interior T = interior S" using gim by blast show "continuous_on (interior T) g" by (metis interior_subset continuous_on_subset contg) qed qed lemma homeomorphic_open_imp_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done proposition homeomorphic_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") case True with assms show ?thesis by auto next case False then have "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior) done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)" shows "(frontier S) homeomorphic (frontier T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp then have fim: "f ` frontier S \ frontier T" apply (simp add: frontier_def) using continuous_image_subset_interior assms(2) assms(3) S by auto have "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp then have gim: "g ` frontier T \ frontier S" apply (simp add: frontier_def) using continuous_image_subset_interior T assms(2) assms(3) by auto show "homeomorphism (frontier S) (frontier T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ frontier S \ g (f x) = x" by (simp add: S assms(2) frontier_def) show fg: "\y. y \ frontier T \ f (g y) = y" by (simp add: T assms(3) frontier_def) have "frontier T \ f ` frontier S" proof fix x assume "x \ frontier T" then have "g x \ frontier S" using gim by blast then show "x \ f ` frontier S" by (metis fg \x \ frontier T\ imageI) qed then show "f ` frontier S = frontier T" using fim by blast show "continuous_on (frontier S) f" by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def) have "frontier S \ g ` frontier T" proof fix x assume "x \ frontier S" then have "f x \ frontier T" using fim by blast then show "x \ g ` frontier T" by (metis gf \x \ frontier S\ imageI) qed then show "g ` frontier T = frontier S" using gim by blast show "continuous_on (frontier T) g" by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def) qed qed lemma homeomorphic_frontiers: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" "interior S = {} \ interior T = {}" shows "(frontier S) homeomorphic (frontier T)" proof (cases "interior T = {}") case True then show ?thesis by (metis Diff_empty assms closure_eq frontier_def) next case False show ?thesis apply (rule homeomorphic_frontiers_same_dimension) apply (simp_all add: assms) using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast qed lemma continuous_image_subset_rel_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and TS: "aff_dim T \ aff_dim S" shows "f ` (rel_interior S) \ rel_interior(f ` S)" proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "aff_dim (affine hull f ` S) \ aff_dim (affine hull S)" by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans) show "f ` rel_interior S \ affine hull f ` S" by (meson \f ` rel_interior S \ f ` S\ hull_subset order_trans) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast qed auto qed lemma homeomorphic_rel_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(rel_interior S) homeomorphic (rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (rel_interior S) (rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast have "rel_interior T \ f ` rel_interior S" proof fix x assume "x \ rel_interior T" then have "g x \ rel_interior S" using gim by blast then show "x \ f ` rel_interior S" by (metis fg \x \ rel_interior T\ imageI) qed moreover have "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) ultimately show "f ` rel_interior S = rel_interior T" by blast show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast have "rel_interior S \ g ` rel_interior T" proof fix x assume "x \ rel_interior S" then have "f x \ rel_interior T" using fim by blast then show "x \ g ` rel_interior T" by (metis gf \x \ rel_interior S\ imageI) qed then show "g ` rel_interior T = rel_interior S" using gim by blast show "continuous_on (rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(rel_interior S) homeomorphic (rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_rel_boundaries_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ S - rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ T - rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast show "f ` (S - rel_interior S) = T - rel_interior T" using S fST fim gim by auto show "continuous_on (S - rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "g ` (T - rel_interior T) = S - rel_interior S" using T gTS gim fim by auto show "continuous_on (T - rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_boundaries: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_boundaries_same_dimension [OF \S homeomorphic T\]) qed proposition uniformly_continuous_homeomorphism_UNIV_trivial: fixes f :: "'a::euclidean_space \ 'a" assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g" shows "S = UNIV" proof (cases "S = {}") case True then show ?thesis by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI) next case False have "inj g" by (metis UNIV_I hom homeomorphism_apply2 injI) then have "open (g ` UNIV)" by (blast intro: invariance_of_domain hom homeomorphism_cont2) then have "open S" using hom homeomorphism_image2 by blast moreover have "complete S" unfolding complete_def proof clarify fix \ assume \: "\n. \ n \ S" and "Cauchy \" have "Cauchy (f o \)" using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf by blast then obtain l where "(f \ \) \ l" by (auto simp: convergent_eq_Cauchy [symmetric]) show "\l\S. \ \ l" proof show "g l \ S" using hom homeomorphism_image2 by blast have "(g \ (f \ \)) \ g l" by (meson UNIV_I \(f \ \) \ l\ continuous_on_sequentially hom homeomorphism_cont2) then show "\ \ g l" proof - have "\n. \ n = (g \ (f \ \)) n" by (metis (no_types) \ comp_eq_dest_lhs hom homeomorphism_apply1) then show ?thesis by (metis (no_types) LIMSEQ_iff \(g \ (f \ \)) \ g l\) qed qed qed then have "closed S" by (simp add: complete_eq_closed) ultimately show ?thesis using clopen [of S] False by simp qed subsection\Formulation of loop homotopy in terms of maps out of type complex\ lemma homotopic_circlemaps_imp_homotopic_loops: assumes "homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "homotopic_loops S (f \ exp \ (\t. 2 * of_real pi * of_real t * \)) (g \ exp \ (\t. 2 * of_real pi * of_real t * \))" proof - have "homotopic_with_canon (\f. True) {z. cmod z = 1} S f g" using assms by (auto simp: sphere_def) moreover have "continuous_on {0..1} (exp \ (\t. 2 * of_real pi * of_real t * \))" by (intro continuous_intros) moreover have "(exp \ (\t. 2 * of_real pi * of_real t * \)) ` {0..1} \ {z. cmod z = 1}" by (auto simp: norm_mult) ultimately show ?thesis apply (simp add: homotopic_loops_def comp_assoc) apply (rule homotopic_with_compose_continuous_right) apply (auto simp: pathstart_def pathfinish_def) done qed lemma homotopic_loops_imp_homotopic_circlemaps: assumes "homotopic_loops S p q" shows "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. (Arg2pi z / (2 * pi)))) (q \ (\z. (Arg2pi z / (2 * pi))))" proof - obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "(\x. h (0, x) = p x)" and h1: "(\x. h (1, x) = q x)" and h01: "(\t\{0..1}. h (t, 1) = h (t, 0)) " using assms by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def) define j where "j \ \z. if 0 \ Im (snd z) then h (fst z, Arg2pi (snd z) / (2 * pi)) else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))" have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \ Arg2pi y = 0 \ Arg2pi (cnj y) = 0" if "cmod y = 1" for y using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps) show ?thesis proof (simp add: homotopic_with; intro conjI ballI exI) show "continuous_on ({0..1} \ sphere 0 1) (\w. h (fst w, Arg2pi (snd w) / (2 * pi)))" proof (rule continuous_on_eq) show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \ {0..1} \ sphere 0 1" for x using Arg2pi_eq that h01 by (force simp: j_def) have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" by auto have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) apply (auto simp: Arg2pi) apply (meson Arg2pi_lt_2pi linear not_le) done have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) apply (auto simp: Arg2pi) apply (meson Arg2pi_lt_2pi linear not_le) done show "continuous_on ({0..1} \ sphere 0 1) j" apply (simp add: j_def) apply (subst eq) apply (rule continuous_on_cases_local) apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2) using Arg2pi_eq h01 by force qed have "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ h ` ({0..1} \ {0..1})" by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le) also have "... \ S" using him by blast finally show "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ S" . qed (auto simp: h0 h1) qed lemma simply_connected_homotopic_loops: "simply_connected S \ (\p q. homotopic_loops S p p \ homotopic_loops S q q \ homotopic_loops S p q)" unfolding simply_connected_def using homotopic_loops_refl by metis lemma simply_connected_eq_homotopic_circlemaps1: fixes f :: "complex \ 'a::topological_space" and g :: "complex \ 'a" assumes S: "simply_connected S" and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \ S" and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" proof - have "homotopic_loops S (f \ exp \ (\t. of_real(2 * pi * t) * \)) (g \ exp \ (\t. of_real(2 * pi * t) * \))" apply (rule S [unfolded simply_connected_homotopic_loops, rule_format]) - apply (simp add: homotopic_circlemaps_imp_homotopic_loops homotopic_with_refl contf fim contg gim) + apply (simp add: homotopic_circlemaps_imp_homotopic_loops contf fim contg gim) done then show ?thesis apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) done qed lemma simply_connected_eq_homotopic_circlemaps2a: fixes h :: "complex \ 'a::topological_space" assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \ S" and hom: "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" apply (rule_tac x="h 1" in exI) apply (rule hom) - using assms - by (auto simp: continuous_on_const) + using assms by (auto) lemma simply_connected_eq_homotopic_circlemaps2b: fixes S :: "'a::real_normed_vector set" assumes "\f g::complex \ 'a. \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" shows "path_connected S" proof (clarsimp simp add: path_connected_eq_homotopic_points) fix a b assume "a \ S" "b \ S" then show "homotopic_loops S (linepath a a) (linepath b b)" using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\x. a" "\x. b"]] - by (auto simp: o_def continuous_on_const linepath_def) + by (auto simp: o_def linepath_def) qed lemma simply_connected_eq_homotopic_circlemaps3: fixes h :: "complex \ 'a::real_normed_vector" assumes "path_connected S" and hom: "\f::complex \ 'a. \continuous_on (sphere 0 1) f; f `(sphere 0 1) \ S\ \ \a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)" shows "simply_connected S" proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms) fix p assume p: "path p" "path_image p \ S" "pathfinish p = pathstart p" then have "homotopic_loops S p p" by (simp add: homotopic_loops_refl) then obtain a where homp: "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. Arg2pi z / (2 * pi))) (\x. a)" by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom) show "\a. a \ S \ homotopic_loops S p (linepath a a)" proof (intro exI conjI) show "a \ S" using homotopic_with_imp_subset2 [OF homp] by (metis dist_0_norm image_subset_iff mem_sphere norm_one) have teq: "\t. \0 \ t; t \ 1\ \ t = Arg2pi (exp (2 * of_real pi * of_real t * \)) / (2 * pi) \ t=1 \ Arg2pi (exp (2 * of_real pi * of_real t * \)) = 0" apply (rule disjCI) using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp) done have "homotopic_loops S p (p \ (\z. Arg2pi z / (2 * pi)) \ exp \ (\t. 2 * complex_of_real pi * complex_of_real t * \))" apply (rule homotopic_loops_eq [OF p]) using p teq apply (fastforce simp: pathfinish_def pathstart_def) done then show "homotopic_loops S p (linepath a a)" by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]]) qed qed proposition simply_connected_eq_homotopic_circlemaps: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ (\f g::complex \ 'a. continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" apply (rule iffI) apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1) by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) proposition simply_connected_eq_contractible_circlemap: fixes S :: "'a::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\f::complex \ 'a. continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" apply (rule iffI) apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b) using simply_connected_eq_homotopic_circlemaps3 by blast corollary homotopy_eqv_simple_connectedness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) subsection\Homeomorphism of simple closed curves to circles\ proposition homeomorphic_simple_path_image_circle: fixes a :: complex and \ :: "real \ 'a::t2_space" assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" shows "(path_image \) homeomorphic sphere a r" proof - have "homotopic_loops (path_image \) \ \" by (simp add: assms homotopic_loops_refl simple_path_imp_path) then have hom: "homotopic_with_canon (\h. True) (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) (\ \ (\z. Arg2pi z / (2*pi)))" by (rule homotopic_loops_imp_homotopic_circlemaps) have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) g" proof (rule homeomorphism_compact) show "continuous_on (sphere 0 1) (\ \ (\z. Arg2pi z / (2*pi)))" using hom homotopic_with_imp_continuous by blast show "inj_on (\ \ (\z. Arg2pi z / (2*pi))) (sphere 0 1)" proof fix x y assume xy: "x \ sphere 0 1" "y \ sphere 0 1" and eq: "(\ \ (\z. Arg2pi z / (2*pi))) x = (\ \ (\z. Arg2pi z / (2*pi))) y" then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))" proof - have "(Arg2pi x / (2*pi)) \ {0..1}" "(Arg2pi y / (2*pi)) \ {0..1}" using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+ with eq show ?thesis using \simple_path \\ Arg2pi_lt_2pi unfolding simple_path_def o_def by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) qed with xy show "x = y" by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) qed have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg2pi z / (2*pi)) = \ x" by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) moreover have "\z\sphere 0 1. \ x = \ (Arg2pi z / (2*pi))" if "0 \ x" "x \ 1" for x proof (cases "x=1") case True with Arg2pi_of_real [of 1] loop show ?thesis by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \0 \ x\) next case False then have *: "(Arg2pi (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" using that by (auto simp: Arg2pi_exp field_split_simps) show ?thesis by (rule_tac x="exp(\ * of_real(2*pi*x))" in bexI) (auto simp: *) qed ultimately show "(\ \ (\z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \" by (auto simp: path_image_def image_iff) qed auto then have "path_image \ homeomorphic sphere (0::complex) 1" using homeomorphic_def homeomorphic_sym by blast also have "... homeomorphic sphere a r" by (simp add: assms homeomorphic_spheres) finally show ?thesis . qed lemma homeomorphic_simple_path_images: fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" shows "(path_image \1) homeomorphic (path_image \2)" by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) subsection\Dimension-based conditions for various homeomorphisms\ lemma homeomorphic_subspaces_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "subspace S" "subspace T" shows "S homeomorphic T \ dim S = dim T" proof assume "S homeomorphic T" then obtain f g where hom: "homeomorphism S T f g" using homeomorphic_def by blast show "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le) show "dim T \ dim S" by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le) qed next assume "dim S = dim T" then show "S homeomorphic T" by (simp add: assms homeomorphic_subspaces) qed lemma homeomorphic_affine_sets_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "affine S" "affine T" shows "S homeomorphic T \ aff_dim S = aff_dim T" proof (cases "S = {} \ T = {}") case True then show ?thesis using assms homeomorphic_affine_sets by force next case False then obtain a b where "a \ S" "b \ T" by blast then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ then show ?thesis by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets) qed lemma homeomorphic_hyperplanes_eq: fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('a) = DIM('b))" apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) by (metis DIM_positive Suc_pred) lemma homeomorphic_UNIV_UNIV: shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \ DIM('a::euclidean_space) = DIM('b::euclidean_space)" by (simp add: homeomorphic_subspaces_eq) lemma simply_connected_sphere_gen: assumes "convex S" "bounded S" and 3: "3 \ aff_dim S" shows "simply_connected(rel_frontier S)" proof - have pa: "path_connected (rel_frontier S)" using assms by (simp add: path_connected_sphere_gen) show ?thesis proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa) fix f assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \ rel_frontier S" have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)" by simp have "convex (cball (0::complex) 1)" by (rule convex_cball) then obtain c where "homotopic_with_canon (\z. True) (sphere (0::complex) 1) (rel_frontier S) f (\x. c)" apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \convex S\ \bounded S\, where f=f]) using f 3 apply (auto simp: aff_dim_cball) done then show "\a. homotopic_with_canon (\h. True) (sphere 0 1) (rel_frontier S) f (\x. a)" by blast qed qed subsection\more invariance of domain\(*FIX ME title? *) proposition invariance_of_domain_sphere_affine_set_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and U: "bounded U" "convex U" and "affine T" and affTU: "aff_dim T < aff_dim U" and ope: "openin (top_of_set (rel_frontier U)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "rel_frontier U = {}") case True then show ?thesis using ope openin_subset by force next case False obtain b c where b: "b \ rel_frontier U" and c: "c \ rel_frontier U" and "b \ c" using \bounded U\ rel_frontier_not_sing [of U] subset_singletonD False by fastforce obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1" proof (rule choose_affine_subset [OF affine_UNIV]) show "- 1 \ aff_dim U - 1" by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le) show "aff_dim U - 1 \ aff_dim (UNIV::'a set)" by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq) qed auto have SU: "S \ rel_frontier U" using ope openin_imp_subset by auto have homb: "rel_frontier U - {b} homeomorphic V" and homc: "rel_frontier U - {c} homeomorphic V" using homeomorphic_punctured_sphere_affine_gen [of U _ V] by (simp_all add: \affine V\ affV U b c) then obtain g h j k where gh: "homeomorphism (rel_frontier U - {b}) V g h" and jk: "homeomorphism (rel_frontier U - {c}) V j k" by (auto simp: homeomorphic_def) with SU have hgsub: "(h ` g ` (S - {b})) \ S" and kjsub: "(k ` j ` (S - {c})) \ S" by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (g ` (S - {b}))" apply (rule homeomorphism_imp_open_map [OF gh]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (g ` (S - {b})) (f \ h)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) using contf continuous_on_subset hgsub by blast show "inj_on (f \ h) (g ` (S - {b}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (j ` (S - {c}))" apply (rule homeomorphism_imp_open_map [OF jk]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (j ` (S - {c})) (f \ k)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) using contf continuous_on_subset kjsub by blast show "inj_on (f \ k) (j ` (S - {c}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) ultimately have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" by (rule openin_Un) moreover have "(f \ h) ` g ` (S - {b}) = f ` (S - {b})" proof - have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" proof show "k ` j ` (S - {c}) \ S - {c}" using homeomorphism_apply1 [OF jk] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {c} \ k ` j ` (S - {c})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "f ` (S - {b}) \ f ` (S - {c}) = f ` (S)" using \b \ c\ by blast ultimately show ?thesis by simp qed lemma invariance_of_domain_sphere_affine_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and "r \ 0" "affine T" and affTU: "aff_dim T < DIM('a)" and ope: "openin (top_of_set (sphere a r)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "sphere a r = {}") case True then show ?thesis using ope openin_subset by force next case False show ?thesis proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \affine T\]) show "aff_dim T < aff_dim (cball a r)" by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty) show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed lemma no_embedding_sphere_lowdim: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0" shows "DIM('a) \ DIM('b)" proof - have "False" if "DIM('a) > DIM('b)" proof - have "compact (f ` sphere a r)" using compact_continuous_image by (simp add: compact_continuous_image contf) then have "\ open (f ` sphere a r)" using compact_open by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty) then show False using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \r > 0\ by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that) qed then show ?thesis using not_less by blast qed lemma simply_connected_sphere: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by simp next case equal then show ?thesis by (auto simp: convex_imp_simply_connected) next case greater then show ?thesis using simply_connected_sphere_gen [of "cball a r"] assms by (simp add: aff_dim_cball) qed lemma simply_connected_sphere_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(sphere a r) \ 3 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True have "simply_connected (sphere a r)" apply (rule convex_imp_simply_connected) using True less_eq_real_def by auto with True show ?thesis by auto next case False show ?thesis proof assume L: ?lhs have "False" if "DIM('a) = 1 \ DIM('a) = 2" using that proof assume "DIM('a) = 1" with L show False using connected_sphere_eq simply_connected_imp_connected by (metis False Suc_1 not_less_eq_eq order_refl) next assume "DIM('a) = 2" then have "sphere a r homeomorphic sphere (0::complex) 1" by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one) then have "simply_connected(sphere (0::complex) 1)" using L homeomorphic_simply_connected_eq by blast then obtain a::complex where "homotopic_with_canon (\h. True) (sphere 0 1) (sphere 0 1) id (\x. a)" apply (simp add: simply_connected_eq_contractible_circlemap) by (metis continuous_on_id' id_apply image_id subset_refl) then show False using contractible_sphere contractible_def not_one_le_zero by blast qed with False show ?rhs apply simp by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) next assume ?rhs with False show ?lhs by (simp add: simply_connected_sphere) qed qed lemma simply_connected_punctured_universe_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(- {a}) \ 3 \ DIM('a)" proof - have [simp]: "a \ rel_interior (cball a 1)" by (simp add: rel_interior_nonempty_interior) have [simp]: "affine hull cball a 1 - {a} = -{a}" by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one) have "simply_connected(- {a}) \ simply_connected(sphere a 1)" apply (rule sym) apply (rule homotopy_eqv_simple_connectedness) using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto done also have "... \ 3 \ DIM('a)" by (simp add: simply_connected_sphere_eq) finally show ?thesis . qed lemma not_simply_connected_circle: fixes a :: complex shows "0 < r \ \ simply_connected(sphere a r)" by (simp add: simply_connected_sphere_eq) proposition simply_connected_punctured_convex: fixes a :: "'a::euclidean_space" assumes "convex S" and 3: "3 \ aff_dim S" shows "simply_connected(S - {a})" proof (cases "a \ rel_interior S") case True then obtain e where "a \ S" "0 < e" and e: "cball a e \ affine hull S \ S" by (auto simp: rel_interior_cball) have con: "convex (cball a e \ affine hull S)" by (simp add: convex_Int) have bo: "bounded (cball a e \ affine hull S)" by (simp add: bounded_Int) have "affine hull S \ interior (cball a e) \ {}" using \0 < e\ \a \ S\ hull_subset by fastforce then have "3 \ aff_dim (affine hull S \ cball a e)" by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull]) also have "... = aff_dim (cball a e \ affine hull S)" by (simp add: Int_commute) finally have "3 \ aff_dim (cball a e \ affine hull S)" . moreover have "rel_frontier (cball a e \ affine hull S) homotopy_eqv S - {a}" proof (rule homotopy_eqv_rel_frontier_punctured_convex) show "a \ rel_interior (cball a e \ affine hull S)" by (meson IntI Int_mono \a \ S\ \0 < e\ e \cball a e \ affine hull S \ S\ ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball) have "closed (cball a e \ affine hull S)" by blast then show "rel_frontier (cball a e \ affine hull S) \ S" apply (simp add: rel_frontier_def) using e by blast show "S \ affine hull (cball a e \ affine hull S)" by (metis (no_types, lifting) IntI \a \ S\ \0 < e\ affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI) qed (auto simp: assms con bo) ultimately show ?thesis using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo] by blast next case False show ?thesis apply (rule contractible_imp_simply_connected) apply (rule contractible_convex_tweak_boundary_points [OF \convex S\]) apply (simp add: False rel_interior_subset subset_Diff_insert) by (meson Diff_subset closure_subset subset_trans) qed corollary simply_connected_punctured_universe: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(- {a})" proof - have [simp]: "affine hull cball a 1 = UNIV" apply auto by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq) have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})" apply (rule homotopy_eqv_simple_connectedness) apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull) apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+ done then show ?thesis using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV) qed subsection\The power, squaring and exponential functions as covering maps\ proposition covering_space_power_punctured_plane: assumes "0 < n" shows "covering_space (- {0}) (\z::complex. z^n) (- {0})" proof - consider "n = 1" | "2 \ n" using assms by linarith then obtain e where "0 < e" and e: "\w z. cmod(w - z) < e * cmod z \ (w^n = z^n \ w = z)" proof cases assume "n = 1" then show ?thesis by (rule_tac e=1 in that) auto next assume "2 \ n" have eq_if_pow_eq: "w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z" and eq: "w^n = z^n" for w z proof (cases "z = 0") case True with eq assms show ?thesis by (auto simp: power_0_left) next case False then have "z \ 0" by auto have "(w/z)^n = 1" by (metis False divide_self_if eq power_divide power_one) then obtain j where j: "w / z = exp (2 * of_real pi * \ * j / n)" and "j < n" using Suc_leI assms \2 \ n\ complex_roots_unity [THEN eqset_imp_iff, of n "w/z"] by force have "cmod (w/z - 1) < 2 * sin (pi / real n)" using lt assms \z \ 0\ by (simp add: field_split_simps norm_divide) then have "cmod (exp (\ * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)" by (simp add: j field_simps) then have "2 * \sin((2 * pi * j / n) / 2)\ < 2 * sin (pi / real n)" by (simp only: dist_exp_i_1) then have sin_less: "sin((pi * j / n)) < sin (pi / real n)" by (simp add: field_simps) then have "w / z = 1" proof (cases "j = 0") case True then show ?thesis by (auto simp: j) next case False then have "sin (pi / real n) \ sin((pi * j / n))" proof (cases "j / n \ 1/2") case True show ?thesis apply (rule sin_monotone_2pi_le) using \j \ 0 \ \j < n\ True apply (auto simp: field_simps intro: order_trans [of _ 0]) done next case False then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)" using \j < n\ by (simp add: algebra_simps diff_divide_distrib of_nat_diff) show ?thesis apply (simp only: seq) apply (rule sin_monotone_2pi_le) using \j < n\ False apply (auto simp: field_simps intro: order_trans [of _ 0]) done qed with sin_less show ?thesis by force qed then show ?thesis by simp qed show ?thesis apply (rule_tac e = "2 * sin(pi / n)" in that) apply (force simp: \2 \ n\ sin_pi_divide_n_gt_0) apply (meson eq_if_pow_eq) done qed have zn1: "continuous_on (- {0}) (\z::complex. z^n)" by (rule continuous_intros)+ have zn2: "(\z::complex. z^n) ` (- {0}) = - {0}" using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n]) have zn3: "\T. z^n \ T \ open T \ 0 \ T \ (\v. \v = -{0} \ (\z. z ^ n) -` T \ (\u\v. open u \ 0 \ u) \ pairwise disjnt v \ (\u\v. Ex (homeomorphism u T (\z. z^n))))" if "z \ 0" for z::complex proof - define d where "d \ min (1/2) (e/4) * norm z" have "0 < d" by (simp add: d_def \0 < e\ \z \ 0\) have iff_x_eq_y: "x^n = y^n \ x = y" if eq: "w^n = z^n" and x: "x \ ball w d" and y: "y \ ball w d" for w x y proof - have [simp]: "norm z = norm w" using that by (simp add: assms power_eq_imp_eq_norm) show ?thesis proof (cases "w = 0") case True with \z \ 0\ assms eq show ?thesis by (auto simp: power_0_left) next case False have "cmod (x - y) < 2*d" using x y by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add) also have "... \ 2 * e / 4 * norm w" using \e > 0\ by (simp add: d_def min_mult_distrib_right) also have "... = e * (cmod w / 2)" by simp also have "... \ e * cmod y" apply (rule mult_left_mono) using \e > 0\ y apply (simp_all add: dist_norm d_def min_mult_distrib_right del: divide_const_simps) apply (metis dist_0_norm dist_complex_def dist_triangle_half_l linorder_not_less order_less_irrefl) done finally have "cmod (x - y) < e * cmod y" . then show ?thesis by (rule e) qed qed then have inj: "inj_on (\w. w^n) (ball z d)" by (simp add: inj_on_def) have cont: "continuous_on (ball z d) (\w. w ^ n)" by (intro continuous_intros) have noncon: "\ (\w::complex. w^n) constant_on UNIV" by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power) have im_eq: "(\w. w^n) ` ball z' d = (\w. w^n) ` ball z d" if z': "z'^n = z^n" for z' proof - have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast have "(w \ (\w. w^n) ` ball z' d) = (w \ (\w. w^n) ` ball z d)" for w proof (cases "w=0") case True with assms show ?thesis by (simp add: image_def ball_def nz') next case False have "z' \ 0" using \z \ 0\ nz' by force have [simp]: "(z*x / z')^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have [simp]: "cmod (z - z * x / z') = cmod (z' - x)" if "x \ 0" for x proof - have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')" by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) divide_complex_def mult.right_neutral norm_mult right_diff_distrib') also have "... = cmod z' * cmod (1 - x / z')" by (simp add: nz') also have "... = cmod (z' - x)" by (simp add: \z' \ 0\ diff_divide_eq_iff norm_divide) finally show ?thesis . qed have [simp]: "(z'*x / z)^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have [simp]: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \ 0" for x proof - have "cmod (z * (1 - x * inverse z)) = cmod (z - x)" by (metis \z \ 0\ diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7)) then show ?thesis by (metis (no_types) mult.assoc divide_complex_def mult.right_neutral norm_mult nz' right_diff_distrib') qed show ?thesis unfolding image_def ball_def apply safe apply simp_all apply (rule_tac x="z/z' * x" in exI) using assms False apply (simp add: dist_norm) apply (rule_tac x="z'/z * x" in exI) using assms False apply (simp add: dist_norm) done qed then show ?thesis by blast qed have ex_ball: "\B. (\z'. B = ball z' d \ z'^n = z^n) \ x \ B" if "x \ 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w proof - have "w \ 0" by (metis assms power_eq_0_iff that(1) that(2)) have [simp]: "cmod x = cmod w" using assms power_eq_imp_eq_norm eq by blast have [simp]: "cmod (x * z / w - x) = cmod (z - w)" proof - have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)" by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right) also have "... = cmod w * cmod (z / w - 1)" by simp also have "... = cmod (z - w)" by (simp add: \w \ 0\ divide_diff_eq_iff nonzero_norm_divide) finally show ?thesis . qed show ?thesis apply (rule_tac x="ball (z / w * x) d" in exI) using \d > 0\ that apply (simp add: ball_eq_ball_iff) apply (simp add: \z \ 0\ \w \ 0\ field_simps) apply (simp add: dist_norm) done qed show ?thesis proof (rule exI, intro conjI) show "z ^ n \ (\w. w ^ n) ` ball z d" using \d > 0\ by simp show "open ((\w. w ^ n) ` ball z d)" by (rule invariance_of_domain [OF cont open_ball inj]) show "0 \ (\w. w ^ n) ` ball z d" using \z \ 0\ assms by (force simp: d_def) show "\v. \v = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d \ (\u\v. open u \ 0 \ u) \ disjoint v \ (\u\v. Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n)))" proof (rule exI, intro ballI conjI) show "\{ball z' d |z'. z'^n = z^n} = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d" (is "?l = ?r") proof show "?l \ ?r" apply auto apply (simp add: assms d_def power_eq_imp_eq_norm that) by (metis im_eq image_eqI mem_ball) show "?r \ ?l" by auto (meson ex_ball) qed show "\u. u \ {ball z' d |z'. z' ^ n = z ^ n} \ 0 \ u" by (force simp add: assms d_def power_eq_imp_eq_norm that) show "disjoint {ball z' d |z'. z' ^ n = z ^ n}" proof (clarsimp simp add: pairwise_def disjnt_iff) fix \ \ x assume "\^n = z^n" "\^n = z^n" "ball \ d \ ball \ d" and "dist \ x < d" "dist \ x < d" then have "dist \ \ < d+d" using dist_triangle_less_add by blast then have "cmod (\ - \) < 2*d" by (simp add: dist_norm) also have "... \ e * cmod z" using mult_right_mono \0 < e\ that by (auto simp: d_def) finally have "cmod (\ - \) < e * cmod z" . with e have "\ = \" by (metis \\^n = z^n\ \\^n = z^n\ assms power_eq_imp_eq_norm) then show "False" using \ball \ d \ ball \ d\ by blast qed show "Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" if "u \ {ball z' d |z'. z' ^ n = z ^ n}" for u proof (rule invariance_of_domain_homeomorphism [of "u" "\z. z^n"]) show "open u" using that by auto show "continuous_on u (\z. z ^ n)" by (intro continuous_intros) show "inj_on (\z. z ^ n) u" using that by (auto simp: iff_x_eq_y inj_on_def) show "\g. homeomorphism u ((\z. z ^ n) ` u) (\z. z ^ n) g \ Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" using im_eq that by clarify metis qed auto qed auto qed qed show ?thesis using assms apply (simp add: covering_space_def zn1 zn2) apply (subst zn2 [symmetric]) apply (simp add: openin_open_eq open_Compl) apply (blast intro: zn3) done qed corollary covering_space_square_punctured_plane: "covering_space (- {0}) (\z::complex. z^2) (- {0})" by (simp add: covering_space_power_punctured_plane) proposition covering_space_exp_punctured_plane: "covering_space UNIV (\z::complex. exp z) (- {0})" proof (simp add: covering_space_def, intro conjI ballI) show "continuous_on UNIV (\z::complex. exp z)" by (rule continuous_on_exp [OF continuous_on_id]) show "range exp = - {0::complex}" by auto (metis exp_Ln range_eqI) show "\T. z \ T \ openin (top_of_set (- {0})) T \ (\v. \v = exp -` T \ (\u\v. open u) \ disjoint v \ (\u\v. \q. homeomorphism u T exp q))" if "z \ - {0::complex}" for z proof - have "z \ 0" using that by auto have inj_exp: "inj_on exp (ball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: ball_subset_ball_iff) define \ where "\ \ range (\n. (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1))" show ?thesis proof (intro exI conjI) show "z \ exp ` (ball(Ln z) 1)" by (metis \z \ 0\ centre_in_ball exp_Ln rev_image_eqI zero_less_one) have "open (- {0::complex})" by blast moreover have "inj_on exp (ball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: ball_subset_ball_iff) ultimately show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id]) show "\\ = exp -` exp ` ball (Ln z) 1" by (force simp: \_def Complex_Transcendental.exp_eq image_iff) show "\V\\. open V" by (auto simp: \_def inj_on_def continuous_intros invariance_of_domain) have xy: "2 \ cmod (2 * of_int x * of_real pi * \ - 2 * of_int y * of_real pi * \)" if "x < y" for x y proof - have "1 \ abs (x - y)" using that by linarith then have "1 \ cmod (of_int x - of_int y) * 1" by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff) also have "... \ cmod (of_int x - of_int y) * of_real pi" apply (rule mult_left_mono) using pi_ge_two by auto also have "... \ cmod ((of_int x - of_int y) * of_real pi * \)" by (simp add: norm_mult) also have "... \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" by (simp add: algebra_simps) finally have "1 \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" . then have "2 * 1 \ cmod (2 * (of_int x * of_real pi * \ - of_int y * of_real pi * \))" by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral) then show ?thesis by (simp add: algebra_simps) qed show "disjoint \" apply (clarsimp simp add: \_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y] image_add_ball ball_eq_ball_iff) apply (rule disjoint_ballI) apply (auto simp: dist_norm neq_iff) by (metis norm_minus_commute xy)+ show "\u\\. \q. homeomorphism u (exp ` ball (Ln z) 1) exp q" proof fix u assume "u \ \" then obtain n where n: "u = (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1)" by (auto simp: \_def) have "compact (cball (Ln z) 1)" by simp moreover have "continuous_on (cball (Ln z) 1) exp" by (rule continuous_on_exp [OF continuous_on_id]) moreover have "inj_on exp (cball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: cball_subset_ball_iff) ultimately obtain \ where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \" using homeomorphism_compact by blast have eq1: "exp ` u = exp ` ball (Ln z) 1" unfolding n apply (auto simp: algebra_simps) apply (rename_tac w) apply (rule_tac x = "w + \ * (of_int n * (of_real pi * 2))" in image_eqI) apply (auto simp: image_iff) done have \exp: "\ (exp x) + 2 * of_int n * of_real pi * \ = x" if "x \ u" for x proof - have "exp x = exp (x - 2 * of_int n * of_real pi * \)" by (simp add: exp_eq) then have "\ (exp x) = \ (exp (x - 2 * of_int n * of_real pi * \))" by simp also have "... = x - 2 * of_int n * of_real pi * \" apply (rule homeomorphism_apply1 [OF hom]) using \x \ u\ by (auto simp: n) finally show ?thesis by simp qed have exp2n: "exp (\ (exp x) + 2 * of_int n * complex_of_real pi * \) = exp x" if "dist (Ln z) x < 1" for x using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom]) have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + 2 * of_int n * complex_of_real pi * \)" apply (intro continuous_intros) apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom]]) apply (force simp:) done show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" apply (rule_tac x="(\x. x + of_real(2 * n * pi) * \) \ \" in exI) unfolding homeomorphism_def apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id]) apply (auto simp: \exp exp2n cont n) apply (simp add: homeomorphism_apply1 [OF hom]) using hom homeomorphism_apply1 apply (force simp: image_iff) done qed qed qed qed subsection\Hence the Borsukian results about mappings into circles\(*FIX ME title *) lemma inessential_eq_continuous_logarithm: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x)))" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane) next assume ?rhs then obtain g where contg: "continuous_on S g" and f: "\x. x \ S \ f x = exp(g x)" by metis obtain a where "homotopic_with_canon (\h. True) S (- {of_real 0}) (exp \ g) (\x. a)" proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV]) show "continuous_on (UNIV::complex set) exp" by (intro continuous_intros) show "range exp \ - {0}" by auto qed force then have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using f homotopic_with_eq by fastforce then show ?lhs .. qed corollary inessential_imp_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" assumes "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" obtains g where "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" proof - have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using assms homotopic_with_subset_right by fastforce then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma inessential_eq_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(\ * of_real(g x))))" (is "?lhs \ ?rhs") proof assume L: ?lhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(g x)" using inessential_imp_continuous_logarithm_circle by blast have "f ` S \ sphere 0 1" by (metis L homotopic_with_imp_subset1) then have "\x. x \ S \ Re (g x) = 0" using g by auto then show ?rhs apply (rule_tac x="Im \ g" in exI) apply (intro conjI contg continuous_intros) apply (auto simp: Euler g) done next assume ?rhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(\* of_real(g x))" by metis obtain a where "homotopic_with_canon (\h. True) S (sphere 0 1) ((exp \ (\z. \*z)) \ (of_real \ g)) (\x. a)" proof (rule nullhomotopic_through_contractible) show "continuous_on S (complex_of_real \ g)" by (intro conjI contg continuous_intros) show "(complex_of_real \ g) ` S \ \" by auto show "continuous_on \ (exp \ (*)\)" by (intro continuous_intros) show "(exp \ (*)\) ` \ \ sphere 0 1" by (auto simp: complex_is_Real_iff) qed (auto simp: convex_Reals convex_imp_contractible) moreover have "\x. x \ S \ (exp \ (*)\ \ (complex_of_real \ g)) x = f x" by (simp add: g) ultimately have "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" using homotopic_with_eq by force then show ?lhs .. qed proposition homotopic_with_sphere_times: fixes f :: "'a::real_normed_vector \ complex" assumes hom: "homotopic_with_canon (\x. True) S (sphere 0 1) f g" and conth: "continuous_on S h" and hin: "\x. x \ S \ h x \ sphere 0 1" shows "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x * h x) (\x. g x * h x)" proof - obtain k where contk: "continuous_on ({0..1::real} \ S) k" and kim: "k ` ({0..1} \ S) \ sphere 0 1" and k0: "\x. k(0, x) = f x" and k1: "\x. k(1, x) = g x" using hom by (auto simp: homotopic_with_def) show ?thesis apply (simp add: homotopic_with) apply (rule_tac x="\z. k z*(h \ snd)z" in exI) apply (intro conjI contk continuous_intros) apply (simp add: conth) using kim hin apply (force simp: norm_mult k0 k1)+ done qed proposition homotopic_circlemaps_divide: fixes f :: "'a::real_normed_vector \ complex" shows "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c))" proof - have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" if "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)" for c proof - have "S = {} \ path_component (sphere 0 1) 1 c" using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1] by (auto simp: path_connected_component) then have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. 1) (\x. c)" by (simp add: homotopic_constant_maps) then show ?thesis using homotopic_with_symD homotopic_with_trans that by blast qed then have *: "(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)) \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" by auto have "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" (is "?lhs \ ?rhs") proof assume L: ?lhs have geq1 [simp]: "\x. x \ S \ cmod (g x) = 1" using homotopic_with_imp_subset2 [OF L] by (simp add: image_subset_iff) have cont: "continuous_on S (inverse \ g)" apply (rule continuous_intros) using homotopic_with_imp_continuous [OF L] apply blast apply (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse]) - apply (auto simp: continuous_on_id) + apply (auto) done have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" using homotopic_with_sphere_times [OF L cont] apply (rule homotopic_with_eq) apply (auto simp: division_ring_class.divide_inverse norm_inverse) by (metis geq1 norm_zero right_inverse zero_neq_one) with L show ?rhs by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2) next assume ?rhs then show ?lhs by (force simp: elim: homotopic_with_eq dest: homotopic_with_sphere_times [where h=g]) qed then show ?thesis by (simp add: *) qed subsection\Upper and lower hemicontinuous functions\ text\And relation in the case of preimage map to open and closed maps, and fact that upper and lower hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the function gives a bounded and nonempty set).\ text\Many similar proofs below.\ lemma upper_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}) \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ T - U}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ T - U} = {x \ S. f x \ U \ {}}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ (T - U) \ {}} = S - {x \ S. f x \ U}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma lower_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}) \ (\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ T-U}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ T-U} = S - {x \ S. f x \ U \ {}}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ (T - U) \ {}} = {x \ S. f x \ U}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma open_map_iff_lower_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)) \ (\U. closedin (top_of_set S) U \ closedin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "T - (f ` (S - U)) = {y \ T. {x \ S. f x = y} \ U}" using assms by blast ultimately show "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ U}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "{y \ T. {x \ S. f x = y} \ S - U} = T - (f ` U)" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) (f ` U)" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) qed lemma closed_map_iff_upper_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)) \ (\U. openin (top_of_set S) U \ openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "f ` (S - U) = T - {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and cloSU: "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "(f ` U) = T - {y \ T. {x \ S. f x = y} \ S - U}" using assms closedin_imp_subset [OF cloSU] by auto ultimately show "closedin (top_of_set T) (f ` U)" by (simp add: openin_closedin_eq) qed proposition upper_lower_hemicontinuous_explicit: fixes T :: "('b::{real_normed_vector,heine_borel}) set" assumes fST: "\x. x \ S \ f x \ T" and ope: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}" and clo: "\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}" and "x \ S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \ {}" obtains d where "0 < d" "\x'. \x' \ S; dist x x' < d\ \ (\y \ f x. \y'. y' \ f x' \ dist y y' < e) \ (\y' \ f x'. \y. y \ f x \ dist y' y < e)" proof - have "openin (top_of_set T) (T \ (\a\f x. \b\ball 0 e. {a + b}))" by (auto simp: open_sums openin_open_Int) with ope have "openin (top_of_set S) {u \ S. f u \ T \ (\a\f x. \b\ball 0 e. {a + b})}" by blast with \0 < e\ \x \ S\ obtain d1 where "d1 > 0" and d1: "\x'. \x' \ S; dist x' x < d1\ \ f x' \ T \ f x' \ (\a \ f x. \b \ ball 0 e. {a + b})" by (force simp: openin_euclidean_subtopology_iff dest: fST) have oo: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}" apply (rule lower_hemicontinuous [THEN iffD1, rule_format]) using fST clo by auto have "compact (closure(f x))" by (simp add: bofx) moreover have "closure(f x) \ (\a \ f x. ball a (e/2))" using \0 < e\ by (force simp: closure_approachable simp del: divide_const_simps) ultimately obtain C where "C \ f x" "finite C" "closure(f x) \ (\a \ C. ball a (e/2))" apply (rule compactE, force) by (metis finite_subset_image) then have fx_cover: "f x \ (\a \ C. ball a (e/2))" by (meson closure_subset order_trans) with fx_ne have "C \ {}" by blast have xin: "x \ (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" using \x \ S\ \0 < e\ fST \C \ f x\ by force have "openin (top_of_set S) {x \ S. f x \ (T \ ball a (e/2)) \ {}}" for a by (simp add: openin_open_Int oo) then have "openin (top_of_set S) (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" by (simp add: Int_assoc openin_INT2 [OF \finite C\ \C \ {}\]) with xin obtain d2 where "d2>0" and d2: "\u v. \u \ S; dist u x < d2; v \ C\ \ f u \ T \ ball v (e/2) \ {}" unfolding openin_euclidean_subtopology_iff using xin by fastforce show ?thesis proof (intro that conjI ballI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by linarith next fix x' y assume "x' \ S" "dist x x' < min d1 d2" "y \ f x" then have dd2: "dist x' x < d2" by (auto simp: dist_commute) obtain a where "a \ C" "y \ ball a (e/2)" using fx_cover \y \ f x\ by auto then show "\y'. y' \ f x' \ dist y y' < e" using d2 [OF \x' \ S\ dd2] dist_triangle_half_r by fastforce next fix x' y' assume "x' \ S" "dist x x' < min d1 d2" "y' \ f x'" then have "dist x' x < d1" by (auto simp: dist_commute) then have "y' \ (\a\f x. \b\ball 0 e. {a + b})" using d1 [OF \x' \ S\] \y' \ f x'\ by force then show "\y. y \ f x \ dist y' y < e" apply auto by (metis add_diff_cancel_left' dist_norm) qed qed subsection\Complex logs exist on various "well-behaved" sets\ lemma continuous_logarithm_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" proof - obtain c where hom: "homotopic_with_canon (\h. True) S (-{0}) f (\x. c)" using nullhomotopic_from_contractible assms by (metis imageE subset_Compl_singleton) then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma continuous_logarithm_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" using covering_space_lift [OF covering_space_exp_punctured_plane S contf] by (metis (full_types) f imageE subset_Compl_singleton) lemma continuous_logarithm_on_cball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (cball a r) f" and "\z. z \ cball a r \ f z \ 0" obtains h where "continuous_on (cball a r) h" "\z. z \ cball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_logarithm_on_ball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (ball a r) f" and "\z. z \ ball a r \ f z \ 0" obtains h where "continuous_on (ball a r) h" "\z. z \ ball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_sqrt_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" and "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_contractible [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed lemma continuous_sqrt_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_simply_connected [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed subsection\Another simple case where sphere maps are nullhomotopic\ lemma inessential_spheremap_2_aux: fixes f :: "'a::euclidean_space \ complex" assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere 0 1)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) f (\x. c)" proof - obtain g where contg: "continuous_on (sphere a r) g" and feq: "\x. x \ sphere a r \ f x = exp(g x)" proof (rule continuous_logarithm_on_simply_connected [OF contf]) show "simply_connected (sphere a r)" using 2 by (simp add: simply_connected_sphere_eq) show "locally path_connected (sphere a r)" by (simp add: locally_path_connected_sphere) show "\z. z \ sphere a r \ f z \ 0" using fim by force qed auto have "\g. continuous_on (sphere a r) g \ (\x\sphere a r. f x = exp (\ * complex_of_real (g x)))" proof (intro exI conjI) show "continuous_on (sphere a r) (Im \ g)" by (intro contg continuous_intros continuous_on_compose) show "\x\sphere a r. f x = exp (\ * complex_of_real ((Im \ g) x))" using exp_eq_polar feq fim norm_exp_eq_Re by auto qed with inessential_eq_continuous_logarithm_circle that show ?thesis by metis qed lemma inessential_spheremap_2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis using contf contractible_sphere fim nullhomotopic_into_contractible that by blast next case False then have "sphere b s homeomorphic sphere (0::complex) 1" using assms by (simp add: homeomorphic_spheres_gen) then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k" by (auto simp: homeomorphic_def) then have conth: "continuous_on (sphere b s) h" and contk: "continuous_on (sphere 0 1) k" and him: "h ` sphere b s \ sphere 0 1" and kim: "k ` sphere 0 1 \ sphere b s" by (simp_all add: homeomorphism_def) obtain c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) (h \ f) (\x. c)" proof (rule inessential_spheremap_2_aux [OF a2]) show "continuous_on (sphere a r) (h \ f)" by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim) show "(h \ f) ` sphere a r \ sphere 0 1" using fim him by force qed auto then have "homotopic_with_canon (\f. True) (sphere a r) (sphere b s) (k \ (h \ f)) (k \ (\x. c))" by (rule homotopic_compose_continuous_left [OF _ contk kim]) then have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" apply (rule homotopic_with_eq, auto) by (metis fim hk homeomorphism_def image_subset_iff mem_sphere) then show ?thesis by (metis that) qed subsection\Holomorphic logarithms and square roots\ lemma contractible_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_contractible [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" apply (subst Lim_at_zero) apply (simp add: DERIV_D cong: if_cong Lim_cong_within) done qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" apply (rule eventually_mono) apply (auto simp: exp_eq dist_norm norm_mult) done then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed (*Identical proofs*) lemma simply_connected_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_simply_connected [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" apply (subst Lim_at_zero) apply (simp add: DERIV_D cong: if_cong Lim_cong_within) done qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" apply (rule eventually_mono) apply (auto simp: exp_eq dist_norm norm_mult) done then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed lemma contractible_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using contractible_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" apply (auto simp: feq) by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) qed qed lemma simply_connected_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using simply_connected_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" apply (auto simp: feq) by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) qed qed text\ Related theorems about holomorphic inverse cosines.\ lemma contractible_imp_holomorphic_arccos: assumes holf: "f holomorphic_on S" and S: "contractible S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = cos(g z)" proof - have hol1f: "(\z. 1 - f z ^ 2) holomorphic_on S" by (intro holomorphic_intros holf) obtain g where holg: "g holomorphic_on S" and eq: "\z. z \ S \ 1 - (f z)\<^sup>2 = (g z)\<^sup>2" using contractible_imp_holomorphic_sqrt [OF hol1f S] by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff) have holfg: "(\z. f z + \*g z) holomorphic_on S" by (intro holf holg holomorphic_intros) have "\z. z \ S \ f z + \*g z \ 0" by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff) then obtain h where holh: "h holomorphic_on S" and fgeq: "\z. z \ S \ f z + \*g z = exp (h z)" using contractible_imp_holomorphic_log [OF holfg S] by metis show ?thesis proof show "(\z. -\*h z) holomorphic_on S" by (intro holh holomorphic_intros) show "f z = cos (- \*h z)" if "z \ S" for z proof - have "(f z + \*g z)*(f z - \*g z) = 1" using that eq by (auto simp: algebra_simps power2_eq_square) then have "f z - \*g z = inverse (f z + \*g z)" using inverse_unique by force also have "... = exp (- h z)" by (simp add: exp_minus fgeq that) finally have "f z = exp (- h z) + \*g z" by (simp add: diff_eq_eq) then show ?thesis apply (simp add: cos_exp_eq) by (metis fgeq add.assoc mult_2_right that) qed qed qed lemma contractible_imp_holomorphic_arccos_bounded: assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \ S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "norm(g a) \ pi + norm(f a)" "\z. z \ S \ f z = cos(g z)" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = cos (g z)" using contractible_imp_holomorphic_arccos [OF holf S non1] by blast obtain b where "cos b = f a" "norm b \ pi + norm (f a)" using cos_Arccos norm_Arccos_bounded by blast then have "cos b = cos (g a)" by (simp add: \a \ S\ feq) then consider n where "n \ \" "b = g a + of_real(2*n*pi)" | n where "n \ \" "b = -g a + of_real(2*n*pi)" by (auto simp: complex_cos_eq) then show ?thesis proof cases case 1 show ?thesis proof show "(\z. g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "1" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed next case 2 show ?thesis proof show "(\z. -g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (-g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "2" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (-g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed qed qed subsection\The "Borsukian" property of sets\ text\This doesn't have a standard name. Kuratowski uses ``contractible with respect to \[S\<^sup>1]\'' while Whyburn uses ``property b''. It's closely related to unicoherence.\ definition\<^marker>\tag important\ Borsukian where "Borsukian S \ \f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\a. homotopic_with_canon (\h. True) S (- {0}) f (\x. a))" lemma Borsukian_retraction_gen: assumes "Borsukian S" "continuous_on S h" "h ` S = T" "continuous_on T k" "k ` T \ S" "\y. y \ T \ h(k y) = y" shows "Borsukian T" proof - interpret R: Retracts S h T k using assms by (simp add: Retracts.intro) show ?thesis using assms apply (simp add: Borsukian_def, clarify) apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto) done qed lemma retract_of_Borsukian: "\Borsukian T; S retract_of T\ \ Borsukian S" apply (auto simp: retract_of_def retraction_def) apply (erule (1) Borsukian_retraction_gen) apply (meson retraction retraction_def) - apply (auto simp: continuous_on_id) + apply (auto) done lemma homeomorphic_Borsukian: "\Borsukian S; S homeomorphic T\ \ Borsukian T" using Borsukian_retraction_gen order_refl by (fastforce simp add: homeomorphism_def homeomorphic_def) lemma homeomorphic_Borsukian_eq: "S homeomorphic T \ Borsukian S \ Borsukian T" by (meson homeomorphic_Borsukian homeomorphic_sym) lemma Borsukian_translation: fixes S :: "'a::real_normed_vector set" shows "Borsukian (image (\x. a + x) S) \ Borsukian S" apply (rule homeomorphic_Borsukian_eq) using homeomorphic_translation homeomorphic_sym by blast lemma Borsukian_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "Borsukian(f ` S) \ Borsukian S" apply (rule homeomorphic_Borsukian_eq) using assms homeomorphic_sym linear_homeomorphic_image by blast lemma homotopy_eqv_Borsukianness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "S homotopy_eqv T" shows "(Borsukian S \ Borsukian T)" by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null) lemma Borsukian_alt: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f g. continuous_on S f \ f ` S \ -{0} \ continuous_on S g \ g ` S \ -{0} \ homotopic_with_canon (\h. True) S (- {0::complex}) f g)" unfolding Borsukian_def homotopic_triviality by (simp add: path_connected_punctured_universe) lemma Borsukian_continuous_logarithm: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" by (simp add: Borsukian_def inessential_eq_continuous_logarithm) lemma Borsukian_continuous_logarithm_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp: Borsukian_continuous_logarithm) next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on S f" and 0: "0 \ f ` S" then have "continuous_on S (\x. f x / complex_of_real (cmod (f x)))" by (intro continuous_intros) auto moreover have "(\x. f x / complex_of_real (cmod (f x))) ` S \ sphere 0 1" using 0 by (auto simp: norm_divide) ultimately obtain g where contg: "continuous_on S g" and fg: "\x \ S. f x / complex_of_real (cmod (f x)) = exp(g x)" using RHS [of "\x. f x / of_real(norm(f x))"] by auto show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" proof (intro exI ballI conjI) show "continuous_on S (\x. (Ln \ of_real \ norm \ f)x + g x)" by (intro continuous_intros contf contg conjI) (use "0" in auto) show "f x = exp ((Ln \ complex_of_real \ cmod \ f) x + g x)" if "x \ S" for x using 0 that apply (clarsimp simp: exp_add) apply (subst exp_Ln, force) by (metis eq_divide_eq exp_not_eq_zero fg mult.commute) qed qed qed lemma Borsukian_continuous_logarithm_circle_real: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S (complex_of_real \ g) \ (\x \ S. f x = exp(\ * of_real(g x)))))" (is "?lhs = ?rhs") proof assume LHS: ?lhs show ?rhs proof (clarify) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" using LHS by (auto simp: Borsukian_continuous_logarithm_circle) then have "\x\S. f x = exp (\ * complex_of_real ((Im \ g) x))" using f01 apply (simp add: image_iff subset_iff) by (metis cis_conv_exp exp_eq_polar mult.left_neutral norm_exp_eq_Re of_real_1) then show "\g. continuous_on S (complex_of_real \ g) \ (\x\S. f x = exp (\ * complex_of_real (g x)))" by (rule_tac x="Im \ g" in exI) (force intro: continuous_intros contg) qed next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm_circle) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S (complex_of_real \ g)" and "\x. x \ S \ f x = exp(\ * of_real(g x))" by (metis RHS) then show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" by (rule_tac x="\x. \* of_real(g x)" in exI) (auto simp: continuous_intros contg) qed qed lemma Borsukian_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\a. homotopic_with_canon (\h. True) S (sphere (0::complex) 1) f (\x. a)))" by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real) lemma contractible_imp_Borsukian: "contractible S \ Borsukian S" by (meson Borsukian_def nullhomotopic_from_contractible) lemma simply_connected_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "\simply_connected S; locally path_connected S\ \ Borsukian S" apply (simp add: Borsukian_continuous_logarithm) by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff) lemma starlike_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "starlike S \ Borsukian S" by (simp add: contractible_imp_Borsukian starlike_imp_contractible) lemma Borsukian_empty: "Borsukian {}" by (auto simp: contractible_imp_Borsukian) lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)" by (auto simp: contractible_imp_Borsukian) lemma convex_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "convex S \ Borsukian S" by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible) proposition Borsukian_sphere: fixes a :: "'a::euclidean_space" shows "3 \ DIM('a) \ Borsukian (sphere a r)" apply (rule simply_connected_imp_Borsukian) using simply_connected_sphere apply blast using ENR_imp_locally_path_connected ENR_sphere by blast proposition Borsukian_open_Un: fixes S :: "'a::real_normed_vector set" assumes opeS: "openin (top_of_set (S \ T)) S" and opeT: "openin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" apply (rule continuous_on_cases_local_open [OF opeS opeT contg conth]) using True by blast show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) apply (meson contg continuous_on_subset inf_le1) by (meson conth continuous_on_subset inf_sup_ord(2)) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed 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 (auto simp: constant_on_def) 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 qed qed text\The proof is a duplicate of that of \Borsukian_open_Un\.\ lemma Borsukian_closed_Un: fixes S :: "'a::real_normed_vector set" assumes cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" apply (rule continuous_on_cases_local [OF cloS cloT contg conth]) using True by blast show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) apply (meson contg continuous_on_subset inf_le1) by (meson conth continuous_on_subset inf_sup_ord(2)) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed 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 (auto simp: constant_on_def) 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 qed qed lemma Borsukian_separation_compact: fixes S :: "complex set" assumes "compact S" shows "Borsukian S \ connected(- S)" by (simp add: Borsuk_separation_theorem Borsukian_circle assms) lemma Borsukian_monotone_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and conn: "\y. y \ T \ connected {x. x \ S \ f x = y}" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix g :: "'b \ complex" assume contg: "continuous_on T g" and 0: "0 \ g ` T" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ -{0}" using fim 0 by auto ultimately obtain h where conth: "continuous_on S h" and gfh: "\x. x \ S \ (g \ f) x = exp(h x)" using \Borsukian S\ by (auto simp: Borsukian_continuous_logarithm) have "\y. \x. y \ T \ x \ S \ f x = y" using fim by auto then obtain f' where f': "\y. y \ T \ f' y \ S \ f (f' y) = y" by metis have *: "(\x. h x - h(f' y)) constant_on {x. x \ S \ f x = y}" if "y \ T" for y 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 - have "h u = h x" if "u \ S" "f u = y" "cmod (h u - h x) < 2 * pi" for u proof (rule exp_complex_eqI) have "\Im (h u) - Im (h x)\ \ cmod (h u - h x)" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (h u) - Im (h x)\ < 2 * pi" using that by linarith show "exp (h u) = exp (h x)" by (simp add: gfh [symmetric] x that) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed have "h x = h (f' (f x))" if "x \ S" for x using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq) moreover have "\x. x \ T \ \u. u \ S \ x = f u \ h (f' x) = h u" using f' by fastforce ultimately have eq: "((\x. (x, (h \ f') x)) ` T) = {p. \x. x \ S \ (x, p) \ (S \ UNIV) \ ((\z. snd z - ((f \ fst) z, (h \ fst) z)) -` {0})}" using fim by (auto simp: image_iff) show "\h. continuous_on T h \ (\x\T. g x = exp (h x))" proof (intro exI conjI) show "continuous_on T (h \ f')" proof (rule continuous_from_closed_graph [of "h ` S"]) show "compact (h ` S)" by (simp add: \compact S\ compact_continuous_image conth) show "(h \ f') ` T \ h ` S" by (auto simp: f') show "closed ((\x. (x, (h \ f') x)) ` T)" apply (subst eq) apply (intro closed_compact_projection [OF \compact S\] continuous_closed_preimage continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth]) apply (auto simp: \compact S\ closed_Times compact_imp_closed) done qed qed (use f' gfh in fastforce) qed lemma Borsukian_open_map_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and ope: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real) fix g :: "'b \ complex" assume contg: "continuous_on T g" and gim: "g ` T \ sphere 0 1" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ sphere 0 1" using fim gim by auto ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \ h)" and gfh: "\x. x \ S \ (g \ f) x = exp(\ * of_real(h x))" using \Borsukian S\ Borsukian_continuous_logarithm_circle_real by metis then have conth: "continuous_on S h" by simp have "\x. x \ S \ f x = y \ (\x' \ S. f x' = y \ h x \ h x')" if "y \ T" for y proof - have 1: "compact (h ` {x \ S. f x = y})" proof (rule compact_continuous_image) show "continuous_on {x \ S. f x = y} h" by (rule continuous_on_subset [OF conth]) auto have "compact (S \ f -` {y})" by (rule proper_map_from_compact [OF contf _ \compact S\, of T]) (simp_all add: fim that) then show "compact {x \ S. f x = y}" by (auto simp: vimage_def Int_def) qed have 2: "h ` {x \ S. f x = y} \ {}" using fim that by auto have "\s \ h ` {x \ S. f x = y}. \t \ h ` {x \ S. f x = y}. s \ t" using compact_attains_inf [OF 1 2] by blast then show ?thesis by auto qed then obtain k where kTS: "\y. y \ T \ k y \ S" and fk: "\y. y \ T \ f (k y) = y " and hle: "\x' y. \y \ T; x' \ S; f x' = y\ \ h (k y) \ h x'" by metis have "continuous_on T (h \ k)" proof (clarsimp simp add: continuous_on_iff) fix y and e::real assume "y \ T" "0 < e" moreover have "uniformly_continuous_on S (complex_of_real \ h)" using \compact S\ cont_cxh compact_uniformly_continuous by blast ultimately obtain d where "0 < d" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (h x') (h x) < e" by (force simp: uniformly_continuous_on_def) obtain \ where "0 < \" and \: "\x'. \x' \ T; dist y x' < \\ \ (\v \ {z \ S. f z = y}. \v'. v' \ {z \ S. f z = x'} \ dist v v' < d) \ (\v' \ {z \ S. f z = x'}. \v. v \ {z \ S. f z = y} \ dist v' v < d)" proof (rule upper_lower_hemicontinuous_explicit [of T "\y. {z \ S. f z = y}" S]) show "\U. openin (top_of_set S) U \ openin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]] by (simp add: Abstract_Topology_2.continuous_imp_closed_map \compact S\ contf fim) show "\U. closedin (top_of_set S) U \ closedin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] by meson show "bounded {z \ S. f z = y}" by (metis (no_types, lifting) compact_imp_bounded [OF \compact S\] bounded_subset mem_Collect_eq subsetI) qed (use \y \ T\ \0 < d\ fk kTS in \force+\) have "dist (h (k y')) (h (k y)) < e" if "y' \ T" "dist y y' < \" for y' proof - have k1: "k y \ S" "f (k y) = y" and k2: "k y' \ S" "f (k y') = y'" by (auto simp: \y \ T\ \y' \ T\ kTS fk) have 1: "\v. \v \ S; f v = y\ \ \v'. v' \ {z \ S. f z = y'} \ dist v v' < d" and 2: "\v'. \v' \ S; f v' = y'\ \ \v. v \ {z \ S. f z = y} \ dist v' v < d" using \ [OF that] by auto then obtain w' w where "w' \ S" "f w' = y'" "dist (k y) w' < d" and "w \ S" "f w = y" "dist (k y') w < d" using 1 [OF k1] 2 [OF k2] by auto then show ?thesis using d [of w "k y'"] d [of w' "k y"] k1 k2 \y' \ T\ \y \ T\ hle by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps) qed then show "\d>0. \x'\T. dist x' y < d \ dist (h (k x')) (h (k y)) < e" using \0 < \\ by (auto simp: dist_commute) qed then show "\h. continuous_on T h \ (\x\T. g x = exp (\ * complex_of_real (h x)))" using fk gfh kTS by force qed text\If two points are separated by a closed set, there's a minimal one.\ proposition closed_irreducible_separator: fixes a :: "'a::real_normed_vector" assumes "closed S" and ab: "\ connected_component (- S) a b" obtains T where "T \ S" "closed T" "T \ {}" "\ connected_component (- T) a b" "\U. U \ T \ connected_component (- U) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis proof assume *: "a \ S" show ?thesis proof show "{a} \ S" using * by blast show "\ connected_component (- {a}) a b" using connected_component_in by auto show "\U. U \ {a} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto next assume *: "b \ S" show ?thesis proof show "{b} \ S" using * by blast show "\ connected_component (- {b}) a b" using connected_component_in by auto show "\U. U \ {b} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto qed next case False define A where "A \ connected_component_set (- S) a" define B where "B \ connected_component_set (- (closure A)) b" have "a \ A" using False A_def by auto have "b \ B" unfolding A_def B_def closure_Un_frontier using ab False \closed S\ frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force have "frontier B \ frontier (connected_component_set (- closure A) b)" using B_def by blast also have frsub: "... \ frontier A" proof - have "\A. closure (- closure (- A)) \ closure A" by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl) then show ?thesis by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans) qed finally have frBA: "frontier B \ frontier A" . show ?thesis proof show "frontier B \ S" proof - have "frontier S \ S" by (simp add: \closed S\ frontier_subset_closed) then show ?thesis using frsub frontier_complement frontier_of_connected_component_subset unfolding A_def B_def by blast qed show "closed (frontier B)" by simp show "\ connected_component (- frontier B) a b" unfolding connected_component_def proof clarify fix T assume "connected T" and TB: "T \ - frontier B" and "a \ T" and "b \ T" have "a \ B" by (metis A_def B_def ComplD \a \ A\ assms(1) closed_open connected_component_subset in_closure_connected_component subsetD) have "T \ B \ {}" using \b \ B\ \b \ T\ by blast moreover have "T - B \ {}" using \a \ B\ \a \ T\ by blast ultimately show "False" using connected_Int_frontier [of T B] TB \connected T\ by blast qed moreover have "connected_component (- frontier B) a b" if "frontier B = {}" apply (simp add: that) using connected_component_eq_UNIV by blast ultimately show "frontier B \ {}" by blast show "connected_component (- U) a b" if "U \ frontier B" for U proof - obtain p where Usub: "U \ frontier B" and p: "p \ frontier B" "p \ U" using \U \ frontier B\ by blast show ?thesis unfolding connected_component_def proof (intro exI conjI) have "connected ((insert p A) \ (insert p B))" proof (rule connected_Un) show "connected (insert p A)" by (metis A_def IntD1 frBA \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI) show "connected (insert p B)" by (metis B_def IntD1 \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI) qed blast then show "connected (insert p (B \ A))" by (simp add: sup.commute) have "A \ - U" using A_def Usub \frontier B \ S\ connected_component_subset by fastforce moreover have "B \ - U" using B_def Usub connected_component_subset frBA frontier_closures by fastforce ultimately show "insert p (B \ A) \ - U" using p by auto qed (auto simp: \a \ A\ \b \ B\) qed qed qed lemma frontier_minimal_separating_closed_pointwise: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" "a \ S" and nconn: "\ connected_component (- S) a b" and conn: "\T. \closed T; T \ S\ \ connected_component (- T) a b" shows "frontier(connected_component_set (- S) a) = S" (is "?F = S") proof - have "?F \ S" by (simp add: S componentsI frontier_of_components_closed_complement) moreover have False if "?F \ S" proof - have "connected_component (- ?F) a b" by (simp add: conn that) then obtain T where "connected T" "T \ -?F" "a \ T" "b \ T" by (auto simp: connected_component_def) moreover have "T \ ?F \ {}" proof (rule connected_Int_frontier [OF \connected T\]) show "T \ connected_component_set (- S) a \ {}" using \a \ S\ \a \ T\ by fastforce show "T - connected_component_set (- S) a \ {}" using \b \ T\ nconn by blast qed ultimately show ?thesis by blast qed ultimately show ?thesis by blast qed subsection\Unicoherence (closed)\ definition\<^marker>\tag important\ unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ closedin (top_of_set U) S \ closedin (top_of_set U) T \ connected (S \ T)" lemma unicoherentI [intro?]: assumes "\S T. \connected S; connected T; U = S \ T; closedin (top_of_set U) S; closedin (top_of_set U) T\ \ connected (S \ T)" shows "unicoherent U" using assms unfolding unicoherent_def by blast lemma unicoherentD: assumes "unicoherent U" "connected S" "connected T" "U = S \ T" "closedin (top_of_set U) S" "closedin (top_of_set U) T" shows "connected (S \ T)" using assms unfolding unicoherent_def by blast proposition homeomorphic_unicoherent: assumes ST: "S homeomorphic T" and S: "unicoherent S" shows "unicoherent T" proof - obtain f g where gf: "\x. x \ S \ g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S" and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g" using ST by (auto simp: homeomorphic_def homeomorphism_def) show ?thesis proof fix U V assume "connected U" "connected V" and T: "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" have "f ` (g ` U \ g ` V) \ U" "f ` (g ` U \ g ` V) \ V" using gf fim T by auto (metis UnCI image_iff)+ moreover have "U \ V \ f ` (g ` U \ g ` V)" using gf fim by (force simp: image_iff T) ultimately have "U \ V = f ` (g ` U \ g ` V)" by blast moreover have "connected (f ` (g ` U \ g ` V))" proof (rule connected_continuous_image) show "continuous_on (g ` U \ g ` V) f" apply (rule continuous_on_subset [OF contf]) using T fim gfim by blast show "connected (g ` U \ g ` V)" proof (intro conjI unicoherentD [OF S]) show "connected (g ` U)" "connected (g ` V)" using \connected U\ cloU \connected V\ cloV by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+ show "S = g ` U \ g ` V" using T fim gfim by auto have hom: "homeomorphism T S g f" by (simp add: contf contg fim gf gfim homeomorphism_def) have "closedin (top_of_set T) U" "closedin (top_of_set T) V" by (simp_all add: cloU cloV) then show "closedin (top_of_set S) (g ` U)" "closedin (top_of_set S) (g ` V)" by (blast intro: homeomorphism_imp_closed_map [OF hom])+ qed qed ultimately show "connected (U \ V)" by metis qed qed lemma homeomorphic_unicoherent_eq: "S homeomorphic T \ (unicoherent S \ unicoherent T)" by (meson homeomorphic_sym homeomorphic_unicoherent) lemma unicoherent_translation: fixes S :: "'a::real_normed_vector set" shows "unicoherent (image (\x. a + x) S) \ unicoherent S" using homeomorphic_translation homeomorphic_unicoherent_eq by blast lemma unicoherent_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(unicoherent(f ` S) \ unicoherent S)" using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast lemma Borsukian_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "Borsukian U" shows "unicoherent U" unfolding unicoherent_def proof clarify fix S T assume "connected S" "connected T" "U = S \ T" and cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" show "connected (S \ T)" unfolding connected_closedin_eq proof clarify fix V W assume "closedin (top_of_set (S \ T)) V" and "closedin (top_of_set (S \ T)) W" and VW: "V \ W = S \ T" "V \ W = {}" and "V \ {}" "W \ {}" then have cloV: "closedin (top_of_set U) V" and cloW: "closedin (top_of_set U) W" using \U = S \ T\ cloS cloT closedin_trans by blast+ obtain q where contq: "continuous_on U q" and q01: "\x. x \ U \ q x \ {0..1::real}" and qV: "\x. x \ V \ q x = 0" and qW: "\x. x \ W \ q x = 1" by (rule Urysohn_local [OF cloV cloW \V \ W = {}\, of 0 1]) (fastforce simp: closed_segment_eq_real_ivl) let ?h = "\x. if x \ S then exp(pi * \ * q x) else 1 / exp(pi * \ * q x)" have eqST: "exp(pi * \ * q x) = 1 / exp(pi * \ * q x)" if "x \ S \ T" for x proof - have "x \ V \ W" using that \V \ W = S \ T\ by blast with qV qW show ?thesis by force qed obtain g where contg: "continuous_on U g" and circle: "g ` U \ sphere 0 1" and S: "\x. x \ S \ g x = exp(pi * \ * q x)" and T: "\x. x \ T \ g x = 1 / exp(pi * \ * q x)" proof show "continuous_on U ?h" unfolding \U = S \ T\ proof (rule continuous_on_cases_local [OF cloS cloT]) show "continuous_on S (\x. exp (pi * \ * q x))" apply (intro continuous_intros) using \U = S \ T\ continuous_on_subset contq by blast show "continuous_on T (\x. 1 / exp (pi * \ * q x))" apply (intro continuous_intros) using \U = S \ T\ continuous_on_subset contq by auto qed (use eqST in auto) qed (use eqST in \auto simp: norm_divide\) then obtain h where conth: "continuous_on U h" and heq: "\x. x \ U \ g x = exp (h x)" by (metis Borsukian_continuous_logarithm_circle assms) obtain v w where "v \ V" "w \ W" using \V \ {}\ \W \ {}\ by blast then have vw: "v \ S \ T" "w \ S \ T" using VW by auto have iff: "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 1 \ abs (m - n)" for m n proof - have "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 2 * pi \ cmod ((2 * pi * \) * (of_int m - of_int n))" by (simp add: algebra_simps) also have "... \ 2 * pi \ 2 * pi * cmod (of_int m - of_int n)" by (simp add: norm_mult) also have "... \ 1 \ abs (m - n)" by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff) finally show ?thesis . qed have *: "\n::int. h x - (pi * \ * q x) = (of_int(2*n) * pi) * \" if "x \ S" for x using that S \U = S \ T\ heq exp_eq [symmetric] by (simp add: algebra_simps) moreover have "(\x. h x - (pi * \ * q x)) constant_on S" proof (rule continuous_discrete_range_constant [OF \connected S\]) have "continuous_on S h" "continuous_on S q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on S (\x. h x - (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" if "x \ S" "y \ S" and ne: "h y - (pi * \ * q y) \ h x - (pi * \ * q x)" for x y using * [OF \x \ S\] * [OF \y \ S\] ne by (auto simp: iff) then show "\x. x \ S \ \e>0. \y. y \ S \ h y - (pi * \ * q y) \ h x - (pi * \ * q x) \ e \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain m where m: "\x. x \ S \ h x - (pi * \ * q x) = (of_int(2*m) * pi) * \" using vw by (force simp: constant_on_def) have *: "\n::int. h x = - (pi * \ * q x) + (of_int(2*n) * pi) * \" if "x \ T" for x unfolding exp_eq [symmetric] using that T \U = S \ T\ by (simp add: exp_minus field_simps heq [symmetric]) moreover have "(\x. h x + (pi * \ * q x)) constant_on T" proof (rule continuous_discrete_range_constant [OF \connected T\]) have "continuous_on T h" "continuous_on T q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on T (\x. h x + (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" if "x \ T" "y \ T" and ne: "h y + (pi * \ * q y) \ h x + (pi * \ * q x)" for x y using * [OF \x \ T\] * [OF \y \ T\] ne by (auto simp: iff) then show "\x. x \ T \ \e>0. \y. y \ T \ h y + (pi * \ * q y) \ h x + (pi * \ * q x) \ e \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain n where n: "\x. x \ T \ h x + (pi * \ * q x) = (of_int(2*n) * pi) * \" using vw by (force simp: constant_on_def) show "False" using m [of v] m [of w] n [of v] n [of w] vw by (auto simp: algebra_simps \v \ V\ \w \ W\ qV qW) qed qed corollary contractible_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "contractible U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian) corollary convex_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "convex U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian) text\If the type class constraint can be relaxed, I don't know how!\ corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)" by (simp add: convex_imp_unicoherent) lemma unicoherent_monotone_image_compact: fixes T :: "'b :: t2_space set" assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T" and conn: "\y. y \ T \ connected (S \ f -` {y})" shows "unicoherent T" proof fix U V assume UV: "connected U" "connected V" "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" moreover have "compact T" using \compact S\ compact_continuous_image contf fim by blast ultimately have "closed U" "closed V" by (auto simp: closedin_closed_eq compact_imp_closed) let ?SUV = "(S \ f -` U) \ (S \ f -` V)" have UV_eq: "f ` ?SUV = U \ V" using \T = U \ V\ fim by force+ have "connected (f ` ?SUV)" proof (rule connected_continuous_image) show "continuous_on ?SUV f" by (meson contf continuous_on_subset inf_le1) show "connected ?SUV" proof (rule unicoherentD [OF \unicoherent S\, of "S \ f -` U" "S \ f -` V"]) have "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" by (metis \compact S\ closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono) then show "connected (S \ f -` U)" "connected (S \ f -` V)" using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim]) show "S = (S \ f -` U) \ (S \ f -` V)" using UV fim by blast show "closedin (top_of_set S) (S \ f -` U)" "closedin (top_of_set S) (S \ f -` V)" by (auto simp: continuous_on_imp_closedin cloU cloV contf fim) qed qed with UV_eq show "connected (U \ V)" by simp qed subsection\Several common variants of unicoherence\ lemma connected_frontier_simple: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected(- S)" shows "connected(frontier S)" unfolding frontier_closures apply (rule unicoherentD [OF unicoherent_UNIV]) apply (simp_all add: assms connected_imp_connected_closure) by (simp add: closure_def) lemma connected_frontier_component_complement: fixes S :: "'a :: euclidean_space set" assumes "connected S" and C: "C \ components(- S)" shows "connected(frontier C)" apply (rule connected_frontier_simple) using C in_components_connected apply blast by (metis Compl_eq_Diff_UNIV connected_UNIV assms top_greatest component_complement_connected) lemma connected_frontier_disjoint: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \ frontier T" shows "connected(frontier S)" proof (cases "S = UNIV") case True then show ?thesis by simp next case False then have "-S \ {}" by blast then obtain C where C: "C \ components(- S)" and "T \ C" by (metis ComplI disjnt_iff subsetI exists_component_superset \disjnt S T\ \connected T\) moreover have "frontier S = frontier C" proof - have "frontier C \ frontier S" using C frontier_complement frontier_of_components_subset by blast moreover have "x \ frontier C" if "x \ frontier S" for x proof - have "x \ closure C" using that unfolding frontier_def by (metis (no_types) Diff_eq ST \T \ C\ closure_mono contra_subsetD frontier_def le_inf_iff that) moreover have "x \ interior C" using that unfolding frontier_def by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono) ultimately show ?thesis by (auto simp: frontier_def) qed ultimately show ?thesis by blast qed ultimately show ?thesis using \connected S\ connected_frontier_component_complement by auto qed subsection\Some separation results\ lemma separation_by_component_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected_component (- S) a b" obtains C where "C \ components S" "\ connected_component(- C) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis using connected_component_in componentsI that by fastforce next case False obtain T where "T \ S" "closed T" "T \ {}" and nab: "\ connected_component (- T) a b" and conn: "\U. U \ T \ connected_component (- U) a b" using closed_irreducible_separator [OF assms] by metis moreover have "connected T" proof - have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T" using frontier_minimal_separating_closed_pointwise by (metis False \T \ S\ \closed T\ connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+ have "connected (frontier (connected_component_set (- T) a))" proof (rule connected_frontier_disjoint) show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)" unfolding disjnt_iff by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab) show "frontier (connected_component_set (- T) a) \ frontier (connected_component_set (- T) b)" by (simp add: ab) qed auto with ab \closed T\ show ?thesis by simp qed ultimately obtain C where "C \ components S" "T \ C" using exists_component_superset [of T S] by blast then show ?thesis by (meson Compl_anti_mono connected_component_of_subset nab that) qed lemma separation_by_component_closed: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain x y where "closed S" "x \ S" "y \ S" and "\ connected_component (- S) x y" using assms by (auto simp: connected_iff_connected_component) then obtain C where "C \ components S" "\ connected_component(- C) x y" using separation_by_component_closed_pointwise by metis then show "thesis" apply (clarify elim!: componentsE) by (metis Compl_iff \C \ components S\ \x \ S\ \y \ S\ connected_component_eq connected_component_eq_eq connected_iff_connected_component that) qed lemma separation_by_Un_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b" shows "connected_component (- (S \ T)) a b" proof (rule ccontr) have "a \ S" "b \ S" "a \ T" "b \ T" using conS conT connected_component_in by auto assume "\ connected_component (- (S \ T)) a b" then obtain C where "C \ components (S \ T)" and C: "\ connected_component(- C) a b" using separation_by_component_closed_pointwise assms by blast then have "C \ S \ C \ T" proof - have "connected C" "C \ S \ T" using \C \ components (S \ T)\ in_components_subset by (blast elim: componentsE)+ moreover then have "C \ T = {} \ C \ S = {}" by (metis Int_empty_right ST inf.commute connected_closed) ultimately show ?thesis by blast qed then show False by (meson Compl_anti_mono C conS conT connected_component_of_subset) qed lemma separation_by_Un_closed: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected(- S)" and conT: "connected(- T)" shows "connected(- (S \ T))" using assms separation_by_Un_closed_pointwise by (fastforce simp add: connected_iff_connected_component) lemma open_unicoherent_UNIV: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "connected S" "connected T" "S \ T = UNIV" shows "connected(S \ T)" proof - have "connected(- (-S \ -T))" by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms) then show ?thesis by simp qed lemma separation_by_component_open_aux: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and "S \ {}" "T \ {}" obtains C where "C \ components(-(S \ T))" "C \ {}" "frontier C \ S \ {}" "frontier C \ T \ {}" proof (rule ccontr) let ?S = "S \ \{C \ components(- (S \ T)). frontier C \ S}" let ?T = "T \ \{C \ components(- (S \ T)). frontier C \ T}" assume "\ thesis" with that have *: "frontier C \ S = {} \ frontier C \ T = {}" if C: "C \ components (- (S \ T))" "C \ {}" for C using C by blast have "\A B::'a set. closed A \ closed B \ UNIV \ A \ B \ A \ B = {} \ A \ {} \ B \ {}" proof (intro exI conjI) have "frontier (\{C \ components (- S \ - T). frontier C \ S}) \ S" apply (rule subset_trans [OF frontier_Union_subset_closure]) by (metis (no_types, lifting) SUP_least \closed S\ closure_minimal mem_Collect_eq) then have "frontier ?S \ S" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?S" using frontier_subset_eq by fastforce have "frontier (\{C \ components (- S \ - T). frontier C \ T}) \ T" apply (rule subset_trans [OF frontier_Union_subset_closure]) by (metis (no_types, lifting) SUP_least \closed T\ closure_minimal mem_Collect_eq) then have "frontier ?T \ T" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?T" using frontier_subset_eq by fastforce have "UNIV \ (S \ T) \ \(components(- (S \ T)))" using Union_components by blast also have "... \ ?S \ ?T" proof - have "C \ components (-(S \ T)) \ frontier C \ S \ C \ components (-(S \ T)) \ frontier C \ T" if "C \ components (- (S \ T))" "C \ {}" for C using * [OF that] that by clarify (metis (no_types, lifting) UnE \closed S\ \closed T\ closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE) then show ?thesis by blast qed finally show "UNIV \ ?S \ ?T" . have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} \ - (S \ T)" using in_components_subset by fastforce moreover have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} = {}" proof - have "C \ C' = {}" if "C \ components (- (S \ T))" "frontier C \ S" "C' \ components (- (S \ T))" "frontier C' \ T" for C C' proof - have NUN: "- S \ - T \ UNIV" using \T \ {}\ by blast have "C \ C'" proof assume "C = C'" with that have "frontier C' \ S \ T" by simp also have "... = {}" using \S \ T = {}\ by blast finally have "C' = {} \ C' = UNIV" using frontier_eq_empty by auto then show False using \C = C'\ NUN that by (force simp: dest: in_components_nonempty in_components_subset) qed with that show ?thesis by (simp add: components_nonoverlap [of _ "-(S \ T)"]) qed then show ?thesis by blast qed ultimately show "?S \ ?T = {}" using ST by blast show "?S \ {}" "?T \ {}" using \S \ {}\ \T \ {}\ by blast+ qed then show False by (metis Compl_disjoint connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI) qed proposition separation_by_component_open: fixes S :: "'a :: euclidean_space set" assumes "open S" and non: "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain T U where "closed T" "closed U" and TU: "T \ U = - S" "T \ U = {}" "T \ {}" "U \ {}" using assms by (auto simp: connected_closed_set closed_def) then obtain C where C: "C \ components(-(T \ U))" "C \ {}" and "frontier C \ T \ {}" "frontier C \ U \ {}" using separation_by_component_open_aux [OF \closed T\ \closed U\ \T \ U = {}\] by force show "thesis" proof show "C \ components S" using C(1) TU(1) by auto show "\ connected (- C)" proof assume "connected (- C)" then have "connected (frontier C)" using connected_frontier_simple [of C] \C \ components S\ in_components_connected by blast then show False unfolding connected_closed by (metis C(1) TU(2) \closed T\ \closed U\ \frontier C \ T \ {}\ \frontier C \ U \ {}\ closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute) qed qed qed lemma separation_by_Un_open: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "S \ T = {}" and cS: "connected(-S)" and cT: "connected(-T)" shows "connected(- (S \ T))" using assms unicoherent_UNIV unfolding unicoherent_def by force lemma nonseparation_by_component_eq: fixes S :: "'a :: euclidean_space set" assumes "open S \ closed S" shows "((\C \ components S. connected(-C)) \ connected(- S))" (is "?lhs = ?rhs") proof assume ?lhs with assms show ?rhs by (meson separation_by_component_closed separation_by_component_open) next assume ?rhs with assms show ?lhs using component_complement_connected by force qed text\Another interesting equivalent of an inessential mapping into C-{0}\ proposition inessential_eq_extensible: fixes f :: "'a::euclidean_space \ complex" assumes "closed S" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on UNIV g \ (\x \ S. g x = f x) \ (\x. g x \ 0))" (is "?lhs = ?rhs") proof assume ?lhs then obtain a where a: "homotopic_with_canon (\h. True) S (-{0}) f (\t. a)" .. show ?rhs proof (cases "S = {}") case True - with a show ?thesis - using continuous_on_const by force + with a show ?thesis by force next case False have anr: "ANR (-{0::complex})" by (simp add: ANR_delete open_Compl open_imp_ANR) obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \ -{0}" and gf: "\x. x \ S \ g x = f x" proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]]) show "closedin (top_of_set UNIV) S" using assms by auto show "range (\t. a) \ - {0}" using a homotopic_with_imp_subset2 False by blast qed (use anr that in \force+\) then show ?thesis by force qed next assume ?rhs then obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ S \ g x = f x" and non0: "\x. g x \ 0" by metis obtain h k::"'a\'a" where hk: "homeomorphism (ball 0 1) UNIV h k" using homeomorphic_ball01_UNIV homeomorphic_def by blast then have "continuous_on (ball 0 1) (g \ h)" by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest) then obtain j where contj: "continuous_on (ball 0 1) j" and j: "\z. z \ ball 0 1 \ exp(j z) = (g \ h) z" by (metis (mono_tags, hide_lams) continuous_logarithm_on_ball comp_apply non0) have [simp]: "\x. x \ S \ h (k x) = x" using hk homeomorphism_apply2 by blast have "\\. continuous_on S \\ (\x\S. f x = exp (\ x))" proof (intro exI conjI ballI) show "continuous_on S (j \ k)" proof (rule continuous_on_compose) show "continuous_on S k" by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest) show "continuous_on (k ` S) j" apply (rule continuous_on_subset [OF contj]) using homeomorphism_image2 [OF hk] continuous_on_subset [OF contj] by blast qed show "f x = exp ((j \ k) x)" if "x \ S" for x proof - have "f x = (g \ h) (k x)" by (simp add: gf that) also have "... = exp (j (k x))" by (metis rangeI homeomorphism_image2 [OF hk] j) finally show ?thesis by simp qed qed then show ?lhs by (simp add: inessential_eq_continuous_logarithm) qed lemma inessential_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes T: "path_connected T" and "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and hom: "\S. S \ \ \ \a. homotopic_with_canon (\x. True) S T f (\x. a)" obtains a where "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (cases "\\ = {}") case True with that show ?thesis by force next case False then obtain C where "C \ \" "C \ {}" by blast then obtain a where clo: "closedin (top_of_set (\\)) C" and ope: "openin (top_of_set (\\)) C" and "homotopic_with_canon (\x. True) C T f (\x. a)" using assms by blast with \C \ {}\ have "f ` C \ T" "a \ T" using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+ have "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (rule homotopic_on_clopen_Union) show "\S. S \ \ \ closedin (top_of_set (\\)) S" "\S. S \ \ \ openin (top_of_set (\\)) S" by (simp_all add: assms) show "homotopic_with_canon (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain b where "b \ S" by blast obtain c where c: "homotopic_with_canon (\x. True) S T f (\x. c)" using \S \ \\ hom by blast then have "c \ T" using \b \ S\ homotopic_with_imp_subset2 by blast then have "homotopic_with_canon (\x. True) S T (\x. a) (\x. c)" using T \a \ T\ homotopic_constant_maps path_connected_component by (simp add: homotopic_constant_maps path_connected_component) then show ?thesis using c homotopic_with_symD homotopic_with_trans by blast qed qed then show ?thesis .. qed proposition Janiszewski_dual: fixes S :: "complex set" assumes "compact S" "compact T" "connected S" "connected T" "connected(- (S \ T))" shows "connected(S \ T)" proof - have ST: "compact (S \ T)" by (simp add: assms compact_Un) with Borsukian_imp_unicoherent [of "S \ T"] ST assms show ?thesis by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def) qed end diff --git a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy --- a/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy +++ b/src/HOL/Analysis/Henstock_Kurzweil_Integration.thy @@ -1,7627 +1,7627 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) Huge cleanup by LCP *) section \Henstock-Kurzweil Gauge Integration in Many Dimensions\ theory Henstock_Kurzweil_Integration imports Lebesgue_Measure Tagged_Division begin lemma norm_diff2: "\y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \ e1; norm(y2 - x2) \ e2\ \ norm(y-x) \ e" using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] by (simp add: add_diff_add) lemma setcomp_dot1: "{z. P (z \ (i,0))} = {(x,y). P(x \ i)}" by auto lemma setcomp_dot2: "{z. P (z \ (0,i))} = {(x,y). P(y \ i)}" by auto lemma Sigma_Int_Paircomp1: "(Sigma A B) \ {(x, y). P x} = Sigma (A \ {x. P x}) B" by blast lemma Sigma_Int_Paircomp2: "(Sigma A B) \ {(x, y). P y} = Sigma A (\z. B z \ {y. P y})" by blast (* END MOVE *) subsection \Content (length, area, volume...) of an interval\ abbreviation content :: "'a::euclidean_space set \ real" where "content s \ measure lborel s" lemma content_cbox_cases: "content (cbox a b) = (if \i\Basis. a\i \ b\i then prod (\i. b\i - a\i) Basis else 0)" by (simp add: measure_lborel_cbox_eq inner_diff) lemma content_cbox: "\i\Basis. a\i \ b\i \ content (cbox a b) = (\i\Basis. b\i - a\i)" unfolding content_cbox_cases by simp lemma content_cbox': "cbox a b \ {} \ content (cbox a b) = (\i\Basis. b\i - a\i)" by (simp add: box_ne_empty inner_diff) lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \i\Basis. b\i - a\i)" by (simp add: content_cbox') lemma content_cbox_cart: "cbox a b \ {} \ content(cbox a b) = prod (\i. b$i - a$i) UNIV" by (simp add: content_cbox_if Basis_vec_def cart_eq_inner_axis axis_eq_axis prod.UNION_disjoint) lemma content_cbox_if_cart: "content(cbox a b) = (if cbox a b = {} then 0 else prod (\i. b$i - a$i) UNIV)" by (simp add: content_cbox_cart) lemma content_division_of: assumes "K \ \" "\ division_of S" shows "content K = (\i \ Basis. interval_upperbound K \ i - interval_lowerbound K \ i)" proof - obtain a b where "K = cbox a b" using cbox_division_memE assms by metis then show ?thesis using assms by (force simp: division_of_def content_cbox') qed lemma content_real: "a \ b \ content {a..b} = b - a" by simp lemma abs_eq_content: "\y - x\ = (if x\y then content {x..y} else content {y..x})" by (auto simp: content_real) lemma content_singleton: "content {a} = 0" by simp lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1" by simp lemma content_pos_le [iff]: "0 \ content X" by simp corollary\<^marker>\tag unimportant\ content_nonneg [simp]: "\ content (cbox a b) < 0" using not_le by blast lemma content_pos_lt: "\i\Basis. a\i < b\i \ 0 < content (cbox a b)" by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos) lemma content_eq_0: "content (cbox a b) = 0 \ (\i\Basis. b\i \ a\i)" by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl) lemma content_eq_0_interior: "content (cbox a b) = 0 \ interior(cbox a b) = {}" unfolding content_eq_0 interior_cbox box_eq_empty by auto lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \ (\i\Basis. a\i < b\i)" by (auto simp add: content_cbox_cases less_le prod_nonneg) lemma content_empty [simp]: "content {} = 0" by simp lemma content_real_if [simp]: "content {a..b} = (if a \ b then b - a else 0)" by (simp add: content_real) lemma content_subset: "cbox a b \ cbox c d \ content (cbox a b) \ content (cbox c d)" unfolding measure_def by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq) lemma content_lt_nz: "0 < content (cbox a b) \ content (cbox a b) \ 0" unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" unfolding measure_lborel_cbox_eq Basis_prod_def apply (subst prod.union_disjoint) apply (auto simp: bex_Un ball_Un) apply (subst (1 2) prod.reindex_nontrivial) apply auto done lemma content_cbox_pair_eq0_D: "content (cbox (a,c) (b,d)) = 0 \ content (cbox a b) = 0 \ content (cbox c d) = 0" by (simp add: content_Pair) lemma content_cbox_plus: fixes x :: "'a::euclidean_space" shows "content(cbox x (x + h *\<^sub>R One)) = (if h \ 0 then h ^ DIM('a) else 0)" by (simp add: algebra_simps content_cbox_if box_eq_empty) lemma content_0_subset: "content(cbox a b) = 0 \ s \ cbox a b \ content s = 0" using emeasure_mono[of s "cbox a b" lborel] by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq) lemma content_split: fixes a :: "'a::euclidean_space" assumes "k \ Basis" shows "content (cbox a b) = content(cbox a b \ {x. x\k \ c}) + content(cbox a b \ {x. x\k \ c})" \ \Prove using measure theory\ proof (cases "\i\Basis. a \ i \ b \ i") case True have 1: "\X Y Z. (\i\Basis. Z i (if i = k then X else Y i)) = Z k X * (\i\Basis-{k}. Z i (Y i))" by (simp add: if_distrib prod.delta_remove assms) note simps = interval_split[OF assms] content_cbox_cases have 2: "(\i\Basis. b\i - a\i) = (\i\Basis-{k}. b\i - a\i) * (b\k - a\k)" by (metis (no_types, lifting) assms finite_Basis mult.commute prod.remove) have "\x. min (b \ k) c = max (a \ k) c \ x * (b\k - a\k) = x * (max (a \ k) c - a \ k) + x * (b \ k - max (a \ k) c)" by (auto simp add: field_simps) moreover have **: "(\i\Basis. ((\i\Basis. (if i = k then min (b \ k) c else b \ i) *\<^sub>R i) \ i - a \ i)) = (\i\Basis. (if i = k then min (b \ k) c else b \ i) - a \ i)" "(\i\Basis. b \ i - ((\i\Basis. (if i = k then max (a \ k) c else a \ i) *\<^sub>R i) \ i)) = (\i\Basis. b \ i - (if i = k then max (a \ k) c else a \ i))" by (auto intro!: prod.cong) have "\ a \ k \ c \ \ c \ b \ k \ False" unfolding not_le using True assms by auto ultimately show ?thesis using assms unfolding simps ** 1[of "\i x. b\i - x"] 1[of "\i x. x - a\i"] 2 by auto next case False then have "cbox a b = {}" unfolding box_eq_empty by (auto simp: not_le) then show ?thesis by (auto simp: not_le) qed lemma division_of_content_0: assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \ d" shows "content K = 0" unfolding forall_in_division[OF assms(2)] by (meson assms content_0_subset division_of_def) lemma sum_content_null: assumes "content (cbox a b) = 0" and "p tagged_division_of (cbox a b)" shows "(\(x,K)\p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" proof (rule sum.neutral, rule) fix y assume y: "y \ p" obtain x K where xk: "y = (x, K)" using surj_pair[of y] by blast then obtain c d where k: "K = cbox c d" "K \ cbox a b" by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y) have "(\(x',K'). content K' *\<^sub>R f x') y = content K *\<^sub>R f x" unfolding xk by auto also have "\ = 0" using assms(1) content_0_subset k(2) by auto finally show "(\(x, k). content k *\<^sub>R f x) y = 0" . qed global_interpretation sum_content: operative plus 0 content rewrites "comm_monoid_set.F plus 0 = sum" proof - interpret operative plus 0 content by standard (auto simp add: content_split [symmetric] content_eq_0_interior) show "operative plus 0 content" by standard show "comm_monoid_set.F plus 0 = sum" by (simp add: sum_def) qed lemma additive_content_division: "d division_of (cbox a b) \ sum content d = content (cbox a b)" by (fact sum_content.division) lemma additive_content_tagged_division: "d tagged_division_of (cbox a b) \ sum (\(x,l). content l) d = content (cbox a b)" by (fact sum_content.tagged_division) lemma subadditive_content_division: assumes "\ division_of S" "S \ cbox a b" shows "sum content \ \ content(cbox a b)" proof - have "\ division_of \\" "\\ \ cbox a b" using assms by auto then obtain \' where "\ \ \'" "\' division_of cbox a b" using partial_division_extend_interval by metis then have "sum content \ \ sum content \'" using sum_mono2 by blast also have "... \ content(cbox a b)" by (simp add: \\' division_of cbox a b\ additive_content_division less_eq_real_def) finally show ?thesis . qed lemma content_real_eq_0: "content {a..b::real} = 0 \ a \ b" by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) lemma property_empty_interval: "\a b. content (cbox a b) = 0 \ P (cbox a b) \ P {}" using content_empty unfolding empty_as_interval by auto lemma interval_bounds_nz_content [simp]: assumes "content (cbox a b) \ 0" shows "interval_upperbound (cbox a b) = b" and "interval_lowerbound (cbox a b) = a" by (metis assms content_empty interval_bounds')+ subsection \Gauge integral\ text \Case distinction to define it first on compact intervals first, then use a limit. This is only much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\ definition has_integral :: "('n::euclidean_space \ 'b::real_normed_vector) \ 'b \ 'n set \ bool" (infixr "has'_integral" 46) where "(f has_integral I) s \ (if \a b. s = cbox a b then ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter s) else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\p. \(x,k)\p. content k *\<^sub>R (if x \ s then f x else 0)) \ z) (division_filter (cbox a b)) \ norm (z - I) < e)))" lemma has_integral_cbox: "(f has_integral I) (cbox a b) \ ((\p. \(x,k)\p. content k *\<^sub>R f x) \ I) (division_filter (cbox a b))" by (auto simp add: has_integral_def) lemma has_integral: "(f has_integral y) (cbox a b) \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff) lemma has_integral_real: "(f has_integral y) {a..b::real} \ (\e>0. \\. gauge \ \ (\\. \ tagged_division_of {a..b} \ \ fine \ \ norm (sum (\(x,k). content(k) *\<^sub>R f x) \ - y) < e))" unfolding box_real[symmetric] by (rule has_integral) lemma has_integralD[dest]: assumes "(f has_integral y) (cbox a b)" and "e > 0" obtains \ where "gauge \" and "\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,k)\\. content k *\<^sub>R f x) - y) < e" using assms unfolding has_integral by auto lemma has_integral_alt: "(f has_integral y) i \ (if \a b. i = cbox a b then (f has_integral y) i else (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f x else 0) has_integral z) (cbox a b) \ norm (z - y) < e)))" by (subst has_integral_def) (auto simp add: has_integral_cbox) lemma has_integral_altD: assumes "(f has_integral y) i" and "\ (\a b. i = cbox a b)" and "e>0" obtains B where "B > 0" and "\a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ i then f(x) else 0) has_integral z) (cbox a b) \ norm(z - y) < e)" using assms has_integral_alt[of f y i] by auto definition integrable_on (infixr "integrable'_on" 46) where "f integrable_on i \ (\y. (f has_integral y) i)" definition "integral i f = (SOME y. (f has_integral y) i \ \ f integrable_on i \ y=0)" lemma integrable_integral[intro]: "f integrable_on i \ (f has_integral (integral i f)) i" unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) lemma not_integrable_integral: "\ f integrable_on i \ integral i f = 0" unfolding integrable_on_def integral_def by blast lemma has_integral_integrable[dest]: "(f has_integral i) s \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_integral: "f integrable_on s \ (f has_integral (integral s f)) s" by auto subsection \Basic theorems about integrals\ lemma has_integral_eq_rhs: "(f has_integral j) S \ i = j \ (f has_integral i) S" by (rule forw_subst) lemma has_integral_unique_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" shows "(f has_integral k1) (cbox a b) \ (f has_integral k2) (cbox a b) \ k1 = k2" by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty]) lemma has_integral_unique: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k1) i" "(f has_integral k2) i" shows "k1 = k2" proof (rule ccontr) let ?e = "norm (k1 - k2)/2" let ?F = "(\x. if x \ i then f x else 0)" assume "k1 \ k2" then have e: "?e > 0" by auto have nonbox: "\ (\a b. i = cbox a b)" using \k1 \ k2\ assms has_integral_unique_cbox by blast obtain B1 where B1: "0 < B1" "\a b. ball 0 B1 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k1) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast obtain B2 where B2: "0 < B2" "\a b. ball 0 B2 \ cbox a b \ \z. (?F has_integral z) (cbox a b) \ norm (z - k2) < norm (k1 - k2)/2" by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast obtain a b :: 'n where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox_symmetric) obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2" using B1(2)[OF ab(1)] by blast obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2" using B2(2)[OF ab(2)] by blast have "z = w" using has_integral_unique_cbox[OF w(1) z(1)] by auto then have "norm (k1 - k2) \ norm (z - k2) + norm (w - k1)" using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] by (auto simp add: norm_minus_commute) also have "\ < norm (k1 - k2)/2 + norm (k1 - k2)/2" by (metis add_strict_mono z(2) w(2)) finally show False by auto qed lemma integral_unique [intro]: "(f has_integral y) k \ integral k f = y" unfolding integral_def by (rule some_equality) (auto intro: has_integral_unique) lemma has_integral_iff: "(f has_integral i) S \ (f integrable_on S \ integral S f = i)" by blast lemma eq_integralD: "integral k f = y \ (f has_integral y) k \ \ f integrable_on k \ y=0" unfolding integral_def integrable_on_def apply (erule subst) apply (rule someI_ex) by blast lemma has_integral_const [intro]: fixes a b :: "'a::euclidean_space" shows "((\x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" using eventually_division_filter_tagged_division[of "cbox a b"] additive_content_tagged_division[of _ a b] by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric] elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const]) lemma has_integral_const_real [intro]: fixes a b :: real shows "((\x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}" by (metis box_real(2) has_integral_const) lemma has_integral_integrable_integral: "(f has_integral i) s \ f integrable_on s \ integral s f = i" by blast lemma integral_const [simp]: fixes a b :: "'a::euclidean_space" shows "integral (cbox a b) (\x. c) = content (cbox a b) *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) lemma integral_const_real [simp]: fixes a b :: real shows "integral {a..b} (\x. c) = content {a..b} *\<^sub>R c" by (metis box_real(2) integral_const) lemma has_integral_is_0_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ cbox a b \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] assms by (subst tendsto_cong[where g="\_. 0"]) (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval) lemma has_integral_is_0: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "\x. x \ S \ f x = 0" shows "(f has_integral 0) S" proof (cases "(\a b. S = cbox a b)") case True with assms has_integral_is_0_cbox show ?thesis by blast next case False have *: "(\x. if x \ S then f x else 0) = (\x. 0)" by (auto simp add: assms) show ?thesis using has_integral_is_0_cbox False by (subst has_integral_alt) (force simp add: *) qed lemma has_integral_0[simp]: "((\x::'n::euclidean_space. 0) has_integral 0) S" by (rule has_integral_is_0) auto lemma has_integral_0_eq[simp]: "((\x. 0) has_integral i) S \ i = 0" using has_integral_unique[OF has_integral_0] by auto lemma has_integral_linear_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) (cbox a b)" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) (cbox a b)" proof - interpret bounded_linear h using h . show ?thesis unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]] by (simp add: sum scaleR split_beta') qed lemma has_integral_linear: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral y) S" and h: "bounded_linear h" shows "((h \ f) has_integral (h y)) S" proof (cases "(\a b. S = cbox a b)") case True with f h has_integral_linear_cbox show ?thesis by blast next case False interpret bounded_linear h using h . from pos_bounded obtain B where B: "0 < B" "\x. norm (h x) \ norm x * B" by blast let ?S = "\f x. if x \ S then f x else 0" show ?thesis proof (subst has_integral_alt, clarsimp simp: False) fix e :: real assume e: "e > 0" have *: "0 < e/B" using e B(1) by simp obtain M where M: "M > 0" "\a b. ball 0 M \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - y) < e/B" using has_integral_altD[OF f False *] by blast show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e)" proof (rule exI, intro allI conjI impI) show "M > 0" using M by metis next fix a b::'n assume sb: "ball 0 M \ cbox a b" obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B" using M(2)[OF sb] by blast have *: "?S(h \ f) = h \ ?S f" using zero by auto show "\z. (?S(h \ f) has_integral z) (cbox a b) \ norm (z - h y) < e" apply (rule_tac x="h z" in exI) apply (simp add: * has_integral_linear_cbox[OF z(1) h]) apply (metis B diff le_less_trans pos_less_divide_eq z(2)) done qed qed qed lemma has_integral_scaleR_left: "(f has_integral y) S \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) S" using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) lemma integrable_on_scaleR_left: assumes "f integrable_on A" shows "(\x. f x *\<^sub>R y) integrable_on A" using assms has_integral_scaleR_left unfolding integrable_on_def by blast lemma has_integral_mult_left: fixes c :: "_ :: real_normed_algebra" shows "(f has_integral y) S \ ((\x. f x * c) has_integral (y * c)) S" using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) lemma has_integral_divide: fixes c :: "_ :: real_normed_div_algebra" shows "(f has_integral y) S \ ((\x. f x / c) has_integral (y / c)) S" unfolding divide_inverse by (simp add: has_integral_mult_left) text\The case analysis eliminates the condition \<^term>\f integrable_on S\ at the cost of the type class constraint \division_ring\\ corollary integral_mult_left [simp]: fixes c:: "'a::{real_normed_algebra,division_ring}" shows "integral S (\x. f x * c) = integral S f * c" proof (cases "f integrable_on S \ c = 0") case True then show ?thesis by (force intro: has_integral_mult_left) next case False then have "\ (\x. f x * c) integrable_on S" using has_integral_mult_left [of "(\x. f x * c)" _ S "inverse c"] by (auto simp add: mult.assoc) with False show ?thesis by (simp add: not_integrable_integral) qed corollary integral_mult_right [simp]: fixes c:: "'a::{real_normed_field}" shows "integral S (\x. c * f x) = c * integral S f" by (simp add: mult.commute [of c]) corollary integral_divide [simp]: fixes z :: "'a::real_normed_field" shows "integral S (\x. f x / z) = integral S (\x. f x) / z" using integral_mult_left [of S f "inverse z"] by (simp add: divide_inverse_commute) lemma has_integral_mult_right: fixes c :: "'a :: real_normed_algebra" shows "(f has_integral y) i \ ((\x. c * f x) has_integral (c * y)) i" using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) lemma has_integral_cmul: "(f has_integral k) S \ ((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) S" unfolding o_def[symmetric] by (metis has_integral_linear bounded_linear_scaleR_right) lemma has_integral_cmult_real: fixes c :: real assumes "c \ 0 \ (f has_integral x) A" shows "((\x. c * f x) has_integral c * x) A" proof (cases "c = 0") case True then show ?thesis by simp next case False from has_integral_cmul[OF assms[OF this], of c] show ?thesis unfolding real_scaleR_def . qed lemma has_integral_neg: "(f has_integral k) S \ ((\x. -(f x)) has_integral -k) S" by (drule_tac c="-1" in has_integral_cmul) auto lemma has_integral_neg_iff: "((\x. - f x) has_integral k) S \ (f has_integral - k) S" using has_integral_neg[of f "- k"] has_integral_neg[of "\x. - f x" k] by auto lemma has_integral_add_cbox: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)" shows "((\x. f x + g x) has_integral (k + l)) (cbox a b)" using assms unfolding has_integral_cbox by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add) lemma has_integral_add: fixes f :: "'n::euclidean_space \ 'a::real_normed_vector" assumes f: "(f has_integral k) S" and g: "(g has_integral l) S" shows "((\x. f x + g x) has_integral (k + l)) S" proof (cases "\a b. S = cbox a b") case True with has_integral_add_cbox assms show ?thesis by blast next let ?S = "\f x. if x \ S then f x else 0" case False then show ?thesis proof (subst has_integral_alt, clarsimp, goal_cases) case (1 e) then have e2: "e/2 > 0" by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?S f has_integral z) (cbox a b) \ norm (z - k) < e/2" using has_integral_altD[OF f False e2] by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ (cbox a b) \ \z. (?S g has_integral z) (cbox a b) \ norm (z - l) < e/2" using has_integral_altD[OF g False e2] by blast show ?case proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \0 < Bf\) fix a b assume "ball 0 (max Bf Bg) \ cbox a (b::'n)" then have fs: "ball 0 Bf \ cbox a (b::'n)" and gs: "ball 0 Bg \ cbox a (b::'n)" by auto obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2" using Bf[OF fs] by blast obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2" using Bg[OF gs] by blast have *: "\x. (if x \ S then f x + g x else 0) = (?S f x) + (?S g x)" by auto show "\z. (?S(\x. f x + g x) has_integral z) (cbox a b) \ norm (z - (k + l)) < e" apply (rule_tac x="w + z" in exI) apply (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) apply (auto simp add: field_simps) done qed qed qed lemma has_integral_diff: "(f has_integral k) S \ (g has_integral l) S \ ((\x. f x - g x) has_integral (k - l)) S" using has_integral_add[OF _ has_integral_neg, of f k S g l] by (auto simp: algebra_simps) lemma integral_0 [simp]: "integral S (\x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" by (rule integral_unique has_integral_0)+ lemma integral_add: "f integrable_on S \ g integrable_on S \ integral S (\x. f x + g x) = integral S f + integral S g" by (rule integral_unique) (metis integrable_integral has_integral_add) lemma integral_cmul [simp]: "integral S (\x. c *\<^sub>R f x) = c *\<^sub>R integral S f" proof (cases "f integrable_on S \ c = 0") case True with has_integral_cmul integrable_integral show ?thesis by fastforce next case False then have "\ (\x. c *\<^sub>R f x) integrable_on S" using has_integral_cmul [of "(\x. c *\<^sub>R f x)" _ S "inverse c"] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_mult: fixes K::real shows "f integrable_on X \ K * integral X f = integral X (\x. K * f x)" unfolding real_scaleR_def[symmetric] integral_cmul .. lemma integral_neg [simp]: "integral S (\x. - f x) = - integral S f" proof (cases "f integrable_on S") case True then show ?thesis by (simp add: has_integral_neg integrable_integral integral_unique) next case False then have "\ (\x. - f x) integrable_on S" using has_integral_neg [of "(\x. - f x)" _ S ] by auto with False show ?thesis by (simp add: not_integrable_integral) qed lemma integral_diff: "f integrable_on S \ g integrable_on S \ integral S (\x. f x - g x) = integral S f - integral S g" by (rule integral_unique) (metis integrable_integral has_integral_diff) lemma integrable_0: "(\x. 0) integrable_on S" unfolding integrable_on_def using has_integral_0 by auto lemma integrable_add: "f integrable_on S \ g integrable_on S \ (\x. f x + g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_add) lemma integrable_cmul: "f integrable_on S \ (\x. c *\<^sub>R f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_cmul) lemma integrable_on_scaleR_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c *\<^sub>R f x) integrable_on S \ f integrable_on S" using integrable_cmul[of "\x. c *\<^sub>R f x" S "1 / c"] integrable_cmul[of f S c] \c \ 0\ by auto lemma integrable_on_cmult_iff [simp]: fixes c :: real assumes "c \ 0" shows "(\x. c * f x) integrable_on S \ f integrable_on S" using integrable_on_scaleR_iff [of c f] assms by simp lemma integrable_on_cmult_left: assumes "f integrable_on S" shows "(\x. of_real c * f x) integrable_on S" using integrable_cmul[of f S "of_real c"] assms by (simp add: scaleR_conv_of_real) lemma integrable_neg: "f integrable_on S \ (\x. -f(x)) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_neg) lemma integrable_neg_iff: "(\x. -f(x)) integrable_on S \ f integrable_on S" using integrable_neg by fastforce lemma integrable_diff: "f integrable_on S \ g integrable_on S \ (\x. f x - g x) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_diff) lemma integrable_linear: "f integrable_on S \ bounded_linear h \ (h \ f) integrable_on S" unfolding integrable_on_def by(auto intro: has_integral_linear) lemma integral_linear: "f integrable_on S \ bounded_linear h \ integral S (h \ f) = h (integral S f)" apply (rule has_integral_unique [where i=S and f = "h \ f"]) apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) done lemma integrable_on_cnj_iff: "(\x. cnj (f x)) integrable_on A \ f integrable_on A" using integrable_linear[OF _ bounded_linear_cnj, of f A] integrable_linear[OF _ bounded_linear_cnj, of "cnj \ f" A] by (auto simp: o_def) lemma integral_cnj: "cnj (integral A f) = integral A (\x. cnj (f x))" by (cases "f integrable_on A") (simp_all add: integral_linear[OF _ bounded_linear_cnj, symmetric] o_def integrable_on_cnj_iff not_integrable_integral) lemma integral_component_eq[simp]: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "f integrable_on S" shows "integral S (\x. f x \ k) = integral S f \ k" unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] .. lemma has_integral_sum: assumes "finite T" and "\a. a \ T \ ((f a) has_integral (i a)) S" shows "((\x. sum (\a. f a x) T) has_integral (sum i T)) S" using assms(1) subset_refl[of T] proof (induct rule: finite_subset_induct) case empty then show ?case by auto next case (insert x F) with assms show ?case by (simp add: has_integral_add) qed lemma integral_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ integral S (\x. \a\I. f a x) = (\a\I. integral S (f a))" by (simp add: has_integral_sum integrable_integral integral_unique) lemma integrable_sum: "\finite I; \a. a \ I \ f a integrable_on S\ \ (\x. \a\I. f a x) integrable_on S" unfolding integrable_on_def using has_integral_sum[of I] by metis lemma has_integral_eq: assumes "\x. x \ s \ f x = g x" and "(f has_integral k) s" shows "(g has_integral k) s" using has_integral_diff[OF assms(2), of "\x. f x - g x" 0] using has_integral_is_0[of s "\x. f x - g x"] using assms(1) by auto lemma integrable_eq: "\f integrable_on s; \x. x \ s \ f x = g x\ \ g integrable_on s" unfolding integrable_on_def using has_integral_eq[of s f g] has_integral_eq by blast lemma has_integral_cong: assumes "\x. x \ s \ f x = g x" shows "(f has_integral i) s = (g has_integral i) s" using has_integral_eq[of s f g] has_integral_eq[of s g f] assms by auto lemma integral_cong: assumes "\x. x \ s \ f x = g x" shows "integral s f = integral s g" unfolding integral_def by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) lemma integrable_on_cmult_left_iff [simp]: assumes "c \ 0" shows "(\x. of_real c * f x) integrable_on s \ f integrable_on s" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. of_real (1 / c) * (of_real c * f x)) integrable_on s" using integrable_cmul[of "\x. of_real c * f x" s "1 / of_real c"] by (simp add: scaleR_conv_of_real) then have "(\x. (of_real (1 / c) * of_real c * f x)) integrable_on s" by (simp add: algebra_simps) with \c \ 0\ show ?rhs by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) qed (blast intro: integrable_on_cmult_left) lemma integrable_on_cmult_right: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "f integrable_on s" shows "(\x. f x * of_real c) integrable_on s" using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) lemma integrable_on_cmult_right_iff [simp]: fixes f :: "_ \ 'b :: {comm_ring,real_algebra_1,real_normed_vector}" assumes "c \ 0" shows "(\x. f x * of_real c) integrable_on s \ f integrable_on s" using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) lemma integrable_on_cdivide: fixes f :: "_ \ 'b :: real_normed_field" assumes "f integrable_on s" shows "(\x. f x / of_real c) integrable_on s" by (simp add: integrable_on_cmult_right divide_inverse assms flip: of_real_inverse) lemma integrable_on_cdivide_iff [simp]: fixes f :: "_ \ 'b :: real_normed_field" assumes "c \ 0" shows "(\x. f x / of_real c) integrable_on s \ f integrable_on s" by (simp add: divide_inverse assms flip: of_real_inverse) lemma has_integral_null [intro]: "content(cbox a b) = 0 \ (f has_integral 0) (cbox a b)" unfolding has_integral_cbox using eventually_division_filter_tagged_division[of "cbox a b"] by (subst tendsto_cong[where g="\_. 0"]) (auto elim: eventually_mono intro: sum_content_null) lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \ (f has_integral 0) {a..b}" by (metis box_real(2) has_integral_null) lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \ (f has_integral i) (cbox a b) \ i = 0" by (auto simp add: has_integral_null dest!: integral_unique) lemma integral_null [simp]: "content (cbox a b) = 0 \ integral (cbox a b) f = 0" by (metis has_integral_null integral_unique) lemma integrable_on_null [intro]: "content (cbox a b) = 0 \ f integrable_on (cbox a b)" by (simp add: has_integral_integrable) lemma has_integral_empty[intro]: "(f has_integral 0) {}" by (meson ex_in_conv has_integral_is_0) lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \ i = 0" by (auto simp add: has_integral_empty has_integral_unique) lemma integrable_on_empty[intro]: "f integrable_on {}" unfolding integrable_on_def by auto lemma integral_empty[simp]: "integral {} f = 0" by (rule integral_unique) (rule has_integral_empty) lemma has_integral_refl[intro]: fixes a :: "'a::euclidean_space" shows "(f has_integral 0) (cbox a a)" and "(f has_integral 0) {a}" proof - show "(f has_integral 0) (cbox a a)" by (rule has_integral_null) simp then show "(f has_integral 0) {a}" by simp qed lemma integrable_on_refl[intro]: "f integrable_on cbox a a" unfolding integrable_on_def by auto lemma integral_refl [simp]: "integral (cbox a a) f = 0" by (rule integral_unique) auto lemma integral_singleton [simp]: "integral {a} f = 0" by auto lemma integral_blinfun_apply: assumes "f integrable_on s" shows "integral s (\x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) lemma blinfun_apply_integral: assumes "f integrable_on s" shows "blinfun_apply (integral s f) x = integral s (\y. blinfun_apply (f y) x)" by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) lemma has_integral_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(f has_integral y) A \ (\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" proof safe fix b :: 'b assume "(f has_integral y) A" from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] show "((\x. f x \ b) has_integral (y \ b)) A" by (simp add: o_def) next assume "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral ((y \ b) *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. (y \ b) *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "(f has_integral y) A" by (simp add: euclidean_representation) qed lemma has_integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ ((\x. f x \ b) has_integral (y \ b)) A) \ (f has_integral y) A" by (subst has_integral_componentwise_iff) blast lemma integrable_componentwise_iff: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "f integrable_on A \ (\b\Basis. (\x. f x \ b) integrable_on A)" proof assume "f integrable_on A" then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) hence "(\b\Basis. ((\x. f x \ b) has_integral (y \ b)) A)" by (subst (asm) has_integral_componentwise_iff) thus "(\b\Basis. (\x. f x \ b) integrable_on A)" by (auto simp: integrable_on_def) next assume "(\b\Basis. (\x. f x \ b) integrable_on A)" then obtain y where "\b\Basis. ((\x. f x \ b) has_integral y b) A" unfolding integrable_on_def by (subst (asm) bchoice_iff) blast hence "\b\Basis. (((\x. x *\<^sub>R b) \ (\x. f x \ b)) has_integral (y b *\<^sub>R b)) A" by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) hence "((\x. \b\Basis. (f x \ b) *\<^sub>R b) has_integral (\b\Basis. y b *\<^sub>R b)) A" by (intro has_integral_sum) (simp_all add: o_def) thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) qed lemma integrable_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" shows "(\b. b \ Basis \ (\x. f x \ b) integrable_on A) \ f integrable_on A" by (subst integrable_componentwise_iff) blast lemma integral_componentwise: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "f integrable_on A" shows "integral A f = (\b\Basis. integral A (\x. (f x \ b) *\<^sub>R b))" proof - from assms have integrable: "\b\Basis. (\x. x *\<^sub>R b) \ (\x. (f x \ b)) integrable_on A" by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) (simp_all add: bounded_linear_scaleR_left) have "integral A f = integral A (\x. \b\Basis. (f x \ b) *\<^sub>R b)" by (simp add: euclidean_representation) also from integrable have "\ = (\a\Basis. integral A (\x. (f x \ a) *\<^sub>R a))" by (subst integral_sum) (simp_all add: o_def) finally show ?thesis . qed lemma integrable_component: "f integrable_on A \ (\x. f x \ (y :: 'b :: euclidean_space)) integrable_on A" by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) subsection \Cauchy-type criterion for integrability\ proposition integrable_Cauchy: fixes f :: "'n::euclidean_space \ 'a::{real_normed_vector,complete_space}" shows "f integrable_on cbox a b \ (\e>0. \\. gauge \ \ (\\1 \2. \1 tagged_division_of (cbox a b) \ \ fine \1 \ \2 tagged_division_of (cbox a b) \ \ fine \2 \ norm ((\(x,K)\\1. content K *\<^sub>R f x) - (\(x,K)\\2. content K *\<^sub>R f x)) < e))" (is "?l = (\e>0. \\. ?P e \)") proof (intro iffI allI impI) assume ?l then obtain y where y: "\e. e > 0 \ \\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" by (auto simp: integrable_on_def has_integral) show "\\. ?P e \" if "e > 0" for e proof - have "e/2 > 0" using that by auto with y obtain \ where "gauge \" and \: "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - y) < e/2" by meson show ?thesis apply (rule_tac x=\ in exI, clarsimp simp: \gauge \\) by (blast intro!: \ dist_triangle_half_l[where y=y,unfolded dist_norm]) qed next assume "\e>0. \\. ?P e \" then have "\n::nat. \\. ?P (1 / (n + 1)) \" by auto then obtain \ :: "nat \ 'n \ 'n set" where \: "\m. gauge (\ m)" "\m \1 \2. \\1 tagged_division_of cbox a b; \ m fine \1; \2 tagged_division_of cbox a b; \ m fine \2\ \ norm ((\(x,K) \ \1. content K *\<^sub>R f x) - (\(x,K) \ \2. content K *\<^sub>R f x)) < 1 / (m + 1)" by metis have "\n. gauge (\x. \{\ i x |i. i \ {0..n}})" apply (rule gauge_Inter) using \ by auto then have "\n. \p. p tagged_division_of (cbox a b) \ (\x. \{\ i x |i. i \ {0..n}}) fine p" by (meson fine_division_exists) then obtain p where p: "\z. p z tagged_division_of cbox a b" "\z. (\x. \{\ i x |i. i \ {0..z}}) fine p z" by meson have dp: "\i n. i\n \ \ i fine p n" using p unfolding fine_Inter using atLeastAtMost_iff by blast have "Cauchy (\n. sum (\(x,K). content K *\<^sub>R (f x)) (p n))" proof (rule CauchyI) fix e::real assume "0 < e" then obtain N where "N \ 0" and N: "inverse (real N) < e" using real_arch_inverse[of e] by blast show "\M. \m\M. \n\M. norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" proof (intro exI allI impI) fix m n assume mn: "N \ m" "N \ n" have "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < 1 / (real N + 1)" by (simp add: p(1) dp mn \) also have "... < e" using N \N \ 0\ \0 < e\ by (auto simp: field_simps) finally show "norm ((\(x,K) \ p m. content K *\<^sub>R f x) - (\(x,K) \ p n. content K *\<^sub>R f x)) < e" . qed qed then obtain y where y: "\no. \n\no. norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D) show ?l unfolding integrable_on_def has_integral proof (rule_tac x=y in exI, clarify) fix e :: real assume "e>0" then have e2: "e/2 > 0" by auto then obtain N1::nat where N1: "N1 \ 0" "inverse (real N1) < e/2" using real_arch_inverse by blast obtain N2::nat where N2: "\n. n \ N2 \ norm ((\(x,K) \ p n. content K *\<^sub>R f x) - y) < e/2" using y[OF e2] by metis show "\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - y) < e)" proof (intro exI conjI allI impI) show "gauge (\ (N1+N2))" using \ by auto show "norm ((\(x,K) \ q. content K *\<^sub>R f x) - y) < e" if "q tagged_division_of cbox a b \ \ (N1+N2) fine q" for q proof (rule norm_triangle_half_r) have "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < 1 / (real (N1+N2) + 1)" by (rule \; simp add: dp p that) also have "... < e/2" using N1 \0 < e\ by (auto simp: field_simps intro: less_le_trans) finally show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - (\(x,K) \ q. content K *\<^sub>R f x)) < e/2" . show "norm ((\(x,K) \ p (N1+N2). content K *\<^sub>R f x) - y) < e/2" using N2 le_add_same_cancel2 by blast qed qed qed qed subsection \Additivity of integral on abutting intervals\ lemma tagged_division_split_left_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. x \ k \ c}) = {}" by (metis tagged_division_split_left_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed lemma tagged_division_split_right_inj_content: assumes \: "\ tagged_division_of S" and "(x1, K1) \ \" "(x2, K2) \ \" "K1 \ K2" "K1 \ {x. x\k \ c} = K2 \ {x. x\k \ c}" "k \ Basis" shows "content (K1 \ {x. x\k \ c}) = 0" proof - from tagged_division_ofD(4)[OF \ \(x1, K1) \ \\] obtain a b where K1: "K1 = cbox a b" by auto then have "interior (K1 \ {x. c \ x \ k}) = {}" by (metis tagged_division_split_right_inj assms) then show ?thesis unfolding K1 interval_split[OF \k \ Basis\] by (auto simp: content_eq_0_interior) qed proposition has_integral_split: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes fi: "(f has_integral i) (cbox a b \ {x. x\k \ c})" and fj: "(f has_integral j) (cbox a b \ {x. x\k \ c})" and k: "k \ Basis" shows "(f has_integral (i + j)) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "0 < e" then have e: "e/2 > 0" by auto obtain \1 where \1: "gauge \1" and \1norm: "\\. \\ tagged_division_of cbox a b \ {x. x \ k \ c}; \1 fine \\ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - i) < e/2" apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done obtain \2 where \2: "gauge \2" and \2norm: "\\. \\ tagged_division_of cbox a b \ {x. c \ x \ k}; \2 fine \\ \ norm ((\(x, k) \ \. content k *\<^sub>R f x) - j) < e/2" apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) apply (simp add: interval_split[symmetric] k) done let ?\ = "\x. if x\k = c then (\1 x \ \2 x) else ball x \x\k - c\ \ \1 x \ \2 x" have "gauge ?\" using \1 \2 unfolding gauge_def by auto then show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - (i + j)) < e)" proof (rule_tac x="?\" in exI, safe) fix p assume p: "p tagged_division_of (cbox a b)" and "?\ fine p" have ab_eqp: "cbox a b = \{K. \x. (x, K) \ p}" using p by blast have xk_le_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have xk_ge_c: "x\k \ c" if as: "(x,K) \ p" and K: "K \ {x. x\k \ c} \ {}" for x K proof (rule ccontr) assume **: "\ x \ k \ c" then have "K \ ball x \x \ k - c\" using \?\ fine p\ as by (fastforce simp: not_le algebra_simps) with K obtain y where y: "y \ ball x \x \ k - c\" "y\k \ c" by blast then have "\x \ k - y \ k\ < \x \ k - c\" using Basis_le_norm[OF k, of "x - y"] by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) with y show False using ** by (auto simp add: field_simps) qed have fin_finite: "finite {(x,f K) | x K. (x,K) \ s \ P x K}" if "finite s" for s and f :: "'a set \ 'a set" and P :: "'a \ 'a set \ bool" proof - from that have "finite ((\(x,K). (x, f K)) ` s)" by auto then show ?thesis by (rule rev_finite_subset) auto qed { fix \ :: "'a set \ 'a set" fix i :: "'a \ 'a set" assume "i \ (\(x, k). (x, \ k)) ` p - {(x, \ k) |x k. (x, k) \ p \ \ k \ {}}" then obtain x K where xk: "i = (x, \ K)" "(x,K) \ p" "(x, \ K) \ {(x, \ K) |x K. (x,K) \ p \ \ K \ {}}" by auto have "content (\ K) = 0" using xk using content_empty by auto then have "(\(x,K). content K *\<^sub>R f x) i = 0" unfolding xk split_conv by auto } note [simp] = this have "finite p" using p by blast let ?M1 = "{(x, K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \1_fine: "\1 fine ?M1" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M1. content k *\<^sub>R f x) - i) < e/2" proof (rule \1norm [OF tagged_division_ofI \1_fine]) show "finite ?M1" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M1} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M1" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_le_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M1" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed moreover let ?M2 = "{(x,K \ {x. x\k \ c}) |x K. (x,K) \ p \ K \ {x. x\k \ c} \ {}}" have \2_fine: "\2 fine ?M2" using \?\ fine p\ by (fastforce simp: fine_def split: if_split_asm) have "norm ((\(x, k)\?M2. content k *\<^sub>R f x) - j) < e/2" proof (rule \2norm [OF tagged_division_ofI \2_fine]) show "finite ?M2" by (rule fin_finite) (use p in blast) show "\{k. \x. (x, k) \ ?M2} = cbox a b \ {x. x\k \ c}" by (auto simp: ab_eqp) fix x L assume xL: "(x, L) \ ?M2" then obtain x' L' where xL': "x = x'" "L = L' \ {x. x \ k \ c}" "(x', L') \ p" "L' \ {x. x \ k \ c} \ {}" by blast then obtain a' b' where ab': "L' = cbox a' b'" using p by blast show "x \ L" "L \ cbox a b \ {x. x \ k \ c}" using p xk_ge_c xL' by auto show "\a b. L = cbox a b" using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) fix y R assume yR: "(y, R) \ ?M2" then obtain y' R' where yR': "y = y'" "R = R' \ {x. x \ k \ c}" "(y', R') \ p" "R' \ {x. x \ k \ c} \ {}" by blast assume as: "(x, L) \ (y, R)" show "interior L \ interior R = {}" proof (cases "L' = R' \ x' = y'") case False have "interior R' = {}" by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) then show ?thesis using yR' by simp next case True then have "L' \ R'" using as unfolding xL' yR' by auto have "interior L' \ interior R' = {}" by (metis (no_types) Pair_inject \L' \ R'\ p tagged_division_ofD(5) xL'(3) yR'(3)) then show ?thesis using xL'(2) yR'(2) by auto qed qed ultimately have "norm (((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j)) < e/2 + e/2" using norm_add_less by blast moreover have "((\(x,K) \ ?M1. content K *\<^sub>R f x) - i) + ((\(x,K) \ ?M2. content K *\<^sub>R f x) - j) = (\(x, ka)\p. content ka *\<^sub>R f x) - (i + j)" proof - have eq0: "\x y. x = (0::real) \ x *\<^sub>R (y::'b) = 0" by auto have cont_eq: "\g. (\(x,l). content l *\<^sub>R f x) \ (\(x,l). (x,g l)) = (\(x,l). content (g l) *\<^sub>R f x)" by auto have *: "\\ :: 'a set \ 'a set. (\(x,K)\{(x, \ K) |x K. (x,K) \ p \ \ K \ {}}. content K *\<^sub>R f x) = (\(x,K)\(\(x,K). (x, \ K)) ` p. content K *\<^sub>R f x)" by (rule sum.mono_neutral_left) (auto simp: \finite p\) have "((\(x, k)\?M1. content k *\<^sub>R f x) - i) + ((\(x, k)\?M2. content k *\<^sub>R f x) - j) = (\(x, k)\?M1. content k *\<^sub>R f x) + (\(x, k)\?M2. content k *\<^sub>R f x) - (i + j)" by auto moreover have "\ = (\(x,K) \ p. content (K \ {x. x \ k \ c}) *\<^sub>R f x) + (\(x,K) \ p. content (K \ {x. c \ x \ k}) *\<^sub>R f x) - (i + j)" unfolding * apply (subst (1 2) sum.reindex_nontrivial) apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content simp: cont_eq \finite p\) done moreover have "\x. x \ p \ (\(a,B). content (B \ {a. a \ k \ c}) *\<^sub>R f a) x + (\(a,B). content (B \ {a. c \ a \ k}) *\<^sub>R f a) x = (\(a,B). content B *\<^sub>R f a) x" proof clarify fix a B assume "(a, B) \ p" with p obtain u v where uv: "B = cbox u v" by blast then show "content (B \ {x. x \ k \ c}) *\<^sub>R f a + content (B \ {x. c \ x \ k}) *\<^sub>R f a = content B *\<^sub>R f a" by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c]) qed ultimately show ?thesis by (auto simp: sum.distrib[symmetric]) qed ultimately show "norm ((\(x, k)\p. content k *\<^sub>R f x) - (i + j)) < e" by auto qed qed subsection \A sort of converse, integrability on subintervals\ lemma has_integral_separate_sides: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes f: "(f has_integral i) (cbox a b)" and "e > 0" and k: "k \ Basis" obtains d where "gauge d" "\p1 p2. p1 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p1 \ p2 tagged_division_of (cbox a b \ {x. x\k \ c}) \ d fine p2 \ norm ((sum (\(x,k). content k *\<^sub>R f x) p1 + sum (\(x,k). content k *\<^sub>R f x) p2) - i) < e" proof - obtain \ where d: "gauge \" "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e" using has_integralD[OF f \e > 0\] by metis { fix p1 p2 assume tdiv1: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" and "\ fine p1" note p1=tagged_division_ofD[OF this(1)] assume tdiv2: "p2 tagged_division_of (cbox a b) \ {x. c \ x \ k}" and "\ fine p2" note p2=tagged_division_ofD[OF this(1)] note tagged_division_Un_interval[OF tdiv1 tdiv2] note p12 = tagged_division_ofD[OF this] this { fix a b assume ab: "(a, b) \ p1 \ p2" have "(a, b) \ p1" using ab by auto obtain u v where uv: "b = cbox u v" using \(a, b) \ p1\ p1(4) by moura have "b \ {x. x\k = c}" using ab p1(3)[of a b] p2(3)[of a b] by fastforce moreover have "interior {x::'a. x \ k = c} = {}" proof (rule ccontr) assume "\ ?thesis" then obtain x where x: "x \ interior {x::'a. x\k = c}" by auto then obtain \ where "0 < \" and \: "ball x \ \ {x. x \ k = c}" using mem_interior by metis have x: "x\k = c" using x interior_subset by fastforce have *: "\i. i \ Basis \ \(x - (x + (\/2) *\<^sub>R k)) \ i\ = (if i = k then \/2 else 0)" using \0 < \\ k by (auto simp: inner_simps inner_not_same_Basis) have "(\i\Basis. \(x - (x + (\/2 ) *\<^sub>R k)) \ i\) = (\i\Basis. (if i = k then \/2 else 0))" using "*" by (blast intro: sum.cong) also have "\ < \" by (subst sum.delta) (use \0 < \\ in auto) finally have "x + (\/2) *\<^sub>R k \ ball x \" unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) then have "x + (\/2) *\<^sub>R k \ {x. x\k = c}" using \ by auto then show False using \0 < \\ x k by (auto simp: inner_simps) qed ultimately have "content b = 0" unfolding uv content_eq_0_interior using interior_mono by blast then have "content b *\<^sub>R f a = 0" by auto } then have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) = norm ((\(x, k)\p1 \ p2. content k *\<^sub>R f x) - i)" by (subst sum.union_inter_neutral) (auto simp: p1 p2) also have "\ < e" using d(2) p12 by (simp add: fine_Un k \\ fine p1\ \\ fine p2\) finally have "norm ((\(x, k)\p1. content k *\<^sub>R f x) + (\(x, k)\p2. content k *\<^sub>R f x) - i) < e" . } then show ?thesis using d(1) that by auto qed lemma integrable_split [intro]: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on cbox a b" and k: "k \ Basis" shows "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis1) and "f integrable_on (cbox a b \ {x. x\k \ c})" (is ?thesis2) proof - obtain y where y: "(f has_integral y) (cbox a b)" using f by blast define a' where "a' = (\i\Basis. (if i = k then max (a\k) c else a\i)*\<^sub>R i)" define b' where "b' = (\i\Basis. (if i = k then min (b\k) c else b\i)*\<^sub>R i)" have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a' b]) fix p assume "p tagged_division_of cbox a' b" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis1 by (simp add: interval_split[OF k] integrable_Cauchy) have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p1 \ p2 tagged_division_of cbox a b \ {x. x \ k \ c} \ d fine p2 \ norm ((\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)) < e)" if "e > 0" for e proof - have "e/2 > 0" using that by auto with has_integral_separate_sides[OF y this k, of c] obtain d where "gauge d" and d: "\p1 p2. \p1 tagged_division_of cbox a b \ {x. x \ k \ c}; d fine p1; p2 tagged_division_of cbox a b \ {x. c \ x \ k}; d fine p2\ \ norm ((\(x,K)\p1. content K *\<^sub>R f x) + (\(x,K)\p2. content K *\<^sub>R f x) - y) < e/2" by metis show ?thesis proof (rule_tac x=d in exI, clarsimp simp add: \gauge d\) fix p1 p2 assume as: "p1 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p1" "p2 tagged_division_of (cbox a b) \ {x. x \ k \ c}" "d fine p2" show "norm ((\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R f x)) < e" proof (rule fine_division_exists[OF \gauge d\, of a b']) fix p assume "p tagged_division_of cbox a b'" "d fine p" then show ?thesis using as norm_triangle_half_l[OF d[of p p1] d[of p p2]] unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] by (auto simp add: algebra_simps) qed qed qed with f show ?thesis2 by (simp add: interval_split[OF k] integrable_Cauchy) qed lemma operative_integralI: fixes f :: "'a::euclidean_space \ 'b::banach" shows "operative (lift_option (+)) (Some 0) (\i. if f integrable_on i then Some (integral i f) else None)" proof - interpret comm_monoid "lift_option plus" "Some (0::'b)" by (rule comm_monoid_lift_option) (rule add.comm_monoid_axioms) show ?thesis proof fix a b c fix k :: 'a assume k: "k \ Basis" show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = lift_option (+) (if f integrable_on cbox a b \ {x. x \ k \ c} then Some (integral (cbox a b \ {x. x \ k \ c}) f) else None) (if f integrable_on cbox a b \ {x. c \ x \ k} then Some (integral (cbox a b \ {x. c \ x \ k}) f) else None)" proof (cases "f integrable_on cbox a b") case True with k show ?thesis apply (simp add: integrable_split) apply (rule integral_unique [OF has_integral_split[OF _ _ k]]) apply (auto intro: integrable_integral) done next case False have "\ (f integrable_on cbox a b \ {x. x \ k \ c}) \ \ ( f integrable_on cbox a b \ {x. c \ x \ k})" proof (rule ccontr) assume "\ ?thesis" then have "f integrable_on cbox a b" unfolding integrable_on_def apply (rule_tac x="integral (cbox a b \ {x. x \ k \ c}) f + integral (cbox a b \ {x. x \ k \ c}) f" in exI) apply (rule has_integral_split[OF _ _ k]) apply (auto intro: integrable_integral) done then show False using False by auto qed then show ?thesis using False by auto qed next fix a b :: 'a assume "box a b = {}" then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" using has_integral_null_eq by (auto simp: integrable_on_null content_eq_0_interior) qed qed subsection \Bounds on the norm of Riemann sums and the integral itself\ lemma dsum_bound: assumes "p division_of (cbox a b)" and "norm c \ e" shows "norm (sum (\l. content l *\<^sub>R c) p) \ e * content(cbox a b)" proof - have sumeq: "(\i\p. \content i\) = sum content p" apply (rule sum.cong) using assms apply simp apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4)) done have e: "0 \ e" using assms(2) norm_ge_zero order_trans by blast have "norm (sum (\l. content l *\<^sub>R c) p) \ (\i\p. norm (content i *\<^sub>R c))" using norm_sum by blast also have "... \ e * (\i\p. \content i\)" by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg) also have "... \ e * content (cbox a b)" apply (rule mult_left_mono [OF _ e]) apply (simp add: sumeq) using additive_content_division assms(1) eq_iff apply blast done finally show ?thesis . qed lemma rsum_bound: assumes p: "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ e * content (cbox a b)" proof (cases "cbox a b = {}") case True show ?thesis using p unfolding True tagged_division_of_trivial by auto next case False then have e: "e \ 0" by (meson ex_in_conv assms(2) norm_ge_zero order_trans) have sum_le: "sum (content \ snd) p \ content (cbox a b)" unfolding additive_content_tagged_division[OF p, symmetric] split_def by (auto intro: eq_refl) have con: "\xk. xk \ p \ 0 \ content (snd xk)" using tagged_division_ofD(4) [OF p] content_pos_le by force have norm: "\xk. xk \ p \ norm (f (fst xk)) \ e" unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms by (metis prod.collapse subset_eq) have "norm (sum (\(x,k). content k *\<^sub>R f x) p) \ (\i\p. norm (case i of (x, k) \ content k *\<^sub>R f x))" by (rule norm_sum) also have "... \ e * content (cbox a b)" unfolding split_def norm_scaleR apply (rule order_trans[OF sum_mono]) apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) apply (metis norm) unfolding sum_distrib_right[symmetric] using con sum_le apply (auto simp: mult.commute intro: mult_left_mono [OF _ e]) done finally show ?thesis . qed lemma rsum_diff_bound: assumes "p tagged_division_of (cbox a b)" and "\x\cbox a b. norm (f x - g x) \ e" shows "norm (sum (\(x,k). content k *\<^sub>R f x) p - sum (\(x,k). content k *\<^sub>R g x) p) \ e * content (cbox a b)" apply (rule order_trans[OF _ rsum_bound[OF assms]]) apply (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) done lemma has_integral_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof (rule ccontr) assume "\ ?thesis" then have "norm i - B * content (cbox a b) > 0" by auto with f[unfolded has_integral] obtain \ where "gauge \" and \: "\p. \p tagged_division_of cbox a b; \ fine p\ \ norm ((\(x, K)\p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)" by metis then obtain p where p: "p tagged_division_of cbox a b" and "\ fine p" using fine_division_exists by blast have "\s B. norm s \ B \ \ norm (s - i) < norm i - B" unfolding not_less by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans) then show False using \ [OF p \\ fine p\] rsum_bound[OF p] assms by metis qed corollary integrable_bound: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "f integrable_on (cbox a b)" and "\x. x\cbox a b \ norm (f x) \ B" shows "norm (integral (cbox a b) f) \ B * content (cbox a b)" by (metis integrable_integral has_integral_bound assms) subsection \Similar theorems about relationship among components\ lemma rsum_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes p: "p tagged_division_of (cbox a b)" and "\x. x \ cbox a b \ (f x)\i \ (g x)\i" shows "(\(x, K)\p. content K *\<^sub>R f x) \ i \ (\(x, K)\p. content K *\<^sub>R g x) \ i" unfolding inner_sum_left proof (rule sum_mono, clarify) fix x K assume ab: "(x, K) \ p" with p obtain u v where K: "K = cbox u v" by blast then show "(content K *\<^sub>R f x) \ i \ (content K *\<^sub>R g x) \ i" by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval) qed lemma has_integral_component_le: fixes f g :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" assumes "(f has_integral i) S" "(g has_integral j) S" and f_le_g: "\x. x \ S \ (f x)\k \ (g x)\k" shows "i\k \ j\k" proof - have ik_le_jk: "i\k \ j\k" if f_i: "(f has_integral i) (cbox a b)" and g_j: "(g has_integral j) (cbox a b)" and le: "\x\cbox a b. (f x)\k \ (g x)\k" for a b i and j :: 'b and f g :: "'a \ 'b" proof (rule ccontr) assume "\ ?thesis" then have *: "0 < (i\k - j\k) / 3" by auto obtain \1 where "gauge \1" and \1: "\p. \p tagged_division_of cbox a b; \1 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < (i \ k - j \ k) / 3" using f_i[unfolded has_integral,rule_format,OF *] by fastforce obtain \2 where "gauge \2" and \2: "\p. \p tagged_division_of cbox a b; \2 fine p\ \ norm ((\(x, k)\p. content k *\<^sub>R g x) - j) < (i \ k - j \ k) / 3" using g_j[unfolded has_integral,rule_format,OF *] by fastforce obtain p where p: "p tagged_division_of cbox a b" and "\1 fine p" "\2 fine p" using fine_division_exists[OF gauge_Int[OF \gauge \1\ \gauge \2\], of a b] unfolding fine_Int by metis then have "\((\(x, k)\p. content k *\<^sub>R f x) - i) \ k\ < (i \ k - j \ k) / 3" "\((\(x, k)\p. content k *\<^sub>R g x) - j) \ k\ < (i \ k - j \ k) / 3" using le_less_trans[OF Basis_le_norm[OF k]] k \1 \2 by metis+ then show False unfolding inner_simps using rsum_component_le[OF p] le by (fastforce simp add: abs_real_def split: if_split_asm) qed show ?thesis proof (cases "\a b. S = cbox a b") case True with ik_le_jk assms show ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ i\k \ j\k" then have ij: "(i\k - j\k) / 3 > 0" by auto obtain B1 where "0 < B1" and B1: "\a b. ball 0 B1 \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast obtain B2 where "0 < B2" and B2: "\a b. ball 0 B2 \ cbox a b \ \z. ((\x. if x \ S then g x else 0) has_integral z) (cbox a b) \ norm (z - j) < (i \ k - j \ k) / 3" using has_integral_altD[OF _ False ij] assms by blast have "bounded (ball 0 B1 \ ball (0::'a) B2)" unfolding bounded_Un by(rule conjI bounded_ball)+ from bounded_subset_cbox_symmetric[OF this] obtain a b::'a where ab: "ball 0 B1 \ cbox a b" "ball 0 B2 \ cbox a b" by (meson Un_subset_iff) then obtain w1 w2 where int_w1: "((\x. if x \ S then f x else 0) has_integral w1) (cbox a b)" and norm_w1: "norm (w1 - i) < (i \ k - j \ k) / 3" and int_w2: "((\x. if x \ S then g x else 0) has_integral w2) (cbox a b)" and norm_w2: "norm (w2 - j) < (i \ k - j \ k) / 3" using B1 B2 by blast have *: "\w1 w2 j i::real .\w1 - i\ < (i - j) / 3 \ \w2 - j\ < (i - j) / 3 \ w1 \ w2 \ False" by (simp add: abs_real_def split: if_split_asm) have "\(w1 - i) \ k\ < (i \ k - j \ k) / 3" "\(w2 - j) \ k\ < (i \ k - j \ k) / 3" using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+ moreover have "w1\k \ w2\k" using ik_le_jk int_w1 int_w2 f_le_g by auto ultimately show False unfolding inner_simps by(rule *) qed qed qed lemma integral_component_le: fixes g f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "f integrable_on S" "g integrable_on S" and "\x. x \ S \ (f x)\k \ (g x)\k" shows "(integral S f)\k \ (integral S g)\k" apply (rule has_integral_component_le) using integrable_integral assms apply auto done lemma has_integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ i\k" using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] using assms(3-) by auto lemma integral_component_nonneg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "\x. x \ S \ 0 \ (f x)\k" shows "0 \ (integral S f)\k" proof (cases "f integrable_on S") case True show ?thesis apply (rule has_integral_component_nonneg) using assms True apply auto done next case False then show ?thesis by (simp add: not_integrable_integral) qed lemma has_integral_component_neg: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "k \ Basis" and "(f has_integral i) S" and "\x. x \ S \ (f x)\k \ 0" shows "i\k \ 0" using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) by auto lemma has_integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ i\k" using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\i\Basis. B *\<^sub>R i)::'b"] assms(2-) by (auto simp add: field_simps) lemma has_integral_component_ubound: fixes f::"'a::euclidean_space => 'b::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "i\k \ B * content (cbox a b)" using has_integral_component_le[OF assms(3,1) has_integral_const, of "\i\Basis. B *\<^sub>R i"] assms(2-) by (auto simp add: field_simps) lemma integral_component_lbound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. B \ f(x)\k" and "k \ Basis" shows "B * content (cbox a b) \ (integral(cbox a b) f)\k" apply (rule has_integral_component_lbound) using assms unfolding has_integral_integral apply auto done lemma integral_component_lbound_real: assumes "f integrable_on {a ::real..b}" and "\x\{a..b}. B \ f(x)\k" and "k \ Basis" shows "B * content {a..b} \ (integral {a..b} f)\k" using assms by (metis box_real(2) integral_component_lbound) lemma integral_component_ubound: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f integrable_on cbox a b" and "\x\cbox a b. f x\k \ B" and "k \ Basis" shows "(integral (cbox a b) f)\k \ B * content (cbox a b)" apply (rule has_integral_component_ubound) using assms unfolding has_integral_integral apply auto done lemma integral_component_ubound_real: fixes f :: "real \ 'a::euclidean_space" assumes "f integrable_on {a..b}" and "\x\{a..b}. f x\k \ B" and "k \ Basis" shows "(integral {a..b} f)\k \ B * content {a..b}" using assms by (metis box_real(2) integral_component_ubound) subsection \Uniform limit of integrable functions is integrable\ lemma real_arch_invD: "0 < (e::real) \ (\n::nat. n \ 0 \ 0 < inverse (real n) \ inverse (real n) < e)" by (subst(asm) real_arch_inverse) lemma integrable_uniform_limit: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "\e. e > 0 \ \g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" shows "f integrable_on cbox a b" proof (cases "content (cbox a b) > 0") case False then show ?thesis using has_integral_null by (simp add: content_lt_nz integrable_on_def) next case True have "1 / (real n + 1) > 0" for n by auto then have "\g. (\x\cbox a b. norm (f x - g x) \ 1 / (real n + 1)) \ g integrable_on cbox a b" for n using assms by blast then obtain g where g_near_f: "\n x. x \ cbox a b \ norm (f x - g n x) \ 1 / (real n + 1)" and int_g: "\n. g n integrable_on cbox a b" by metis then obtain h where h: "\n. (g n has_integral h n) (cbox a b)" unfolding integrable_on_def by metis have "Cauchy h" unfolding Cauchy_def proof clarify fix e :: real assume "e>0" then have "e/4 / content (cbox a b) > 0" using True by (auto simp: field_simps) then obtain M where "M \ 0" and M: "1 / (real M) < e/4 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) show "\M. \m\M. \n\M. dist (h m) (h n) < e" proof (rule exI [where x=M], clarify) fix m n assume m: "M \ m" and n: "M \ n" have "e/4>0" using \e>0\ by auto then obtain gm gn where "gauge gm" "gauge gn" and gm: "\\. \ tagged_division_of cbox a b \ gm fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g m x) - h m) < e/4" and gn: "\\. \ tagged_division_of cbox a b \ gn fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g n x) - h n) < e/4" using h[unfolded has_integral] by meson then obtain \ where \: "\ tagged_division_of cbox a b" "(\x. gm x \ gn x) fine \" by (metis (full_types) fine_division_exists gauge_Int) have triangle3: "norm (i1 - i2) < e" if no: "norm(s2 - s1) \ e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b proof - have "norm (i1 - i2) \ norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed have finep: "gm fine \" "gn fine \" using fine_Int \ by auto have norm_le: "norm (g n x - g m x) \ 2 / real M" if x: "x \ cbox a b" for x proof - have "norm (f x - g n x) + norm (f x - g m x) \ 1 / (real n + 1) + 1 / (real m + 1)" using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp also have "\ \ 1 / (real M) + 1 / (real M)" apply (rule add_mono) using \M \ 0\ m n by (auto simp: field_split_simps) also have "\ = 2 / real M" by auto finally show "norm (g n x - g m x) \ 2 / real M" using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] by (auto simp: algebra_simps simp add: norm_minus_commute) qed have "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ 2 / real M * content (cbox a b)" by (blast intro: norm_le rsum_diff_bound[OF \(1), where e="2 / real M"]) also have "... \ e/2" using M True by (auto simp: field_simps) finally have le_e2: "norm ((\(x,K) \ \. content K *\<^sub>R g n x) - (\(x,K) \ \. content K *\<^sub>R g m x)) \ e/2" . then show "dist (h m) (h n) < e" unfolding dist_norm using gm gn \ finep by (auto intro!: triangle3) qed qed then obtain s where s: "h \ s" using convergent_eq_Cauchy[symmetric] by blast show ?thesis unfolding integrable_on_def has_integral proof (rule_tac x=s in exI, clarify) fix e::real assume e: "0 < e" then have "e/3 > 0" by auto then obtain N1 where N1: "\n\N1. norm (h n - s) < e/3" using LIMSEQ_D [OF s] by metis from e True have "e/3 / content (cbox a b) > 0" by (auto simp: field_simps) then obtain N2 :: nat where "N2 \ 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)" by (metis inverse_eq_divide real_arch_inverse) obtain g' where "gauge g'" and g': "\\. \ tagged_division_of cbox a b \ g' fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" by (metis h has_integral \e/3 > 0\) have *: "norm (sf - s) < e" if no: "norm (sf - sg) \ e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h proof - have "norm (sf - s) \ norm (sf - sg) + norm (sg - h) + norm (h - s)" using norm_triangle_ineq[of "sf - sg" "sg - s"] using norm_triangle_ineq[of "sg - h" " h - s"] by (auto simp: algebra_simps) also have "\ < e" using no by (auto simp: algebra_simps norm_minus_commute) finally show ?thesis . qed { fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "g' fine \" then have norm_less: "norm ((\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" using g' by blast have "content (cbox a b) < e/3 * (of_nat N2)" using \N2 \ 0\ N2 using True by (auto simp: field_split_simps) moreover have "e/3 * of_nat N2 \ e/3 * (of_nat (N1 + N2) + 1)" using \e>0\ by auto ultimately have "content (cbox a b) < e/3 * (of_nat (N1 + N2) + 1)" by linarith then have le_e3: "1 / (real (N1 + N2) + 1) * content (cbox a b) \ e/3" unfolding inverse_eq_divide by (auto simp: field_simps) have ne3: "norm (h (N1 + N2) - s) < e/3" using N1 by auto have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - (\(x,K) \ \. content K *\<^sub>R g (N1 + N2) x)) \ 1 / (real (N1 + N2) + 1) * content (cbox a b)" by (blast intro: g_near_f rsum_diff_bound[OF ptag]) then have "norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e" by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less]) } then show "\d. gauge d \ (\\. \ tagged_division_of cbox a b \ d fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - s) < e)" by (blast intro: g' \gauge g'\) qed qed lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] subsection \Negligible sets\ definition "negligible (s:: 'a::euclidean_space set) \ (\a b. ((indicator s :: 'a\real) has_integral 0) (cbox a b))" subsubsection \Negligibility of hyperplane\ lemma content_doublesplit: fixes a :: "'a::euclidean_space" assumes "0 < e" and k: "k \ Basis" obtains d where "0 < d" and "content (cbox a b \ {x. \x\k - c\ \ d}) < e" proof cases assume *: "a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j)" define a' where "a' d = (\j\Basis. (if j = k then max (a\j) (c - d) else a\j) *\<^sub>R j)" for d define b' where "b' d = (\j\Basis. (if j = k then min (b\j) (c + d) else b\j) *\<^sub>R j)" for d have "((\d. \j\Basis. (b' d - a' d) \ j) \ (\j\Basis. (b' 0 - a' 0) \ j)) (at_right 0)" by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros) also have "(\j\Basis. (b' 0 - a' 0) \ j) = 0" using k * by (intro prod_zero bexI[OF _ k]) (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong) also have "((\d. \j\Basis. (b' d - a' d) \ j) \ 0) (at_right 0) = ((\d. content (cbox a b \ {x. \x\k - c\ \ d})) \ 0) (at_right 0)" proof (intro tendsto_cong eventually_at_rightI) fix d :: real assume d: "d \ {0<..<1}" have "cbox a b \ {x. \x\k - c\ \ d} = cbox (a' d) (b' d)" for d using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def) moreover have "j \ Basis \ a' d \ j \ b' d \ j" for j using * d k by (auto simp: a'_def b'_def) ultimately show "(\j\Basis. (b' d - a' d) \ j) = content (cbox a b \ {x. \x\k - c\ \ d})" by simp qed simp finally have "((\d. content (cbox a b \ {x. \x \ k - c\ \ d})) \ 0) (at_right 0)" . from order_tendstoD(2)[OF this \0] obtain d' where "0 < d'" and d': "\y. y > 0 \ y < d' \ content (cbox a b \ {x. \x \ k - c\ \ y}) < e" by (subst (asm) eventually_at_right[of _ 1]) auto show ?thesis by (rule that[of "d'/2"], insert \0 d'[of "d'/2"], auto) next assume *: "\ (a \ k \ c \ c \ b \ k \ (\j\Basis. a \ j \ b \ j))" then have "(\j\Basis. b \ j < a \ j) \ (c < a \ k \ b \ k < c)" by (auto simp: not_le) show thesis proof cases assume "\j\Basis. b \ j < a \ j" then have [simp]: "cbox a b = {}" using box_ne_empty(1)[of a b] by auto show ?thesis by (rule that[of 1]) (simp_all add: \0) next assume "\ (\j\Basis. b \ j < a \ j)" with * have "c < a \ k \ b \ k < c" by auto then show thesis proof assume c: "c < a \ k" moreover have "x \ cbox a b \ c \ x \ k" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (a \ k - c)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(a \ k - c)/2"] show ?thesis by auto next assume c: "b \ k < c" moreover have "x \ cbox a b \ x \ k \ c" for x using k c by (auto simp: cbox_def) ultimately have "cbox a b \ {x. \x \ k - c\ \ (c - b \ k)/2} = {}" using k by (auto simp: cbox_def) with \0 c that[of "(c - b \ k)/2"] show ?thesis by auto qed qed qed proposition negligible_standard_hyperplane[intro]: fixes k :: "'a::euclidean_space" assumes k: "k \ Basis" shows "negligible {x. x\k = c}" unfolding negligible_def has_integral proof clarsimp fix a b and e::real assume "e > 0" with k obtain d where "0 < d" and d: "content (cbox a b \ {x. \x \ k - c\ \ d}) < e" by (metis content_doublesplit) let ?i = "indicator {x::'a. x\k = c} :: 'a\real" show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K * ?i x\ < e)" proof (intro exI, safe) show "gauge (\x. ball x d)" using \0 < d\ by blast next fix \ assume p: "\ tagged_division_of (cbox a b)" "(\x. ball x d) fine \" have "content L = content (L \ {x. \x \ k - c\ \ d})" if "(x, L) \ \" "?i x \ 0" for x L proof - have xk: "x\k = c" using that by (simp add: indicator_def split: if_split_asm) have "L \ {x. \x \ k - c\ \ d}" proof fix y assume y: "y \ L" have "L \ ball x d" using p(2) that(1) by auto then have "norm (x - y) < d" by (simp add: dist_norm subset_iff y) then have "\(x - y) \ k\ < d" using k norm_bound_Basis_lt by blast then show "y \ {x. \x \ k - c\ \ d}" unfolding inner_simps xk by auto qed then show "content L = content (L \ {x. \x \ k - c\ \ d})" by (metis inf.orderE) qed then have *: "(\(x,K)\\. content K * ?i x) = (\(x,K)\\. content (K \ {x. \x\k - c\ \ d}) *\<^sub>R ?i x)" by (force simp add: split_paired_all intro!: sum.cong [OF refl]) note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)] have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * indicator {x. x \ k = c} x) < e" proof - have "(\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) \ (\(x,K)\\. content (K \ {x. \x \ k - c\ \ d}))" by (force simp add: indicator_def intro!: sum_mono) also have "\ < e" proof (subst sum.over_tagged_division_lemma[OF p(1)]) fix u v::'a assume "box u v = {}" then have *: "content (cbox u v) = 0" unfolding content_eq_0_interior by simp have "cbox u v \ {x. \x \ k - c\ \ d} \ cbox u v" by auto then have "content (cbox u v \ {x. \x \ k - c\ \ d}) \ content (cbox u v)" unfolding interval_doublesplit[OF k] by (rule content_subset) then show "content (cbox u v \ {x. \x \ k - c\ \ d}) = 0" unfolding * interval_doublesplit[OF k] by (blast intro: antisym) next have "(\l\snd ` \. content (l \ {x. \x \ k - c\ \ d})) = sum content ((\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}})" proof (subst (2) sum.reindex_nontrivial) fix x y assume "x \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "y \ {l \ snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}" "x \ y" and eq: "x \ {x. \x \ k - c\ \ d} = y \ {x. \x \ k - c\ \ d}" then obtain x' y' where "(x', x) \ \" "x \ {x. \x \ k - c\ \ d} \ {}" "(y', y) \ \" "y \ {x. \x \ k - c\ \ d} \ {}" by (auto) from p'(5)[OF \(x', x) \ \\ \(y', y) \ \\] \x \ y\ have "interior (x \ y) = {}" by auto moreover have "interior ((x \ {x. \x \ k - c\ \ d}) \ (y \ {x. \x \ k - c\ \ d})) \ interior (x \ y)" by (auto intro: interior_mono) ultimately have "interior (x \ {x. \x \ k - c\ \ d}) = {}" by (auto simp: eq) then show "content (x \ {x. \x \ k - c\ \ d}) = 0" using p'(4)[OF \(x', x) \ \\] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) qed (insert p'(1), auto intro!: sum.mono_neutral_right) also have "\ \ norm (\l\(\l. l \ {x. \x \ k - c\ \ d})`{l\snd ` \. l \ {x. \x \ k - c\ \ d} \ {}}. content l *\<^sub>R 1::real)" by simp also have "\ \ 1 * content (cbox a b \ {x. \x \ k - c\ \ d})" using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto also have "\ < e" using d by simp finally show "(\K\snd ` \. content (K \ {x. \x \ k - c\ \ d})) < e" . qed finally show "(\(x, K)\\. content (K \ {x. \x \ k - c\ \ d}) * ?i x) < e" . qed then show "\\(x, K)\\. content K * ?i x\ < e" unfolding * apply (subst abs_of_nonneg) using measure_nonneg by (force simp add: indicator_def intro: sum_nonneg)+ qed qed corollary negligible_standard_hyperplane_cart: fixes k :: "'a::finite" shows "negligible {x. x$k = (0::real)}" by (simp add: cart_eq_inner_axis negligible_standard_hyperplane) subsubsection \Hence the main theorem about negligible sets\ lemma has_integral_negligible_cbox: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and 0: "\x. x \ S \ f x = 0" shows "(f has_integral 0) (cbox a b)" unfolding has_integral proof clarify fix e::real assume "e > 0" then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n by simp then have "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n))" for n using negs [unfolded negligible_def has_integral] by auto then obtain \ where gd: "\n. gauge (\ n)" and \: "\n \. \\ tagged_division_of cbox a b; \ n fine \\ \ \\(x,K) \ \. content K *\<^sub>R indicator S x\ < e/2 / ((real n + 1) * 2 ^ n)" by metis show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e)" proof (intro exI, safe) show "gauge (\x. \ (nat \norm (f x)\) x)" using gd by (auto simp: gauge_def) show "norm ((\(x,K) \ \. content K *\<^sub>R f x) - 0) < e" if "\ tagged_division_of (cbox a b)" "(\x. \ (nat \norm (f x)\) x) fine \" for \ proof (cases "\ = {}") case True with \0 < e\ show ?thesis by simp next case False obtain N where "Max ((\(x, K). norm (f x)) ` \) \ real N" using real_arch_simple by blast then have N: "\x. x \ (\(x, K). norm (f x)) ` \ \ x \ real N" by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite) have "\i. \q. q tagged_division_of (cbox a b) \ (\ i) fine q \ (\(x,K) \ \. K \ (\ i) x \ (x, K) \ q)" by (auto intro: tagged_division_finer[OF that(1) gd]) from choice[OF this] obtain q where q: "\n. q n tagged_division_of cbox a b" "\n. \ n fine q n" "\n x K. \(x, K) \ \; K \ \ n x\ \ (x, K) \ q n" by fastforce have "finite \" using that(1) by blast then have sum_le_inc: "\finite T; \x y. (x,y) \ T \ (0::real) \ g(x,y); \y. y\\ \ \x. (x,y) \ T \ f(y) \ g(x,y)\ \ sum f \ \ sum g T" for f g T by (rule sum_le_included[of \ T g snd f]; force) have "norm (\(x,K) \ \. content K *\<^sub>R f x) \ (\(x,K) \ \. norm (content K *\<^sub>R f x))" unfolding split_def by (rule norm_sum) also have "... \ (\(i, j) \ Sigma {..N + 1} q. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" proof (rule sum_le_inc, safe) show "finite (Sigma {..N+1} q)" by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1)) next fix x K assume xk: "(x, K) \ \" define n where "n = nat \norm (f x)\" have *: "norm (f x) \ (\(x, K). norm (f x)) ` \" using xk by auto have nfx: "real n \ norm (f x)" "norm (f x) \ real n + 1" unfolding n_def by auto then have "n \ {0..N + 1}" using N[OF *] by auto moreover have "K \ \ (nat \norm (f x)\) x" using that(2) xk by auto moreover then have "(x, K) \ q (nat \norm (f x)\)" by (simp add: q(3) xk) moreover then have "(x, K) \ q n" using n_def by blast moreover have "norm (content K *\<^sub>R f x) \ (real n + 1) * (content K * indicator S x)" proof (cases "x \ S") case False then show ?thesis by (simp add: 0) next case True have *: "content K \ 0" using tagged_division_ofD(4)[OF that(1) xk] by auto moreover have "content K * norm (f x) \ content K * (real n + 1)" by (simp add: mult_left_mono nfx(2)) ultimately show ?thesis using nfx True by (auto simp: field_simps) qed ultimately show "\y. (y, x, K) \ (Sigma {..N + 1} q) \ norm (content K *\<^sub>R f x) \ (real y + 1) * (content K *\<^sub>R indicator S x)" by force qed auto also have "... = (\i\N + 1. \j\q i. (real i + 1) * (case j of (x, K) \ content K *\<^sub>R indicator S x))" apply (rule sum_Sigma_product [symmetric]) using q(1) apply auto done also have "... \ (\i\N + 1. (real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\)" by (rule sum_mono) (simp add: sum_distrib_left [symmetric]) also have "... \ (\i\N + 1. e/2/2 ^ i)" proof (rule sum_mono) show "(real i + 1) * \\(x,K) \ q i. content K *\<^sub>R indicator S x\ \ e/2/2 ^ i" if "i \ {..N + 1}" for i using \[of "q i" i] q by (simp add: divide_simps mult.left_commute) qed also have "... = e/2 * (\i\N + 1. (1/2) ^ i)" unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over) also have "\ < e/2 * 2" proof (rule mult_strict_left_mono) have "sum (power (1/2)) {..N + 1} = sum (power (1/2::real)) {..0 < e\ in auto) finally show ?thesis by auto qed qed qed proposition has_integral_negligible: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes negs: "negligible S" and "\x. x \ (T - S) \ f x = 0" shows "(f has_integral 0) T" proof (cases "\a b. T = cbox a b") case True then have "((\x. if x \ T then f x else 0) has_integral 0) T" using assms by (auto intro!: has_integral_negligible_cbox) then show ?thesis by (rule has_integral_eq [rotated]) auto next case False let ?f = "(\x. if x \ T then f x else 0)" have "((\x. if x \ T then f x else 0) has_integral 0) T" apply (auto simp: False has_integral_alt [of ?f]) apply (rule_tac x=1 in exI, auto) apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms) done then show ?thesis by (rule_tac f="?f" in has_integral_eq) auto qed lemma assumes "negligible S" shows integrable_negligible: "f integrable_on S" and integral_negligible: "integral S f = 0" using has_integral_negligible [OF assms] by (auto simp: has_integral_iff) lemma has_integral_spike: fixes f :: "'b::euclidean_space \ 'a::real_normed_vector" assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" and fint: "(f has_integral y) T" shows "(g has_integral y) T" proof - have *: "(g has_integral y) (cbox a b)" if "(f has_integral y) (cbox a b)" "\x \ cbox a b - S. g x = f x" for a b f and g:: "'b \ 'a" and y proof - have "((\x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" using that by (intro has_integral_add has_integral_negligible) (auto intro!: \negligible S\) then show ?thesis by auto qed show ?thesis using fint gf apply (subst has_integral_alt) apply (subst (asm) has_integral_alt) apply (simp split: if_split_asm) apply (blast dest: *) apply (erule_tac V = "\a b. T \ cbox a b" in thin_rl) apply (elim all_forward imp_forward ex_forward all_forward conj_forward asm_rl) apply (auto dest!: *[where f="\x. if x\T then f x else 0" and g="\x. if x \ T then g x else 0"]) done qed lemma has_integral_spike_eq: assumes "negligible S" and gf: "\x. x \ T - S \ g x = f x" shows "(f has_integral y) T \ (g has_integral y) T" using has_integral_spike [OF \negligible S\] gf by metis lemma integrable_spike: assumes "f integrable_on T" "negligible S" "\x. x \ T - S \ g x = f x" shows "g integrable_on T" using assms unfolding integrable_on_def by (blast intro: has_integral_spike) lemma integral_spike: assumes "negligible S" and "\x. x \ T - S \ g x = f x" shows "integral T f = integral T g" using has_integral_spike_eq[OF assms] by (auto simp: integral_def integrable_on_def) subsection \Some other trivialities about negligible sets\ lemma negligible_subset: assumes "negligible s" "t \ s" shows "negligible t" unfolding negligible_def by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2)) lemma negligible_diff[intro?]: assumes "negligible s" shows "negligible (s - t)" using assms by (meson Diff_subset negligible_subset) lemma negligible_Int: assumes "negligible s \ negligible t" shows "negligible (s \ t)" using assms negligible_subset by force lemma negligible_Un: assumes "negligible S" and T: "negligible T" shows "negligible (S \ T)" proof - have "(indicat_real (S \ T) has_integral 0) (cbox a b)" if S0: "(indicat_real S has_integral 0) (cbox a b)" and "(indicat_real T has_integral 0) (cbox a b)" for a b proof (subst has_integral_spike_eq[OF T]) show "indicat_real S x = indicat_real (S \ T) x" if "x \ cbox a b - T" for x by (metis Diff_iff Un_iff indicator_def that) show "(indicat_real S has_integral 0) (cbox a b)" by (simp add: S0) qed with assms show ?thesis unfolding negligible_def by blast qed lemma negligible_Un_eq[simp]: "negligible (s \ t) \ negligible s \ negligible t" using negligible_Un negligible_subset by blast lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}" using negligible_standard_hyperplane[OF SOME_Basis, of "a \ (SOME i. i \ Basis)"] negligible_subset by blast lemma negligible_insert[simp]: "negligible (insert a s) \ negligible s" apply (subst insert_is_Un) unfolding negligible_Un_eq apply auto done lemma negligible_empty[iff]: "negligible {}" using negligible_insert by blast text\Useful in this form for backchaining\ lemma empty_imp_negligible: "S = {} \ negligible S" by simp lemma negligible_finite[intro]: assumes "finite s" shows "negligible s" using assms by (induct s) auto lemma negligible_Union[intro]: assumes "finite \" and "\t. t \ \ \ negligible t" shows "negligible(\\)" using assms by induct auto lemma negligible: "negligible S \ (\T. (indicat_real S has_integral 0) T)" proof (intro iffI allI) fix T assume "negligible S" then show "(indicator S has_integral 0) T" by (meson Diff_iff has_integral_negligible indicator_simps(2)) qed (simp add: negligible_def) subsection \Finite case of the spike theorem is quite commonly needed\ lemma has_integral_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "(f has_integral y) T" shows "(g has_integral y) T" using assms has_integral_spike negligible_finite by blast lemma has_integral_spike_finite_eq: assumes "finite S" and "\x. x \ T - S \ g x = f x" shows "((f has_integral y) T \ (g has_integral y) T)" by (metis assms has_integral_spike_finite) lemma integrable_spike_finite: assumes "finite S" and "\x. x \ T - S \ g x = f x" and "f integrable_on T" shows "g integrable_on T" using assms has_integral_spike_finite by blast lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed corollary has_integral_bound_real: fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" "finite S" and "(f has_integral i) {a..b}" and "\x. x \ {a..b} - S \ norm (f x) \ B" shows "norm i \ B * content {a..b}" by (metis assms box_real(2) has_integral_bound_spike_finite) subsection \In particular, the boundary of an interval is negligible\ lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" proof - let ?A = "\((\k. {x. x\k = a\k} \ {x::'a. x\k = b\k}) ` Basis)" have "negligible ?A" by (force simp add: negligible_Union[OF finite_imageI]) moreover have "cbox a b - box a b \ ?A" by (force simp add: mem_box) ultimately show ?thesis by (rule negligible_subset) qed lemma has_integral_spike_interior: assumes f: "(f has_integral y) (cbox a b)" and gf: "\x. x \ box a b \ g x = f x" shows "(g has_integral y) (cbox a b)" apply (rule has_integral_spike[OF negligible_frontier_interval _ f]) using gf by auto lemma has_integral_spike_interior_eq: assumes "\x. x \ box a b \ g x = f x" shows "(f has_integral y) (cbox a b) \ (g has_integral y) (cbox a b)" by (metis assms has_integral_spike_interior) lemma integrable_spike_interior: assumes "\x. x \ box a b \ g x = f x" and "f integrable_on cbox a b" shows "g integrable_on cbox a b" using assms has_integral_spike_interior_eq by blast subsection \Integrability of continuous functions\ lemma operative_approximableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i)" proof - interpret comm_monoid conj True by standard auto show ?thesis proof (standard, safe) fix a b :: 'b show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if "box a b = {}" for a b apply (rule_tac x=f in exI) using assms that by (auto simp: content_eq_0_interior) { fix c g and k :: 'b assume fg: "\x\cbox a b. norm (f x - g x) \ e" and g: "g integrable_on cbox a b" assume k: "k \ Basis" show "\g. (\x\cbox a b \ {x. x \ k \ c}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. x \ k \ c}" "\g. (\x\cbox a b \ {x. c \ x \ k}. norm (f x - g x) \ e) \ g integrable_on cbox a b \ {x. c \ x \ k}" apply (rule_tac[!] x=g in exI) using fg integrable_split[OF g k] by auto } show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" if fg1: "\x\cbox a b \ {x. x \ k \ c}. norm (f x - g1 x) \ e" and g1: "g1 integrable_on cbox a b \ {x. x \ k \ c}" and fg2: "\x\cbox a b \ {x. c \ x \ k}. norm (f x - g2 x) \ e" and g2: "g2 integrable_on cbox a b \ {x. c \ x \ k}" and k: "k \ Basis" for c k g1 g2 proof - let ?g = "\x. if x\k = c then f x else if x\k \ c then g1 x else g2 x" show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" proof (intro exI conjI ballI) show "norm (f x - ?g x) \ e" if "x \ cbox a b" for x by (auto simp: that assms fg1 fg2) show "?g integrable_on cbox a b" proof - have "?g integrable_on cbox a b \ {x. x \ k \ c}" "?g integrable_on cbox a b \ {x. x \ k \ c}" by(rule integrable_spike[OF _ negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+ with has_integral_split[OF _ _ k] show ?thesis unfolding integrable_on_def by blast qed qed qed qed qed lemma comm_monoid_set_F_and: "comm_monoid_set.F (\) True f s \ (finite s \ (\x\s. f x))" proof - interpret bool: comm_monoid_set "(\)" True proof qed auto show ?thesis by (induction s rule: infinite_finite_induct) auto qed lemma approximable_on_division: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" and d: "d division_of (cbox a b)" and f: "\i\d. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" obtains g where "\x\cbox a b. norm (f x - g x) \ e" "g integrable_on cbox a b" proof - interpret operative conj True "\i. \g. (\x\i. norm (f x - g (x::'b)) \ e) \ g integrable_on i" using \0 \ e\ by (rule operative_approximableI) from f local.division [OF d] that show thesis by auto qed lemma integrable_continuous: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" proof (rule integrable_uniform_limit) fix e :: real assume e: "e > 0" then obtain d where "0 < d" and d: "\x x'. \x \ cbox a b; x' \ cbox a b; dist x' x < d\ \ dist (f x') (f x) < e" using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x d) fine p" using fine_division_exists[OF gauge_ball[OF \0 < d\], of a b] . have *: "\i\snd ` p. \g. (\x\i. norm (f x - g x) \ e) \ g integrable_on i" proof (safe, unfold snd_conv) fix x l assume as: "(x, l) \ p" obtain a b where l: "l = cbox a b" using as ptag by blast then have x: "x \ cbox a b" using as ptag by auto show "\g. (\x\l. norm (f x - g x) \ e) \ g integrable_on l" apply (rule_tac x="\y. f x" in exI) proof safe show "(\y. f x) integrable_on l" unfolding integrable_on_def l by blast next fix y assume y: "y \ l" then have "y \ ball x d" using as finep by fastforce then show "norm (f y - f x) \ e" using d x y as l by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3)) qed qed from e have "e \ 0" by auto from approximable_on_division[OF this division_of_tagged_division[OF ptag] *] show "\g. (\x\cbox a b. norm (f x - g x) \ e) \ g integrable_on cbox a b" by metis qed lemma integrable_continuous_interval: fixes f :: "'b::ordered_euclidean_space \ 'a::banach" assumes "continuous_on {a..b} f" shows "f integrable_on {a..b}" by (metis assms integrable_continuous interval_cbox) lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real] lemma integrable_continuous_closed_segment: fixes f :: "real \ 'a::banach" assumes "continuous_on (closed_segment a b) f" shows "f integrable_on (closed_segment a b)" using assms by (auto intro!: integrable_continuous_interval simp: closed_segment_eq_real_ivl) subsection \Specialization of additivity to one dimension\ subsection \A useful lemma allowing us to factor out the content size\ lemma has_integral_factor_content: "(f has_integral i) (cbox a b) \ (\e>0. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content (cbox a b)))" proof (cases "content (cbox a b) = 0") case True have "\e p. p tagged_division_of cbox a b \ norm ((\(x, k)\p. content k *\<^sub>R f x)) \ e * content (cbox a b)" unfolding sum_content_null[OF True] True by force moreover have "i = 0" if "\e. e > 0 \ \d. gauge d \ (\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) \ e * content (cbox a b))" using that [of 1] by (force simp add: True sum_content_null[OF True] intro: fine_division_exists[of _ a b]) ultimately show ?thesis unfolding has_integral_null_eq[OF True] by (force simp add: ) next case False then have F: "0 < content (cbox a b)" using zero_less_measure_iff by blast let ?P = "\e opp. \d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ opp (norm ((\(x, k)\p. content k *\<^sub>R f x) - i)) e)" show ?thesis apply (subst has_integral) proof safe fix e :: real assume e: "e > 0" { assume "\e>0. ?P e (<)" then show "?P (e * content (cbox a b)) (\)" apply (rule allE [where x="e * content (cbox a b)"]) apply (elim impE ex_forward conj_forward) using F e apply (auto simp add: algebra_simps) done } { assume "\e>0. ?P (e * content (cbox a b)) (\)" then show "?P e (<)" apply (rule allE [where x="e/2 / content (cbox a b)"]) apply (elim impE ex_forward conj_forward) using F e apply (auto simp add: algebra_simps) done } qed qed lemma has_integral_factor_content_real: "(f has_integral i) {a..b::real} \ (\e>0. \d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm (sum (\(x,k). content k *\<^sub>R f x) p - i) \ e * content {a..b} ))" unfolding box_real[symmetric] by (rule has_integral_factor_content) subsection \Fundamental theorem of calculus\ lemma interval_bounds_real: fixes q b :: real assumes "a \ b" shows "Sup {a..b} = b" and "Inf {a..b} = a" using assms by auto theorem fundamental_theorem_of_calculus: fixes f :: "real \ 'a::banach" assumes "a \ b" and vecd: "\x. x \ {a..b} \ (f has_vector_derivative f' x) (at x within {a..b})" shows "(f' has_integral (f b - f a)) {a..b}" unfolding has_integral_factor_content box_real[symmetric] proof safe fix e :: real assume "e > 0" then have "\x. \d>0. x \ {a..b} \ (\y\{a..b}. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x))" using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast then obtain d where d: "\x. 0 < d x" "\x y. \x \ {a..b}; y \ {a..b}; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e * norm (y-x)" by metis show "\d. gauge d \ (\p. p tagged_division_of (cbox a b) \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b))" proof (rule exI, safe) show "gauge (\x. ball x (d x))" using d(1) gauge_ball_dependent by blast next fix p assume ptag: "p tagged_division_of cbox a b" and finep: "(\x. ball x (d x)) fine p" have ba: "b - a = (\(x,K)\p. Sup K - Inf K)" "f b - f a = (\(x,K)\p. f(Sup K) - f(Inf K))" using additive_tagged_division_1[where f= "\x. x"] additive_tagged_division_1[where f= f] \a \ b\ ptag by auto have "norm (\(x, K) \ p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, k) \ Sup k - Inf k))" proof (rule sum_norm_le,safe) fix x K assume "(x, K) \ p" then have "x \ K" and kab: "K \ cbox a b" using ptag by blast+ then obtain u v where k: "K = cbox u v" and "x \ K" and kab: "K \ cbox a b" using ptag \(x, K) \ p\ by auto have "u \ v" using \x \ K\ unfolding k by auto have ball: "\y\K. y \ ball x (d x)" using finep \(x, K) \ p\ by blast have "u \ K" "v \ K" by (simp_all add: \u \ v\ k) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))" by (auto simp add: algebra_simps) also have "... \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" by (rule norm_triangle_ineq4) finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \ norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" . also have "\ \ e * norm (u - x) + e * norm (v - x)" proof (rule add_mono) show "norm (f u - f x - (u - x) *\<^sub>R f' x) \ e * norm (u - x)" apply (rule d(2)[of x u]) using \x \ K\ kab \u \ K\ ball dist_real_def by (auto simp add:dist_real_def) show "norm (f v - f x - (v - x) *\<^sub>R f' x) \ e * norm (v - x)" apply (rule d(2)[of x v]) using \x \ K\ kab \v \ K\ ball dist_real_def by (auto simp add:dist_real_def) qed also have "\ \ e * (Sup K - Inf K)" using \x \ K\ by (auto simp: k interval_bounds_real[OF \u \ v\] field_simps) finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (Sup K - Inf K)" using \u \ v\ by (simp add: k) qed with \a \ b\ show "norm ((\(x, K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content (cbox a b)" by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left) qed qed lemma ident_has_integral: fixes a::real assumes "a \ b" shows "((\x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}" proof - have "((\x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}" apply (rule fundamental_theorem_of_calculus [OF assms]) unfolding power2_eq_square by (rule derivative_eq_intros | simp)+ then show ?thesis by (simp add: field_simps) qed lemma integral_ident [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} (\x. x) = (if a \ b then (b\<^sup>2 - a\<^sup>2)/2 else 0)" by (metis assms ident_has_integral integral_unique) lemma ident_integrable_on: fixes a::real shows "(\x. x) integrable_on {a..b}" by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) lemma integral_sin [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} sin = cos a - cos b" proof - have "(sin has_integral (- cos b - - cos a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "((\a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma integral_cos [simp]: fixes a::real assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" proof - have "(cos has_integral (sin b - sin a)) {a..b}" proof (rule fundamental_theorem_of_calculus) show "(sin has_vector_derivative cos x) (at x within {a..b})" for x unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (rule derivative_eq_intros | force)+ qed (use assms in auto) then show ?thesis by (simp add: integral_unique) qed lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False apply (simp only: has_vector_derivative_def) apply (rule derivative_eq_intros | force)+ done qed auto then show ?thesis by simp qed auto lemma integral_sin_nx: "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" using has_integral_sin_nx [of n] by (force simp: mult.commute) lemma has_integral_cos_nx: "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" proof (cases "n = 0") case True then show ?thesis using has_integral_const_real [of "1::real" "-pi" pi] by auto next case False have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" if "x \ {-pi..pi}" for x :: real using that False apply (simp only: has_vector_derivative_def) apply (rule derivative_eq_intros | force)+ done qed auto with False show ?thesis by (simp add: mult.commute) qed lemma integral_cos_nx: "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" using has_integral_cos_nx [of n] by (force simp: mult.commute) subsection \Taylor series expansion\ lemma mvt_integral: fixes f::"'a::real_normed_vector\'b::banach" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes line_in: "\t. t \ {0..1} \ x + t *\<^sub>R y \ S" shows "f (x + y) - f x = integral {0..1} (\t. f' (x + t *\<^sub>R y) y)" (is ?th1) proof - from assms have subset: "(\xa. x + xa *\<^sub>R y) ` {0..1} \ S" by auto note [derivative_intros] = has_derivative_subset[OF _ subset] has_derivative_in_compose[where f="(\xa. x + xa *\<^sub>R y)" and g = f] note [continuous_intros] = continuous_on_compose2[where f="(\xa. x + xa *\<^sub>R y)"] continuous_on_subset[OF _ subset] have "\t. t \ {0..1} \ ((\t. f (x + t *\<^sub>R y)) has_vector_derivative f' (x + t *\<^sub>R y) y) (at t within {0..1})" using assms by (auto simp: has_vector_derivative_def linear_cmul[OF has_derivative_linear[OF f'], symmetric] intro!: derivative_eq_intros) from fundamental_theorem_of_calculus[rule_format, OF _ this] show ?th1 by (auto intro!: integral_unique[symmetric]) qed lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative: assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and g0: "Dg 0 = g" and Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" and ivl: "a \ t" "t \ b" shows "((\t. \iR prod (Df i t) (Dg (p - Suc i) t)) has_vector_derivative prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) (at t within {a..b})" using assms proof cases assume p: "p \ 1" define p' where "p' = p - 2" from assms p have p': "{..i. i \ p' \ Suc (Suc p' - i) = (Suc (Suc p') - i)" by auto let ?f = "\i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = (\i\(Suc p'). ?f i - ?f (Suc i))" by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) also note sum_telescope finally have "(\iR (prod (Df i t) (Dg (Suc (p - Suc i)) t) + prod (Df (Suc i) t) (Dg (p - Suc i) t))) = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" unfolding p'[symmetric] by (simp add: assms) thus ?thesis using assms by (auto intro!: derivative_eq_intros has_vector_derivative) qed (auto intro!: derivative_eq_intros has_vector_derivative) lemma fixes f::"real\'a::banach" assumes "p>0" and f0: "Df 0 = f" and Df: "\m t. m < p \ a \ t \ t \ b \ (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})" and ivl: "a \ b" defines "i \ \x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" shows Taylor_has_integral: "(i has_integral f b - (\iR Df i a)) {a..b}" and Taylor_integral: "f b = (\iR Df i a) + integral {a..b} i" and Taylor_integrable: "i integrable_on {a..b}" proof goal_cases case 1 interpret bounded_bilinear "scaleR::real\'a\'a" by (rule bounded_bilinear_scaleR) define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s define Dg where [abs_def]: "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s have g0: "Dg 0 = g" using \p > 0\ by (auto simp add: Dg_def field_split_simps g_def split: if_split_asm) { fix m assume "p > Suc m" hence "p - Suc m = Suc (p - Suc (Suc m))" by auto hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" by auto } note fact_eq = this have Dg: "\m t. m < p \ a \ t \ t \ b \ (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})" unfolding Dg_def by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq field_split_simps) let ?sum = "\t. \iR Dg i t *\<^sub>R Df (p - Suc i) t" from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, OF \p > 0\ g0 Dg f0 Df] have deriv: "\t. a \ t \ t \ b \ (?sum has_vector_derivative g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})" by auto from fundamental_theorem_of_calculus[rule_format, OF \a \ b\ deriv] have "(i has_integral ?sum b - ?sum a) {a..b}" using atLeastatMost_empty'[simp del] by (simp add: i_def g_def Dg_def) also have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" and "{.. {i. p = Suc i} = {p - 1}" for p' using \p > 0\ by (auto simp: power_mult_distrib[symmetric]) then have "?sum b = f b" using Suc_pred'[OF \p > 0\] by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib if_distribR sum.If_cases f0) also have "{..x. p - x - 1) ` {.. (\x. p - x - 1) ` {..iR Df i a)" by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one) finally show c: ?case . case 2 show ?case using c integral_unique by (metis (lifting) add.commute diff_eq_eq integral_unique) case 3 show ?case using c by force qed subsection \Only need trivial subintervals if the interval itself is trivial\ proposition division_of_nontrivial: fixes \ :: "'a::euclidean_space set set" assumes sdiv: "\ division_of (cbox a b)" and cont0: "content (cbox a b) \ 0" shows "{k. k \ \ \ content k \ 0} division_of (cbox a b)" using sdiv proof (induction "card \" arbitrary: \ rule: less_induct) case less note \ = division_ofD[OF less.prems] { presume *: "{k \ \. content k \ 0} \ \ \ ?case" then show ?case using less.prems by fastforce } assume noteq: "{k \ \. content k \ 0} \ \" then obtain K c d where "K \ \" and contk: "content K = 0" and keq: "K = cbox c d" using \(4) by blast then have "card \ > 0" unfolding card_gt_0_iff using less by auto then have card: "card (\ - {K}) < card \" using less \K \ \\ by (simp add: \(1)) have closed: "closed (\(\ - {K}))" using less.prems by auto have "x islimpt \(\ - {K})" if "x \ K" for x unfolding islimpt_approachable proof (intro allI impI) fix e::real assume "e > 0" obtain i where i: "c\i = d\i" "i\Basis" using contk \(3) [OF \K \ \\] unfolding box_ne_empty keq by (meson content_eq_0 dual_order.antisym) then have xi: "x\i = d\i" using \x \ K\ unfolding keq mem_box by (metis antisym) define y where "y = (\j\Basis. (if j = i then if c\i \ (a\i + b\i)/2 then c\i + min e (b\i - c\i)/2 else c\i - min e (c\i - a\i)/2 else x\j) *\<^sub>R j)" show "\x'\\(\ - {K}). x' \ x \ dist x' x < e" proof (intro bexI conjI) have "d \ cbox c d" using \(3)[OF \K \ \\] by (simp add: box_ne_empty(1) keq mem_box(2)) then have "d \ cbox a b" using \(2)[OF \K \ \\] by (auto simp: keq) then have di: "a \ i \ d \ i \ d \ i \ b \ i" using \i \ Basis\ mem_box(2) by blast then have xyi: "y\i \ x\i" unfolding y_def i xi using \e > 0\ cont0 \i \ Basis\ by (auto simp: content_eq_0 elim!: ballE[of _ _ i]) then show "y \ x" unfolding euclidean_eq_iff[where 'a='a] using i by auto have "norm (y-x) \ (\b\Basis. \(y - x) \ b\)" by (rule norm_le_l1) also have "... = \(y - x) \ i\ + (\b \ Basis - {i}. \(y - x) \ b\)" by (meson finite_Basis i(2) sum.remove) also have "... < e + sum (\i. 0) Basis" proof (rule add_less_le_mono) show "\(y-x) \ i\ < e" using di \e > 0\ y_def i xi by (auto simp: inner_simps) show "(\i\Basis - {i}. \(y-x) \ i\) \ (\i\Basis. 0)" unfolding y_def by (auto simp: inner_simps) qed finally have "norm (y-x) < e + sum (\i. 0) Basis" . then show "dist y x < e" unfolding dist_norm by auto have "y \ K" unfolding keq mem_box using i(1) i(2) xi xyi by fastforce moreover have "y \ \\" using subsetD[OF \(2)[OF \K \ \\] \x \ K\] \e > 0\ di i by (auto simp: \ mem_box y_def field_simps elim!: ballE[of _ _ i]) ultimately show "y \ \(\ - {K})" by auto qed qed then have "K \ \(\ - {K})" using closed closed_limpt by blast then have "\(\ - {K}) = cbox a b" unfolding \(6)[symmetric] by auto then have "\ - {K} division_of cbox a b" by (metis Diff_subset less.prems division_of_subset \(6)) then have "{ka \ \ - {K}. content ka \ 0} division_of (cbox a b)" using card less.hyps by blast moreover have "{ka \ \ - {K}. content ka \ 0} = {K \ \. content K \ 0}" using contk by auto ultimately show ?case by auto qed subsection \Integrability on subintervals\ lemma operative_integrableI: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "0 \ e" shows "operative conj True (\i. f integrable_on i)" proof - interpret comm_monoid conj True by standard auto have 1: "\a b. box a b = {} \ f integrable_on cbox a b" by (simp add: content_eq_0_interior integrable_on_null) have 2: "\a b c k. \k \ Basis; f integrable_on cbox a b \ {x. x \ k \ c}; f integrable_on cbox a b \ {x. c \ x \ k}\ \ f integrable_on cbox a b" unfolding integrable_on_def by (auto intro!: has_integral_split) show ?thesis apply standard using 1 2 by blast+ qed lemma integrable_subinterval: fixes f :: "'b::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and cd: "cbox c d \ cbox a b" shows "f integrable_on cbox c d" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) show ?thesis proof (cases "cbox c d = {}") case True then show ?thesis using division [symmetric] f by (auto simp: comm_monoid_set_F_and) next case False then show ?thesis by (metis cd comm_monoid_set_F_and division division_of_finite f partial_division_extend_1) qed qed lemma integrable_subinterval_real: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "{c..d} \ {a..b}" shows "f integrable_on {c..d}" by (metis assms box_real(2) integrable_subinterval) subsection \Combining adjacent intervals in 1 dimension\ lemma has_integral_combine: fixes a b c :: real and j :: "'a::banach" assumes "a \ c" and "c \ b" and ac: "(f has_integral i) {a..c}" and cb: "(f has_integral j) {c..b}" shows "(f has_integral (i + j)) {a..b}" proof - interpret operative_real "lift_option plus" "Some 0" "\i. if f integrable_on i then Some (integral i f) else None" using operative_integralI by (rule operative_realI) from \a \ c\ \c \ b\ ac cb coalesce_less_eq have *: "lift_option (+) (if f integrable_on {a..c} then Some (integral {a..c} f) else None) (if f integrable_on {c..b} then Some (integral {c..b} f) else None) = (if f integrable_on {a..b} then Some (integral {a..b} f) else None)" by (auto simp: split: if_split_asm) then have "f integrable_on cbox a b" using ac cb by (auto split: if_split_asm) with * show ?thesis using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) qed lemma integral_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and ab: "f integrable_on {a..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" proof - have "(f has_integral integral {a..c} f) {a..c}" using ab \c \ b\ integrable_subinterval_real by fastforce moreover have "(f has_integral integral {c..b} f) {c..b}" using ab \a \ c\ integrable_subinterval_real by fastforce ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}" using \a \ c\ \c \ b\ has_integral_combine by blast then show ?thesis by (simp add: has_integral_integrable_integral) qed lemma integrable_combine: fixes f :: "real \ 'a::banach" assumes "a \ c" and "c \ b" and "f integrable_on {a..c}" and "f integrable_on {c..b}" shows "f integrable_on {a..b}" using assms unfolding integrable_on_def by (auto intro!: has_integral_combine) lemma integral_minus_sets: fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {c .. max a b} \ integral {c .. a} f - integral {c .. b} f = (if a \ b then - integral {a .. b} f else integral {b .. a} f)" using integral_combine[of c a b f] integral_combine[of c b a f] by (auto simp: algebra_simps max_def) lemma integral_minus_sets': fixes f::"real \ 'a::banach" shows "c \ a \ c \ b \ f integrable_on {min a b .. c} \ integral {a .. c} f - integral {b .. c} f = (if a \ b then integral {a .. b} f else - integral {b .. a} f)" using integral_combine[of b a c f] integral_combine[of a b c f] by (auto simp: algebra_simps min_def) subsection \Reduce integrability to "local" integrability\ lemma integrable_on_little_subintervals: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "\x\cbox a b. \d>0. \u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v" shows "f integrable_on cbox a b" proof - interpret operative conj True "\i. f integrable_on i" using order_refl by (rule operative_integrableI) have "\x. \d>0. x\cbox a b \ (\u v. x \ cbox u v \ cbox u v \ ball x d \ cbox u v \ cbox a b \ f integrable_on cbox u v)" using assms by (metis zero_less_one) then obtain d where d: "\x. 0 < d x" "\x u v. \x \ cbox a b; x \ cbox u v; cbox u v \ ball x (d x); cbox u v \ cbox a b\ \ f integrable_on cbox u v" by metis obtain p where p: "p tagged_division_of cbox a b" "(\x. ball x (d x)) fine p" using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast then have sndp: "snd ` p division_of cbox a b" by (metis division_of_tagged_division) have "f integrable_on k" if "(x, k) \ p" for x k using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto then show ?thesis unfolding division [symmetric, OF sndp] comm_monoid_set_F_and by auto qed subsection \Second FTC or existence of antiderivative\ lemma integrable_const[intro]: "(\x. c) integrable_on cbox a b" unfolding integrable_on_def by blast lemma integral_has_vector_derivative_continuous_at: fixes f :: "real \ 'a::banach" assumes f: "f integrable_on {a..b}" and x: "x \ {a..b} - S" and "finite S" and fx: "continuous (at x within ({a..b} - S)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - S; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" if y: "y \ {a..b} - S" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" x y] False apply simp done have "\xa. y - x < d \ (\x'. a \ x' \ x' \ b \ x' \ S \ \x' - x\ < d \ norm (f x' - f x) \ e) \ 0 < e \ xa \ S \ a \ x \ x \ S \ y \ b \ y \ S \ x \ xa \ xa \ y \ norm (f xa - f x) \ e" using assms by auto show ?thesis using False apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) using yx False d x y \e>0\ assms by (auto simp: Idiff fux_int) next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" y x] True apply simp done have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" using True apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real[where f="(\u. f u - f x)"]) using yx True d x y \e>0\ assms by (auto simp: Idiff fux_int) then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - S. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})" using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] by (fastforce simp: continuous_on_eq_continuous_within) lemma integral_has_real_derivative: assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {a..x} g) has_real_derivative g t) (at t within {a..b})" using integral_has_vector_derivative[of a b g t] assms by (auto simp: has_field_derivative_iff_has_vector_derivative) lemma antiderivative_continuous: fixes q b :: real assumes "continuous_on {a..b} f" obtains g where "\x. x \ {a..b} \ (g has_vector_derivative (f x::_::banach)) (at x within {a..b})" using integral_has_vector_derivative[OF assms] by auto subsection \Combined fundamental theorem of calculus\ lemma antiderivative_integral_continuous: fixes f :: "real \ 'a::banach" assumes "continuous_on {a..b} f" obtains g where "\u\{a..b}. \v \ {a..b}. u \ v \ (f has_integral (g v - g u)) {u..v}" proof - obtain g where g: "\x. x \ {a..b} \ (g has_vector_derivative f x) (at x within {a..b})" using antiderivative_continuous[OF assms] by metis have "(f has_integral g v - g u) {u..v}" if "u \ {a..b}" "v \ {a..b}" "u \ v" for u v proof - have "\x. x \ cbox u v \ (g has_vector_derivative f x) (at x within cbox u v)" by (metis atLeastAtMost_iff atLeastatMost_subset_iff box_real(2) g has_vector_derivative_within_subset subsetCE that(1,2)) then show ?thesis by (metis box_real(2) that(3) fundamental_theorem_of_calculus) qed then show ?thesis using that by blast qed subsection \General "twiddling" for interval-to-interval function image\ lemma has_integral_twiddle: assumes "0 < r" and hg: "\x. h(g x) = x" and gh: "\x. g(h x) = x" and contg: "\x. continuous (at x) g" and g: "\u v. \w z. g ` cbox u v = cbox w z" and h: "\u v. \w z. h ` cbox u v = cbox w z" and r: "\u v. content(g ` cbox u v) = r * content (cbox u v)" and intfi: "(f has_integral i) (cbox a b)" shows "((\x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" proof (cases "cbox a b = {}") case True then show ?thesis using intfi by auto next case False obtain w z where wz: "h ` cbox a b = cbox w z" using h by blast have inj: "inj g" "inj h" using hg gh injI by metis+ from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast have "\d. gauge d \ (\p. p tagged_division_of h ` cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" if "e > 0" for e proof - have "e * r > 0" using that \0 < r\ by simp with intfi[unfolded has_integral] obtain d where "gauge d" and d: "\p. p tagged_division_of cbox a b \ d fine p \ norm ((\(x, k)\p. content k *\<^sub>R f x) - i) < e * r" by metis define d' where "d' x = g -` d (g x)" for x show ?thesis proof (rule_tac x=d' in exI, safe) show "gauge d'" using \gauge d\ continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def) next fix p assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p" note p = tagged_division_ofD[OF ptag] have gab: "g y \ cbox a b" if "y \ K" "(x, K) \ p" for x y K by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that) have gimp: "(\(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \ d fine (\(x, k). (g x, g ` k)) ` p" unfolding tagged_division_of proof safe show "finite ((\(x, k). (g x, g ` k)) ` p)" using ptag by auto show "d fine (\(x, k). (g x, g ` k)) ` p" using finep unfolding fine_def d'_def by auto next fix x k assume xk: "(x, k) \ p" show "g x \ g ` k" using p(2)[OF xk] by auto show "\u v. g ` k = cbox u v" using p(4)[OF xk] using assms(5-6) by auto fix x' K' u assume xk': "(x', K') \ p" and u: "u \ interior (g ` k)" "u \ interior (g ` K')" have "interior k \ interior K' \ {}" proof assume "interior k \ interior K' = {}" moreover have "u \ g ` (interior k \ interior K')" using interior_image_subset[OF \inj g\ contg] u unfolding image_Int[OF inj(1)] by blast ultimately show False by blast qed then have same: "(x, k) = (x', K')" using ptag xk' xk by blast then show "g x = g x'" by auto show "g u \ g ` K'"if "u \ k" for u using that same by auto show "g u \ g ` k"if "u \ K'" for u using that same by auto next fix x assume "x \ cbox a b" then have "h x \ \{k. \x. (x, k) \ p}" using p(6) by auto then obtain X y where "h x \ X" "(y, X) \ p" by blast then show "x \ \{k. \x. (x, k) \ (\(x, k). (g x, g ` k)) ` p}" apply clarsimp by (metis (no_types, lifting) assms(3) image_eqI pair_imageI) qed (use gab in auto) have *: "inj_on (\(x, k). (g x, g ` k)) p" using inj(1) unfolding inj_on_def by fastforce have "(\(x, k)\(\(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\(x, k)\p. content k *\<^sub>R f (g x)) - i" (is "?l = _") using r apply (simp only: algebra_simps add_left_cancel scaleR_right.sum) apply (subst sum.reindex_bij_betw[symmetric, where h="\(x, k). (g x, g ` k)" and S=p]) apply (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) done also have "\ = r *\<^sub>R ((\(x, k)\p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") using \0 < r\ by (auto simp: scaleR_diff_right) finally have eq: "?l = ?r" . show "norm ((\(x,K)\p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" using d[OF gimp] \0 < r\ by (auto simp add: eq) qed qed then show ?thesis by (auto simp: h_eq has_integral) qed subsection \Special case of a basic affine transformation\ lemma AE_lborel_inner_neq: assumes k: "k \ Basis" shows "AE x in lborel. x \ k \ c" proof - interpret finite_product_sigma_finite "\_. lborel" Basis proof qed simp have "emeasure lborel {x\space lborel. x \ k = c} = emeasure (\\<^sub>M j::'a\Basis. lborel) (\\<^sub>E j\Basis. if j = k then {c} else UNIV)" using k by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure]) (auto simp: space_PiM PiE_iff extensional_def split: if_split_asm) also have "\ = (\j\Basis. emeasure lborel (if j = k then {c} else UNIV))" by (intro measure_times) auto also have "\ = 0" by (intro prod_zero bexI[OF _ k]) auto finally show ?thesis by (subst AE_iff_measurable[OF _ refl]) auto qed lemma content_image_stretch_interval: fixes m :: "'a::euclidean_space \ real" defines "s f x \ (\k::'a\Basis. (f k * (x\k)) *\<^sub>R k)" shows "content (s m ` cbox a b) = \\k\Basis. m k\ * content (cbox a b)" proof cases have s[measurable]: "s f \ borel \\<^sub>M borel" for f by (auto simp: s_def[abs_def]) assume m: "\k\Basis. m k \ 0" then have s_comp_s: "s (\k. 1 / m k) \ s m = id" "s m \ s (\k. 1 / m k) = id" by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation) then have "inv (s (\k. 1 / m k)) = s m" "bij (s (\k. 1 / m k))" by (auto intro: inv_unique_comp o_bij) then have eq: "s m ` cbox a b = s (\k. 1 / m k) -` cbox a b" using bij_vimage_eq_inv_image[OF \bij (s (\k. 1 / m k))\, of "cbox a b"] by auto show ?thesis using m unfolding eq measure_def by (subst lborel_affine_euclidean[where c=m and t=0]) (simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg) next assume "\ (\k\Basis. m k \ 0)" then obtain k where k: "k \ Basis" "m k = 0" by auto then have [simp]: "(\k\Basis. m k) = 0" by (intro prod_zero) auto have "emeasure lborel {x\space lborel. x \ s m ` cbox a b} = 0" proof (rule emeasure_eq_0_AE) show "AE x in lborel. x \ s m ` cbox a b" using AE_lborel_inner_neq[OF \k\Basis\] proof eventually_elim show "x \ k \ 0 \ x \ s m ` cbox a b " for x using k by (auto simp: s_def[abs_def] cbox_def) qed qed then show ?thesis by (simp add: measure_def) qed lemma interval_image_affinity_interval: "\u v. (\x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" unfolding image_affinity_cbox by auto lemma content_image_affinity_cbox: "content((\x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = \m\ ^ DIM('a) * content (cbox a b)" (is "?l = ?r") proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False show ?thesis proof (cases "m \ 0") case True with \cbox a b \ {}\ have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \ {}" unfolding box_ne_empty apply (intro ballI) apply (erule_tac x=i in ballE) apply (auto simp: inner_simps mult_left_mono) done moreover from True have *: "\i. (m *\<^sub>R b + c) \ i - (m *\<^sub>R a + c) \ i = m *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis by (simp add: image_affinity_cbox True content_cbox' prod.distrib prod_constant inner_diff_left) next case False with \cbox a b \ {}\ have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \ {}" unfolding box_ne_empty apply (intro ballI) apply (erule_tac x=i in ballE) apply (auto simp: inner_simps mult_left_mono) done moreover from False have *: "\i. (m *\<^sub>R a + c) \ i - (m *\<^sub>R b + c) \ i = (-m) *\<^sub>R (b-a) \ i" by (simp add: inner_simps field_simps) ultimately show ?thesis using False by (simp add: image_affinity_cbox content_cbox' prod.distrib[symmetric] inner_diff_left flip: prod_constant) qed qed lemma has_integral_affinity: fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m \ 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral ((1 / (\m\ ^ DIM('a))) *\<^sub>R i)) ((\x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)" apply (rule has_integral_twiddle) using assms apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox) apply (rule zero_less_power) unfolding scaleR_right_distrib apply auto done lemma integrable_affinity: assumes "f integrable_on cbox a b" and "m \ 0" shows "(\x. f(m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" using assms unfolding integrable_on_def apply safe apply (drule has_integral_affinity) apply auto done lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed subsection \Special case of stretching coordinate axes separately\ lemma has_integral_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "(f has_integral i) (cbox a b)" and "\k\Basis. m k \ 0" shows "((\x. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) has_integral ((1/ \prod m Basis\) *\<^sub>R i)) ((\x. (\k\Basis. (1 / m k * (x\k))*\<^sub>R k)) ` cbox a b)" apply (rule has_integral_twiddle[where f=f]) unfolding zero_less_abs_iff content_image_stretch_interval unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] using assms by auto lemma has_integral_stretch_real: fixes f :: "real \ 'b::real_normed_vector" assumes "(f has_integral i) {a..b}" and "m \ 0" shows "((\x. f (m * x)) has_integral (1 / \m\) *\<^sub>R i) ((\x. x / m) ` {a..b})" using has_integral_stretch [of f i a b "\b. m"] assms by simp lemma integrable_stretch: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "f integrable_on cbox a b" and "\k\Basis. m k \ 0" shows "(\x::'a. f (\k\Basis. (m k * (x\k))*\<^sub>R k)) integrable_on ((\x. \k\Basis. (1 / m k * (x\k))*\<^sub>R k) ` cbox a b)" using assms unfolding integrable_on_def by (force dest: has_integral_stretch) lemma vec_lambda_eq_sum: shows "(\ k. f k (x $ k)) = (\k\Basis. (f (axis_index k) (x \ k)) *\<^sub>R k)" apply (simp add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp add: vec_eq_iff axis_def if_distrib cong: if_cong) done lemma has_integral_stretch_cart: fixes m :: "'n::finite \ real" assumes f: "(f has_integral i) (cbox a b)" and m: "\k. m k \ 0" shows "((\x. f(\ k. m k * x$k)) has_integral i /\<^sub>R \prod m UNIV\) ((\x. \ k. x$k / m k) ` (cbox a b))" proof - have *: "\k:: real^'n \ Basis. m (axis_index k) \ 0" using axis_index by (simp add: m) have eqp: "(\k:: real^'n \ Basis. m (axis_index k)) = prod m UNIV" by (simp add: Basis_vec_def UNION_singleton_eq_range prod.reindex axis_eq_axis inj_on_def) show ?thesis using has_integral_stretch [OF f *] vec_lambda_eq_sum [where f="\i x. m i * x"] vec_lambda_eq_sum [where f="\i x. x / m i"] by (simp add: field_simps eqp) qed lemma image_stretch_interval_cart: fixes m :: "'n::finite \ real" shows "(\x. \ k. m k * x$k) ` cbox a b = (if cbox a b = {} then {} else cbox (\ k. min (m k * a$k) (m k * b$k)) (\ k. max (m k * a$k) (m k * b$k)))" proof - have *: "(\k\Basis. min (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. min (m k * a $ k) (m k * b $ k))" "(\k\Basis. max (m (axis_index k) * (a \ k)) (m (axis_index k) * (b \ k)) *\<^sub>R k) = (\ k. max (m k * a $ k) (m k * b $ k))" apply (simp_all add: Basis_vec_def cart_eq_inner_axis UNION_singleton_eq_range sum.reindex axis_eq_axis inj_on_def) apply (simp_all add: vec_eq_iff axis_def if_distrib cong: if_cong) done show ?thesis by (simp add: vec_lambda_eq_sum [where f="\i x. m i * x"] image_stretch_interval eq_cbox *) qed subsection \even more special cases\ lemma uminus_interval_vector[simp]: fixes a b :: "'a::euclidean_space" shows "uminus ` cbox a b = cbox (-b) (-a)" apply safe apply (simp add: mem_box(2)) apply (rule_tac x="-x" in image_eqI) apply (auto simp add: mem_box) done lemma has_integral_reflect_lemma[intro]: assumes "(f has_integral i) (cbox a b)" shows "((\x. f(-x)) has_integral i) (cbox (-b) (-a))" using has_integral_affinity[OF assms, of "-1" 0] by auto lemma has_integral_reflect_lemma_real[intro]: assumes "(f has_integral i) {a..b::real}" shows "((\x. f(-x)) has_integral i) {-b .. -a}" using assms unfolding box_real[symmetric] by (rule has_integral_reflect_lemma) lemma has_integral_reflect[simp]: "((\x. f (-x)) has_integral i) (cbox (-b) (-a)) \ (f has_integral i) (cbox a b)" by (auto dest: has_integral_reflect_lemma) lemma has_integral_reflect_real[simp]: fixes a b::real shows "((\x. f (-x)) has_integral i) {-b..-a} \ (f has_integral i) {a..b}" by (metis has_integral_reflect interval_cbox) lemma integrable_reflect[simp]: "(\x. f(-x)) integrable_on cbox (-b) (-a) \ f integrable_on cbox a b" unfolding integrable_on_def by auto lemma integrable_reflect_real[simp]: "(\x. f(-x)) integrable_on {-b .. -a} \ f integrable_on {a..b::real}" unfolding box_real[symmetric] by (rule integrable_reflect) lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\x. f (-x)) = integral (cbox a b) f" unfolding integral_def by auto lemma integral_reflect_real[simp]: "integral {-b .. -a} (\x. f (-x)) = integral {a..b::real} f" unfolding box_real[symmetric] by (rule integral_reflect) subsection \Stronger form of FCT; quite a tedious proof\ lemma split_minus[simp]: "(\(x, k). f x k) x - (\(x, k). g x k) x = (\(x, k). f x k - g x k) x" by (simp add: split_def) theorem fundamental_theorem_of_calculus_interior: fixes f :: "real \ 'a::real_normed_vector" assumes "a \ b" and contf: "continuous_on {a..b} f" and derf: "\x. x \ {a <..< b} \ (f has_vector_derivative f' x) (at x)" shows "(f' has_integral (f b - f a)) {a..b}" proof (cases "a = b") case True then have *: "cbox a b = {b}" "f b - f a = 0" by (auto simp add: order_antisym) with True show ?thesis by auto next case False with \a \ b\ have ab: "a < b" by arith show ?thesis unfolding has_integral_factor_content_real proof (intro allI impI) fix e :: real assume e: "e > 0" then have eba8: "(e * (b-a)) / 8 > 0" using ab by (auto simp add: field_simps) note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt] have bounded: "\x. x \ {a<.. bounded_linear (\u. u *\<^sub>R f' x)" using derf_exp by simp have "\x \ box a b. \d>0. \y. norm (y-x) < d \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" (is "\x \ box a b. ?Q x") proof fix x assume x: "x \ box a b" show "?Q x" apply (rule allE [where x="e/2", OF derf_exp [THEN conjunct2, of x]]) using x e by auto qed from this [unfolded bgauge_existence_lemma] obtain d where d: "\x. 0 < d x" "\x y. \x \ box a b; norm (y-x) < d x\ \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have "bounded (f ` cbox a b)" using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image) then obtain B where "0 < B" and B: "\x. x \ f ` cbox a b \ norm x \ B" unfolding bounded_pos by metis obtain da where "0 < da" and da: "\c. \a \ c; {a..c} \ {a..b}; {a..c} \ ball a da\ \ norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ (e * (b-a)) / 4" proof - have "continuous (at a within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm (x-a); norm (x-a) < k\ \ norm (f x - f a) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \ e * (b-a) / 8" proof (cases "f' a = 0") case True with ab e that show ?thesis by auto next case False then show ?thesis apply (rule_tac l="(e * (b-a)) / 8 / norm (f' a)" in that) using ab e apply (auto simp add: field_simps) done qed have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" if "a \ c" "{a..c} \ {a..b}" and bmin: "{a..c} \ ball a (min k l)" for c proof - have minkl: "\a - x\ < min k l" if "x \ {a..c}" for x using bmin dist_real_def that by auto then have lel: "\c - a\ \ \l\" using that by force have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \ norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((c - a) *\<^sub>R f' a) \ norm (l *\<^sub>R f' a)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((c - a) *\<^sub>R f' a) \ e * (b-a) / 8" . next have "norm (f c - f a) < e * (b-a) / 8" proof (cases "a = c") case True then show ?thesis using eba8 by auto next case False show ?thesis by (rule k) (use minkl \a \ c\ that False in auto) qed then show "norm (f c - f a) \ e * (b-a) / 8" by simp qed finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \ e * (b-a) / 4" unfolding content_real[OF \a \ c\] by auto qed then show ?thesis by (rule_tac da="min k l" in that) (auto simp: l \0 < k\) qed obtain db where "0 < db" and db: "\c. \c \ b; {c..b} \ {a..b}; {c..b} \ ball b db\ \ norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ (e * (b-a)) / 4" proof - have "continuous (at b within {a..b}) f" using contf continuous_on_eq_continuous_within by force with eba8 obtain k where "0 < k" and k: "\x. \x \ {a..b}; 0 < norm(x-b); norm(x-b) < k\ \ norm (f b - f x) < e * (b-a) / 8" unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \ (e * (b-a)) / 8" proof (cases "f' b = 0") case True thus ?thesis using ab e that by auto next case False then show ?thesis apply (rule_tac l="(e * (b-a)) / 8 / norm (f' b)" in that) using ab e by (auto simp add: field_simps) qed have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" if "c \ b" "{c..b} \ {a..b}" and bmin: "{c..b} \ ball b (min k l)" for c proof - have minkl: "\b - x\ < min k l" if "x \ {c..b}" for x using bmin dist_real_def that by auto then have lel: "\b - c\ \ \l\" using that by force have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \ norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" by (rule norm_triangle_ineq4) also have "\ \ e * (b-a) / 8 + e * (b-a) / 8" proof (rule add_mono) have "norm ((b - c) *\<^sub>R f' b) \ norm (l *\<^sub>R f' b)" by (auto intro: mult_right_mono [OF lel]) also have "... \ e * (b-a) / 8" by (rule l) finally show "norm ((b - c) *\<^sub>R f' b) \ e * (b-a) / 8" . next have "norm (f b - f c) < e * (b-a) / 8" proof (cases "b = c") case True with eba8 show ?thesis by auto next case False show ?thesis by (rule k) (use minkl \c \ b\ that False in auto) qed then show "norm (f b - f c) \ e * (b-a) / 8" by simp qed finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \ e * (b-a) / 4" unfolding content_real[OF \c \ b\] by auto qed then show ?thesis by (rule_tac db="min k l" in that) (auto simp: l \0 < k\) qed let ?d = "(\x. ball x (if x=a then da else if x=b then db else d x))" show "\d. gauge d \ (\p. p tagged_division_of {a..b} \ d fine p \ norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b})" proof (rule exI, safe) show "gauge ?d" using ab \db > 0\ \da > 0\ d(1) by (auto intro: gauge_ball_dependent) next fix p assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p" let ?A = "{t. fst t \ {a, b}}" note p = tagged_division_ofD[OF ptag] have pA: "p = (p \ ?A) \ (p - ?A)" "finite (p \ ?A)" "finite (p - ?A)" "(p \ ?A) \ (p - ?A) = {}" using ptag fine by auto have le_xz: "\w x y z::real. y \ z/2 \ w - x \ z/2 \ w + y \ x + z" by arith have non: False if xk: "(x,K) \ p" and "x \ a" "x \ b" and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" for x K proof - obtain u v where k: "K = cbox u v" using p(4) xk by blast then have "u \ v" and uv: "{u, v} \ cbox u v" using p(2)[OF xk] by auto then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using less[unfolded k box_real interval_bounds_real content_real] by auto then have "x \ box a b" using p(2) p(3) \x \ a\ \x \ b\ xk by fastforce with d have *: "\y. norm (y-x) < d x \ norm (f y - f x - (y-x) *\<^sub>R f' x) \ e/2 * norm (y-x)" by metis have xd: "norm (u - x) < d x" "norm (v - x) < d x" using fineD[OF fine xk] \x \ a\ \x \ b\ uv by (auto simp add: k subset_eq dist_commute dist_real_def) have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))" by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff) also have "\ \ e/2 * norm (u - x) + e/2 * norm (v - x)" by (metis norm_triangle_le_diff add_mono * xd) also have "\ \ e/2 * norm (v - u)" using p(2)[OF xk] by (auto simp add: field_simps k) also have "\ < norm ((v - u) *\<^sub>R f' x - (f v - f u))" using result by (simp add: \u \ v\) finally have "e * (v - u)/2 < e * (v - u)/2" using uv by auto then show False by auto qed have "norm (\(x, K)\p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\(x, K)\p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" by (auto intro: sum_norm_le) also have "... \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k)/2)" using non by (fastforce intro: sum_mono) finally have I: "norm (\(x, k)\p - ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" by (simp add: sum_divide_distrib) have II: "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - (\n\p \ ?A. e * (case n of (x, k) \ Sup k - Inf k)) \ (\n\p - ?A. e * (case n of (x, k) \ Sup k - Inf k))/2" proof - have ge0: "0 \ e * (Sup k - Inf k)" if xkp: "(x, k) \ p \ ?A" for x k proof - obtain u v where uv: "k = cbox u v" by (meson Int_iff xkp p(4)) with p(2) that uv have "cbox u v \ {}" by blast then show "0 \ e * ((Sup k) - (Inf k))" unfolding uv using e by (auto simp add: field_simps) qed let ?B = "\x. {t \ p. fst t = x \ content (snd t) \ 0}" let ?C = "{t \ p. fst t \ {a, b} \ content (snd t) \ 0}" have "norm (\(x, k)\p \ {t. fst t \ {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" proof - have *: "\S f e. sum f S = sum f (p \ ?C) \ norm (sum f (p \ ?C)) \ e \ norm (sum f S) \ e" by auto have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" if "(x,K) \ p \ {t. fst t \ {a, b}} - p \ ?C" for x K proof - have xk: "(x,K) \ p" and k0: "content K = 0" using that by auto then obtain u v where uv: "K = cbox u v" using p(4) by blast then have "u = v" using xk k0 p(2) by force then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" using xk unfolding uv by auto qed have 2: "norm(\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b-a)/2" proof - have norm_le: "norm (sum f S) \ e" if "\x y. \x \ S; y \ S\ \ x = y" "\x. x \ S \ norm (f x) \ e" "e > 0" for S f and e :: real proof (cases "S = {}") case True with that show ?thesis by auto next case False then obtain x where "x \ S" by auto then have "S = {x}" using that(1) by auto then show ?thesis using \x \ S\ that(2) by auto qed have *: "p \ ?C = ?B a \ ?B b" by blast then have "norm (\(x,K)\p \ ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) = norm (\(x,K) \ ?B a \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" by simp also have "... = norm ((\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) + (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" apply (subst sum.union_disjoint) using p(1) ab e by auto also have "... \ e * (b - a) / 4 + e * (b - a) / 4" proof (rule norm_triangle_le [OF add_mono]) have pa: "\v. k = cbox a v \ a \ v" if "(a, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(a, K) \ p" "(a, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pa obtain v v' where v: "K = cbox a v" "a \ v" and v': "K' = cbox a v'" "a \ v'" by blast let ?v = "min v v'" have "box a ?v \ K \ K'" unfolding v v' by (auto simp add: mem_box) then have "interior (box a (min v v')) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have "(a + ?v)/2 \ box a ?v" using ne0 unfolding v v' content_eq_0 not_le by (auto simp add: mem_box) ultimately have "(a + ?v)/2 \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(a, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(a, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox a v" and "a \ v" using pa[OF \(a, c) \ p\] by metis then have "a \ {a..v}" "v \ b" using p(3)[OF \(a, c) \ p\] by auto moreover have "{a..v} \ ball a da" using fineD[OF \?d fine p\ \(a, c) \ p\] by (simp add: v split: if_split_asm) ultimately show ?thesis unfolding v interval_bounds_real[OF \a \ v\] box_real using da \a \ v\ by auto qed qed (use ab e in auto) next have pb: "\v. k = cbox v b \ b \ v" if "(b, k) \ p" for k using p(2) p(3) p(4) that by fastforce show "norm (\(x,K) \ ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (b - a) / 4" proof (intro norm_le; clarsimp) fix K K' assume K: "(b, K) \ p" "(b, K') \ p" and ne0: "content K \ 0" "content K' \ 0" with pb obtain v v' where v: "K = cbox v b" "v \ b" and v': "K' = cbox v' b" "v' \ b" by blast let ?v = "max v v'" have "box ?v b \ K \ K'" unfolding v v' by (auto simp: mem_box) then have "interior (box (max v v') b) \ interior K \ interior K'" using interior_Int interior_mono by blast moreover have " ((b + ?v)/2) \ box ?v b" using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) ultimately have "((b + ?v)/2) \ interior K \ interior K'" unfolding interior_open[OF open_box] by auto then show "K = K'" using p(5)[OF K] by auto next fix K assume K: "(b, K) \ p" and ne0: "content K \ 0" show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \ e * (b-a)" if "(b, c) \ p" and ne0: "content c \ 0" for c proof - obtain v where v: "c = cbox v b" and "v \ b" using \(b, c) \ p\ pb by blast then have "v \ a""b \ {v.. b}" using p(3)[OF \(b, c) \ p\] by auto moreover have "{v..b} \ ball b db" using fineD[OF \?d fine p\ \(b, c) \ p\] box_real(2) v False by force ultimately show ?thesis using db v by auto qed qed (use ab e in auto) qed also have "... = e * (b-a)/2" by simp finally show "norm (\(x,k)\p \ ?C. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \ e * (b-a)/2" . qed show "norm (\(x, k)\p \ ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \ e * (b-a)/2" apply (rule * [OF sum.mono_neutral_right[OF pA(2)]]) using 1 2 by (auto simp: split_paired_all) qed also have "... = (\n\p. e * (case n of (x, k) \ Sup k - Inf k))/2" unfolding sum_distrib_left[symmetric] apply (subst additive_tagged_division_1[OF \a \ b\ ptag]) by auto finally have norm_le: "norm (\(x,K)\p \ {t. fst t \ {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ (\n\p. e * (case n of (x, K) \ Sup K - Inf K))/2" . have le2: "\x s1 s2::real. 0 \ s1 \ x \ (s1 + s2)/2 \ x - s1 \ s2/2" by auto show ?thesis apply (rule le2 [OF sum_nonneg]) using ge0 apply force unfolding sum.union_disjoint[OF pA(2-), symmetric] pA(1)[symmetric] by (metis norm_le) qed note * = additive_tagged_division_1[OF assms(1) ptag, symmetric] have "norm (\(x,K)\p \ ?A \ (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \ e * (\(x,K)\p \ ?A \ (p - ?A). Sup K - Inf K)" unfolding sum_distrib_left unfolding sum.union_disjoint[OF pA(2-)] using le_xz norm_triangle_le I II by blast then show "norm ((\(x,K)\p. content K *\<^sub>R f' x) - (f b - f a)) \ e * content {a..b}" by (simp only: content_real[OF \a \ b\] *[of "\x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric]) qed qed qed subsection \Stronger form with finite number of exceptional points\ lemma fundamental_theorem_of_calculus_interior_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..continuous_on {a..b} f\ \a < x\ \x < b\ continuous_on_subset by (force simp: intro!: insert)+ then have "(f' has_integral f x - f a + (f b - f x)) {a..b}" using \a < x\ \x < b\ has_integral_combine less_imp_le by blast then show ?thesis by simp qed qed corollary fundamental_theorem_of_calculus_strong: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" and vec: "\x. x \ {a..b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a..b} f" shows "(f' has_integral (f b - f a)) {a..b}" by (rule fundamental_theorem_of_calculus_interior_strong [OF \finite S\]) (force simp: assms)+ proposition indefinite_integral_continuous_left: fixes f:: "real \ 'a::banach" assumes intf: "f integrable_on {a..b}" and "a < c" "c \ b" "e > 0" obtains d where "d > 0" and "\t. c - d < t \ t \ c \ norm (integral {a..c} f - integral {a..t} f) < e" proof - obtain w where "w > 0" and w: "\t. \c - w < t; t < c\ \ norm (f c) * norm(c - t) < e/3" proof (cases "f c = 0") case False hence e3: "0 < e/3 / norm (f c)" using \e>0\ by simp moreover have "norm (f c) * norm (c - t) < e/3" if "t < c" and "c - e/3 / norm (f c) < t" for t proof - have "norm (c - t) < e/3 / norm (f c)" using that by auto then show "norm (f c) * norm (c - t) < e/3" by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff) qed ultimately show ?thesis using that by auto next case True then show ?thesis using \e > 0\ that by auto qed let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" have e3: "e/3 > 0" using \e > 0\ by auto have "f integrable_on {a..c}" apply (rule integrable_subinterval_real[OF intf]) using \a < c\ \c \ b\ by auto then obtain d1 where "gauge d1" and d1: "\p. \p tagged_division_of {a..c}; d1 fine p\ \ norm (?SUM p - integral {a..c} f) < e/3" using integrable_integral has_integral_real e3 by metis define d where [abs_def]: "d x = ball x w \ d1 x" for x have "gauge d" unfolding d_def using \w > 0\ \gauge d1\ by auto then obtain k where "0 < k" and k: "ball c k \ d c" by (meson gauge_def open_contains_ball) let ?d = "min k (c - a)/2" show thesis proof (intro that[of ?d] allI impI, safe) show "?d > 0" using \0 < k\ \a < c\ by auto next fix t assume t: "c - ?d < t" "t \ c" show "norm (integral ({a..c}) f - integral ({a..t}) f) < e" proof (cases "t < c") case False with \t \ c\ show ?thesis by (simp add: \e > 0\) next case True have "f integrable_on {a..t}" apply (rule integrable_subinterval_real[OF intf]) using \t < c\ \c \ b\ by auto then obtain d2 where d2: "gauge d2" "\p. p tagged_division_of {a..t} \ d2 fine p \ norm (?SUM p - integral {a..t} f) < e/3" using integrable_integral has_integral_real e3 by metis define d3 where "d3 x = (if x \ t then d1 x \ d2 x else d1 x)" for x have "gauge d3" using \gauge d1\ \gauge d2\ unfolding d3_def gauge_def by auto then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p" by (metis box_real(2) fine_division_exists) note p' = tagged_division_ofD[OF ptag] have pt: "(x,K)\p \ x \ t" for x K by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE) with pfine have "d2 fine p" unfolding fine_def d3_def by fastforce then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3" using d2(2) ptag by auto have eqs: "{a..c} \ {x. x \ t} = {a..t}" "{a..c} \ {x. x \ t} = {t..c}" using t by (auto simp add: field_simps) have "p \ {(c, {t..c})} tagged_division_of {a..c}" apply (rule tagged_division_Un_interval_real[of _ _ _ 1 "t"]) using \t \ c\ by (auto simp: eqs ptag tagged_division_of_self_real) moreover have "d1 fine p \ {(c, {t..c})}" unfolding fine_def proof safe fix x K y assume "(x,K) \ p" and "y \ K" then show "y \ d1 x" by (metis Int_iff d3_def subsetD fineD pfine) next fix x assume "x \ {t..c}" then have "dist c x < k" using t(1) by (auto simp add: field_simps dist_real_def) with k show "x \ d1 c" unfolding d_def by auto qed ultimately have d1_fin: "norm (?SUM(p \ {(c, {t..c})}) - integral {a..c} f) < e/3" using d1 by metis have SUMEQ: "?SUM(p \ {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p" proof - have "?SUM(p \ {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p" proof (subst sum.union_disjoint) show "p \ {(c, {t..c})} = {}" using \t < c\ pt by force qed (use p'(1) in auto) also have "... = (c - t) *\<^sub>R f c + ?SUM p" using \t \ c\ by auto finally show ?thesis . qed have "c - k < t" using \k>0\ t(1) by (auto simp add: field_simps) moreover have "k \ w" proof (rule ccontr) assume "\ k \ w" then have "c + (k + w) / 2 \ d c" by (auto simp add: field_simps not_le not_less dist_real_def d_def) then have "c + (k + w) / 2 \ ball c k" using k by blast then show False using \0 < w\ \\ k \ w\ dist_real_def by auto qed ultimately have cwt: "c - w < t" by (auto simp add: field_simps) have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) - integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c" by auto have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3" unfolding eq proof (intro norm_triangle_lt add_strict_mono) show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3" by (metis SUMEQ d1_fin norm_minus_cancel) show "norm (?SUM p - integral {a..t} f) < e/3" using d2_fin by blast show "norm ((c - t) *\<^sub>R f c) < e/3" using w cwt \t < c\ by simp (simp add: field_simps) qed then show ?thesis by simp qed qed qed lemma indefinite_integral_continuous_right: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" and "a \ c" and "c < b" and "e > 0" obtains d where "0 < d" and "\t. c \ t \ t < c + d \ norm (integral {a..c} f - integral {a..t} f) < e" proof - have intm: "(\x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \ - a" using assms by auto from indefinite_integral_continuous_left[OF intm \e>0\] obtain d where "0 < d" and d: "\t. \- c - d < t; t \ -c\ \ norm (integral {- b..- c} (\x. f (-x)) - integral {- b..t} (\x. f (-x))) < e" by metis let ?d = "min d (b - c)" show ?thesis proof (intro that[of "?d"] allI impI) show "0 < ?d" using \0 < d\ \c < b\ by auto fix t :: real assume t: "c \ t \ t < c + ?d" have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f" "integral {a..t} f = integral {a..b} f - integral {t..b} f" apply (simp_all only: algebra_simps) using assms t by (auto simp: integral_combine) have "(- c) - d < (- t)" "- t \ - c" using t by auto from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e" by (auto simp add: algebra_simps norm_minus_commute *) qed qed lemma indefinite_integral_continuous_1: fixes f :: "real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {a..x} f)" proof - have "\d>0. \x'\{a..b}. dist x' x < d \ dist (integral {a..x'} f) (integral {a..x} f) < e" if x: "x \ {a..b}" and "e > 0" for x e :: real proof (cases "a = b") case True with that show ?thesis by force next case False with x have "a < b" by force with x consider "x = a" | "x = b" | "a < x" "x < b" by force then show ?thesis proof cases case 1 show ?thesis apply (rule indefinite_integral_continuous_right [OF assms _ \a < b\ \e > 0\], force) using \x = a\ apply (force simp: dist_norm algebra_simps) done next case 2 show ?thesis apply (rule indefinite_integral_continuous_left [OF assms \a < b\ _ \e > 0\], force) using \x = b\ apply (force simp: dist_norm norm_minus_commute algebra_simps) done next case 3 obtain d1 where "0 < d1" and d1: "\t. \x - d1 < t; t \ x\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \a < x\ _ \e > 0\]) obtain d2 where "0 < d2" and d2: "\t. \x \ t; t < x + d2\ \ norm (integral {a..x} f - integral {a..t} f) < e" using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \x < b\ \e > 0\]) show ?thesis proof (intro exI ballI conjI impI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by simp show "dist (integral {a..y} f) (integral {a..x} f) < e" if "y \ {a..b}" "dist y x < min d1 d2" for y proof (cases "y < x") case True with that d1 show ?thesis by (auto simp: dist_commute dist_norm) next case False with that d2 show ?thesis by (auto simp: dist_commute dist_norm) qed qed qed qed then show ?thesis by (auto simp: continuous_on_iff) qed lemma indefinite_integral_continuous_1': fixes f::"real \ 'a::banach" assumes "f integrable_on {a..b}" shows "continuous_on {a..b} (\x. integral {x..b} f)" proof - have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \ {a..b}" for x using integral_combine[OF _ _ assms, of x] that by (auto simp: algebra_simps) with _ show ?thesis by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms) qed theorem integral_has_vector_derivative': fixes f :: "real \ 'b::banach" assumes "continuous_on {a..b} f" and "x \ {a..b}" shows "((\u. integral {u..b} f) has_vector_derivative - f x) (at x within {a..b})" proof - have *: "integral {x..b} f = integral {a .. b} f - integral {a .. x} f" if "a \ x" "x \ b" for x using integral_combine[of a x b for x, OF that integrable_continuous_real[OF assms(1)]] by (simp add: algebra_simps) show ?thesis using \x \ _\ * by (rule has_vector_derivative_transform) (auto intro!: derivative_eq_intros assms integral_has_vector_derivative) qed lemma integral_has_real_derivative': assumes "continuous_on {a..b} g" assumes "t \ {a..b}" shows "((\x. integral {x..b} g) has_real_derivative -g t) (at t within {a..b})" using integral_has_vector_derivative'[OF assms] by (auto simp: has_field_derivative_iff_has_vector_derivative) subsection \This doesn't directly involve integration, but that gives an easy proof\ lemma has_derivative_zero_unique_strong_interval: fixes f :: "real \ 'a::banach" assumes "finite k" and contf: "continuous_on {a..b} f" and "f a = y" and fder: "\x. x \ {a..b} - k \ (f has_derivative (\h. 0)) (at x within {a..b})" and x: "x \ {a..b}" shows "f x = y" proof - have "a \ b" "a \ x" using assms by auto have "((\x. 0::'a) has_integral f x - f a) {a..x}" proof (rule fundamental_theorem_of_calculus_interior_strong[OF \finite k\ \a \ x\]; clarify?) have "{a..x} \ {a..b}" using x by auto then show "continuous_on {a..x} f" by (rule continuous_on_subset[OF contf]) show "(f has_vector_derivative 0) (at z)" if z: "z \ {a<.. k" for z unfolding has_vector_derivative_def proof (simp add: at_within_open[OF z, symmetric]) show "(f has_derivative (\x. 0)) (at z within {a<..Generalize a bit to any convex set\ lemma has_derivative_zero_unique_strong_convex: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "convex S" "finite K" and contf: "continuous_on S f" and "c \ S" "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof (cases "x = c") case True with \f c = y\ show ?thesis by blast next case False let ?\ = "\u. (1 - u) *\<^sub>R c + u *\<^sub>R x" have contf': "continuous_on {0 ..1} (f \ ?\)" proof (rule continuous_intros continuous_on_subset[OF contf])+ show "(\u. (1 - u) *\<^sub>R c + u *\<^sub>R x) ` {0..1} \ S" using \convex S\ \x \ S\ \c \ S\ by (auto simp add: convex_alt algebra_simps) qed have "t = u" if "?\ t = ?\ u" for t u proof - from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c" by (auto simp add: algebra_simps) then show ?thesis using \x \ c\ by auto qed then have eq: "(SOME t. ?\ t = ?\ u) = u" for u by blast then have "(?\ -` K) \ (\z. SOME t. ?\ t = z) ` K" by (clarsimp simp: image_iff) (metis (no_types) eq) then have fin: "finite (?\ -` K)" by (rule finite_surj[OF \finite K\]) have derf': "((\u. f (?\ u)) has_derivative (\h. 0)) (at t within {0..1})" if "t \ {0..1} - {t. ?\ t \ K}" for t proof - have df: "(f has_derivative (\h. 0)) (at (?\ t) within ?\ ` {0..1})" apply (rule has_derivative_within_subset [OF derf]) using \convex S\ \x \ S\ \c \ S\ that by (auto simp add: convex_alt algebra_simps) have "(f \ ?\ has_derivative (\x. 0) \ (\z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})" by (rule derivative_eq_intros df | simp)+ then show ?thesis unfolding o_def . qed have "(f \ ?\) 1 = y" apply (rule has_derivative_zero_unique_strong_interval[OF fin contf']) unfolding o_def using \f c = y\ derf' by auto then show ?thesis by auto qed text \Also to any open connected set with finite set of exceptions. Could generalize to locally convex set with limpt-free set of exceptions.\ lemma has_derivative_zero_unique_strong_connected: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite K" and contf: "continuous_on S f" and "c \ S" and "f c = y" and derf: "\x. x \ S - K \ (f has_derivative (\h. 0)) (at x within S)" and "x \ S" shows "f x = y" proof - have "\e>0. ball x e \ (S \ f -` {f x})" if "x \ S" for x proof - obtain e where "0 < e" and e: "ball x e \ S" using \x \ S\ \open S\ open_contains_ball by blast have "ball x e \ {u \ S. f u \ {f x}}" proof safe fix y assume y: "y \ ball x e" then show "y \ S" using e by auto show "f y = f x" proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \finite K\]) show "continuous_on (ball x e) f" using contf continuous_on_subset e by blast show "(f has_derivative (\h. 0)) (at u within ball x e)" if "u \ ball x e - K" for u by (metis Diff_iff contra_subsetD derf e has_derivative_within_subset that) qed (use y e \0 < e\ in auto) qed then show "\e>0. ball x e \ (S \ f -` {f x})" using \0 < e\ by blast qed then have "openin (top_of_set S) (S \ f -` {y})" by (auto intro!: open_openin_trans[OF \open S\] simp: open_contains_ball) moreover have "closedin (top_of_set S) (S \ f -` {y})" by (force intro!: continuous_closedin_preimage [OF contf]) ultimately have "(S \ f -` {y}) = {} \ (S \ f -` {y}) = S" using \connected S\ by (simp add: connected_clopen) then show ?thesis using \x \ S\ \f c = y\ \c \ S\ by auto qed lemma has_derivative_zero_connected_constant: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "connected S" and "open S" and "finite k" and "continuous_on S f" and "\x\(S - k). (f has_derivative (\h. 0)) (at x within S)" obtains c where "\x. x \ S \ f(x) = c" proof (cases "S = {}") case True then show ?thesis by (metis empty_iff that) next case False then obtain c where "c \ S" by (metis equals0I) then show ?thesis by (metis has_derivative_zero_unique_strong_connected assms that) qed lemma DERIV_zero_connected_constant: fixes f :: "'a::{real_normed_field,euclidean_space} \ 'a" assumes "connected S" and "open S" and "finite K" and "continuous_on S f" and "\x\(S - K). DERIV f x :> 0" obtains c where "\x. x \ S \ f(x) = c" using has_derivative_zero_connected_constant [OF assms(1-4)] assms by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def) subsection \Integrating characteristic function of an interval\ lemma has_integral_restrict_open_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes intf: "(f has_integral i) (cbox c d)" and cb: "cbox c d \ cbox a b" shows "((\x. if x \ box c d then f x else 0) has_integral i) (cbox a b)" proof (cases "cbox c d = {}") case True then have "box c d = {}" by (metis bot.extremum_uniqueI box_subset_cbox) then show ?thesis using True intf by auto next case False then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \ p" using cb partial_division_extend_1 by blast define g where [abs_def]: "g x = (if x \box c d then f x else 0)" for x interpret operative "lift_option plus" "Some (0 :: 'b)" "\i. if g integrable_on i then Some (integral i g) else None" by (fact operative_integralI) note operat = division [OF pdiv, symmetric] show ?thesis proof (cases "(g has_integral i) (cbox a b)") case True then show ?thesis by (simp add: g_def) next case False have iterate:"F (\i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0" proof (intro neutral ballI) fix x assume x: "x \ p - {cbox c d}" then have "x \ p" by auto then obtain u v where uv: "x = cbox u v" using pdiv by blast have "interior x \ interior (cbox c d) = {}" using pdiv inp x by blast then have "(g has_integral 0) x" unfolding uv using has_integral_spike_interior[where f="\x. 0"] by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox) then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" by auto qed interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)" by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms) have intg: "g integrable_on cbox c d" using integrable_spike_interior[where f=f] by (meson g_def has_integral_integrable intf) moreover have "integral (cbox c d) g = i" proof (rule has_integral_unique[OF has_integral_spike_interior intf]) show "\x. x \ box c d \ f x = g x" by (auto simp: g_def) show "(g has_integral integral (cbox c d) g) (cbox c d)" by (rule integrable_integral[OF intg]) qed ultimately have "F (\A. if g integrable_on A then Some (integral A g) else None) p = Some i" by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) then have "(g has_integral i) (cbox a b)" by (metis integrable_on_def integral_unique operat option.inject option.simps(3)) with False show ?thesis by blast qed qed lemma has_integral_restrict_closed_subinterval: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) (cbox c d)" and "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)" proof - note has_integral_restrict_open_subinterval[OF assms] note * = has_integral_spike[OF negligible_frontier_interval _ this] show ?thesis by (rule *[of c d]) (use box_subset_cbox[of c d] in auto) qed lemma has_integral_restrict_closed_subintervals_eq: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "cbox c d \ cbox a b" shows "((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b) \ (f has_integral i) (cbox c d)" (is "?l = ?r") proof (cases "cbox c d = {}") case False let ?g = "\x. if x \ cbox c d then f x else 0" show ?thesis proof assume ?l then have "?g integrable_on cbox c d" using assms has_integral_integrable integrable_subinterval by blast then have "f integrable_on cbox c d" by (rule integrable_eq) auto moreover then have "i = integral (cbox c d) f" by (meson \((\x. if x \ cbox c d then f x else 0) has_integral i) (cbox a b)\ assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral) ultimately show ?r by auto next assume ?r then show ?l by (rule has_integral_restrict_closed_subinterval[OF _ assms]) qed qed auto text \Hence we can apply the limit process uniformly to all integrals.\ lemma has_integral': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) S \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ S then f(x) else 0) has_integral z) (cbox a b) \ norm(z - i) < e))" (is "?l \ (\e>0. ?r e)") proof (cases "\a b. S = cbox a b") case False then show ?thesis by (simp add: has_integral_alt) next case True then obtain a b where S: "S = cbox a b" by blast obtain B where " 0 < B" and B: "\x. x \ cbox a b \ norm x \ B" using bounded_cbox[unfolded bounded_pos] by blast show ?thesis proof safe fix e :: real assume ?l and "e > 0" have "((\x. if x \ S then f x else 0) has_integral i) (cbox c d)" if "ball 0 (B+1) \ cbox c d" for c d unfolding S using B that by (force intro: \?l\[unfolded S] has_integral_restrict_closed_subinterval) then show "?r e" apply (rule_tac x="B+1" in exI) using \B>0\ \e>0\ by force next assume as: "\e>0. ?r e" then obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b)" by (meson zero_less_one) define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" "i \ Basis" for x i using that and Basis_le_norm[OF \i\Basis\, of x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (meson B mem_box(2) subsetI) have "c \ i \ x \ i \ x \ i \ d \ i" if x: "norm (0 - x) < C" and i: "i \ Basis" for x i using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain y where y: "(f has_integral y) (cbox a b)" using c_d has_integral_restrict_closed_subintervals_eq S by blast have "y = i" proof (rule ccontr) assume "y \ i" then have "0 < norm (y - i)" by auto from as[rule_format,OF this] obtain C where C: "\a b. ball 0 C \ cbox a b \ \z. ((\x. if x \ S then f x else 0) has_integral z) (cbox a b) \ norm (z-i) < norm (y-i)" by auto define c :: 'n where "c = (\i\Basis. (- max B C) *\<^sub>R i)" define d :: 'n where "d = (\i\Basis. max B C *\<^sub>R i)" have "c \ i \ x \ i \ x \ i \ d \ i" if "norm x \ B" and "i \ Basis" for x i using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def) then have c_d: "cbox a b \ cbox c d" by (simp add: B mem_box(2) subset_eq) have "c \ i \ x \ i \ x \ i \ d \ i" if "norm (0 - x) < C" and "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def) then have "ball 0 C \ cbox c d" by (auto simp: mem_box dist_norm) with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)" using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast moreover then have "z = y" by (blast intro: has_integral_unique[OF _ y]) ultimately show False by auto qed then show ?l using y by (auto simp: S) qed qed lemma has_integral_le: fixes f :: "'n::euclidean_space \ real" assumes fg: "(f has_integral i) S" "(g has_integral j) S" and le: "\x. x \ S \ f x \ g x" shows "i \ j" using has_integral_component_le[OF _ fg, of 1] le by auto lemma integral_le: fixes f :: "'n::euclidean_space \ real" assumes "f integrable_on S" and "g integrable_on S" and "\x. x \ S \ f x \ g x" shows "integral S f \ integral S g" by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) lemma has_integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes "(f has_integral i) S" and "\x. x \ S \ 0 \ f x" shows "0 \ i" using has_integral_component_nonneg[of 1 f i S] unfolding o_def using assms by auto lemma integral_nonneg: fixes f :: "'n::euclidean_space \ real" assumes f: "f integrable_on S" and 0: "\x. x \ S \ 0 \ f x" shows "0 \ integral S f" by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0]) text \Hence a general restriction property.\ lemma has_integral_restrict [simp]: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) S" proof - have *: "\x. (if x \ T then if x \ S then f x else 0 else 0) = (if x\S then f x else 0)" using assms by auto show ?thesis apply (subst(2) has_integral') apply (subst has_integral') apply (simp add: *) done qed corollary has_integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "((\x. if x \ s then f x else 0) has_integral i) UNIV \ (f has_integral i) s" by auto lemma has_integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "((\x. if x \ S then f x else 0) has_integral i) T \ (f has_integral i) (S \ T)" proof - have "((\x. if x \ T then if x \ S then f x else 0 else 0) has_integral i) UNIV = ((\x. if x \ S \ T then f x else 0) has_integral i) UNIV" by (rule has_integral_cong) auto then show ?thesis using has_integral_restrict_UNIV by fastforce qed lemma integral_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral T (\x. if x \ S then f x else 0) = integral (S \ T) f" by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral) lemma integrable_restrict_Int: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(\x. if x \ S then f x else 0) integrable_on T \ f integrable_on (S \ T)" using has_integral_restrict_Int by fastforce lemma has_integral_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" and "\x. x \ S \ f x = 0" and "S \ T" shows "(f has_integral i) T" proof - have "(\x. if x \ S then f x else 0) = (\x. if x \ T then f x else 0)" using assms by fastforce with f show ?thesis by (simp only: has_integral_restrict_UNIV [symmetric, of f]) qed lemma integrable_on_superset: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\x. x \ S \ f x = 0" and "S \ t" shows "f integrable_on t" using assms unfolding integrable_on_def by (auto intro:has_integral_on_superset) lemma integral_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "integral UNIV (\x. if x \ S then f x else 0) = integral S f" by (simp add: integral_restrict_Int) lemma integrable_restrict_UNIV: fixes f :: "'n::euclidean_space \ 'a::banach" shows "(\x. if x \ s then f x else 0) integrable_on UNIV \ f integrable_on s" unfolding integrable_on_def by auto lemma has_integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes k: "k \ Basis" and as: "S \ T" "(f has_integral i) S" "(f has_integral j) T" "\x. x\T \ 0 \ f(x)\k" shows "i\k \ j\k" proof - have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" by (simp_all add: assms) then show ?thesis apply (rule has_integral_component_le[OF k]) using as by auto qed subsection\Integrals on set differences\ lemma has_integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes S: "(f has_integral i) S" and T: "(f has_integral j) T" and neg: "negligible (T - S)" shows "(f has_integral (i - j)) (S - T)" proof - show ?thesis unfolding has_integral_restrict_UNIV [symmetric, of f] proof (rule has_integral_spike [OF neg]) have eq: "(\x. (if x \ S then f x else 0) - (if x \ T then f x else 0)) = (\x. if x \ T - S then - f x else if x \ S - T then f x else 0)" by (force simp add: ) have "((\x. if x \ S then f x else 0) has_integral i) UNIV" "((\x. if x \ T then f x else 0) has_integral j) UNIV" using S T has_integral_restrict_UNIV by auto from has_integral_diff [OF this] show "((\x. if x \ T - S then - f x else if x \ S - T then f x else 0) has_integral i-j) UNIV" by (simp add: eq) qed force qed lemma integral_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "f integrable_on S" "f integrable_on T" "negligible(T - S)" shows "integral (S - T) f = integral S f - integral T f" by (rule integral_unique) (simp add: assms has_integral_setdiff integrable_integral) lemma integrable_setdiff: fixes f :: "'a::euclidean_space \ 'b::banach" assumes "(f has_integral i) S" "(f has_integral j) T" "negligible (T - S)" shows "f integrable_on (S - T)" using has_integral_setdiff [OF assms] by (simp add: has_integral_iff) lemma negligible_setdiff [simp]: "T \ S \ negligible (T - S)" by (metis Diff_eq_empty_iff negligible_empty) lemma negligible_on_intervals: "negligible s \ (\a b. negligible(s \ cbox a b))" (is "?l \ ?r") proof assume ?r show ?l unfolding negligible_def proof safe fix a b show "(indicator s has_integral 0) (cbox a b)" apply (rule has_integral_negligible[OF \?r\[rule_format,of a b]]) unfolding indicator_def apply auto done qed qed (simp add: negligible_Int) lemma negligible_translation: assumes "negligible S" shows "negligible ((+) c ` S)" proof - have inj: "inj ((+) c)" by simp show ?thesis using assms proof (clarsimp simp: negligible_def) fix a b assume "\x y. (indicator S has_integral 0) (cbox x y)" then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) have eq: "indicator ((+) c ` S) = (\x. indicator S (x - c))" by (force simp add: indicator_def) show "(indicator ((+) c ` S) has_integral 0) (cbox a b)" using has_integral_affinity [OF *, of 1 "-c"] cbox_translation [of "c" "-c+a" "-c+b"] by (simp add: eq) (simp add: ac_simps) qed qed lemma negligible_translation_rev: assumes "negligible ((+) c ` S)" shows "negligible S" by (metis negligible_translation [OF assms, of "-c"] translation_galois) lemma has_integral_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "(f has_integral y) S \ (f has_integral y) T" proof - have "((\x. if x \ S then f x else 0) has_integral y) UNIV = ((\x. if x \ T then f x else 0) has_integral y) UNIV" proof (rule has_integral_spike_eq) show "negligible ({x \ S - T. f x \ 0} \ {x \ T - S. f x \ 0})" by (rule negligible_Un [OF assms]) qed auto then show ?thesis by (simp add: has_integral_restrict_UNIV) qed corollary integral_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "integral S f = integral T f" using has_integral_spike_set_eq [OF assms] by (metis eq_integralD integral_unique) lemma integrable_spike_set: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and neg: "negligible {x \ S - T. f x \ 0}" "negligible {x \ T - S. f x \ 0}" shows "f integrable_on T" using has_integral_spike_set_eq [OF neg] f by blast lemma integrable_spike_set_eq: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "negligible ((S - T) \ (T - S))" shows "f integrable_on S \ f integrable_on T" by (blast intro: integrable_spike_set assms negligible_subset) lemma integrable_on_insert_iff: "f integrable_on (insert x X) \ f integrable_on X" for f::"_ \ 'a::banach" by (rule integrable_spike_set_eq) (auto simp: insert_Diff_if) lemma has_integral_interior: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (interior S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF empty_imp_negligible negligible_subset]) (use interior_subset in \auto simp: frontier_def closure_def\) lemma has_integral_closure: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "negligible(frontier S) \ (f has_integral y) (closure S) \ (f has_integral y) S" by (rule has_integral_spike_set_eq [OF negligible_subset empty_imp_negligible]) (auto simp: closure_Un_frontier ) lemma has_integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "(f has_integral y) (box a b) \ (f has_integral y) (cbox a b)" unfolding interior_cbox [symmetric] by (metis frontier_cbox has_integral_interior negligible_frontier_interval) lemma integrable_on_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "f integrable_on box a b \ f integrable_on cbox a b" by (simp add: has_integral_open_interval integrable_on_def) lemma integral_open_interval: fixes f :: "'a :: euclidean_space \ 'b :: banach" shows "integral(box a b) f = integral(cbox a b) f" by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) subsection \More lemmas that are useful later\ lemma has_integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "(f has_integral i) s" and "(f has_integral j) t" and "\x\t. 0 \ f x" shows "i \ j" using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] using assms by auto lemma integral_subset_component_le: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "k \ Basis" and "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x \ k" shows "(integral s f)\k \ (integral t f)\k" apply (rule has_integral_subset_component_le) using assms apply auto done lemma integral_subset_le: fixes f :: "'n::euclidean_space \ real" assumes "s \ t" and "f integrable_on s" and "f integrable_on t" and "\x \ t. 0 \ f x" shows "integral s f \ integral t f" apply (rule has_integral_subset_le) using assms apply auto done lemma has_integral_alt': fixes f :: "'n::euclidean_space \ 'a::banach" shows "(f has_integral i) s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e)" (is "?l = ?r") proof assume rhs: ?r show ?l proof (subst has_integral', intro allI impI) fix e::real assume "e > 0" from rhs[THEN conjunct2,rule_format,OF this] show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e)" apply (rule ex_forward) using rhs by blast qed next let ?\ = "\e a b. \z. ((\x. if x \ s then f x else 0) has_integral z) (cbox a b) \ norm (z - i) < e" assume ?l then have lhs: "\B>0. \a b. ball 0 B \ cbox a b \ ?\ e a b" if "e > 0" for e using that has_integral'[of f] by auto let ?f = "\x. if x \ s then f x else 0" show ?r proof (intro conjI allI impI) fix a b :: 'n from lhs[OF zero_less_one] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ ?\ 1 a b" by blast let ?a = "\i\Basis. min (a\i) (-B) *\<^sub>R i::'n" let ?b = "\i\Basis. max (b\i) B *\<^sub>R i::'n" show "?f integrable_on cbox a b" proof (rule integrable_subinterval[of _ ?a ?b]) have "?a \ i \ x \ i \ x \ i \ ?b \ i" if "norm (0 - x) < B" "i \ Basis" for x i using Basis_le_norm[of i x] that by (auto simp add:field_simps) then have "ball 0 B \ cbox ?a ?b" by (auto simp: mem_box dist_norm) then show "?f integrable_on cbox ?a ?b" unfolding integrable_on_def using B by blast show "cbox a b \ cbox ?a ?b" by (force simp: mem_box) qed fix e :: real assume "e > 0" with lhs show "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - i) < e" by (metis (no_types, lifting) has_integral_integrable_integral) qed qed subsection \Continuity of the integral (for a 1-dimensional interval)\ lemma integrable_alt: fixes f :: "'n::euclidean_space \ 'a::banach" shows "f integrable_on s \ (\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e)" (is "?l = ?r") proof let ?F = "\x. if x \ s then f x else 0" assume ?l then obtain y where intF: "\a b. ?F integrable_on cbox a b" and y: "\e. 0 < e \ \B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e" unfolding integrable_on_def has_integral_alt'[of f] by auto show ?r proof (intro conjI allI impI intF) fix e::real assume "e > 0" then have "e/2 > 0" by auto obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - y) < e/2" using \0 < e/2\ y by blast show "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" proof (intro conjI exI impI allI, rule \0 < B\) fix a b c d::'n assume sub: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" using sub by (auto intro: norm_triangle_half_l dest: B) qed qed next let ?F = "\x. if x \ s then f x else 0" assume rhs: ?r let ?cube = "\n. cbox (\i\Basis. - real n *\<^sub>R i::'n) (\i\Basis. real n *\<^sub>R i)" have "Cauchy (\n. integral (?cube n) ?F)" unfolding Cauchy_def proof (intro allI impI) fix e::real assume "e > 0" with rhs obtain B where "0 < B" and B: "\a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" by blast obtain N where N: "B \ real N" using real_arch_simple by blast have "ball 0 B \ ?cube n" if n: "n \ N" for n proof - have "sum ((*\<^sub>R) (- real n)) Basis \ i \ x \ i \ x \ i \ sum ((*\<^sub>R) (real n)) Basis \ i" if "norm x < B" "i \ Basis" for x i::'n using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf) then show ?thesis by (auto simp: mem_box dist_norm) qed then show "\M. \m\M. \n\M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e" by (fastforce simp add: dist_norm intro!: B) qed then obtain i where i: "(\n. integral (?cube n) ?F) \ i" using convergent_eq_Cauchy by blast have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?F - i) < e" if "e > 0" for e proof - have *: "e/2 > 0" using that by auto then obtain N where N: "\n. N \ n \ norm (i - integral (?cube n) ?F) < e/2" using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson obtain B where "0 < B" and B: "\a b c d. \ball 0 B \ cbox a b; ball 0 B \ cbox c d\ \ norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2" using rhs * by meson let ?B = "max (real N) B" show ?thesis proof (intro exI conjI allI impI) show "0 < ?B" using \B > 0\ by auto fix a b :: 'n assume "ball 0 ?B \ cbox a b" moreover obtain n where n: "max (real N) B \ real n" using real_arch_simple by blast moreover have "ball 0 B \ ?cube n" proof fix x :: 'n assume x: "x \ ball 0 B" have "\norm (0 - x) < B; i \ Basis\ \ sum ((*\<^sub>R) (-n)) Basis \ i\ x \ i \ x \ i \ sum ((*\<^sub>R) n) Basis \ i" for i using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf) then show "x \ ?cube n" using x by (auto simp: mem_box dist_norm) qed ultimately show "norm (integral (cbox a b) ?F - i) < e" using norm_triangle_half_l [OF B N] by force qed qed then show ?l unfolding integrable_on_def has_integral_alt'[of f] using rhs by blast qed lemma integrable_altD: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on s" shows "\a b. (\x. if x \ s then f x else 0) integrable_on cbox a b" and "\e. e > 0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" using assms[unfolded integrable_alt[of f]] by auto lemma integrable_alt_subset: fixes f :: "'a::euclidean_space \ 'b::banach" shows "f integrable_on S \ (\a b. (\x. if x \ S then f x else 0) integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm(integral (cbox a b) (\x. if x \ S then f x else 0) - integral (cbox c d) (\x. if x \ S then f x else 0)) < e)" (is "_ = ?rhs") proof - let ?g = "\x. if x \ S then f x else 0" have "f integrable_on S \ (\a b. ?g integrable_on cbox a b) \ (\e>0. \B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e)" by (rule integrable_alt) also have "\ = ?rhs" proof - { fix e :: "real" assume e: "\e. e>0 \ \B>0. \a b c d. ball 0 B \ cbox a b \ cbox a b \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" and "e > 0" obtain B where "B > 0" and B: "\a b c d. \ball 0 B \ cbox a b; cbox a b \ cbox c d\ \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e/2" using \e > 0\ e [of "e/2"] by force have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" proof (intro exI allI conjI impI) fix a b c d :: "'a" let ?\ = "\i\Basis. max (a \ i) (c \ i) *\<^sub>R i" let ?\ = "\i\Basis. min (b \ i) (d \ i) *\<^sub>R i" show "norm (integral (cbox a b) ?g - integral (cbox c d) ?g) < e" if ball: "ball 0 B \ cbox a b \ ball 0 B \ cbox c d" proof - have B': "norm (integral (cbox a b \ cbox c d) ?g - integral (cbox x y) ?g) < e/2" if "cbox a b \ cbox c d \ cbox x y" for x y using B [of ?\ ?\ x y] ball that by (simp add: Int_interval [symmetric]) show ?thesis using B' [of a b] B' [of c d] norm_triangle_half_r by blast qed qed (use \B > 0\ in auto)} then show ?thesis by force qed finally show ?thesis . qed lemma integrable_on_subcbox: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on S" and sub: "cbox a b \ S" shows "f integrable_on cbox a b" proof - have "(\x. if x \ S then f x else 0) integrable_on cbox a b" by (simp add: intf integrable_altD(1)) then show ?thesis by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym) qed subsection \A straddling criterion for integrability\ lemma integrable_straddle_interval: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. (g x) \ f x \ f x \ h x)" shows "f integrable_on cbox a b" proof - have "\d. gauge d \ (\p1 p2. p1 tagged_division_of cbox a b \ d fine p1 \ p2 tagged_division_of cbox a b \ d fine p2 \ \(\(x,K)\p1. content K *\<^sub>R f x) - (\(x,K)\p2. content K *\<^sub>R f x)\ < e)" if "e > 0" for e proof - have e: "e/3 > 0" using that by auto then obtain g h i j where ij: "\i - j\ < e/3" and "(g has_integral i) (cbox a b)" and "(h has_integral j) (cbox a b)" and fgh: "\x. x \ cbox a b \ g x \ f x \ f x \ h x" using assms real_norm_def by metis then obtain d1 d2 where "gauge d1" "gauge d2" and d1: "\p. \p tagged_division_of cbox a b; d1 fine p\ \ \(\(x,K)\p. content K *\<^sub>R g x) - i\ < e/3" and d2: "\p. \p tagged_division_of cbox a b; d2 fine p\ \ \(\(x,K) \ p. content K *\<^sub>R h x) - j\ < e/3" by (metis e has_integral real_norm_def) have "\(\(x,K) \ p1. content K *\<^sub>R f x) - (\(x,K) \ p2. content K *\<^sub>R f x)\ < e" if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1" and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2 proof - have *: "\g1 g2 h1 h2 f1 f2. \\g2 - i\ < e/3; \g1 - i\ < e/3; \h2 - j\ < e/3; \h1 - j\ < e/3; g1 - h2 \ f1 - f2; f1 - f2 \ h1 - g2\ \ \f1 - f2\ < e" using \e > 0\ ij by arith have 0: "(\(x, k)\p1. content k *\<^sub>R f x) - (\(x, k)\p1. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p2. content k *\<^sub>R h x) - (\(x, k)\p2. content k *\<^sub>R f x)" "(\(x, k)\p2. content k *\<^sub>R f x) - (\(x, k)\p2. content k *\<^sub>R g x) \ 0" "0 \ (\(x, k)\p1. content k *\<^sub>R h x) - (\(x, k)\p1. content k *\<^sub>R f x)" unfolding sum_subtractf[symmetric] apply (auto intro!: sum_nonneg) apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+ done show ?thesis proof (rule *) show "\(\(x,K) \ p2. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p2 12]) show "\(\(x,K) \ p1. content K *\<^sub>R g x) - i\ < e/3" by (rule d1[OF p1 11]) show "\(\(x,K) \ p2. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p2 22]) show "\(\(x,K) \ p1. content K *\<^sub>R h x) - j\ < e/3" by (rule d2[OF p1 21]) qed (use 0 in auto) qed then show ?thesis by (rule_tac x="\x. d1 x \ d2 x" in exI) (auto simp: fine_Int intro: \gauge d1\ \gauge d2\ d1 d2) qed then show ?thesis by (simp add: integrable_Cauchy) qed lemma integrable_straddle: fixes f :: "'n::euclidean_space \ real" assumes "\e. e>0 \ \g h i j. (g has_integral i) s \ (h has_integral j) s \ \i - j\ < e \ (\x\s. g x \ f x \ f x \ h x)" shows "f integrable_on s" proof - let ?fs = "(\x. if x \ s then f x else 0)" have "?fs integrable_on cbox a b" for a b proof (rule integrable_straddle_interval) fix e::real assume "e > 0" then have *: "e/4 > 0" by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/4" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/4" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/4" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson define c where "c = (\i\Basis. min (a\i) (- (max Bg Bh)) *\<^sub>R i)" define d where "d = (\i\Basis. max (b\i) (max Bg Bh) *\<^sub>R i)" have "\norm (0 - x) < Bg; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBg: "ball 0 Bg \ cbox c d" by (auto simp: mem_box dist_norm) have "\norm (0 - x) < Bh; i \ Basis\ \ c \ i \ x \ i \ x \ i \ d \ i" for x i using Basis_le_norm[of i x] unfolding c_def d_def by auto then have ballBh: "ball 0 Bh \ cbox c d" by (auto simp: mem_box dist_norm) have ab_cd: "cbox a b \ cbox c d" by (auto simp: c_def d_def subset_box_imp) have **: "\ch cg ag ah::real. \\ah - ag\ \ \ch - cg\; \cg - i\ < e/4; \ch - j\ < e/4\ \ \ag - ah\ < e" using ij by arith show "\g h i j. (g has_integral i) (cbox a b) \ (h has_integral j) (cbox a b) \ \i - j\ < e \ (\x\cbox a b. g x \ (if x \ s then f x else 0) \ (if x \ s then f x else 0) \ h x)" proof (intro exI ballI conjI) have eq: "\x f g. (if x \ s then f x else 0) - (if x \ s then g x else 0) = (if x \ s then f x - g x else (0::real))" by auto have int_hg: "(\x. if x \ s then h x - g x else 0) integrable_on cbox a b" "(\x. if x \ s then h x - g x else 0) integrable_on cbox c d" by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+ show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)" "(?hs has_integral integral (cbox a b) ?hs) (cbox a b)" by (intro integrable_integral int_g int_h)+ then have "integral (cbox a b) ?gs \ integral (cbox a b) ?hs" apply (rule has_integral_le) using fgh by force then have "0 \ integral (cbox a b) ?hs - integral (cbox a b) ?gs" by simp then have "\integral (cbox a b) ?hs - integral (cbox a b) ?gs\ \ \integral (cbox c d) ?hs - integral (cbox c d) ?gs\" apply (simp add: integral_diff [symmetric] int_g int_h) apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]]) using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+ done then show "\integral (cbox a b) ?gs - integral (cbox a b) ?hs\ < e" apply (rule **) apply (rule Bg ballBg Bh ballBh)+ done show "\x. x \ cbox a b \ ?gs x \ ?fs x" "\x. x \ cbox a b \ ?fs x \ ?hs x" using fgh by auto qed qed then have int_f: "?fs integrable_on cbox a b" for a b by simp have "\B>0. \a b c d. ball 0 B \ cbox a b \ ball 0 B \ cbox c d \ abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e" if "0 < e" for e proof - have *: "e/3 > 0" using that by auto with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" and ij: "\i - j\ < e/3" and fgh: "\x. x \ s \ g x \ f x \ f x \ h x" by metis let ?gs = "(\x. if x \ s then g x else 0)" let ?hs = "(\x. if x \ s then h x else 0)" obtain Bg where "Bg > 0" and Bg: "\a b. ball 0 Bg \ cbox a b \ \integral (cbox a b) ?gs - i\ < e/3" and int_g: "\a b. ?gs integrable_on cbox a b" using g * unfolding has_integral_alt' real_norm_def by meson obtain Bh where "Bh > 0" and Bh: "\a b. ball 0 Bh \ cbox a b \ \integral (cbox a b) ?hs - j\ < e/3" and int_h: "\a b. ?hs integrable_on cbox a b" using h * unfolding has_integral_alt' real_norm_def by meson { fix a b c d :: 'n assume as: "ball 0 (max Bg Bh) \ cbox a b" "ball 0 (max Bg Bh) \ cbox c d" have **: "ball 0 Bg \ ball (0::'n) (max Bg Bh)" "ball 0 Bh \ ball (0::'n) (max Bg Bh)" by auto have *: "\ga gc ha hc fa fc. \\ga - i\ < e/3; \gc - i\ < e/3; \ha - j\ < e/3; \hc - j\ < e/3; ga \ fa; fa \ ha; gc \ fc; fc \ hc\ \ \fa - fc\ < e" using ij by arith have "abs (integral (cbox a b) (\x. if x \ s then f x else 0) - integral (cbox c d) (\x. if x \ s then f x else 0)) < e" proof (rule *) show "\integral (cbox a b) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox c d) ?gs - i\ < e/3" using "**" Bg as by blast show "\integral (cbox a b) ?hs - j\ < e/3" using "**" Bh as by blast show "\integral (cbox c d) ?hs - j\ < e/3" using "**" Bh as by blast qed (use int_f int_g int_h fgh in \simp_all add: integral_le\) } then show ?thesis apply (rule_tac x="max Bg Bh" in exI) using \Bg > 0\ by auto qed then show ?thesis unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f) qed subsection \Adding integrals over several sets\ lemma has_integral_Un: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "(f has_integral i) S" "(f has_integral j) T" and neg: "negligible (S \ T)" shows "(f has_integral (i + j)) (S \ T)" unfolding has_integral_restrict_UNIV[symmetric, of f] proof (rule has_integral_spike[OF neg]) let ?f = "\x. (if x \ S then f x else 0) + (if x \ T then f x else 0)" show "(?f has_integral i + j) UNIV" by (simp add: f has_integral_add) qed auto lemma integral_Un [simp]: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" "f integrable_on T" "negligible (S \ T)" shows "integral (S \ T) f = integral S f + integral T f" by (simp add: has_integral_Un assms integrable_integral integral_unique) lemma integrable_Un: fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "negligible (A \ B)" "f integrable_on A" "f integrable_on B" shows "f integrable_on (A \ B)" proof - from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" by (auto simp: integrable_on_def) from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) qed lemma integrable_Un': fixes f :: "'a::euclidean_space \ 'b :: banach" assumes "f integrable_on A" "f integrable_on B" "negligible (A \ B)" "C = A \ B" shows "f integrable_on C" using integrable_Un[of A B f] assms by simp lemma has_integral_Union: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "finite \" and int: "\S. S \ \ \ (f has_integral (i S)) S" and neg: "pairwise (\S S'. negligible (S \ S')) \" shows "(f has_integral (sum i \)) (\\)" proof - let ?\ = "((\(a,b). a \ b) ` {(a,b). a \ \ \ b \ {y. y \ \ \ a \ y}})" have "((\x. if x \ \\ then f x else 0) has_integral sum i \) UNIV" proof (rule has_integral_spike) show "negligible (\?\)" proof (rule negligible_Union) have "finite (\ \ \)" by (simp add: \) moreover have "{(a, b). a \ \ \ b \ {y \ \. a \ y}} \ \ \ \" by auto ultimately show "finite ?\" by (blast intro: finite_subset[of _ "\ \ \"]) show "\t. t \ ?\ \ negligible t" using neg unfolding pairwise_def by auto qed next show "(if x \ \\ then f x else 0) = (\A\\. if x \ A then f x else 0)" if "x \ UNIV - (\?\)" for x proof clarsimp fix S assume "S \ \" "x \ S" moreover then have "\b\\. x \ b \ b = S" using that by blast ultimately show "f x = (\A\\. if x \ A then f x else 0)" by (simp add: sum.delta[OF \]) qed next show "((\x. \A\\. if x \ A then f x else 0) has_integral (\A\\. i A)) UNIV" apply (rule has_integral_sum [OF \]) using int by (simp add: has_integral_restrict_UNIV) qed then show ?thesis using has_integral_restrict_UNIV by blast qed text \In particular adding integrals over a division, maybe not of an interval.\ lemma has_integral_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" and "\k. k \ \ \ (f has_integral (i k)) k" shows "(f has_integral (sum i \)) S" proof - note \ = division_ofD[OF assms(1)] have neg: "negligible (S \ s')" if "S \ \" "s' \ \" "S \ s'" for S s' proof - obtain a c b \ where obt: "S = cbox a b" "s' = cbox c \" by (meson \S \ \\ \s' \ \\ \(4)) from \(5)[OF that] show ?thesis unfolding obt interior_cbox by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) qed show ?thesis unfolding \(6)[symmetric] by (auto intro: \ neg assms has_integral_Union pairwiseI) qed lemma integral_combine_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "\ division_of S" "\k. k \ \ \ f integrable_on k" shows "integral S f = sum (\i. integral i f) \" apply (rule integral_unique) by (meson assms has_integral_combine_division has_integral_integrable_integral) lemma has_integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on S" and \: "\ division_of K" and "K \ S" shows "(f has_integral (sum (\i. integral i f) \)) K" proof - have "f integrable_on L" if "L \ \" for L proof - have "L \ S" using \K \ S\ \ that by blast then show "f integrable_on L" using that by (metis (no_types) f \ division_ofD(4) integrable_on_subcbox) qed then show ?thesis by (meson \ has_integral_combine_division has_integral_integrable_integral) qed lemma integral_combine_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on S" and "\ division_of S" shows "integral S f = sum (\i. integral i f) \" apply (rule integral_unique) apply (rule has_integral_combine_division_topdown) using assms apply auto done lemma integrable_combine_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of S" and f: "\i. i \ \ \ f integrable_on i" shows "f integrable_on S" using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \]) lemma integrable_on_subdivision: fixes f :: "'n::euclidean_space \ 'a::banach" assumes \: "\ division_of i" and f: "f integrable_on S" and "i \ S" shows "f integrable_on i" proof - have "f integrable_on i" if "i \ \" for i proof - have "i \ S" using assms that by auto then show "f integrable_on i" using that by (metis (no_types) \ f division_ofD(4) integrable_on_subcbox) qed then show ?thesis using \ integrable_combine_division by blast qed subsection \Also tagged divisions\ lemma has_integral_combine_tagged_division: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of S" and "\(x,k) \ p. (f has_integral (i k)) k" shows "(f has_integral (\(x,k)\p. i k)) S" proof - have *: "(f has_integral (\k\snd`p. integral k f)) S" using assms(2) apply (intro has_integral_combine_division) apply (auto simp: has_integral_integral[symmetric] intro: division_of_tagged_division[OF assms(1)]) apply auto done also have "(\k\snd`p. integral k f) = (\(x, k)\p. integral k f)" by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) (simp add: content_eq_0_interior) finally show ?thesis using assms by (auto simp add: has_integral_iff intro!: sum.cong) qed lemma integral_combine_tagged_division_bottomup: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "p tagged_division_of (cbox a b)" and "\(x,k)\p. f integrable_on k" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" apply (rule integral_unique) apply (rule has_integral_combine_tagged_division[OF assms(1)]) using assms(2) apply auto done lemma has_integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes f: "f integrable_on cbox a b" and p: "p tagged_division_of (cbox a b)" shows "(f has_integral (sum (\(x,K). integral K f) p)) (cbox a b)" proof - have "(f has_integral integral K f) K" if "(x,K) \ p" for x K by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that) then show ?thesis by (metis assms case_prodI2 has_integral_integrable_integral integral_combine_tagged_division_bottomup) qed lemma integral_combine_tagged_division_topdown: fixes f :: "'n::euclidean_space \ 'a::banach" assumes "f integrable_on cbox a b" and "p tagged_division_of (cbox a b)" shows "integral (cbox a b) f = sum (\(x,k). integral k f) p" apply (rule integral_unique [OF has_integral_combine_tagged_division_topdown]) using assms apply auto done subsection \Henstock's lemma\ lemma Henstock_lemma_part1: fixes f :: "'n::euclidean_space \ 'a::banach" assumes intf: "f integrable_on cbox a b" and "e > 0" and "gauge d" and less_e: "\p. \p tagged_division_of (cbox a b); d fine p\ \ norm (sum (\(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e" and p: "p tagged_partial_division_of (cbox a b)" "d fine p" shows "norm (sum (\(x,K). content K *\<^sub>R f x - integral K f) p) \ e" (is "?lhs \ e") proof (rule field_le_epsilon) fix k :: real assume "k > 0" let ?SUM = "\p. (\(x,K) \ p. content K *\<^sub>R f x)" note p' = tagged_partial_division_ofD[OF p(1)] have "\(snd ` p) \ cbox a b" using p'(3) by fastforce then obtain q where q: "snd ` p \ q" and qdiv: "q division_of cbox a b" by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division) note q' = division_ofD[OF qdiv] define r where "r = q - snd ` p" have "snd ` p \ r = {}" unfolding r_def by auto have "finite r" using q' unfolding r_def by auto have "\p. p tagged_division_of i \ d fine p \ norm (?SUM p - integral i f) < k / (real (card r) + 1)" if "i\r" for i proof - have gt0: "k / (real (card r) + 1) > 0" using \k > 0\ by simp have i: "i \ q" using that unfolding r_def by auto then obtain u v where uv: "i = cbox u v" using q'(4) by blast then have "cbox u v \ cbox a b" using i q'(2) by auto then have "f integrable_on cbox u v" by (rule integrable_subinterval[OF intf]) with integrable_integral[OF this, unfolded has_integral[of f]] obtain dd where "gauge dd" and dd: "\\. \\ tagged_division_of cbox u v; dd fine \\ \ norm (?SUM \ - integral (cbox u v) f) < k / (real (card r) + 1)" using gt0 by auto with gauge_Int[OF \gauge d\ \gauge dd\] obtain qq where qq: "qq tagged_division_of cbox u v" "(\x. d x \ dd x) fine qq" using fine_division_exists by blast with dd[of qq] show ?thesis by (auto simp: fine_Int uv) qed then obtain qq where qq: "\i. i \ r \ qq i tagged_division_of i \ d fine qq i \ norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)" by metis let ?p = "p \ \(qq ` r)" have "norm (?SUM ?p - integral (cbox a b) f) < e" proof (rule less_e) show "d fine ?p" by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2)) note ptag = tagged_partial_division_of_Union_self[OF p(1)] have "p \ \(qq ` r) tagged_division_of \(snd ` p) \ \r" proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \finite r\]]) show "\i. i \ r \ qq i tagged_division_of i" using qq by auto show "\i1 i2. \i1 \ r; i2 \ r; i1 \ i2\ \ interior i1 \ interior i2 = {}" by (simp add: q'(5) r_def) show "interior (\(snd ` p)) \ interior (\r) = {}" proof (rule Int_interior_Union_intervals [OF \finite r\]) show "open (interior (\(snd ` p)))" by blast show "\T. T \ r \ \a b. T = cbox a b" by (simp add: q'(4) r_def) have "finite (snd ` p)" by (simp add: p'(1)) then show "\T. T \ r \ interior (\(snd ` p)) \ interior T = {}" apply (subst Int_commute) apply (rule Int_interior_Union_intervals) using r_def q'(5) q(1) apply auto by (simp add: p'(4)) qed qed moreover have "\(snd ` p) \ \r = cbox a b" and "{qq i |i. i \ r} = qq ` r" using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto ultimately show "?p tagged_division_of (cbox a b)" by fastforce qed then have "norm (?SUM p + (?SUM (\(qq ` r))) - integral (cbox a b) f) < e" proof (subst sum.union_inter_neutral[symmetric, OF \finite p\], safe) show "content L *\<^sub>R f x = 0" if "(x, L) \ p" "(x, L) \ qq K" "K \ r" for x K L proof - obtain u v where uv: "L = cbox u v" using \(x,L) \ p\ p'(4) by blast have "L \ K" using qq[OF that(3)] tagged_division_ofD(3) \(x,L) \ qq K\ by metis have "L \ snd ` p" using \(x,L) \ p\ image_iff by fastforce then have "L \ q" "K \ q" "L \ K" using that(1,3) q(1) unfolding r_def by auto with q'(5) have "interior L = {}" using interior_mono[OF \L \ K\] by blast then show "content L *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed show "finite (\(qq ` r))" by (meson finite_UN qq \finite r\ tagged_division_of_finite) qed moreover have "content M *\<^sub>R f x = 0" if x: "(x,M) \ qq K" "(x,M) \ qq L" and KL: "qq K \ qq L" and r: "K \ r" "L \ r" for x M K L proof - note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] obtain u v where uv: "M = cbox u v" using \(x, M) \ qq L\ \L \ r\ kl(2) by blast have empty: "interior (K \ L) = {}" by (metis DiffD1 interior_Int q'(5) r_def KL r) have "interior M = {}" by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r) then show "content M *\<^sub>R f x = 0" unfolding uv content_eq_0_interior[symmetric] by auto qed ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" apply (subst (asm) sum.Union_comp) using qq by (force simp: split_paired_all)+ moreover have "content M *\<^sub>R f x = 0" if "K \ r" "L \ r" "K \ L" "qq K = qq L" "(x, M) \ qq K" for K L x M using tagged_division_ofD(6) qq that by (metis (no_types, lifting)) ultimately have less_e: "norm (?SUM p + sum (?SUM \ qq) r - integral (cbox a b) f) < e" apply (subst (asm) sum.reindex_nontrivial [OF \finite r\]) apply (auto simp: split_paired_all sum.neutral) done have norm_le: "norm (cp - ip) \ e + k" if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" for ir ip i cr cp::'a proof - from that show ?thesis using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] unfolding that(3)[symmetric] norm_minus_cancel by (auto simp add: algebra_simps) qed have "?lhs = norm (?SUM p - (\(x, k)\p. integral k f))" unfolding split_def sum_subtractf .. also have "\ \ e + k" proof (rule norm_le[OF less_e]) have lessk: "k * real (card r) / (1 + real (card r)) < k" using \k>0\ by (auto simp add: field_simps) have "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) \ (\x\r. k / (real (card r) + 1))" unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le) also have "... < k" by (simp add: lessk add.commute mult.commute) finally show "norm (sum (?SUM \ qq) r - (\k\r. integral k f)) < k" . next from q(1) have [simp]: "snd ` p \ q = q" by auto have "integral l f = 0" if inp: "(x, l) \ p" "(y, m) \ p" and ne: "(x, l) \ (y, m)" and "l = m" for x l y m proof - obtain u v where uv: "l = cbox u v" using inp p'(4) by blast have "content (cbox u v) = 0" unfolding content_eq_0_interior using that p(1) uv by auto then show ?thesis using uv by blast qed then have "(\(x, k)\p. integral k f) = (\k\snd ` p. integral k f)" apply (subst sum.reindex_nontrivial [OF \finite p\]) unfolding split_paired_all split_def by auto then show "(\(x, k)\p. integral k f) + (\k\r. integral k f) = integral (cbox a b) f" unfolding integral_combine_division_topdown[OF intf qdiv] r_def using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric] by simp qed finally show "?lhs \ e + k" . qed lemma Henstock_lemma_part2: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d" and less_e: "\\. \\ tagged_division_of (cbox a b); d fine \\ \ norm (sum (\(x,k). content k *\<^sub>R f x) \ - integral (cbox a b) f) < e" and tag: "p tagged_partial_division_of (cbox a b)" and "d fine p" shows "sum (\(x,k). norm (content k *\<^sub>R f x - integral k f)) p \ 2 * real (DIM('n)) * e" proof - have "finite p" using tag by blast then show ?thesis unfolding split_def proof (rule sum_norm_allsubsets_bound) fix Q assume Q: "Q \ p" then have fine: "d fine Q" by (simp add: \d fine p\ fine_subset) show "norm (\x\Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \ e" apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def]) using Q tag tagged_partial_division_subset apply (force simp add: fine)+ done qed qed lemma Henstock_lemma: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes intf: "f integrable_on cbox a b" and "e > 0" obtains \ where "gauge \" and "\p. \p tagged_partial_division_of (cbox a b); \ fine p\ \ sum (\(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" proof - have *: "e/(2 * (real DIM('n) + 1)) > 0" using \e > 0\ by simp with integrable_integral[OF intf, unfolded has_integral] obtain \ where "gauge \" and \: "\\. \\ tagged_division_of cbox a b; \ fine \\ \ norm ((\(x,K)\\. content K *\<^sub>R f x) - integral (cbox a b) f) < e/(2 * (real DIM('n) + 1))" by metis show thesis proof (rule that [OF \gauge \\]) fix p assume p: "p tagged_partial_division_of cbox a b" "\ fine p" have "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) \ 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))" using Henstock_lemma_part2[OF intf * \gauge \\ \ p] by metis also have "... < e" using \e > 0\ by (auto simp add: field_simps) finally show "(\(x,K)\p. norm (content K *\<^sub>R f x - integral K f)) < e" . qed qed subsection \Monotone convergence (bounded interval first)\ lemma bounded_increasing_convergent: fixes f :: "nat \ real" shows "\bounded (range f); \n. f n \ f (Suc n)\ \ \l. f \ l" using Bseq_mono_convergent[of f] incseq_Suc_iff[of f] by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) lemma monotone_convergence_interval: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on cbox a b" and le: "\k x. x \ cbox a b \ (f k x) \ f (Suc k) x" and fg: "\x. x \ cbox a b \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral (cbox a b) (f k)))" shows "g integrable_on cbox a b \ ((\k. integral (cbox a b) (f k)) \ integral (cbox a b) g) sequentially" proof (cases "content (cbox a b) = 0") case True then show ?thesis by auto next case False have fg1: "(f k x) \ (g x)" if x: "x \ cbox a b" for x k proof - have "\\<^sub>F j in sequentially. f k x \ f j x" apply (rule eventually_sequentiallyI [of k]) using le x apply (force intro: transitive_stepwise_le) done then show "f k x \ g x" using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast qed have int_inc: "\n. integral (cbox a b) (f n) \ integral (cbox a b) (f (Suc n))" by (metis integral_le intf le) then obtain i where i: "(\k. integral (cbox a b) (f k)) \ i" using bounded_increasing_convergent bou by blast have "\k. \\<^sub>F x in sequentially. integral (cbox a b) (f k) \ integral (cbox a b) (f x)" unfolding eventually_sequentially by (force intro: transitive_stepwise_le int_inc) then have i': "\k. (integral(cbox a b) (f k)) \ i" using tendsto_le [OF trivial_limit_sequentially i] by blast have "(g has_integral i) (cbox a b)" unfolding has_integral real_norm_def proof clarify fix e::real assume e: "e > 0" have "\k. (\\. gauge \ \ (\\. \ tagged_division_of (cbox a b) \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" using intf e by (auto simp: has_integral_integral has_integral) then obtain c where c: "\x. gauge (c x)" "\x \. \\ tagged_division_of cbox a b; c x fine \\ \ abs ((\(u,K)\\. content K *\<^sub>R f x u) - integral (cbox a b) (f x)) < e/2 ^ (x + 2)" by metis have "\r. \k\r. 0 \ i - (integral (cbox a b) (f k)) \ i - (integral (cbox a b) (f k)) < e/4" proof - have "e/4 > 0" using e by auto show ?thesis using LIMSEQ_D [OF i \e/4 > 0\] i' by auto qed then obtain r where r: "\k. r \ k \ 0 \ i - integral (cbox a b) (f k)" "\k. r \ k \ i - integral (cbox a b) (f k) < e/4" by metis have "\n\r. \k\n. 0 \ (g x) - (f k x) \ (g x) - (f k x) < e/(4 * content(cbox a b))" if "x \ cbox a b" for x proof - have "e/(4 * content (cbox a b)) > 0" by (simp add: False content_lt_nz e) with fg that LIMSEQ_D obtain N where "\n\N. norm (f n x - g x) < e/(4 * content (cbox a b))" by metis then show "\n\r. \k\n. 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" apply (rule_tac x="N + r" in exI) using fg1[OF that] apply (auto simp add: field_simps) done qed then obtain m where r_le_m: "\x. x \ cbox a b \ r \ m x" and m: "\x k. \x \ cbox a b; m x \ k\ \ 0 \ g x - f k x \ g x - f k x < e/(4 * content (cbox a b))" by metis define d where "d x = c (m x) x" for x show "\\. gauge \ \ (\\. \ tagged_division_of cbox a b \ \ fine \ \ abs ((\(x,K)\\. content K *\<^sub>R g x) - i) < e)" proof (rule exI, safe) show "gauge d" using c(1) unfolding gauge_def d_def by auto next fix \ assume ptag: "\ tagged_division_of (cbox a b)" and "d fine \" note p'=tagged_division_ofD[OF ptag] obtain s where s: "\x. x \ \ \ m (fst x) \ s" by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) have *: "\a - d\ < e" if "\a - b\ \ e/4" "\b - c\ < e/2" "\c - d\ < e/4" for a b c d using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"] norm_triangle_lt[of "a - b + (b - c)" "c - d" e] by (auto simp add: algebra_simps) show "\(\(x, k)\\. content k *\<^sub>R g x) - i\ < e" proof (rule *) have "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ (\i\\. \(case i of (x, K) \ content K *\<^sub>R g x) - (case i of (x, K) \ content K *\<^sub>R f (m x) x)\)" by (metis (mono_tags) sum_subtractf sum_abs) also have "... \ (\(x, k)\\. content k * (e/(4 * content (cbox a b))))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xk: "(x,K) \ \" with ptag have x: "x \ cbox a b" by blast then have "abs (content K * (g x - f (m x) x)) \ content K * (e/(4 * content (cbox a b)))" by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl) then show "\content K * g x - content K * f (m x) x\ \ content K * e/(4 * content (cbox a b))" by (simp add: algebra_simps) qed also have "... = (e/(4 * content (cbox a b))) * (\(x, k)\\. content k)" by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute) also have "... \ e/4" by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left) finally show "\(\(x,K)\\. content K *\<^sub>R g x) - (\(x,K)\\. content K *\<^sub>R f (m x) x)\ \ e/4" . next have "norm ((\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))) \ norm (\j = 0..s. \(x,K)\{xk \ \. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))" apply (subst sum.group) using s by (auto simp: sum_subtractf split_def p'(1)) also have "\ < e/2" proof - have "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) \ (\i = 0..s. e/2 ^ (i + 2))" proof (rule sum_norm_le) fix t assume "t \ {0..s}" have "norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) = norm (\(x,k)\{xk \ \. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))" by (force intro!: sum.cong arg_cong[where f=norm]) also have "... \ e/2 ^ (t + 2)" proof (rule Henstock_lemma_part1 [OF intf]) show "{xk \ \. m (fst xk) = t} tagged_partial_division_of cbox a b" apply (rule tagged_partial_division_subset[of \]) using ptag by (auto simp: tagged_division_of_def) show "c t fine {xk \ \. m (fst xk) = t}" using \d fine \\ by (auto simp: fine_def d_def) qed (use c e in auto) finally show "norm (\(x,K)\{xk \ \. m (fst xk) = t}. content K *\<^sub>R f (m x) x - integral K (f (m x))) \ e/2 ^ (t + 2)" . qed also have "... = (e/2/2) * (\i = 0..s. (1/2) ^ i)" by (simp add: sum_distrib_left field_simps) also have "\ < e/2" by (simp add: sum_gp mult_strict_left_mono[OF _ e]) finally show "norm (\j = 0..s. \(x, k)\{xk \ \. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" . qed finally show "\(\(x,K)\\. content K *\<^sub>R f (m x) x) - (\(x,K)\\. integral K (f (m x)))\ < e/2" by simp next have comb: "integral (cbox a b) (f y) = (\(x, k)\\. integral k (f y))" for y using integral_combine_tagged_division_topdown[OF intf ptag] by metis have f_le: "\y m n. \y \ cbox a b; n\m\ \ f m y \ f n y" using le by (auto intro: transitive_stepwise_le) have "(\(x, k)\\. integral k (f r)) \ (\(x, K)\\. integral K (f (m x)))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f r) \ integral K (f (m x))" proof (rule integral_le) show "f r integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f r y \ f (m x) y" if "y \ K" for y using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto qed qed moreover have "(\(x, K)\\. integral K (f (m x))) \ (\(x, k)\\. integral k (f s))" proof (rule sum_mono, simp add: split_paired_all) fix x K assume xK: "(x, K) \ \" show "integral K (f (m x)) \ integral K (f s)" proof (rule integral_le) show "f (m x) integrable_on K" by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) show "f s integrable_on K" by (metis integrable_on_subcbox intf p'(3) p'(4) xK) show "f (m x) y \ f s y" if "y \ K" for y using that s xK f_le p'(3) by fastforce qed qed moreover have "0 \ i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4" using r by auto ultimately show "\(\(x,K)\\. integral K (f (m x))) - i\ < e/4" using comb i'[of s] by auto qed qed qed with i integral_unique show ?thesis by blast qed lemma monotone_convergence_increasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes int_f: "\k. (f k) integrable_on S" and "\k x. x \ S \ (f k x) \ (f (Suc k) x)" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range (\k. integral S (f k)))" shows "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" proof - have lem: "g integrable_on S \ ((\k. integral S (f k)) \ integral S g) sequentially" if f0: "\k x. x \ S \ 0 \ f k x" and int_f: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f k x \ f (Suc k) x" and lim: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" for f :: "nat \ 'n::euclidean_space \ real" and g S proof - have fg: "(f k x) \ (g x)" if "x \ S" for x k apply (rule tendsto_lowerbound [OF lim [OF that]]) apply (rule eventually_sequentiallyI [of k]) using le by (force intro: transitive_stepwise_le that)+ obtain i where i: "(\k. integral S (f k)) \ i" using bounded_increasing_convergent [OF bou] le int_f integral_le by blast have i': "(integral S (f k)) \ i" for k proof - have "\k. \x. x \ S \ \n\k. f k x \ f n x" using le by (force intro: transitive_stepwise_le) then show ?thesis using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially] by (meson int_f integral_le) qed let ?f = "(\k x. if x \ S then f k x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have int: "?f k integrable_on cbox a b" for a b k by (simp add: int_f integrable_altD(1)) have int': "\k a b. f k integrable_on cbox a b \ S" using int by (simp add: Int_commute integrable_restrict_Int) have g: "?g integrable_on cbox a b \ (\k. integral (cbox a b) (?f k)) \ integral (cbox a b) ?g" for a b proof (rule monotone_convergence_interval) have "norm (integral (cbox a b) (?f k)) \ norm (integral S (f k))" for k proof - have "0 \ integral (cbox a b) (?f k)" by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int') moreover have "0 \ integral S (f k)" by (simp add: integral_nonneg f0 int_f) moreover have "integral (S \ cbox a b) (f k) \ integral S (f k)" by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl) ultimately show ?thesis by (simp add: integral_restrict_Int) qed moreover obtain B where "\x. x \ range (\k. integral S (f k)) \ norm x \ B" using bou unfolding bounded_iff by blast ultimately show "bounded (range (\k. integral (cbox a b) (?f k)))" unfolding bounded_iff by (blast intro: order_trans) qed (use int le lim in auto) moreover have "\B>0. \a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) ?g - i) < e" if "0 < e" for e proof - have "e/4>0" using that by auto with LIMSEQ_D [OF i] obtain N where N: "\n. n \ N \ norm (integral S (f n) - i) < e/4" by metis with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]] obtain B where "0 < B" and B: "\a b. ball 0 B \ cbox a b \ norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4" by (meson \0 < e/4\) have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \ cbox a b" for a b proof - obtain M where M: "\n. n \ M \ abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2" using \e > 0\ g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"]) have *: "\\ \ g. \\\ - i\ < e/2; \\ - g\ < e/2; \ \ \; \ \ i\ \ \g - i\ < e" unfolding real_inner_1_right by arith show "norm (integral (cbox a b) ?g - i) < e" unfolding real_norm_def proof (rule *) show "\integral (cbox a b) (?f N) - i\ < e/2" proof (rule abs_triangle_half_l) show "\integral (cbox a b) (?f N) - integral S (f N)\ < e/2/2" using B[OF ab] by simp show "abs (i - integral S (f N)) < e/2/2" using N by (simp add: abs_minus_commute) qed show "\integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\ < e/2" by (metis le_add1 M[of "M + N"]) show "integral (cbox a b) (?f N) \ integral (cbox a b) (?f (M + N))" proof (intro ballI integral_le[OF int int]) fix x assume "x \ cbox a b" have "(f m x) \ (f n x)" if "x \ S" "n \ m" for m n apply (rule transitive_stepwise_le [OF \n \ m\ order_refl]) using dual_order.trans apply blast by (simp add: le \x \ S\) then show "(?f N)x \ (?f (M+N))x" by auto qed have "integral (cbox a b \ S) (f (M + N)) \ integral S (f (M + N))" by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le) then have "integral (cbox a b) (?f (M + N)) \ integral S (f (M + N))" by (metis (no_types) inf_commute integral_restrict_Int) also have "... \ i" using i'[of "M + N"] by auto finally show "integral (cbox a b) (?f (M + N)) \ i" . qed qed then show ?thesis using \0 < B\ by blast qed ultimately have "(g has_integral i) S" unfolding has_integral_alt' by auto then show ?thesis using has_integral_integrable_integral i integral_unique by metis qed have sub: "\k. integral S (\x. f k x - f 0 x) = integral S (f k) - integral S (f 0)" by (simp add: integral_diff int_f) have *: "\x m n. x \ S \ n\m \ f m x \ f n x" using assms(2) by (force intro: transitive_stepwise_le) have gf: "(\x. g x - f 0 x) integrable_on S \ ((\k. integral S (\x. f (Suc k) x - f 0 x)) \ integral S (\x. g x - f 0 x)) sequentially" proof (rule lem) show "\k. (\x. f (Suc k) x - f 0 x) integrable_on S" by (simp add: integrable_diff int_f) show "(\k. f (Suc k) x - f 0 x) \ g x - f 0 x" if "x \ S" for x proof - have "(\n. f (Suc n) x) \ g x" using LIMSEQ_ignore_initial_segment[OF fg[OF \x \ S\], of 1] by simp then show ?thesis by (simp add: tendsto_diff) qed show "bounded (range (\k. integral S (\x. f (Suc k) x - f 0 x)))" proof - obtain B where B: "\k. norm (integral S (f k)) \ B" using bou by (auto simp: bounded_iff) then have "norm (integral S (\x. f (Suc k) x - f 0 x)) \ B + norm (integral S (f 0))" for k unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff) then show ?thesis unfolding bounded_iff by blast qed qed (use * in auto) then have "(\x. integral S (\xa. f (Suc x) xa - f 0 xa) + integral S (f 0)) \ integral S (\x. g x - f 0 x) + integral S (f 0)" by (auto simp add: tendsto_add) moreover have "(\x. g x - f 0 x + f 0 x) integrable_on S" using gf integrable_add int_f [of 0] by metis ultimately show ?thesis by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub) qed lemma has_integral_monotone_convergence_increasing: fixes f :: "nat \ 'a::euclidean_space \ real" assumes f: "\k. (f k has_integral x k) s" assumes "\k x. x \ s \ f k x \ f (Suc k) x" assumes "\x. x \ s \ (\k. f k x) \ g x" assumes "x \ x'" shows "(g has_integral x') s" proof - have x_eq: "x = (\i. integral s (f i))" by (simp add: integral_unique[OF f]) then have x: "range(\k. integral s (f k)) = range x" by auto have *: "g integrable_on s \ (\k. integral s (f k)) \ integral s g" proof (intro monotone_convergence_increasing allI ballI assms) show "bounded (range(\k. integral s (f k)))" using x convergent_imp_bounded assms by metis qed (use f in auto) then have "integral s g = x'" by (intro LIMSEQ_unique[OF _ \x \ x'\]) (simp add: x_eq) with * show ?thesis by (simp add: has_integral_integral) qed lemma monotone_convergence_decreasing: fixes f :: "nat \ 'n::euclidean_space \ real" assumes intf: "\k. (f k) integrable_on S" and le: "\k x. x \ S \ f (Suc k) x \ f k x" and fg: "\x. x \ S \ ((\k. f k x) \ g x) sequentially" and bou: "bounded (range(\k. integral S (f k)))" shows "g integrable_on S \ (\k. integral S (f k)) \ integral S g" proof - have *: "range(\k. integral S (\x. - f k x)) = (*\<^sub>R) (- 1) ` (range(\k. integral S (f k)))" by force have "(\x. - g x) integrable_on S \ (\k. integral S (\x. - f k x)) \ integral S (\x. - g x)" proof (rule monotone_convergence_increasing) show "\k. (\x. - f k x) integrable_on S" by (blast intro: integrable_neg intf) show "\k x. x \ S \ - f k x \ - f (Suc k) x" by (simp add: le) show "\x. x \ S \ (\k. - f k x) \ - g x" by (simp add: fg tendsto_minus) show "bounded (range(\k. integral S (\x. - f k x)))" using "*" bou bounded_scaling by auto qed then show ?thesis by (force dest: integrable_neg tendsto_minus) qed lemma integral_norm_bound_integral: fixes f :: "'n::euclidean_space \ 'a::banach" assumes int_f: "f integrable_on S" and int_g: "g integrable_on S" and le_g: "\x. x \ S \ norm (f x) \ g x" shows "norm (integral S f) \ integral S g" proof - have norm: "norm \ \ y + e" if "norm \ \ x" and "\x - y\ < e/2" and "norm (\ - \) < e/2" for e x y and \ \ :: 'a proof - have "norm (\ - \) < e/2" by (metis norm_minus_commute that(3)) moreover have "x \ y + e/2" using that(2) by linarith ultimately show ?thesis using that(1) le_less_trans[OF norm_triangle_sub[of \ \]] by (auto simp: less_imp_le) qed have lem: "norm (integral(cbox a b) f) \ integral (cbox a b) g" if f: "f integrable_on cbox a b" and g: "g integrable_on cbox a b" and nle: "\x. x \ cbox a b \ norm (f x) \ g x" for f :: "'n \ 'a" and g a b proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto with integrable_integral[OF f,unfolded has_integral[of f]] obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" by meson moreover from integrable_integral[OF g,unfolded has_integral[of g]] e obtain \ where \: "gauge \" "\\. \ tagged_division_of cbox a b \ \ fine \ \ norm ((\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" by meson ultimately have "gauge (\x. \ x \ \ x)" using gauge_Int by blast with fine_division_exists obtain \ where p: "\ tagged_division_of cbox a b" "(\x. \ x \ \ x) fine \" by metis have "\ fine \" "\ fine \" using fine_Int p(2) by blast+ show "norm (integral (cbox a b) f) \ integral (cbox a b) g + e" proof (rule norm) have "norm (content K *\<^sub>R f x) \ content K *\<^sub>R g x" if "(x, K) \ \" for x K proof- have K: "x \ K" "K \ cbox a b" using \(x, K) \ \\ p(1) by blast+ obtain u v where "K = cbox u v" using \(x, K) \ \\ p(1) by blast moreover have "content K * norm (f x) \ content K * g x" by (metis K subsetD dual_order.antisym measure_nonneg mult_zero_left nle not_le real_mult_le_cancel_iff2) then show ?thesis by simp qed then show "norm (\(x, k)\\. content k *\<^sub>R f x) \ (\(x, k)\\. content k *\<^sub>R g x)" by (simp add: sum_norm_le split_def) show "norm ((\(x, k)\\. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" using \\ fine \\ \ p(1) by simp show "\(\(x, k)\\. content k *\<^sub>R g x) - integral (cbox a b) g\ < e/2" using \\ fine \\ \ p(1) by simp qed qed show ?thesis proof (rule field_le_epsilon) fix e :: real assume "e > 0" then have e: "e/2 > 0" by auto let ?f = "(\x. if x \ S then f x else 0)" let ?g = "(\x. if x \ S then g x else 0)" have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b using int_f int_g integrable_altD by auto obtain Bf where "0 < Bf" and Bf: "\a b. ball 0 Bf \ cbox a b \ \z. (?f has_integral z) (cbox a b) \ norm (z - integral S f) < e/2" using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast obtain Bg where "0 < Bg" and Bg: "\a b. ball 0 Bg \ cbox a b \ \z. (?g has_integral z) (cbox a b) \ norm (z - integral S g) < e/2" using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast obtain a b::'n where ab: "ball 0 Bf \ ball 0 Bg \ cbox a b" using ball_max_Un by (metis bounded_ball bounded_subset_cbox_symmetric) have "ball 0 Bf \ cbox a b" using ab by auto with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" by meson have "ball 0 Bg \ cbox a b" using ab by auto with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" by meson show "norm (integral S f) \ integral S g + e" proof (rule norm) show "norm (integral (cbox a b) ?f) \ integral (cbox a b) ?g" by (simp add: le_g lem[OF f g, of a b]) show "\integral (cbox a b) ?g - integral S g\ < e/2" using int_gw integral_unique w by auto show "norm (integral (cbox a b) ?f - integral S f) < e/2" using int_fz integral_unique z by blast qed qed qed lemma continuous_on_imp_absolutely_integrable_on: fixes f::"real \ 'a::banach" shows "continuous_on {a..b} f \ norm (integral {a..b} f) \ integral {a..b} (\x. norm (f x))" by (intro integral_norm_bound_integral integrable_continuous_real continuous_on_norm) auto lemma integral_bound: fixes f::"real \ 'a::banach" assumes "a \ b" assumes "continuous_on {a .. b} f" assumes "\t. t \ {a .. b} \ norm (f t) \ B" shows "norm (integral {a .. b} f) \ B * (b - a)" proof - note continuous_on_imp_absolutely_integrable_on[OF assms(2)] also have "integral {a..b} (\x. norm (f x)) \ integral {a..b} (\_. B)" by (rule integral_le) (auto intro!: integrable_continuous_real continuous_intros assms) also have "\ = B * (b - a)" using assms by simp finally show ?thesis . qed lemma integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "f integrable_on S" and g: "g integrable_on S" and fg: "\x. x \ S \ norm(f x) \ (g x)\k" shows "norm (integral S f) \ (integral S g)\k" proof - have "norm (integral S f) \ integral S ((\x. x \ k) \ g)" apply (rule integral_norm_bound_integral[OF f integrable_linear[OF g]]) apply (simp add: bounded_linear_inner_left) apply (metis fg o_def) done then show ?thesis unfolding o_def integral_component_eq[OF g] . qed lemma has_integral_norm_bound_integral_component: fixes f :: "'n::euclidean_space \ 'a::banach" fixes g :: "'n \ 'b::euclidean_space" assumes f: "(f has_integral i) S" and g: "(g has_integral j) S" and "\x. x \ S \ norm (f x) \ (g x)\k" shows "norm i \ j\k" using integral_norm_bound_integral_component[of f S g k] unfolding integral_unique[OF f] integral_unique[OF g] using assms by auto lemma uniformly_convergent_improper_integral: fixes f :: "'b \ real \ 'a :: {banach}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "\y x. y \ A \ x \ a \ norm (f y x) \ g x" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases) case (1 \) from G have "Cauchy G" by (auto intro!: convergent_Cauchy) with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \" if "m \ M" "n \ M" for m n by (force simp: Cauchy_def) define M' where "M' = max (nat \a\) M" show ?case proof (rule exI[of _ M'], safe, goal_cases) case (1 x m n) have M': "M' \ a" "M' \ M" unfolding M'_def by linarith+ have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}" using 1 M' by (intro fundamental_theorem_of_calculus) (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] intro!: DERIV_subset[OF deriv]) have int_f: "f x integrable_on {a'..real n}" if "a' \ a" for a' using that 1 by (cases "a' \ real n") (auto intro: integrable) have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) = norm (integral {a..real n} (f x) - integral {a..real m} (f x))" by (simp add: dist_norm norm_minus_commute) also have "integral {a..real m} (f x) + integral {real m..real n} (f x) = integral {a..real n} (f x)" using M' and 1 by (intro integral_combine int_f) auto hence "integral {a..real n} (f x) - integral {a..real m} (f x) = integral {real m..real n} (f x)" by (simp add: algebra_simps) also have "norm \ \ integral {real m..real n} g" using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto also from int_g have "integral {real m..real n} g = G (real n) - G (real m)" by (simp add: has_integral_iff) also have "\ \ dist (G m) (G n)" by (simp add: dist_norm) also from 1 and M' have "\ < \" by (intro M) auto finally show ?case . qed qed lemma uniformly_convergent_improper_integral': fixes f :: "'b \ real \ 'a :: {banach, real_normed_algebra}" assumes deriv: "\x. x \ a \ (G has_field_derivative g x) (at x within {a..})" assumes integrable: "\a' b x. x \ A \ a' \ a \ b \ a' \ f x integrable_on {a'..b}" assumes G: "convergent G" assumes le: "eventually (\x. \y\A. norm (f y x) \ g x) at_top" shows "uniformly_convergent_on A (\b x. integral {a..b} (f x))" proof - from le obtain a'' where le: "\y x. y \ A \ x \ a'' \ norm (f y x) \ g x" by (auto simp: eventually_at_top_linorder) define a' where "a' = max a a''" have "uniformly_convergent_on A (\b x. integral {a'..real b} (f x))" proof (rule uniformly_convergent_improper_integral) fix t assume t: "t \ a'" hence "(G has_field_derivative g t) (at t within {a..})" by (intro deriv) (auto simp: a'_def) moreover have "{a'..} \ {a..}" unfolding a'_def by auto ultimately show "(G has_field_derivative g t) (at t within {a'..})" by (rule DERIV_subset) qed (insert le, auto simp: a'_def intro: integrable G) hence "uniformly_convergent_on A (\b x. integral {a..a'} (f x) + integral {a'..real b} (f x))" (is ?P) by (intro uniformly_convergent_add) auto also have "eventually (\x. \y\A. integral {a..a'} (f y) + integral {a'..x} (f y) = integral {a..x} (f y)) sequentially" by (intro eventually_mono [OF eventually_ge_at_top[of "nat \a'\"]] ballI integral_combine) (auto simp: a'_def intro: integrable) hence "?P \ ?thesis" by (intro uniformly_convergent_cong) simp_all finally show ?thesis . qed subsection \differentiation under the integral sign\ lemma integral_continuous_on_param: fixes f::"'a::topological_space \ 'b::euclidean_space \ 'c::banach" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). f x t)" shows "continuous_on U (\x. integral (cbox a b) (f x))" proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis unfolding continuous_on_def proof (safe intro!: tendstoI) fix e'::real and x assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) assume "x \ U" from continuous_on_prod_compactE[OF cont_fx compact_cbox \x \ U\ \0 < e\] obtain X0 where X0: "x \ X0" "open X0" and fx_bound: "\y t. y \ X0 \ U \ t \ cbox a b \ norm (f y t - f x t) \ e" unfolding split_beta fst_conv snd_conv dist_norm by metis have "\\<^sub>F y in at x within U. y \ X0 \ U" using X0(1) X0(2) eventually_at_topological by auto then show "\\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" proof eventually_elim case (elim y) have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = norm (integral (cbox a b) (\t. f y t - f x t))" using elim \x \ U\ unfolding dist_norm by (subst integral_diff) (auto intro!: integrable_continuous continuous_intros) also have "\ \ e * content (cbox a b)" using elim \x \ U\ by (intro integrable_bound) (auto intro!: fx_bound \x \ U \ less_imp_le[OF \0 < e\] integrable_continuous continuous_intros) also have "\ < e'" using \0 < e'\ \e > 0\ by (auto simp: e_def field_split_simps) finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . qed qed qed (auto intro!: continuous_on_const) lemma leibniz_rule: fixes f::"'a::banach \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ f x integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes [intro]: "x0 \ U" assumes "convex U" shows "((\x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" (is "(?F has_derivative ?dF) _") proof cases assume "content (cbox a b) \ 0" then have ne: "cbox a b \ {}" by auto note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] show ?thesis proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) have cont_f1: "\t. t \ cbox a b \ continuous_on U (\x. f x t)" by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) note [continuous_intros] = continuous_on_compose2[OF cont_f1] fix e'::real assume "e' > 0" define e where "e = e' / (content (cbox a b) + 1)" have "e > 0" using \e' > 0\ by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) from continuous_on_prod_compactE[OF cont_fx compact_cbox \x0 \ U\ \e > 0\] obtain X0 where X0: "x0 \ X0" "open X0" and fx_bound: "\x t. x \ X0 \ U \ t \ cbox a b \ norm (fx x t - fx x0 t) \ e" unfolding split_beta fst_conv snd_conv by (metis dist_norm) note eventually_closed_segment[OF \open X0\ \x0 \ X0\, of U] moreover have "\\<^sub>F x in at x0 within U. x \ X0" using \open X0\ \x0 \ X0\ eventually_at_topological by blast moreover have "\\<^sub>F x in at x0 within U. x \ x0" by (auto simp: eventually_at_filter) moreover have "\\<^sub>F x in at x0 within U. x \ U" by (auto simp: eventually_at_filter) ultimately show "\\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" proof eventually_elim case (elim x) from elim have "0 < norm (x - x0)" by simp have "closed_segment x0 x \ U" by (rule \convex U\[unfolded convex_contains_segment, rule_format, OF \x0 \ U\ \x \ U\]) from elim have [intro]: "x \ U" by auto have "?F x - ?F x0 - ?dF (x - x0) = integral (cbox a b) (\y. f x y - f x0 y - fx x0 y (x - x0))" (is "_ = ?id") using \x \ x0\ by (subst blinfun_apply_integral integral_diff, auto intro!: integrable_diff integrable_f2 continuous_intros intro: integrable_continuous)+ also { fix t assume t: "t \ (cbox a b)" have seg: "\t. t \ {0..1} \ x0 + t *\<^sub>R (x - x0) \ X0 \ U" using \closed_segment x0 x \ U\ \closed_segment x0 x \ X0\ by (force simp: closed_segment_def algebra_simps) from t have deriv: "((\x. f x t) has_derivative (fx y t)) (at y within X0 \ U)" if "y \ X0 \ U" for y unfolding has_vector_derivative_def[symmetric] using that \x \ X0\ by (intro has_derivative_within_subset[OF fx]) auto have "\x. x \ X0 \ U \ onorm (blinfun_apply (fx x t) - (fx x0 t)) \ e" using fx_bound t by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) from differentiable_bound_linearization[OF seg deriv this] X0 have "norm (f x t - f x0 t - fx x0 t (x - x0)) \ e * norm (x - x0)" by (auto simp add: ac_simps) } then have "norm ?id \ integral (cbox a b) (\_. e * norm (x - x0))" by (intro integral_norm_bound_integral) (auto intro!: continuous_intros integrable_diff integrable_f2 intro: integrable_continuous) also have "\ = content (cbox a b) * e * norm (x - x0)" by simp also have "\ < e' * norm (x - x0)" using \e' > 0\ apply (intro mult_strict_right_mono[OF _ \0 < norm (x - x0)\]) apply (simp add: divide_simps e_def not_less) done finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . then show ?case by (auto simp: divide_simps) qed qed (rule blinfun.bounded_linear_right) qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) lemma has_vector_derivative_eq_has_derivative_blinfun: "(f has_vector_derivative f') (at x within U) \ (f has_derivative blinfun_scaleR_left f') (at x within U)" by (simp add: has_vector_derivative_def) lemma leibniz_rule_vector_derivative: fixes f::"real \ 'b::euclidean_space \ 'c::banach" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_vector_derivative (fx x t)) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ cbox a b) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_scaleR_left (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_vector_derivative_eq_has_derivative_blinfun apply (rule has_derivative_eq_rhs) apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_scaleR_left (fx x t)"]) using fx cont_fx apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros) done qed lemma has_field_derivative_eq_has_derivative_blinfun: "(f has_field_derivative f') (at x within U) \ (f has_derivative blinfun_mult_right f') (at x within U)" by (simp add: has_field_derivative_def) lemma leibniz_rule_field_derivative: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes fx: "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes integrable_f2: "\x. x \ U \ (f x) integrable_on cbox a b" assumes cont_fx: "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes U: "x0 \ U" "convex U" shows "((\x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" proof - note [continuous_intros] = continuous_on_compose2[OF cont_fx, where f="\y. Pair x y" for x, unfolded split_beta fst_conv snd_conv] have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = integral (cbox a b) (\t. blinfun_mult_right (fx x0 t))" by (subst integral_linear[symmetric]) (auto simp: has_vector_derivative_def o_def intro!: integrable_continuous U continuous_intros bounded_linear_intros) show ?thesis unfolding has_field_derivative_eq_has_derivative_blinfun apply (rule has_derivative_eq_rhs) apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\x t. blinfun_mult_right (fx x t)"]) using fx cont_fx apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros) done qed lemma leibniz_rule_field_differentiable: fixes f::"'a::{real_normed_field, banach} \ 'b::euclidean_space \ 'a" assumes "\x t. x \ U \ t \ cbox a b \ ((\x. f x t) has_field_derivative fx x t) (at x within U)" assumes "\x. x \ U \ (f x) integrable_on cbox a b" assumes "continuous_on (U \ (cbox a b)) (\(x, t). fx x t)" assumes "x0 \ U" "convex U" shows "(\x. integral (cbox a b) (f x)) field_differentiable at x0 within U" using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def) subsection \Exchange uniform limit and integral\ lemma uniform_limit_integral_cbox: fixes f::"'a \ 'b::euclidean_space \ 'c::banach" assumes u: "uniform_limit (cbox a b) f g F" assumes c: "\n. continuous_on (cbox a b) (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) (cbox a b)" "(g has_integral J) (cbox a b)" "(I \ J) F" proof - have fi[simp]: "f n integrable_on (cbox a b)" for n by (auto intro!: integrable_continuous assms) then obtain I where I: "\n. (f n has_integral I n) (cbox a b)" by atomize_elim (auto simp: integrable_on_def intro!: choice) moreover have gi[simp]: "g integrable_on (cbox a b)" by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) then obtain J where J: "(g has_integral J) (cbox a b)" by blast moreover have "(I \ J) F" proof cases assume "content (cbox a b) = 0" hence "I = (\_. 0)" "J = 0" by (auto intro!: has_integral_unique I J) thus ?thesis by simp next assume content_nonzero: "content (cbox a b) \ 0" show ?thesis proof (rule tendstoI) fix e::real assume "e > 0" define e' where "e' = e/2" with \e > 0\ have "e' > 0" by simp then have "\\<^sub>F n in F. \x\cbox a b. norm (f n x - g x) < e' / content (cbox a b)" using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff) then show "\\<^sub>F n in F. dist (I n) J < e" proof eventually_elim case (elim n) have "I n = integral (cbox a b) (f n)" "J = integral (cbox a b) g" using I[of n] J by (simp_all add: integral_unique) then have "dist (I n) J = norm (integral (cbox a b) (\x. f n x - g x))" by (simp add: integral_diff dist_norm) also have "\ \ integral (cbox a b) (\x. (e' / content (cbox a b)))" using elim by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) also have "\ < e" using \0 < e\ by (simp add: e'_def) finally show ?case . qed qed qed ultimately show ?thesis .. qed lemma uniform_limit_integral: fixes f::"'a \ 'b::ordered_euclidean_space \ 'c::banach" assumes u: "uniform_limit {a..b} f g F" assumes c: "\n. continuous_on {a..b} (f n)" assumes [simp]: "F \ bot" obtains I J where "\n. (f n has_integral I n) {a..b}" "(g has_integral J) {a..b}" "(I \ J) F" by (metis interval_cbox assms uniform_limit_integral_cbox) subsection \Integration by parts\ lemma integration_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" proof - interpret bounded_bilinear prod by fact have "((\x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral (prod (f b) (g b) - prod (f a) (g a))) {a..b}" using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) (auto intro!: continuous_intros continuous_on has_vector_derivative) from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps) qed lemma integration_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integration_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}" shows "((\x. prod (f' x) (g x)) has_integral y) {a..b}" by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts_interior_strong: fixes prod :: "_ \ _ \ 'b :: banach" assumes bilinear: "bounded_bilinear (prod)" assumes s: "finite s" and le: "a \ b" assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g" assumes deriv: "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes int: "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" proof - from int obtain I where "((\x. prod (f x) (g' x)) has_integral I) {a..b}" unfolding integrable_on_def by blast hence "((\x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp from integration_by_parts_interior_strong[OF assms(1-7) this] show ?thesis unfolding integrable_on_def by blast qed lemma integrable_by_parts_interior: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a<.. (f has_vector_derivative f' x) (at x)" "\x. x\{a<.. (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) lemma integrable_by_parts: fixes prod :: "_ \ _ \ 'b :: banach" assumes "bounded_bilinear (prod)" "a \ b" "continuous_on {a..b} f" "continuous_on {a..b} g" assumes "\x. x\{a..b} \ (f has_vector_derivative f' x) (at x)" "\x. x\{a..b} \ (g has_vector_derivative g' x) (at x)" assumes "(\x. prod (f x) (g' x)) integrable_on {a..b}" shows "(\x. prod (f' x) (g x)) integrable_on {a..b}" by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all) subsection \Integration by substitution\ lemma has_integral_substitution_general: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1 f integrable_continuous_real)+ have deriv: "(((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b})" if "x \ {a..b} - s" for x proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) show "(g has_vector_derivative g' x) (at x within {a..b})" using deriv has_field_derivative_iff_has_vector_derivative that by blast show "((\x. integral {c..x} f) has_vector_derivative f (g x)) (at (g x) within g ` {a..b})" using that le subset by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x)" if "x \ {a..b} - (s \ {a,b})" for x using deriv[of x] that by (simp add: at_within_Icc_at o_def) have "((\x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using le cont_int s deriv cont_int by (intro fundamental_theorem_of_calculus_interior_strong[of "s \ {a,b}"]) simp_all also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_strong: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" "g a \ g b" and subset: "g ` {a..b} \ {c..d}" and f [continuous_intros]: "continuous_on {c..d} f" and g [continuous_intros]: "continuous_on {a..b} g" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2) by (cases "g a = g b") auto lemma has_integral_substitution: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes "a \ b" "g a \ g b" "g ` {a..b} \ {c..d}" and "continuous_on {c..d} f" and "\x. x \ {a..b} \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}" by (intro has_integral_substitution_strong[of "{}" a b g c d] assms) (auto intro: DERIV_continuous_on assms) lemma integral_shift: fixes f :: "real \ 'a::euclidean_space" assumes cont: "continuous_on {a + c..b + c} f" shows "integral {a..b} (f \ (\x. x + c)) = integral {a + c..b + c} f" proof (cases "a \ b") case True have "((\x. 1 *\<^sub>R f (x + c)) has_integral integral {a+c..b+c} f) {a..b}" using True cont by (intro has_integral_substitution[where c = "a + c" and d = "b + c"]) (auto intro!: derivative_eq_intros) thus ?thesis by (simp add: has_integral_iff o_def) qed auto subsection \Compute a double integral using iterated integrals and switching the order of integration\ lemma continuous_on_imp_integrable_on_Pair1: fixes f :: "_ \ 'b::banach" assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \ cbox a b" shows "(\y. f (x, y)) integrable_on (cbox c d)" proof - have "f \ (\y. (x, y)) integrable_on (cbox c d)" apply (rule integrable_continuous) apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]]) using x apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con) done then show ?thesis by (simp add: o_def) qed lemma integral_integrable_2dim: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) f" shows "(\x. integral (cbox c d) (\y. f (x,y))) integrable_on cbox a b" proof (cases "content(cbox c d) = 0") case True then show ?thesis by (simp add: True integrable_const) next case False have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" by (simp add: assms compact_cbox compact_uniformly_continuous) { fix x::'a and e::real assume x: "x \ cbox a b" and e: "0 < e" then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e" by (auto simp: False content_lt_nz e) then obtain dd where dd: "\x x'. \x\cbox (a, c) (b, d); x'\cbox (a, c) (b, d); norm (x' - x) < dd\ \ norm (f x' - f x) \ e/(2 * content (cbox c d))" "dd>0" using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"] by (auto simp: dist_norm intro: less_imp_le) have "\delta>0. \x'\cbox a b. norm (x' - x) < delta \ norm (integral (cbox c d) (\u. f (x', u) - f (x, u))) < e" apply (rule_tac x=dd in exI) using dd e2_gt assms x apply clarify apply (rule le_less_trans [OF _ e2_less]) apply (rule integrable_bound) apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) done } note * = this show ?thesis apply (rule integrable_continuous) apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) done qed lemma integral_split: fixes f :: "'a::euclidean_space \ 'b::{real_normed_vector,complete_space}" assumes f: "f integrable_on (cbox a b)" and k: "k \ Basis" shows "integral (cbox a b) f = integral (cbox a b \ {x. x\k \ c}) f + integral (cbox a b \ {x. x\k \ c}) f" apply (rule integral_unique [OF has_integral_split [where c=c]]) using k f apply (auto simp: has_integral_integral [symmetric]) done lemma integral_swap_operativeI: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes f: "continuous_on s f" and e: "0 < e" shows "operative conj True (\k. \a b c d. cbox (a,c) (b,d) \ k \ cbox (a,c) (b,d) \ s \ norm(integral (cbox (a,c) (b,d)) f - integral (cbox a b) (\x. integral (cbox c d) (\y. f((x,y))))) \ e * content (cbox (a,c) (b,d)))" proof (standard, auto) fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b assume *: "box (a, c) (b, d) = {}" and cb1: "cbox (u, w) (v, z) \ cbox (a, c) (b, d)" and cb2: "cbox (u, w) (v, z) \ s" then have c0: "content (cbox (a, c) (b, d)) = 0" using * unfolding content_eq_0_interior by simp have c0': "content (cbox (u, w) (v, z)) = 0" by (fact content_0_subset [OF c0 cb1]) show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u,w) (v,z))" using content_cbox_pair_eq0_D [OF c0'] by (force simp add: c0') next fix a::'a and c::'b and b::'a and d::'b and M::real and i::'a and j::'b and u::'a and v::'a and w::'b and z::'b assume ij: "(i,j) \ Basis" and n1: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. x \ (i,j) \ M} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and n2: "\a' b' c' d'. cbox (a',c') (b',d') \ cbox (a,c) (b,d) \ cbox (a',c') (b',d') \ {x. M \ x \ (i,j)} \ cbox (a',c') (b',d') \ s \ norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f (x,y)))) \ e * content (cbox (a',c') (b',d'))" and subs: "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" "cbox (u,w) (v,z) \ s" have fcont: "continuous_on (cbox (u, w) (v, z)) f" using assms(1) continuous_on_subset subs(2) by blast then have fint: "f integrable_on cbox (u, w) (v, z)" by (metis integrable_continuous) consider "i \ Basis" "j=0" | "j \ Basis" "i=0" using ij by (auto simp: Euclidean_Space.Basis_prod_def) then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z))" (is ?normle) proof cases case 1 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v \ {x. x \ i \ M}) * content (cbox w z)) + e * (content (cbox u v \ {x. M \ x \ i}) * content (cbox w z))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v \ {x. x \ i \ M}) (\x. integral (cbox w z) (\y. f (x, y))) + integral (cbox u v \ {x. M \ x \ i}) (\x. integral (cbox w z) (\y. f (x, y)))" using 1 f subs integral_integrable_2dim continuous_on_subset by (blast intro: integral_split) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 1 subs apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\u. M\u"] setcomp_dot1 [of "\u. u\M"] Sigma_Int_Paircomp1) apply (simp_all add: interval_split ij) apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) done next case 2 then have e: "e * content (cbox (u, w) (v, z)) = e * (content (cbox u v) * content (cbox w z \ {x. x \ j \ M})) + e * (content (cbox u v) * content (cbox w z \ {x. M \ x \ j}))" by (simp add: content_split [where c=M] content_Pair algebra_simps) have "(\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) integrable_on cbox u v" "(\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y))) integrable_on cbox u v" using 2 subs apply (simp_all add: interval_split) apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]]) apply (auto simp: cbox_Pair_eq interval_split [symmetric]) done with 2 have *: "integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y))) = integral (cbox u v) (\x. integral (cbox w z \ {x. x \ j \ M}) (\y. f (x, y))) + integral (cbox u v) (\x. integral (cbox w z \ {x. M \ x \ j}) (\y. f (x, y)))" by (simp add: integral_add [symmetric] integral_split [symmetric] continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) show ?normle apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) using 2 subs apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\u. M\u"] setcomp_dot2 [of "\u. u\M"] Sigma_Int_Paircomp2) apply (simp_all add: interval_split ij) apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) done qed qed lemma integral_Pair_const: "integral (cbox (a,c) (b,d)) (\x. k) = integral (cbox a b) (\x. integral (cbox c d) (\y. k))" by (simp add: content_Pair) lemma integral_prod_continuous: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::banach" assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f") shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))" proof (cases "content ?CBOX = 0") case True then show ?thesis by (auto simp: content_Pair) next case False then have cbp: "content ?CBOX > 0" using content_lt_nz by blast have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) = 0" proof (rule dense_eq0_I, simp) fix e :: real assume "0 < e" with \content ?CBOX > 0\ have "0 < e/content ?CBOX" by simp have f_int_acbd: "f integrable_on ?CBOX" by (rule integrable_continuous [OF assms]) { fix p assume p: "p division_of ?CBOX" then have "finite p" by blast define e' where "e' = e/content ?CBOX" with \0 < e\ \0 < e/content ?CBOX\ have "0 < e'" by simp interpret operative conj True "\k. \a' b' c' d'. cbox (a', c') (b', d') \ k \ cbox (a', c') (b', d') \ ?CBOX \ norm (integral (cbox (a', c') (b', d')) f - integral (cbox a' b') (\x. integral (cbox c' d') (\y. f ((x, y))))) \ e' * content (cbox (a', c') (b', d'))" using assms \0 < e'\ by (rule integral_swap_operativeI) have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e' * content ?CBOX" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e' * content (cbox (u, w) (v, z))" using that division [of p "(a, c)" "(b, d)"] p \finite p\ by (auto simp add: comm_monoid_set_F_and) with False have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x, y)))) \ e" if "\t u v w z. t \ p \ cbox (u, w) (v, z) \ t \ cbox (u, w) (v, z) \ ?CBOX \ norm (integral (cbox (u, w) (v, z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x, y)))) \ e * content (cbox (u, w) (v, z)) / content ?CBOX" using that by (simp add: e'_def) } note op_acbd = this { fix k::real and \ and u::'a and v w and z::'b and t1 t2 l assume k: "0 < k" and nf: "\x y u v. \x \ cbox a b; y \ cbox c d; u \ cbox a b; v\cbox c d; norm (u-x, v-y) < k\ \ norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))" and p_acbd: "\ tagged_division_of cbox (a,c) (b,d)" and fine: "(\x. ball x k) fine \" "((t1,t2), l) \ \" and uwvz_sub: "cbox (u,w) (v,z) \ l" "cbox (u,w) (v,z) \ cbox (a,c) (b,d)" have t: "t1 \ cbox a b" "t2 \ cbox c d" by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ have ls: "l \ ball (t1,t2) k" using fine by (simp add: fine_def Ball_def) { fix x1 x2 assume xuvwz: "x1 \ cbox u v" "x2 \ cbox w z" then have x: "x1 \ cbox a b" "x2 \ cbox c d" using uwvz_sub by auto have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)" by (simp add: norm_Pair norm_minus_commute) also have "norm (t1 - x1, t2 - x2) < k" using xuvwz ls uwvz_sub unfolding ball_def by (force simp add: cbox_Pair_eq dist_norm ) finally have "norm (f (x1,x2) - f (t1,t2)) \ e/content ?CBOX/2" using nf [OF t x] by simp } note nf' = this have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" using f_int_acbd uwvz_sub integrable_on_subcbox by blast have f_int_uv: "\x. x \ cbox u v \ (\y. f (x,y)) integrable_on cbox w z" using assms continuous_on_subset uwvz_sub by (blast intro: continuous_on_imp_integrable_on_Pair1) have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\x. f (t1,t2))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) using cbp \0 < e/content ?CBOX\ nf' apply (auto simp: integrable_diff f_int_uwvz integrable_const) done have int_integrable: "(\x. integral (cbox w z) (\y. f (x, y))) integrable_on cbox u v" using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast have normint_wz: "\x. x \ cbox u v \ norm (integral (cbox w z) (\y. f (x, y)) - integral (cbox w z) (\y. f (t1, t2))) \ e * content (cbox w z) / content (cbox (a, c) (b, d))/2" apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) using cbp \0 < e/content ?CBOX\ nf' apply (auto simp: integrable_diff f_int_uv integrable_const) done have "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)) - integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)" apply (rule integrable_bound [OF _ _ normint_wz]) using cbp \0 < e/content ?CBOX\ - apply (auto simp: field_split_simps content_pos_le integrable_diff int_integrable integrable_const) + apply (auto simp: field_split_simps integrable_diff int_integrable integrable_const) done also have "... \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp add: content_Pair field_split_simps) finally have 2: "norm (integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y))) - integral (cbox u v) (\x. integral (cbox w z) (\y. f (t1,t2)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX/2" by (simp only: integral_diff [symmetric] int_integrable integrable_const) have norm_xx: "\x' = y'; norm(x - x') \ e/2; norm(y - y') \ e/2\ \ norm(x - y) \ e" for x::'c and y x' y' e using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] field_sum_of_halves by (simp add: norm_minus_commute) have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\x. integral (cbox w z) (\y. f (x,y)))) \ e * content (cbox (u,w) (v,z)) / content ?CBOX" by (rule norm_xx [OF integral_Pair_const 1 2]) } note * = this have "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" if "\x\?CBOX. \x'\?CBOX. norm (x' - x) < k \ norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k proof - obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" and fine: "(\x. ball x k) fine p" using fine_division_exists \0 < k\ by blast show ?thesis apply (rule op_acbd [OF division_of_tagged_division [OF ptag]]) using that fine ptag \0 < k\ by (auto simp: *) qed then show "norm (integral ?CBOX f - integral (cbox a b) (\x. integral (cbox c d) (\y. f (x,y)))) \ e" using compact_uniformly_continuous [OF assms compact_cbox] apply (simp add: uniformly_continuous_on_def dist_norm) apply (drule_tac x="e/2 / content?CBOX" in spec) using cbp \0 < e\ by (auto simp: zero_less_mult_iff) qed then show ?thesis by simp qed lemma integral_swap_2dim: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox (a, c) (b, d)) (\(x, y). f x y) = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" proof - have "((\(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\(x,y). f y x" "integral (cbox (c, a) (d, b)) (\(x, y). f y x)", simplified]) apply (force simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms)+ done then show ?thesis by force qed theorem integral_swap_continuous: fixes f :: "['a::euclidean_space, 'b::euclidean_space] \ 'c::banach" assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" proof - have "integral (cbox a b) (\x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\(x, y). f x y)" using integral_prod_continuous [OF assms] by auto also have "... = integral (cbox (c, a) (d, b)) (\(x, y). f y x)" by (rule integral_swap_2dim [OF assms]) also have "... = integral (cbox c d) (\y. integral (cbox a b) (\x. f x y))" using integral_prod_continuous [OF swap_continuous] assms by auto finally show ?thesis . qed subsection \Definite integrals for exponential and power function\ lemma has_integral_exp_minus_to_infinity: assumes a: "a > 0" shows "((\x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}" proof - define f where "f = (\k x. if x \ {c..real k} then exp (-a*x) else 0)" { fix k :: nat assume k: "of_nat k \ c" from k a have "((\x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric]) hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def by (subst has_integral_restrict) simp_all } note has_integral_f = this have [simp]: "f k = (\_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) have integral_f: "integral {c..} (f k) = (if real k \ c then exp (-a*c)/a - exp (-a*real k)/a else 0)" for k using integral_unique[OF has_integral_f[of k]] by simp have A: "(\x. exp (-a*x)) integrable_on {c..} \ (\k. integral {c..} (f k)) \ integral {c..} (\x. exp (-a*x))" proof (intro monotone_convergence_increasing allI ballI) fix k ::nat have "(\x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P) unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) hence "(f k) integrable_on {c..of_real k}" by (rule integrable_eq) (simp add: f_def) then show "f k integrable_on {c..}" by (rule integrable_on_superset) (auto simp: f_def) next fix x assume x: "x \ {c..}" have "sequentially \ principal {nat \x\..}" unfolding at_top_def by (simp add: Inf_lower) also have "{nat \x\..} \ {k. x \ real k}" by auto also have "inf (principal \) (principal {k. \x \ real k}) = principal ({k. x \ real k} \ {k. \x \ real k})" by simp also have "{k. x \ real k} \ {k. \x \ real k} = {}" by blast finally have "inf sequentially (principal {k. \x \ real k}) = bot" by (simp add: inf.coboundedI1 bot_unique) with x show "(\k. f k x) \ exp (-a*x)" unfolding f_def by (intro filterlim_If) auto next have "\integral {c..} (f k)\ \ exp (-a*c)/a" for k proof (cases "c > of_nat k") case False hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)" by (simp add: integral_f) also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = exp (- (a * c)) / a - exp (- (a * real k)) / a" using False a by (intro abs_of_nonneg) (simp_all add: field_simps) also have "\ \ exp (- a * c) / a" using a by simp finally show ?thesis . qed (insert a, simp_all add: integral_f) thus "bounded (range(\k. integral {c..} (f k)))" by (intro boundedI[of _ "exp (-a*c)/a"]) auto qed (auto simp: f_def) have "(\k. exp (-a*c)/a - exp (-a * of_nat k)/a) \ exp (-a*c)/a - 0/a" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ (insert a, simp_all) moreover from eventually_gt_at_top[of "nat \c\"] have "eventually (\k. of_nat k > c) sequentially" by eventually_elim linarith hence "eventually (\k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially" by eventually_elim (simp add: integral_f) ultimately have "(\k. integral {c..} (f k)) \ exp (-a*c)/a - 0/a" by (rule Lim_transform_eventually) from LIMSEQ_unique[OF conjunct2[OF A] this] have "integral {c..} (\x. exp (-a*x)) = exp (-a*c)/a" by simp with conjunct1[OF A] show ?thesis by (simp add: has_integral_integral) qed lemma integrable_on_exp_minus_to_infinity: "a > 0 \ (\x. exp (-a*x) :: real) integrable_on {c..}" using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast lemma has_integral_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "((\x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}" proof (cases "c = 0") case False define f where "f = (\k x. if x \ {inverse (of_nat (Suc k))..c} then x powr a else 0)" define F where "F = (\k. if inverse (of_nat (Suc k)) \ c then c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" { fix k :: nat have "(f k has_integral F k) {0..c}" proof (cases "inverse (of_nat (Suc k)) \ c") case True { fix x assume x: "x \ inverse (1 + real k)" have "0 < inverse (1 + real k)" by simp also note x finally have "x > 0" . } note x = this hence "((\x. x powr a) has_integral c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}" using True a by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const - continuous_on_id simp: has_field_derivative_iff_has_vector_derivative [symmetric]) + simp: has_field_derivative_iff_has_vector_derivative [symmetric]) with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all next case False thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto qed } note has_integral_f = this have integral_f: "integral {0..c} (f k) = F k" for k using has_integral_f[of k] by (rule integral_unique) have A: "(\x. x powr a) integrable_on {0..c} \ (\k. integral {0..c} (f k)) \ integral {0..c} (\x. x powr a)" proof (intro monotone_convergence_increasing ballI allI) fix k from has_integral_f[of k] show "f k integrable_on {0..c}" by (auto simp: integrable_on_def) next fix k :: nat and x :: real { assume x: "inverse (real (Suc k)) \ x" have "inverse (real (Suc (Suc k))) \ inverse (real (Suc k))" by (simp add: field_simps) also note x finally have "inverse (real (Suc (Suc k))) \ x" . } thus "f k x \ f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) next fix x assume x: "x \ {0..c}" show "(\k. f k x) \ x powr a" proof (cases "x = 0") case False with x have "x > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. x powr a = f k x) sequentially" by eventually_elim (insert x, simp add: f_def) moreover have "(\_. x powr a) \ x powr a" by simp ultimately show ?thesis by (blast intro: Lim_transform_eventually) qed (simp_all add: f_def) next { fix k from a have "F k \ c powr (a + 1) / (a + 1)" by (auto simp add: F_def divide_simps) also from a have "F k \ 0" by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) hence "F k = abs (F k)" by simp finally have "abs (F k) \ c powr (a + 1) / (a + 1)" . } thus "bounded (range(\k. integral {0..c} (f k)))" by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) qed from False c have "c > 0" by simp from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] have "eventually (\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = integral {0..c} (f k)) sequentially" by eventually_elim (simp add: integral_f F_def) moreover have "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto hence "(\k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) \ c powr (a + 1) / (a + 1)" by simp ultimately have "(\k. integral {0..c} (f k)) \ c powr (a+1) / (a+1)" by (blast intro: Lim_transform_eventually) with A have "integral {0..c} (\x. x powr a) = c powr (a+1) / (a+1)" by (blast intro: LIMSEQ_unique) with A show ?thesis by (simp add: has_integral_integral) qed (simp_all add: has_integral_refl) lemma integrable_on_powr_from_0: assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0..c}" using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast lemma has_integral_powr_to_inf: fixes a e :: real assumes "e < -1" "a > 0" shows "((\x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}" proof - define f :: "nat \ real \ real" where "f = (\n x. if x \ {a..n} then x powr e else 0)" define F where "F = (\x. x powr (e + 1) / (e + 1))" have has_integral_f: "(f n has_integral (F n - F a)) {a..}" if n: "n \ a" for n :: nat proof - from n assms have "((\x. x powr e) has_integral (F n - F a)) {a..n}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def) hence "(f n has_integral (F n - F a)) {a..n}" by (rule has_integral_eq [rotated]) (simp add: f_def) thus "(f n has_integral (F n - F a)) {a..}" by (rule has_integral_on_superset) (auto simp: f_def) qed have integral_f: "integral {a..} (f n) = (if n \ a then F n - F a else 0)" for n :: nat proof (cases "n \ a") case True with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) next case False have "(f n has_integral 0) {a}" by (rule has_integral_refl) hence "(f n has_integral 0) {a..}" by (rule has_integral_on_superset) (insert False, simp_all add: f_def) with False show ?thesis by (simp add: integral_unique) qed have *: "(\x. x powr e) integrable_on {a..} \ (\n. integral {a..} (f n)) \ integral {a..} (\x. x powr e)" proof (intro monotone_convergence_increasing allI ballI) fix n :: nat from assms have "(\x. x powr e) integrable_on {a..n}" by (auto intro!: integrable_continuous_real continuous_intros) hence "f n integrable_on {a..n}" by (rule integrable_eq) (auto simp: f_def) thus "f n integrable_on {a..}" by (rule integrable_on_superset) (auto simp: f_def) next fix n :: nat and x :: real show "f n x \ f (Suc n) x" by (auto simp: f_def) next fix x :: real assume x: "x \ {a..}" from filterlim_real_sequentially have "eventually (\n. real n \ x) at_top" by (simp add: filterlim_at_top) with x have "eventually (\n. f n x = x powr e) at_top" by (auto elim!: eventually_mono simp: f_def) thus "(\n. f n x) \ x powr e" by (simp add: tendsto_eventually) next have "norm (integral {a..} (f n)) \ -F a" for n :: nat proof (cases "n \ a") case True with assms have "a powr (e + 1) \ n powr (e + 1)" by (intro powr_mono2') simp_all with assms show ?thesis by (auto simp: divide_simps F_def integral_f) qed (insert assms, simp add: integral_f F_def field_split_simps) thus "bounded (range(\k. integral {a..} (f k)))" unfolding bounded_iff by (intro exI[of _ "-F a"]) auto qed from filterlim_real_sequentially have "eventually (\n. real n \ a) at_top" by (simp add: filterlim_at_top) hence "eventually (\n. F n - F a = integral {a..} (f n)) at_top" by eventually_elim (simp add: integral_f) moreover have "(\n. F n - F a) \ 0 / (e + 1) - F a" unfolding F_def by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] filterlim_ident filterlim_real_sequentially | simp)+) hence "(\n. F n - F a) \ -F a" by simp ultimately have "(\n. integral {a..} (f n)) \ -F a" by (blast intro: Lim_transform_eventually) from conjunct2[OF *] and this have "integral {a..} (\x. x powr e) = -F a" by (rule LIMSEQ_unique) with conjunct1[OF *] show ?thesis by (simp add: has_integral_integral F_def) qed lemma has_integral_inverse_power_to_inf: fixes a :: real and n :: nat assumes "n > 1" "a > 0" shows "((\x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}" proof - from assms have "real_of_int (-int n) < -1" by simp note has_integral_powr_to_inf[OF this \a > 0\] also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = 1 / (real (n - 1) * a powr (real (n - 1)))" using assms by (simp add: field_split_simps powr_add [symmetric] of_nat_diff) also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" by (intro powr_realpow) finally show ?thesis by (rule has_integral_eq [rotated]) (insert assms, simp_all add: powr_minus powr_realpow field_split_simps) qed subsubsection \Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\ lemma integral_component_eq_cart[simp]: fixes f :: "'n::euclidean_space \ real^'m" assumes "f integrable_on s" shows "integral s (\x. f x $ k) = integral s f $ k" using integral_linear[OF assms(1) bounded_linear_vec_nth,unfolded o_def] . lemma content_closed_interval: fixes a :: "'a::ordered_euclidean_space" assumes "a \ b" shows "content {a..b} = (\i\Basis. b\i - a\i)" using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) lemma integrable_const_ivl[intro]: fixes a::"'a::ordered_euclidean_space" shows "(\x. c) integrable_on {a..b}" unfolding cbox_interval[symmetric] by (rule integrable_const) lemma integrable_on_subinterval: fixes f :: "'n::ordered_euclidean_space \ 'a::banach" assumes "f integrable_on S" "{a..b} \ S" shows "f integrable_on {a..b}" using integrable_on_subcbox[of f S a b] assms by (simp add: cbox_interval) end diff --git a/src/HOL/Analysis/Homeomorphism.thy b/src/HOL/Analysis/Homeomorphism.thy --- a/src/HOL/Analysis/Homeomorphism.thy +++ b/src/HOL/Analysis/Homeomorphism.thy @@ -1,2255 +1,2255 @@ (* Title: HOL/Analysis/Homeomorphism.thy Author: LC Paulson (ported from HOL Light) *) section \Homeomorphism Theorems\ theory Homeomorphism imports Homotopy begin lemma homeomorphic_spheres': fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space" assumes "0 < \" and dimeq: "DIM('a) = DIM('b)" shows "(sphere a \) homeomorphic (sphere b \)" proof - obtain f :: "'a\'b" and g where "linear f" "linear g" and fg: "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g(f x) = x" "\y. f(g y) = y" by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq]) then have "continuous_on UNIV f" "continuous_on UNIV g" using linear_continuous_on linear_linear by blast+ then show ?thesis unfolding homeomorphic_minimal apply(rule_tac x="\x. b + f(x - a)" in exI) apply(rule_tac x="\x. a + g(x - b)" in exI) using assms apply (force intro: continuous_intros continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg) done qed lemma homeomorphic_spheres_gen: fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space" assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)" shows "(sphere a r homeomorphic sphere b s)" apply (rule homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres']) using assms apply auto done subsection \Homeomorphism of all convex compact sets with nonempty interior\ proposition fixes S :: "'a::euclidean_space set" assumes "compact S" and 0: "0 \ rel_interior S" and star: "\x. x \ S \ open_segment 0 x \ rel_interior S" shows starlike_compact_projective1_0: "S - rel_interior S homeomorphic sphere 0 1 \ affine hull S" (is "?SMINUS homeomorphic ?SPHER") and starlike_compact_projective2_0: "S homeomorphic cball 0 1 \ affine hull S" (is "S homeomorphic ?CBALL") proof - have starI: "(u *\<^sub>R x) \ rel_interior S" if "x \ S" "0 \ u" "u < 1" for x u proof (cases "x=0 \ u=0") case True with 0 show ?thesis by force next case False with that show ?thesis by (auto simp: in_segment intro: star [THEN subsetD]) qed have "0 \ S" using assms rel_interior_subset by auto define proj where "proj \ \x::'a. x /\<^sub>R norm x" have eqI: "x = y" if "proj x = proj y" "norm x = norm y" for x y using that by (force simp: proj_def) then have iff_eq: "\x y. (proj x = proj y \ norm x = norm y) \ x = y" by blast have projI: "x \ affine hull S \ proj x \ affine hull S" for x by (metis \0 \ S\ affine_hull_span_0 hull_inc span_mul proj_def) have nproj1 [simp]: "x \ 0 \ norm(proj x) = 1" for x by (simp add: proj_def) have proj0_iff [simp]: "proj x = 0 \ x = 0" for x by (simp add: proj_def) have cont_proj: "continuous_on (UNIV - {0}) proj" unfolding proj_def by (rule continuous_intros | force)+ have proj_spherI: "\x. \x \ affine hull S; x \ 0\ \ proj x \ ?SPHER" by (simp add: projI) have "bounded S" "closed S" using \compact S\ compact_eq_bounded_closed by blast+ have inj_on_proj: "inj_on proj (S - rel_interior S)" proof fix x y assume x: "x \ S - rel_interior S" and y: "y \ S - rel_interior S" and eq: "proj x = proj y" then have xynot: "x \ 0" "y \ 0" "x \ S" "y \ S" "x \ rel_interior S" "y \ rel_interior S" using 0 by auto consider "norm x = norm y" | "norm x < norm y" | "norm x > norm y" by linarith then show "x = y" proof cases assume "norm x = norm y" with iff_eq eq show "x = y" by blast next assume *: "norm x < norm y" have "x /\<^sub>R norm x = norm x *\<^sub>R (x /\<^sub>R norm x) /\<^sub>R norm (norm x *\<^sub>R (x /\<^sub>R norm x))" by force then have "proj ((norm x / norm y) *\<^sub>R y) = proj x" by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) then have [simp]: "(norm x / norm y) *\<^sub>R y = x" by (rule eqI) (simp add: \y \ 0\) have no: "0 \ norm x / norm y" "norm x / norm y < 1" using * by (auto simp: field_split_simps) then show "x = y" using starI [OF \y \ S\ no] xynot by auto next assume *: "norm x > norm y" have "y /\<^sub>R norm y = norm y *\<^sub>R (y /\<^sub>R norm y) /\<^sub>R norm (norm y *\<^sub>R (y /\<^sub>R norm y))" by force then have "proj ((norm y / norm x) *\<^sub>R x) = proj y" by (metis (no_types) divide_inverse local.proj_def eq scaleR_scaleR) then have [simp]: "(norm y / norm x) *\<^sub>R x = y" by (rule eqI) (simp add: \x \ 0\) have no: "0 \ norm y / norm x" "norm y / norm x < 1" using * by (auto simp: field_split_simps) then show "x = y" using starI [OF \x \ S\ no] xynot by auto qed qed have "\surf. homeomorphism (S - rel_interior S) ?SPHER proj surf" proof (rule homeomorphism_compact) show "compact (S - rel_interior S)" using \compact S\ compact_rel_boundary by blast show "continuous_on (S - rel_interior S) proj" using 0 by (blast intro: continuous_on_subset [OF cont_proj]) show "proj ` (S - rel_interior S) = ?SPHER" proof show "proj ` (S - rel_interior S) \ ?SPHER" using 0 by (force simp: hull_inc projI intro: nproj1) show "?SPHER \ proj ` (S - rel_interior S)" proof (clarsimp simp: proj_def) fix x assume "x \ affine hull S" and nox: "norm x = 1" then have "x \ 0" by auto obtain d where "0 < d" and dx: "(d *\<^sub>R x) \ rel_frontier S" and ri: "\e. \0 \ e; e < d\ \ (e *\<^sub>R x) \ rel_interior S" using ray_to_rel_frontier [OF \bounded S\ 0] \x \ affine hull S\ \x \ 0\ by auto show "x \ (\x. x /\<^sub>R norm x) ` (S - rel_interior S)" apply (rule_tac x="d *\<^sub>R x" in image_eqI) using \0 < d\ using dx \closed S\ apply (auto simp: rel_frontier_def field_split_simps nox) done qed qed qed (rule inj_on_proj) then obtain surf where surf: "homeomorphism (S - rel_interior S) ?SPHER proj surf" by blast then have cont_surf: "continuous_on (proj ` (S - rel_interior S)) surf" by (auto simp: homeomorphism_def) have surf_nz: "\x. x \ ?SPHER \ surf x \ 0" by (metis "0" DiffE homeomorphism_def imageI surf) have cont_nosp: "continuous_on (?SPHER) (\x. norm x *\<^sub>R ((surf o proj) x))" apply (rule continuous_intros)+ apply (rule continuous_on_subset [OF cont_proj], force) apply (rule continuous_on_subset [OF cont_surf]) apply (force simp: homeomorphism_image1 [OF surf] dest: proj_spherI) done have surfpS: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ S" by (metis (full_types) DiffE \0 \ S\ homeomorphism_def image_eqI norm_zero proj_spherI real_vector.scale_zero_left scaleR_one surf) have *: "\y. norm y = 1 \ y \ affine hull S \ x = surf (proj y)" if "x \ S" "x \ rel_interior S" for x proof - have "proj x \ ?SPHER" by (metis (full_types) "0" hull_inc proj_spherI that) moreover have "surf (proj x) = x" by (metis Diff_iff homeomorphism_def surf that) ultimately show ?thesis by (metis \\x. x \ ?SPHER \ surf x \ 0\ hull_inc inverse_1 local.proj_def norm_sgn projI scaleR_one sgn_div_norm that(1)) qed have surfp_notin: "\x. \norm x = 1; x \ affine hull S\ \ surf (proj x) \ rel_interior S" by (metis (full_types) DiffE one_neq_zero homeomorphism_def image_eqI norm_zero proj_spherI surf) have no_sp_im: "(\x. norm x *\<^sub>R surf (proj x)) ` (?SPHER) = S - rel_interior S" by (auto simp: surfpS image_def Bex_def surfp_notin *) have inj_spher: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?SPHER" proof fix x y assume xy: "x \ ?SPHER" "y \ ?SPHER" and eq: " norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" then have "norm x = 1" "norm y = 1" "x \ affine hull S" "y \ affine hull S" using 0 by auto with eq show "x = y" by (simp add: proj_def) (metis surf xy homeomorphism_def) qed have co01: "compact ?SPHER" - by (simp add: closed_affine_hull compact_Int_closed) + by (simp add: compact_Int_closed) show "?SMINUS homeomorphic ?SPHER" apply (subst homeomorphic_sym) apply (rule homeomorphic_compact [OF co01 cont_nosp [unfolded o_def] no_sp_im inj_spher]) done have proj_scaleR: "\a x. 0 < a \ proj (a *\<^sub>R x) = proj x" by (simp add: proj_def) have cont_sp0: "continuous_on (affine hull S - {0}) (surf o proj)" apply (rule continuous_on_compose [OF continuous_on_subset [OF cont_proj]], force) apply (rule continuous_on_subset [OF cont_surf]) using homeomorphism_image1 proj_spherI surf by fastforce obtain B where "B>0" and B: "\x. x \ S \ norm x \ B" by (metis compact_imp_bounded \compact S\ bounded_pos_less less_eq_real_def) have cont_nosp: "continuous (at x within ?CBALL) (\x. norm x *\<^sub>R surf (proj x))" if "norm x \ 1" "x \ affine hull S" for x proof (cases "x=0") case True show ?thesis using True apply (simp add: continuous_within) apply (rule lim_null_scaleR_bounded [where B=B]) apply (simp_all add: tendsto_norm_zero eventually_at) apply (rule_tac x=B in exI) using B surfpS proj_def projI apply (auto simp: \B > 0\) done next case False then have "\\<^sub>F x in at x. (x \ affine hull S - {0}) = (x \ affine hull S)" apply (simp add: eventually_at) apply (rule_tac x="norm x" in exI) apply (auto simp: False) done with cont_sp0 have *: "continuous (at x within affine hull S) (\x. surf (proj x))" apply (simp add: continuous_on_eq_continuous_within) apply (drule_tac x=x in bspec, force simp: False that) apply (simp add: continuous_within Lim_transform_within_set) done show ?thesis apply (rule continuous_within_subset [where s = "affine hull S", OF _ Int_lower2]) apply (rule continuous_intros *)+ done qed have cont_nosp2: "continuous_on ?CBALL (\x. norm x *\<^sub>R ((surf o proj) x))" by (simp add: continuous_on_eq_continuous_within cont_nosp) have "norm y *\<^sub>R surf (proj y) \ S" if "y \ cball 0 1" and yaff: "y \ affine hull S" for y proof (cases "y=0") case True then show ?thesis by (simp add: \0 \ S\) next case False then have "norm y *\<^sub>R surf (proj y) = norm y *\<^sub>R surf (proj (y /\<^sub>R norm y))" by (simp add: proj_def) have "norm y \ 1" using that by simp have "surf (proj (y /\<^sub>R norm y)) \ S" apply (rule surfpS) using proj_def projI yaff by (auto simp: False) then have "surf (proj y) \ S" by (simp add: False proj_def) then show "norm y *\<^sub>R surf (proj y) \ S" by (metis dual_order.antisym le_less_linear norm_ge_zero rel_interior_subset scaleR_one starI subset_eq \norm y \ 1\) qed moreover have "x \ (\x. norm x *\<^sub>R surf (proj x)) ` (?CBALL)" if "x \ S" for x proof (cases "x=0") case True with that hull_inc show ?thesis by fastforce next case False then have psp: "proj (surf (proj x)) = proj x" by (metis homeomorphism_def hull_inc proj_spherI surf that) have nxx: "norm x *\<^sub>R proj x = x" by (simp add: False local.proj_def) have affineI: "(1 / norm (surf (proj x))) *\<^sub>R x \ affine hull S" by (metis \0 \ S\ affine_hull_span_0 hull_inc span_clauses(4) that) have sproj_nz: "surf (proj x) \ 0" by (metis False proj0_iff psp) then have "proj x = proj (proj x)" by (metis False nxx proj_scaleR zero_less_norm_iff) moreover have scaleproj: "\a r. r *\<^sub>R proj a = (r / norm a) *\<^sub>R a" by (simp add: divide_inverse local.proj_def) ultimately have "(norm (surf (proj x)) / norm x) *\<^sub>R x \ rel_interior S" by (metis (no_types) sproj_nz divide_self_if hull_inc norm_eq_zero nproj1 projI psp scaleR_one surfp_notin that) then have "(norm (surf (proj x)) / norm x) \ 1" using starI [OF that] by (meson starI [OF that] le_less_linear norm_ge_zero zero_le_divide_iff) then have nole: "norm x \ norm (surf (proj x))" by (simp add: le_divide_eq_1) show ?thesis apply (rule_tac x="inverse(norm(surf (proj x))) *\<^sub>R x" in image_eqI) apply (metis (no_types, hide_lams) mult.commute scaleproj abs_inverse abs_norm_cancel divide_inverse norm_scaleR nxx positive_imp_inverse_positive proj_scaleR psp sproj_nz zero_less_norm_iff) apply (auto simp: field_split_simps nole affineI) done qed ultimately have im_cball: "(\x. norm x *\<^sub>R surf (proj x)) ` ?CBALL = S" by blast have inj_cball: "inj_on (\x. norm x *\<^sub>R surf (proj x)) ?CBALL" proof fix x y assume "x \ ?CBALL" "y \ ?CBALL" and eq: "norm x *\<^sub>R surf (proj x) = norm y *\<^sub>R surf (proj y)" then have x: "x \ affine hull S" and y: "y \ affine hull S" using 0 by auto show "x = y" proof (cases "x=0 \ y=0") case True then show "x = y" using eq proj_spherI surf_nz x y by force next case False with x y have speq: "surf (proj x) = surf (proj y)" by (metis eq homeomorphism_apply2 proj_scaleR proj_spherI surf zero_less_norm_iff) then have "norm x = norm y" by (metis \x \ affine hull S\ \y \ affine hull S\ eq proj_spherI real_vector.scale_cancel_right surf_nz) moreover have "proj x = proj y" by (metis (no_types) False speq homeomorphism_apply2 proj_spherI surf x y) ultimately show "x = y" using eq eqI by blast qed qed have co01: "compact ?CBALL" - by (simp add: closed_affine_hull compact_Int_closed) + by (simp add: compact_Int_closed) show "S homeomorphic ?CBALL" apply (subst homeomorphic_sym) apply (rule homeomorphic_compact [OF co01 cont_nosp2 [unfolded o_def] im_cball inj_cball]) done qed corollary fixes S :: "'a::euclidean_space set" assumes "compact S" and a: "a \ rel_interior S" and star: "\x. x \ S \ open_segment a x \ rel_interior S" shows starlike_compact_projective1: "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" and starlike_compact_projective2: "S homeomorphic cball a 1 \ affine hull S" proof - have 1: "compact ((+) (-a) ` S)" by (meson assms compact_translation) have 2: "0 \ rel_interior ((+) (-a) ` S)" using a rel_interior_translation [of "- a" S] by (simp cong: image_cong_simp) have 3: "open_segment 0 x \ rel_interior ((+) (-a) ` S)" if "x \ ((+) (-a) ` S)" for x proof - have "x+a \ S" using that by auto then have "open_segment a (x+a) \ rel_interior S" by (metis star) then show ?thesis using open_segment_translation [of a 0 x] using rel_interior_translation [of "- a" S] by (fastforce simp add: ac_simps image_iff cong: image_cong_simp) qed have "S - rel_interior S homeomorphic ((+) (-a) ` S) - rel_interior ((+) (-a) ` S)" by (metis rel_interior_translation translation_diff homeomorphic_translation) also have "... homeomorphic sphere 0 1 \ affine hull ((+) (-a) ` S)" by (rule starlike_compact_projective1_0 [OF 1 2 3]) also have "... = (+) (-a) ` (sphere a 1 \ affine hull S)" by (metis affine_hull_translation left_minus sphere_translation translation_Int) also have "... homeomorphic sphere a 1 \ affine hull S" using homeomorphic_translation homeomorphic_sym by blast finally show "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" . have "S homeomorphic ((+) (-a) ` S)" by (metis homeomorphic_translation) also have "... homeomorphic cball 0 1 \ affine hull ((+) (-a) ` S)" by (rule starlike_compact_projective2_0 [OF 1 2 3]) also have "... = (+) (-a) ` (cball a 1 \ affine hull S)" by (metis affine_hull_translation left_minus cball_translation translation_Int) also have "... homeomorphic cball a 1 \ affine hull S" using homeomorphic_translation homeomorphic_sym by blast finally show "S homeomorphic cball a 1 \ affine hull S" . qed corollary starlike_compact_projective_special: assumes "compact S" and cb01: "cball (0::'a::euclidean_space) 1 \ S" and scale: "\x u. \x \ S; 0 \ u; u < 1\ \ u *\<^sub>R x \ S - frontier S" shows "S homeomorphic (cball (0::'a::euclidean_space) 1)" proof - have "ball 0 1 \ interior S" using cb01 interior_cball interior_mono by blast then have 0: "0 \ rel_interior S" by (meson centre_in_ball subsetD interior_subset_rel_interior le_numeral_extra(2) not_le) have [simp]: "affine hull S = UNIV" using \ball 0 1 \ interior S\ by (auto intro!: affine_hull_nonempty_interior) have star: "open_segment 0 x \ rel_interior S" if "x \ S" for x proof fix p assume "p \ open_segment 0 x" then obtain u where "x \ 0" and u: "0 \ u" "u < 1" and p: "u *\<^sub>R x = p" by (auto simp: in_segment) then show "p \ rel_interior S" using scale [OF that u] closure_subset frontier_def interior_subset_rel_interior by fastforce qed show ?thesis using starlike_compact_projective2_0 [OF \compact S\ 0 star] by simp qed lemma homeomorphic_convex_lemma: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "compact S" "convex T" "compact T" and affeq: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T) \ S homeomorphic T" proof (cases "rel_interior S = {} \ rel_interior T = {}") case True then show ?thesis by (metis Diff_empty affeq \convex S\ \convex T\ aff_dim_empty homeomorphic_empty rel_interior_eq_empty aff_dim_empty) next case False then obtain a b where a: "a \ rel_interior S" and b: "b \ rel_interior T" by auto have starS: "\x. x \ S \ open_segment a x \ rel_interior S" using rel_interior_closure_convex_segment a \convex S\ closure_subset subsetCE by blast have starT: "\x. x \ T \ open_segment b x \ rel_interior T" using rel_interior_closure_convex_segment b \convex T\ closure_subset subsetCE by blast let ?aS = "(+) (-a) ` S" and ?bT = "(+) (-b) ` T" have 0: "0 \ affine hull ?aS" "0 \ affine hull ?bT" by (metis a b subsetD hull_inc image_eqI left_minus rel_interior_subset)+ have subs: "subspace (span ?aS)" "subspace (span ?bT)" by (rule subspace_span)+ moreover have "dim (span ((+) (- a) ` S)) = dim (span ((+) (- b) ` T))" by (metis 0 aff_dim_translation_eq aff_dim_zero affeq dim_span nat_int) ultimately obtain f g where "linear f" "linear g" and fim: "f ` span ?aS = span ?bT" and gim: "g ` span ?bT = span ?aS" and fno: "\x. x \ span ?aS \ norm(f x) = norm x" and gno: "\x. x \ span ?bT \ norm(g x) = norm x" and gf: "\x. x \ span ?aS \ g(f x) = x" and fg: "\x. x \ span ?bT \ f(g x) = x" by (rule isometries_subspaces) blast have [simp]: "continuous_on A f" for A using \linear f\ linear_conv_bounded_linear linear_continuous_on by blast have [simp]: "continuous_on B g" for B using \linear g\ linear_conv_bounded_linear linear_continuous_on by blast have eqspanS: "affine hull ?aS = span ?aS" by (metis a affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) have eqspanT: "affine hull ?bT = span ?bT" by (metis b affine_hull_span_0 subsetD hull_inc image_eqI left_minus rel_interior_subset) have "S homeomorphic cball a 1 \ affine hull S" by (rule starlike_compact_projective2 [OF \compact S\ a starS]) also have "... homeomorphic (+) (-a) ` (cball a 1 \ affine hull S)" by (metis homeomorphic_translation) also have "... = cball 0 1 \ (+) (-a) ` (affine hull S)" by (auto simp: dist_norm) also have "... = cball 0 1 \ span ?aS" using eqspanS affine_hull_translation by blast also have "... homeomorphic cball 0 1 \ span ?bT" proof (rule homeomorphicI [where f=f and g=g]) show fim1: "f ` (cball 0 1 \ span ?aS) = cball 0 1 \ span ?bT" apply (rule subset_antisym) using fim fno apply (force simp:, clarify) by (metis IntI fg gim gno image_eqI mem_cball_0) show "g ` (cball 0 1 \ span ?bT) = cball 0 1 \ span ?aS" apply (rule subset_antisym) using gim gno apply (force simp:, clarify) by (metis IntI fim1 gf image_eqI) qed (auto simp: fg gf) also have "... = cball 0 1 \ (+) (-b) ` (affine hull T)" using eqspanT affine_hull_translation by blast also have "... = (+) (-b) ` (cball b 1 \ affine hull T)" by (auto simp: dist_norm) also have "... homeomorphic (cball b 1 \ affine hull T)" by (metis homeomorphic_translation homeomorphic_sym) also have "... homeomorphic T" by (metis starlike_compact_projective2 [OF \compact T\ b starT] homeomorphic_sym) finally have 1: "S homeomorphic T" . have "S - rel_interior S homeomorphic sphere a 1 \ affine hull S" by (rule starlike_compact_projective1 [OF \compact S\ a starS]) also have "... homeomorphic (+) (-a) ` (sphere a 1 \ affine hull S)" by (metis homeomorphic_translation) also have "... = sphere 0 1 \ (+) (-a) ` (affine hull S)" by (auto simp: dist_norm) also have "... = sphere 0 1 \ span ?aS" using eqspanS affine_hull_translation by blast also have "... homeomorphic sphere 0 1 \ span ?bT" proof (rule homeomorphicI [where f=f and g=g]) show fim1: "f ` (sphere 0 1 \ span ?aS) = sphere 0 1 \ span ?bT" apply (rule subset_antisym) using fim fno apply (force simp:, clarify) by (metis IntI fg gim gno image_eqI mem_sphere_0) show "g ` (sphere 0 1 \ span ?bT) = sphere 0 1 \ span ?aS" apply (rule subset_antisym) using gim gno apply (force simp:, clarify) by (metis IntI fim1 gf image_eqI) qed (auto simp: fg gf) also have "... = sphere 0 1 \ (+) (-b) ` (affine hull T)" using eqspanT affine_hull_translation by blast also have "... = (+) (-b) ` (sphere b 1 \ affine hull T)" by (auto simp: dist_norm) also have "... homeomorphic (sphere b 1 \ affine hull T)" by (metis homeomorphic_translation homeomorphic_sym) also have "... homeomorphic T - rel_interior T" by (metis starlike_compact_projective1 [OF \compact T\ b starT] homeomorphic_sym) finally have 2: "S - rel_interior S homeomorphic T - rel_interior T" . show ?thesis using 1 2 by blast qed lemma homeomorphic_convex_compact_sets: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "compact S" "convex T" "compact T" and affeq: "aff_dim S = aff_dim T" shows "S homeomorphic T" using homeomorphic_convex_lemma [OF assms] assms by (auto simp: rel_frontier_def) lemma homeomorphic_rel_frontiers_convex_bounded_sets: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "convex S" "bounded S" "convex T" "bounded T" and affeq: "aff_dim S = aff_dim T" shows "rel_frontier S homeomorphic rel_frontier T" using assms homeomorphic_convex_lemma [of "closure S" "closure T"] by (simp add: rel_frontier_def convex_rel_interior_closure) subsection\Homeomorphisms between punctured spheres and affine sets\ text\Including the famous stereoscopic projection of the 3-D sphere to the complex plane\ text\The special case with centre 0 and radius 1\ lemma homeomorphic_punctured_affine_sphere_affine_01: assumes "b \ sphere 0 1" "affine T" "0 \ T" "b \ T" "affine p" and affT: "aff_dim T = aff_dim p + 1" shows "(sphere 0 1 \ T) - {b} homeomorphic p" proof - have [simp]: "norm b = 1" "b\b = 1" using assms by (auto simp: norm_eq_1) have [simp]: "T \ {v. b\v = 0} \ {}" using \0 \ T\ by auto have [simp]: "\ T \ {v. b\v = 0}" using \norm b = 1\ \b \ T\ by auto define f where "f \ \x. 2 *\<^sub>R b + (2 / (1 - b\x)) *\<^sub>R (x - b)" define g where "g \ \y. b + (4 / (norm y ^ 2 + 4)) *\<^sub>R (y - 2 *\<^sub>R b)" have [simp]: "\x. \x \ T; b\x = 0\ \ f (g x) = x" unfolding f_def g_def by (simp add: algebra_simps field_split_simps add_nonneg_eq_0_iff) have no: "(norm (f x))\<^sup>2 = 4 * (1 + b \ x) / (1 - b \ x)" if "norm x = 1" and "b \ x \ 1" for x using that apply (simp flip: dot_square_norm add: norm_eq_1 nonzero_eq_divide_eq) apply (simp add: f_def vector_add_divide_simps inner_simps) apply (use sum_sqs_eq [of 1 "b \ x"] in \auto simp add: field_split_simps inner_commute\) done have [simp]: "\u::real. 8 + u * (u * 8) = u * 16 \ u=1" by algebra have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ g (f x) = x" unfolding g_def no by (auto simp: f_def field_split_simps) have [simp]: "norm (g x) = 1" if "x \ T" and "b \ x = 0" for x using that apply (simp only: g_def) apply (rule power2_eq_imp_eq) apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps) apply (simp add: algebra_simps inner_commute) done have [simp]: "b \ g x \ 1" if "x \ T" and "b \ x = 0" for x using that unfolding g_def apply (simp_all add: dot_square_norm [symmetric] divide_simps vector_add_divide_simps add_nonneg_eq_0_iff) apply (auto simp: algebra_simps) done have "subspace T" by (simp add: assms subspace_affine) have [simp]: "\x. \x \ T; b \ x = 0\ \ g x \ T" unfolding g_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) have "f ` {x. norm x = 1 \ b\x \ 1} \ {x. b\x = 0}" unfolding f_def using \norm b = 1\ norm_eq_1 by (force simp: field_simps inner_add_right inner_diff_right) moreover have "f ` T \ T" unfolding f_def using assms \subspace T\ by (auto simp add: inner_add_right inner_diff_right mem_affine_3_minus subspace_mul) moreover have "{x. b\x = 0} \ T \ f ` ({x. norm x = 1 \ b\x \ 1} \ T)" apply clarify apply (rule_tac x = "g x" in image_eqI, auto) done ultimately have imf: "f ` ({x. norm x = 1 \ b\x \ 1} \ T) = {x. b\x = 0} \ T" by blast have no4: "\y. b\y = 0 \ norm ((y\y + 4) *\<^sub>R b + 4 *\<^sub>R (y - 2 *\<^sub>R b)) = y\y + 4" apply (rule power2_eq_imp_eq) apply (simp_all add: dot_square_norm [symmetric]) apply (auto simp: power2_eq_square algebra_simps inner_commute) done have [simp]: "\x. \norm x = 1; b \ x \ 1\ \ b \ f x = 0" by (simp add: f_def algebra_simps field_split_simps) have [simp]: "\x. \x \ T; norm x = 1; b \ x \ 1\ \ f x \ T" unfolding f_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) have "g ` {x. b\x = 0} \ {x. norm x = 1 \ b\x \ 1}" unfolding g_def apply (clarsimp simp: no4 vector_add_divide_simps divide_simps add_nonneg_eq_0_iff dot_square_norm [symmetric]) apply (auto simp: algebra_simps) done moreover have "g ` T \ T" unfolding g_def by (blast intro: \subspace T\ \b \ T\ subspace_add subspace_mul subspace_diff) moreover have "{x. norm x = 1 \ b\x \ 1} \ T \ g ` ({x. b\x = 0} \ T)" apply clarify apply (rule_tac x = "f x" in image_eqI, auto) done ultimately have img: "g ` ({x. b\x = 0} \ T) = {x. norm x = 1 \ b\x \ 1} \ T" by blast have aff: "affine ({x. b\x = 0} \ T)" by (blast intro: affine_hyperplane assms) have contf: "continuous_on ({x. norm x = 1 \ b\x \ 1} \ T) f" unfolding f_def by (rule continuous_intros | force)+ have contg: "continuous_on ({x. b\x = 0} \ T) g" unfolding g_def by (rule continuous_intros | force simp: add_nonneg_eq_0_iff)+ have "(sphere 0 1 \ T) - {b} = {x. norm x = 1 \ (b\x \ 1)} \ T" using \norm b = 1\ by (auto simp: norm_eq_1) (metis vector_eq \b\b = 1\) also have "... homeomorphic {x. b\x = 0} \ T" by (rule homeomorphicI [OF imf img contf contg]) auto also have "... homeomorphic p" apply (rule homeomorphic_affine_sets [OF aff \affine p\]) apply (simp add: Int_commute aff_dim_affine_Int_hyperplane [OF \affine T\] affT) done finally show ?thesis . qed theorem homeomorphic_punctured_affine_sphere_affine: fixes a :: "'a :: euclidean_space" assumes "0 < r" "b \ sphere a r" "affine T" "a \ T" "b \ T" "affine p" and aff: "aff_dim T = aff_dim p + 1" shows "(sphere a r \ T) - {b} homeomorphic p" proof - have "a \ b" using assms by auto then have inj: "inj (\x::'a. x /\<^sub>R norm (a - b))" by (simp add: inj_on_def) have "((sphere a r \ T) - {b}) homeomorphic (+) (-a) ` ((sphere a r \ T) - {b})" by (rule homeomorphic_translation) also have "... homeomorphic (*\<^sub>R) (inverse r) ` (+) (- a) ` (sphere a r \ T - {b})" by (metis \0 < r\ homeomorphic_scaling inverse_inverse_eq inverse_zero less_irrefl) also have "... = sphere 0 1 \ ((*\<^sub>R) (inverse r) ` (+) (- a) ` T) - {(b - a) /\<^sub>R r}" using assms by (auto simp: dist_norm norm_minus_commute divide_simps) also have "... homeomorphic p" apply (rule homeomorphic_punctured_affine_sphere_affine_01) using assms affine_translation [symmetric, of "- a"] aff_dim_translation_eq [of "- a"] apply (auto simp: dist_norm norm_minus_commute affine_scaling inj) done finally show ?thesis . qed corollary homeomorphic_punctured_sphere_affine: fixes a :: "'a :: euclidean_space" assumes "0 < r" and b: "b \ sphere a r" and "affine T" and affS: "aff_dim T + 1 = DIM('a)" shows "(sphere a r - {b}) homeomorphic T" using homeomorphic_punctured_affine_sphere_affine [of r b a UNIV T] assms by auto corollary homeomorphic_punctured_sphere_hyperplane: fixes a :: "'a :: euclidean_space" assumes "0 < r" and b: "b \ sphere a r" and "c \ 0" shows "(sphere a r - {b}) homeomorphic {x::'a. c \ x = d}" apply (rule homeomorphic_punctured_sphere_affine) using assms apply (auto simp: affine_hyperplane of_nat_diff) done proposition homeomorphic_punctured_sphere_affine_gen: fixes a :: "'a :: euclidean_space" assumes "convex S" "bounded S" and a: "a \ rel_frontier S" and "affine T" and affS: "aff_dim S = aff_dim T + 1" shows "rel_frontier S - {a} homeomorphic T" proof - obtain U :: "'a set" where "affine U" "convex U" and affdS: "aff_dim U = aff_dim S" using choose_affine_subset [OF affine_UNIV aff_dim_geq] by (meson aff_dim_affine_hull affine_affine_hull affine_imp_convex) have "S \ {}" using assms by auto then obtain z where "z \ U" by (metis aff_dim_negative_iff equals0I affdS) then have bne: "ball z 1 \ U \ {}" by force then have [simp]: "aff_dim(ball z 1 \ U) = aff_dim U" using aff_dim_convex_Int_open [OF \convex U\ open_ball] by (fastforce simp add: Int_commute) have "rel_frontier S homeomorphic rel_frontier (ball z 1 \ U)" by (rule homeomorphic_rel_frontiers_convex_bounded_sets) (auto simp: \affine U\ affine_imp_convex convex_Int affdS assms) also have "... = sphere z 1 \ U" using convex_affine_rel_frontier_Int [of "ball z 1" U] by (simp add: \affine U\ bne) finally have "rel_frontier S homeomorphic sphere z 1 \ U" . then obtain h k where him: "h ` rel_frontier S = sphere z 1 \ U" and kim: "k ` (sphere z 1 \ U) = rel_frontier S" and hcon: "continuous_on (rel_frontier S) h" and kcon: "continuous_on (sphere z 1 \ U) k" and kh: "\x. x \ rel_frontier S \ k(h(x)) = x" and hk: "\y. y \ sphere z 1 \ U \ h(k(y)) = y" unfolding homeomorphic_def homeomorphism_def by auto have "rel_frontier S - {a} homeomorphic (sphere z 1 \ U) - {h a}" proof (rule homeomorphicI) show h: "h ` (rel_frontier S - {a}) = sphere z 1 \ U - {h a}" using him a kh by auto metis show "k ` (sphere z 1 \ U - {h a}) = rel_frontier S - {a}" by (force simp: h [symmetric] image_comp o_def kh) qed (auto intro: continuous_on_subset hcon kcon simp: kh hk) also have "... homeomorphic T" by (rule homeomorphic_punctured_affine_sphere_affine) (use a him in \auto simp: affS affdS \affine T\ \affine U\ \z \ U\\) finally show ?thesis . qed text\ When dealing with AR, ANR and ANR later, it's useful to know that every set is homeomorphic to a closed subset of a convex set, and if the set is locally compact we can take the convex set to be the universe.\ proposition homeomorphic_closedin_convex: fixes S :: "'m::euclidean_space set" assumes "aff_dim S < DIM('n)" obtains U and T :: "'n::euclidean_space set" where "convex U" "U \ {}" "closedin (top_of_set U) T" "S homeomorphic T" proof (cases "S = {}") case True then show ?thesis by (rule_tac U=UNIV and T="{}" in that) auto next case False then obtain a where "a \ S" by auto obtain i::'n where i: "i \ Basis" "i \ 0" using SOME_Basis Basis_zero by force have "0 \ affine hull ((+) (- a) ` S)" by (simp add: \a \ S\ hull_inc) then have "dim ((+) (- a) ` S) = aff_dim ((+) (- a) ` S)" by (simp add: aff_dim_zero) also have "... < DIM('n)" by (simp add: aff_dim_translation_eq_subtract assms cong: image_cong_simp) finally have dd: "dim ((+) (- a) ` S) < DIM('n)" by linarith have span: "span {x. i \ x = 0} = {x. i \ x = 0}" using span_eq_iff [symmetric, of "{x. i \ x = 0}"] subspace_hyperplane [of i] by simp have "dim ((+) (- a) ` S) \ dim {x. i \ x = 0}" using dd by (simp add: dim_hyperplane [OF \i \ 0\]) then obtain T where "subspace T" and Tsub: "T \ {x. i \ x = 0}" and dimT: "dim T = dim ((+) (- a) ` S)" by (rule choose_subspace_of_subspace) (simp add: span) have "subspace (span ((+) (- a) ` S))" using subspace_span by blast then obtain h k where "linear h" "linear k" and heq: "h ` span ((+) (- a) ` S) = T" and keq:"k ` T = span ((+) (- a) ` S)" and hinv [simp]: "\x. x \ span ((+) (- a) ` S) \ k(h x) = x" and kinv [simp]: "\x. x \ T \ h(k x) = x" apply (rule isometries_subspaces [OF _ \subspace T\]) apply (auto simp: dimT) done have hcont: "continuous_on A h" and kcont: "continuous_on B k" for A B using \linear h\ \linear k\ linear_continuous_on linear_conv_bounded_linear by blast+ have ihhhh[simp]: "\x. x \ S \ i \ h (x - a) = 0" using Tsub [THEN subsetD] heq span_superset by fastforce have "sphere 0 1 - {i} homeomorphic {x. i \ x = 0}" apply (rule homeomorphic_punctured_sphere_affine) using i apply (auto simp: affine_hyperplane) by (metis DIM_positive Suc_eq_plus1 add.left_neutral diff_add_cancel not_le not_less_eq_eq of_nat_1 of_nat_diff) then obtain f g where fg: "homeomorphism (sphere 0 1 - {i}) {x. i \ x = 0} f g" by (force simp: homeomorphic_def) have "h ` (+) (- a) ` S \ T" using heq span_superset span_linear_image by blast then have "g ` h ` (+) (- a) ` S \ g ` {x. i \ x = 0}" using Tsub by (simp add: image_mono) also have "... \ sphere 0 1 - {i}" by (simp add: fg [unfolded homeomorphism_def]) finally have gh_sub_sph: "(g \ h) ` (+) (- a) ` S \ sphere 0 1 - {i}" by (metis image_comp) then have gh_sub_cb: "(g \ h) ` (+) (- a) ` S \ cball 0 1" by (metis Diff_subset order_trans sphere_cball) have [simp]: "\u. u \ S \ norm (g (h (u - a))) = 1" using gh_sub_sph [THEN subsetD] by (auto simp: o_def) have ghcont: "continuous_on ((\x. x - a) ` S) (\x. g (h x))" apply (rule continuous_on_compose2 [OF homeomorphism_cont2 [OF fg] hcont], force) done have kfcont: "continuous_on ((\x. g (h (x - a))) ` S) (\x. k (f x))" apply (rule continuous_on_compose2 [OF kcont]) using homeomorphism_cont1 [OF fg] gh_sub_sph apply (force intro: continuous_on_subset, blast) done have "S homeomorphic (+) (- a) ` S" by (fact homeomorphic_translation) also have "\ homeomorphic (g \ h) ` (+) (- a) ` S" apply (simp add: homeomorphic_def homeomorphism_def cong: image_cong_simp) apply (rule_tac x="g \ h" in exI) apply (rule_tac x="k \ f" in exI) apply (auto simp: ghcont kfcont span_base homeomorphism_apply2 [OF fg] image_comp cong: image_cong_simp) done finally have Shom: "S homeomorphic (\x. g (h x)) ` (\x. x - a) ` S" by (simp cong: image_cong_simp) show ?thesis apply (rule_tac U = "ball 0 1 \ image (g o h) ((+) (- a) ` S)" and T = "image (g o h) ((+) (- a) ` S)" in that) apply (rule convex_intermediate_ball [of 0 1], force) using gh_sub_cb apply force apply force apply (simp add: closedin_closed) apply (rule_tac x="sphere 0 1" in exI) apply (auto simp: Shom cong: image_cong_simp) done qed subsection\Locally compact sets in an open set\ text\ Locally compact sets are closed in an open set and are homeomorphic to an absolutely closed set if we have one more dimension to play with.\ lemma locally_compact_open_Int_closure: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains T where "open T" "S = T \ closure S" proof - have "\x\S. \T v u. u = S \ T \ x \ u \ u \ v \ v \ S \ open T \ compact v" by (metis assms locally_compact openin_open) then obtain t v where tv: "\x. x \ S \ v x \ S \ open (t x) \ compact (v x) \ (\u. x \ u \ u \ v x \ u = S \ t x)" by metis then have o: "open (\(t ` S))" by blast have "S = \ (v ` S)" using tv by blast also have "... = \(t ` S) \ closure S" proof show "\(v ` S) \ \(t ` S) \ closure S" apply safe apply (metis Int_iff subsetD UN_iff tv) apply (simp add: closure_def rev_subsetD tv) done have "t x \ closure S \ v x" if "x \ S" for x proof - have "t x \ closure S \ closure (t x \ S)" by (simp add: open_Int_closure_subset that tv) also have "... \ v x" by (metis Int_commute closure_minimal compact_imp_closed that tv) finally show ?thesis . qed then show "\(t ` S) \ closure S \ \(v ` S)" by blast qed finally have e: "S = \(t ` S) \ closure S" . show ?thesis by (rule that [OF o e]) qed lemma locally_compact_closedin_open: fixes S :: "'a :: metric_space set" assumes "locally compact S" obtains T where "open T" "closedin (top_of_set T) S" by (metis locally_compact_open_Int_closure [OF assms] closed_closure closedin_closed_Int) lemma locally_compact_homeomorphism_projection_closed: assumes "locally compact S" obtains T and f :: "'a \ 'a :: euclidean_space \ 'b :: euclidean_space" where "closed T" "homeomorphism S T f fst" proof (cases "closed S") case True then show ?thesis apply (rule_tac T = "S \ {0}" and f = "\x. (x, 0)" in that) apply (auto simp: closed_Times homeomorphism_def continuous_intros) done next case False obtain U where "open U" and US: "U \ closure S = S" by (metis locally_compact_open_Int_closure [OF assms]) with False have Ucomp: "-U \ {}" using closure_eq by auto have [simp]: "closure (- U) = -U" by (simp add: \open U\ closed_Compl) define f :: "'a \ 'a \ 'b" where "f \ \x. (x, One /\<^sub>R setdist {x} (- U))" have "continuous_on U (\x. (x, One /\<^sub>R setdist {x} (- U)))" apply (intro continuous_intros continuous_on_setdist) by (simp add: Ucomp setdist_eq_0_sing_1) then have homU: "homeomorphism U (f`U) f fst" by (auto simp: f_def homeomorphism_def image_iff continuous_intros) have cloS: "closedin (top_of_set U) S" by (metis US closed_closure closedin_closed_Int) have cont: "isCont ((\x. setdist {x} (- U)) o fst) z" for z :: "'a \ 'b" by (rule continuous_at_compose continuous_intros continuous_at_setdist)+ have setdist1D: "setdist {a} (- U) *\<^sub>R b = One \ setdist {a} (- U) \ 0" for a::'a and b::'b by force have *: "r *\<^sub>R b = One \ b = (1 / r) *\<^sub>R One" for r and b::'b by (metis One_non_0 nonzero_divide_eq_eq real_vector.scale_eq_0_iff real_vector.scale_scale scaleR_one) have "f ` U = (\z. (setdist {fst z} (- U) *\<^sub>R snd z)) -` {One}" apply (auto simp: f_def setdist_eq_0_sing_1 field_simps Ucomp) apply (rule_tac x=a in image_eqI) apply (auto simp: * setdist_eq_0_sing_1 dest: setdist1D) done then have clfU: "closed (f ` U)" apply (rule ssubst) apply (rule continuous_closed_vimage) apply (auto intro: continuous_intros cont [unfolded o_def]) done have "closed (f ` S)" apply (rule closedin_closed_trans [OF _ clfU]) apply (rule homeomorphism_imp_closed_map [OF homU cloS]) done then show ?thesis apply (rule that) apply (rule homeomorphism_of_subsets [OF homU]) using US apply auto done qed lemma locally_compact_closed_Int_open: fixes S :: "'a :: euclidean_space set" shows "locally compact S \ (\U u. closed U \ open u \ S = U \ u)" by (metis closed_closure closed_imp_locally_compact inf_commute locally_compact_Int locally_compact_open_Int_closure open_imp_locally_compact) lemma lowerdim_embeddings: assumes "DIM('a) < DIM('b)" obtains f :: "'a::euclidean_space*real \ 'b::euclidean_space" and g :: "'b \ 'a*real" and j :: 'b where "linear f" "linear g" "\z. g (f z) = z" "j \ Basis" "\x. f(x,0) \ j = 0" proof - let ?B = "Basis :: ('a*real) set" have b01: "(0,1) \ ?B" by (simp add: Basis_prod_def) have "DIM('a * real) \ DIM('b)" by (simp add: Suc_leI assms) then obtain basf :: "'a*real \ 'b" where sbf: "basf ` ?B \ Basis" and injbf: "inj_on basf Basis" by (metis finite_Basis card_le_inj) define basg:: "'b \ 'a * real" where "basg \ \i. if i \ basf ` Basis then inv_into Basis basf i else (0,1)" have bgf[simp]: "basg (basf i) = i" if "i \ Basis" for i using inv_into_f_f injbf that by (force simp: basg_def) have sbg: "basg ` Basis \ ?B" by (force simp: basg_def injbf b01) define f :: "'a*real \ 'b" where "f \ \u. \j\Basis. (u \ basg j) *\<^sub>R j" define g :: "'b \ 'a*real" where "g \ \z. (\i\Basis. (z \ basf i) *\<^sub>R i)" show ?thesis proof show "linear f" unfolding f_def by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) show "linear g" unfolding g_def by (intro linear_compose_sum linearI ballI) (auto simp: algebra_simps) have *: "(\a \ Basis. a \ basf b * (x \ basg a)) = x \ b" if "b \ Basis" for x b using sbf that by auto show gf: "g (f x) = x" for x apply (rule euclidean_eqI) apply (simp add: f_def g_def inner_sum_left scaleR_sum_left algebra_simps) apply (simp add: Groups_Big.sum_distrib_left [symmetric] *) done show "basf(0,1) \ Basis" using b01 sbf by auto then show "f(x,0) \ basf(0,1) = 0" for x apply (simp add: f_def inner_sum_left) apply (rule comm_monoid_add_class.sum.neutral) using b01 inner_not_same_Basis by fastforce qed qed proposition locally_compact_homeomorphic_closed: fixes S :: "'a::euclidean_space set" assumes "locally compact S" and dimlt: "DIM('a) < DIM('b)" obtains T :: "'b::euclidean_space set" where "closed T" "S homeomorphic T" proof - obtain U:: "('a*real)set" and h where "closed U" and homU: "homeomorphism S U h fst" using locally_compact_homeomorphism_projection_closed assms by metis obtain f :: "'a*real \ 'b" and g :: "'b \ 'a*real" where "linear f" "linear g" and gf [simp]: "\z. g (f z) = z" using lowerdim_embeddings [OF dimlt] by metis then have "inj f" by (metis injI) have gfU: "g ` f ` U = U" by (simp add: image_comp o_def) have "S homeomorphic U" using homU homeomorphic_def by blast also have "... homeomorphic f ` U" apply (rule homeomorphicI [OF refl gfU]) apply (meson \inj f\ \linear f\ homeomorphism_cont2 linear_homeomorphism_image) using \linear g\ linear_continuous_on linear_conv_bounded_linear apply blast apply (auto simp: o_def) done finally show ?thesis apply (rule_tac T = "f ` U" in that) apply (rule closed_injective_linear_image [OF \closed U\ \linear f\ \inj f\], assumption) done qed lemma homeomorphic_convex_compact_lemma: fixes S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "cball 0 1 \ S" shows "S homeomorphic (cball (0::'a) 1)" proof (rule starlike_compact_projective_special[OF assms(2-3)]) fix x u assume "x \ S" and "0 \ u" and "u < (1::real)" have "open (ball (u *\<^sub>R x) (1 - u))" by (rule open_ball) moreover have "u *\<^sub>R x \ ball (u *\<^sub>R x) (1 - u)" unfolding centre_in_ball using \u < 1\ by simp moreover have "ball (u *\<^sub>R x) (1 - u) \ S" proof fix y assume "y \ ball (u *\<^sub>R x) (1 - u)" then have "dist (u *\<^sub>R x) y < 1 - u" unfolding mem_ball . with \u < 1\ have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ cball 0 1" by (simp add: dist_norm inverse_eq_divide norm_minus_commute) with assms(3) have "inverse (1 - u) *\<^sub>R (y - u *\<^sub>R x) \ S" .. with assms(1) have "(1 - u) *\<^sub>R ((y - u *\<^sub>R x) /\<^sub>R (1 - u)) + u *\<^sub>R x \ S" using \x \ S\ \0 \ u\ \u < 1\ [THEN less_imp_le] by (rule convexD_alt) then show "y \ S" using \u < 1\ by simp qed ultimately have "u *\<^sub>R x \ interior S" .. then show "u *\<^sub>R x \ S - frontier S" using frontier_def and interior_subset by auto qed proposition homeomorphic_convex_compact_cball: fixes e :: real and S :: "'a::euclidean_space set" assumes "convex S" and "compact S" and "interior S \ {}" and "e > 0" shows "S homeomorphic (cball (b::'a) e)" proof - obtain a where "a \ interior S" using assms(3) by auto then obtain d where "d > 0" and d: "cball a d \ S" unfolding mem_interior_cball by auto let ?d = "inverse d" and ?n = "0::'a" have "cball ?n 1 \ (\x. inverse d *\<^sub>R (x - a)) ` S" apply rule apply (rule_tac x="d *\<^sub>R x + a" in image_eqI) defer apply (rule d[unfolded subset_eq, rule_format]) using \d > 0\ unfolding mem_cball dist_norm apply (auto simp add: mult_right_le_one_le) done then have "(\x. inverse d *\<^sub>R (x - a)) ` S homeomorphic cball ?n 1" using homeomorphic_convex_compact_lemma[of "(\x. ?d *\<^sub>R -a + ?d *\<^sub>R x) ` S", OF convex_affinity compact_affinity] using assms(1,2) by (auto simp add: scaleR_right_diff_distrib) then show ?thesis apply (rule_tac homeomorphic_trans[OF _ homeomorphic_balls(2)[of 1 _ ?n]]) apply (rule homeomorphic_trans[OF homeomorphic_affinity[of "?d" S "?d *\<^sub>R -a"]]) using \d>0\ \e>0\ apply (auto simp add: scaleR_right_diff_distrib) done qed corollary homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "convex S" "compact S" "interior S \ {}" and "convex T" "compact T" "interior T \ {}" shows "S homeomorphic T" using assms by (meson zero_less_one homeomorphic_trans homeomorphic_convex_compact_cball homeomorphic_sym) lemma homeomorphic_closed_intervals: fixes a :: "'a::euclidean_space" and b and c :: "'a::euclidean_space" and d assumes "box a b \ {}" and "box c d \ {}" shows "(cbox a b) homeomorphic (cbox c d)" apply (rule homeomorphic_convex_compact) using assms apply auto done lemma homeomorphic_closed_intervals_real: fixes a::real and b and c::real and d assumes "aCovering spaces and lifting results for them\ definition\<^marker>\tag important\ covering_space :: "'a::topological_space set \ ('a \ 'b) \ 'b::topological_space set \ bool" where "covering_space c p S \ continuous_on c p \ p ` c = S \ (\x \ S. \T. x \ T \ openin (top_of_set S) T \ (\v. \v = c \ p -` T \ (\u \ v. openin (top_of_set c) u) \ pairwise disjnt v \ (\u \ v. \q. homeomorphism u T p q)))" lemma covering_space_imp_continuous: "covering_space c p S \ continuous_on c p" by (simp add: covering_space_def) lemma covering_space_imp_surjective: "covering_space c p S \ p ` c = S" by (simp add: covering_space_def) lemma homeomorphism_imp_covering_space: "homeomorphism S T f g \ covering_space S f T" apply (simp add: homeomorphism_def covering_space_def, clarify) apply (rule_tac x=T in exI, simp) apply (rule_tac x="{S}" in exI, auto) done lemma covering_space_local_homeomorphism: assumes "covering_space c p S" "x \ c" obtains T u q where "x \ T" "openin (top_of_set c) T" "p x \ u" "openin (top_of_set S) u" "homeomorphism T u p q" using assms apply (simp add: covering_space_def, clarify) apply (drule_tac x="p x" in bspec, force) by (metis IntI UnionE vimage_eq) lemma covering_space_local_homeomorphism_alt: assumes p: "covering_space c p S" and "y \ S" obtains x T U q where "p x = y" "x \ T" "openin (top_of_set c) T" "y \ U" "openin (top_of_set S) U" "homeomorphism T U p q" proof - obtain x where "p x = y" "x \ c" using assms covering_space_imp_surjective by blast show ?thesis apply (rule covering_space_local_homeomorphism [OF p \x \ c\]) using that \p x = y\ by blast qed proposition covering_space_open_map: fixes S :: "'a :: metric_space set" and T :: "'b :: metric_space set" assumes p: "covering_space c p S" and T: "openin (top_of_set c) T" shows "openin (top_of_set S) (p ` T)" proof - have pce: "p ` c = S" and covs: "\x. x \ S \ \X VS. x \ X \ openin (top_of_set S) X \ \VS = c \ p -` X \ (\u \ VS. openin (top_of_set c) u) \ pairwise disjnt VS \ (\u \ VS. \q. homeomorphism u X p q)" using p by (auto simp: covering_space_def) have "T \ c" by (metis openin_euclidean_subtopology_iff T) have "\X. openin (top_of_set S) X \ y \ X \ X \ p ` T" if "y \ p ` T" for y proof - have "y \ S" using \T \ c\ pce that by blast obtain U VS where "y \ U" and U: "openin (top_of_set S) U" and VS: "\VS = c \ p -` U" and openVS: "\V \ VS. openin (top_of_set c) V" and homVS: "\V. V \ VS \ \q. homeomorphism V U p q" using covs [OF \y \ S\] by auto obtain x where "x \ c" "p x \ U" "x \ T" "p x = y" apply simp using T [unfolded openin_euclidean_subtopology_iff] \y \ U\ \y \ p ` T\ by blast with VS obtain V where "x \ V" "V \ VS" by auto then obtain q where q: "homeomorphism V U p q" using homVS by blast then have ptV: "p ` (T \ V) = U \ q -` (T \ V)" using VS \V \ VS\ by (auto simp: homeomorphism_def) have ocv: "openin (top_of_set c) V" by (simp add: \V \ VS\ openVS) have "openin (top_of_set U) (U \ q -` (T \ V))" apply (rule continuous_on_open [THEN iffD1, rule_format]) using homeomorphism_def q apply blast using openin_subtopology_Int_subset [of c] q T unfolding homeomorphism_def by (metis inf.absorb_iff2 Int_commute ocv openin_euclidean_subtopology_iff) then have os: "openin (top_of_set S) (U \ q -` (T \ V))" using openin_trans [of U] by (simp add: Collect_conj_eq U) show ?thesis apply (rule_tac x = "p ` (T \ V)" in exI) apply (rule conjI) apply (simp only: ptV os) using \p x = y\ \x \ V\ \x \ T\ apply blast done qed with openin_subopen show ?thesis by blast qed lemma covering_space_lift_unique_gen: fixes f :: "'a::topological_space \ 'b::topological_space" fixes g1 :: "'a \ 'c::real_normed_vector" assumes cov: "covering_space c p S" and eq: "g1 a = g2 a" and f: "continuous_on T f" "f ` T \ S" and g1: "continuous_on T g1" "g1 ` T \ c" and fg1: "\x. x \ T \ f x = p(g1 x)" and g2: "continuous_on T g2" "g2 ` T \ c" and fg2: "\x. x \ T \ f x = p(g2 x)" and u_compt: "U \ components T" and "a \ U" "x \ U" shows "g1 x = g2 x" proof - have "U \ T" by (rule in_components_subset [OF u_compt]) define G12 where "G12 \ {x \ U. g1 x - g2 x = 0}" have "connected U" by (rule in_components_connected [OF u_compt]) have contu: "continuous_on U g1" "continuous_on U g2" using \U \ T\ continuous_on_subset g1 g2 by blast+ have o12: "openin (top_of_set U) G12" unfolding G12_def proof (subst openin_subopen, clarify) fix z assume z: "z \ U" "g1 z - g2 z = 0" obtain v w q where "g1 z \ v" and ocv: "openin (top_of_set c) v" and "p (g1 z) \ w" and osw: "openin (top_of_set S) w" and hom: "homeomorphism v w p q" apply (rule_tac x = "g1 z" in covering_space_local_homeomorphism [OF cov]) using \U \ T\ \z \ U\ g1(2) apply blast+ done have "g2 z \ v" using \g1 z \ v\ z by auto have gg: "U \ g -` v = U \ g -` (v \ g ` U)" for g by auto have "openin (top_of_set (g1 ` U)) (v \ g1 ` U)" using ocv \U \ T\ g1 by (fastforce simp add: openin_open) then have 1: "openin (top_of_set U) (U \ g1 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) have "openin (top_of_set (g2 ` U)) (v \ g2 ` U)" using ocv \U \ T\ g2 by (fastforce simp add: openin_open) then have 2: "openin (top_of_set U) (U \ g2 -` v)" unfolding gg by (blast intro: contu continuous_on_open [THEN iffD1, rule_format]) show "\T. openin (top_of_set U) T \ z \ T \ T \ {z \ U. g1 z - g2 z = 0}" using z apply (rule_tac x = "(U \ g1 -` v) \ (U \ g2 -` v)" in exI) apply (intro conjI) apply (rule openin_Int [OF 1 2]) using \g1 z \ v\ \g2 z \ v\ apply (force simp:, clarify) apply (metis \U \ T\ subsetD eq_iff_diff_eq_0 fg1 fg2 hom homeomorphism_def) done qed have c12: "closedin (top_of_set U) G12" unfolding G12_def by (intro continuous_intros continuous_closedin_preimage_constant contu) have "G12 = {} \ G12 = U" by (intro connected_clopen [THEN iffD1, rule_format] \connected U\ conjI o12 c12) with eq \a \ U\ have "\x. x \ U \ g1 x - g2 x = 0" by (auto simp: G12_def) then show ?thesis using \x \ U\ by force qed proposition covering_space_lift_unique: fixes f :: "'a::topological_space \ 'b::topological_space" fixes g1 :: "'a \ 'c::real_normed_vector" assumes "covering_space c p S" "g1 a = g2 a" "continuous_on T f" "f ` T \ S" "continuous_on T g1" "g1 ` T \ c" "\x. x \ T \ f x = p(g1 x)" "continuous_on T g2" "g2 ` T \ c" "\x. x \ T \ f x = p(g2 x)" "connected T" "a \ T" "x \ T" shows "g1 x = g2 x" using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv by blast lemma covering_space_locally: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes loc: "locally \ C" and cov: "covering_space C p S" and pim: "\T. \T \ C; \ T\ \ \(p ` T)" shows "locally \ S" proof - have "locally \ (p ` C)" apply (rule locally_open_map_image [OF loc]) using cov covering_space_imp_continuous apply blast using cov covering_space_imp_surjective covering_space_open_map apply blast by (simp add: pim) then show ?thesis using covering_space_imp_surjective [OF cov] by metis qed proposition covering_space_locally_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and pim: "\T. \T \ C; \ T\ \ \(p ` T)" and qim: "\q U. \U \ S; continuous_on U q; \ U\ \ \(q ` U)" shows "locally \ S \ locally \ C" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (rule locallyI) fix V x assume V: "openin (top_of_set C) V" and "x \ V" have "p x \ p ` C" by (metis IntE V \x \ V\ imageI openin_open) then obtain T \ where "p x \ T" and opeT: "openin (top_of_set S) T" and veq: "\\ = C \ p -` T" and ope: "\U\\. openin (top_of_set C) U" and hom: "\U\\. \q. homeomorphism U T p q" using cov unfolding covering_space_def by (blast intro: that) have "x \ \\" using V veq \p x \ T\ \x \ V\ openin_imp_subset by fastforce then obtain U where "x \ U" "U \ \" by blast then obtain q where opeU: "openin (top_of_set C) U" and q: "homeomorphism U T p q" using ope hom by blast with V have "openin (top_of_set C) (U \ V)" by blast then have UV: "openin (top_of_set S) (p ` (U \ V))" using cov covering_space_open_map by blast obtain W W' where opeW: "openin (top_of_set S) W" and "\ W'" "p x \ W" "W \ W'" and W'sub: "W' \ p ` (U \ V)" using locallyE [OF L UV] \x \ U\ \x \ V\ by blast then have "W \ T" by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans) show "\U Z. openin (top_of_set C) U \ \ Z \ x \ U \ U \ Z \ Z \ V" proof (intro exI conjI) have "openin (top_of_set T) W" by (meson opeW opeT openin_imp_subset openin_subset_trans \W \ T\) then have "openin (top_of_set U) (q ` W)" by (meson homeomorphism_imp_open_map homeomorphism_symD q) then show "openin (top_of_set C) (q ` W)" using opeU openin_trans by blast show "\ (q ` W')" by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \\ W'\ continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim) show "x \ q ` W" by (metis \p x \ W\ \x \ U\ homeomorphism_def imageI q) show "q ` W \ q ` W'" using \W \ W'\ by blast have "W' \ p ` V" using W'sub by blast then show "q ` W' \ V" using W'sub homeomorphism_apply1 [OF q] by auto qed qed next assume ?rhs then show ?lhs using cov covering_space_locally pim by blast qed lemma covering_space_locally_compact_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally compact S \ locally compact C" apply (rule covering_space_locally_eq [OF assms]) apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous) using compact_continuous_image by blast lemma covering_space_locally_connected_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally connected S \ locally connected C" apply (rule covering_space_locally_eq [OF assms]) apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) using connected_continuous_image by blast lemma covering_space_locally_path_connected_eq: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "covering_space C p S" shows "locally path_connected S \ locally path_connected C" apply (rule covering_space_locally_eq [OF assms]) apply (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous) using path_connected_continuous_image by blast lemma covering_space_locally_compact: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally compact C" "covering_space C p S" shows "locally compact S" using assms covering_space_locally_compact_eq by blast lemma covering_space_locally_connected: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally connected C" "covering_space C p S" shows "locally connected S" using assms covering_space_locally_connected_eq by blast lemma covering_space_locally_path_connected: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "locally path_connected C" "covering_space C p S" shows "locally path_connected S" using assms covering_space_locally_path_connected_eq by blast proposition covering_space_lift_homotopy: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and h :: "real \ 'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and conth: "continuous_on ({0..1} \ U) h" and him: "h ` ({0..1} \ U) \ S" and heq: "\y. y \ U \ h (0,y) = p(f y)" and contf: "continuous_on U f" and fim: "f ` U \ C" obtains k where "continuous_on ({0..1} \ U) k" "k ` ({0..1} \ U) \ C" "\y. y \ U \ k(0, y) = f y" "\z. z \ {0..1} \ U \ h z = p(k z)" proof - have "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..1} \ V) k \ k ` ({0..1} \ V) \ C \ (\z \ V. k(0, z) = f z) \ (\z \ {0..1} \ V. h z = p(k z))" if "y \ U" for y proof - obtain UU where UU: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s) \ (\\. \\ = C \ p -` UU s \ (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U (UU s) p q))" using cov unfolding covering_space_def by (metis (mono_tags)) then have ope: "\s. s \ S \ s \ (UU s) \ openin (top_of_set S) (UU s)" by blast have "\k n i. open k \ open n \ t \ k \ y \ n \ i \ S \ h ` (({0..1} \ k) \ (U \ n)) \ UU i" if "t \ {0..1}" for t proof - have hinS: "h (t, y) \ S" using \y \ U\ him that by blast then have "(t,y) \ ({0..1} \ U) \ h -` UU(h(t, y))" using \y \ U\ \t \ {0..1}\ by (auto simp: ope) moreover have ope_01U: "openin (top_of_set ({0..1} \ U)) (({0..1} \ U) \ h -` UU(h(t, y)))" using hinS ope continuous_on_open_gen [OF him] conth by blast ultimately obtain V W where opeV: "open V" and "t \ {0..1} \ V" "t \ {0..1} \ V" and opeW: "open W" and "y \ U" "y \ W" and VW: "({0..1} \ V) \ (U \ W) \ (({0..1} \ U) \ h -` UU(h(t, y)))" by (rule Times_in_interior_subtopology) (auto simp: openin_open) then show ?thesis using hinS by blast qed then obtain K NN X where K: "\t. t \ {0..1} \ open (K t)" and NN: "\t. t \ {0..1} \ open (NN t)" and inUS: "\t. t \ {0..1} \ t \ K t \ y \ NN t \ X t \ S" and him: "\t. t \ {0..1} \ h ` (({0..1} \ K t) \ (U \ NN t)) \ UU (X t)" by (metis (mono_tags)) obtain \ where "\ \ ((\i. K i \ NN i)) ` {0..1}" "finite \" "{0::real..1} \ {y} \ \\" proof (rule compactE) show "compact ({0::real..1} \ {y})" by (simp add: compact_Times) show "{0..1} \ {y} \ (\i\{0..1}. K i \ NN i)" using K inUS by auto show "\B. B \ (\i. K i \ NN i) ` {0..1} \ open B" using K NN by (auto simp: open_Times) qed blast then obtain tk where "tk \ {0..1}" "finite tk" and tk: "{0::real..1} \ {y} \ (\i \ tk. K i \ NN i)" by (metis (no_types, lifting) finite_subset_image) then have "tk \ {}" by auto define n where "n = \(NN ` tk)" have "y \ n" "open n" using inUS NN \tk \ {0..1}\ \finite tk\ by (auto simp: n_def open_INT subset_iff) obtain \ where "0 < \" and \: "\T. \T \ {0..1}; diameter T < \\ \ \B\K ` tk. T \ B" proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"]) show "K ` tk \ {}" using \tk \ {}\ by auto show "{0..1} \ \(K ` tk)" using tk by auto show "\B. B \ K ` tk \ open B" using \tk \ {0..1}\ K by auto qed auto obtain N::nat where N: "N > 1 / \" using reals_Archimedean2 by blast then have "N > 0" using \0 < \\ order.asym by force have *: "\V k. openin (top_of_set U) V \ y \ V \ continuous_on ({0..of_nat n / N} \ V) k \ k ` ({0..of_nat n / N} \ V) \ C \ (\z\V. k (0, z) = f z) \ (\z\{0..of_nat n / N} \ V. h z = p (k z))" if "n \ N" for n using that proof (induction n) case 0 show ?case apply (rule_tac x=U in exI) apply (rule_tac x="f \ snd" in exI) apply (intro conjI \y \ U\ continuous_intros continuous_on_subset [OF contf]) using fim apply (auto simp: heq) done next case (Suc n) then obtain V k where opeUV: "openin (top_of_set U) V" and "y \ V" and contk: "continuous_on ({0..real n / real N} \ V) k" and kim: "k ` ({0..real n / real N} \ V) \ C" and keq: "\z. z \ V \ k (0, z) = f z" and heq: "\z. z \ {0..real n / real N} \ V \ h z = p (k z)" using Suc_leD by auto have "n \ N" using Suc.prems by auto obtain t where "t \ tk" and t: "{real n / real N .. (1 + real n) / real N} \ K t" proof (rule bexE [OF \]) show "{real n / real N .. (1 + real n) / real N} \ {0..1}" - using Suc.prems by (auto simp: field_split_simps algebra_simps) + using Suc.prems by (auto simp: field_split_simps) show diameter_less: "diameter {real n / real N .. (1 + real n) / real N} < \" - using \0 < \\ N by (auto simp: field_split_simps algebra_simps) + using \0 < \\ N by (auto simp: field_split_simps) qed blast have t01: "t \ {0..1}" using \t \ tk\ \tk \ {0..1}\ by blast obtain \ where \: "\\ = C \ p -` UU (X t)" and opeC: "\U. U \ \ \ openin (top_of_set C) U" and "pairwise disjnt \" and homuu: "\U. U \ \ \ \q. homeomorphism U (UU (X t)) p q" using inUS [OF t01] UU by meson have n_div_N_in: "real n / real N \ {real n / real N .. (1 + real n) / real N}" - using N by (auto simp: field_split_simps algebra_simps) + using N by (auto simp: field_split_simps) with t have nN_in_kkt: "real n / real N \ K t" by blast have "k (real n / real N, y) \ C \ p -` UU (X t)" proof (simp, rule conjI) show "k (real n / real N, y) \ C" using \y \ V\ kim keq by force have "p (k (real n / real N, y)) = h (real n / real N, y)" by (simp add: \y \ V\ heq) also have "... \ h ` (({0..1} \ K t) \ (U \ NN t))" apply (rule imageI) using \y \ V\ t01 \n \ N\ apply (simp add: nN_in_kkt \y \ U\ inUS field_split_simps) done also have "... \ UU (X t)" using him t01 by blast finally show "p (k (real n / real N, y)) \ UU (X t)" . qed with \ have "k (real n / real N, y) \ \\" by blast then obtain W where W: "k (real n / real N, y) \ W" and "W \ \" by blast then obtain p' where opeC': "openin (top_of_set C) W" and hom': "homeomorphism W (UU (X t)) p p'" using homuu opeC by blast then have "W \ C" using openin_imp_subset by blast define W' where "W' = UU(X t)" have opeVW: "openin (top_of_set V) (V \ (k \ Pair (n / N)) -` W)" apply (rule continuous_openin_preimage [OF _ _ opeC']) apply (intro continuous_intros continuous_on_subset [OF contk]) using kim apply (auto simp: \y \ V\ W) done obtain N' where opeUN': "openin (top_of_set U) N'" and "y \ N'" and kimw: "k ` ({(real n / real N)} \ N') \ W" apply (rule_tac N' = "(V \ (k \ Pair (n / N)) -` W)" in that) apply (fastforce simp: \y \ V\ W intro!: openin_trans [OF opeVW opeUV])+ done obtain Q Q' where opeUQ: "openin (top_of_set U) Q" and cloUQ': "closedin (top_of_set U) Q'" and "y \ Q" "Q \ Q'" and Q': "Q' \ (U \ NN(t)) \ N' \ V" proof - obtain VO VX where "open VO" "open VX" and VO: "V = U \ VO" and VX: "N' = U \ VX" using opeUV opeUN' by (auto simp: openin_open) then have "open (NN(t) \ VO \ VX)" using NN t01 by blast then obtain e where "e > 0" and e: "cball y e \ NN(t) \ VO \ VX" by (metis Int_iff \N' = U \ VX\ \V = U \ VO\ \y \ N'\ \y \ V\ inUS open_contains_cball t01) show ?thesis proof show "openin (top_of_set U) (U \ ball y e)" by blast show "closedin (top_of_set U) (U \ cball y e)" using e by (auto simp: closedin_closed) qed (use \y \ U\ \e > 0\ VO VX e in auto) qed then have "y \ Q'" "Q \ (U \ NN(t)) \ N' \ V" by blast+ have neq: "{0..real n / real N} \ {real n / real N..(1 + real n) / real N} = {0..(1 + real n) / real N}" apply (auto simp: field_split_simps) by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1) then have neqQ': "{0..real n / real N} \ Q' \ {real n / real N..(1 + real n) / real N} \ Q' = {0..(1 + real n) / real N} \ Q'" by blast have cont: "continuous_on ({0..(1 + real n) / real N} \ Q') (\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" unfolding neqQ' [symmetric] proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply) show "closedin (top_of_set ({0..(1 + real n) / real N} \ Q')) ({0..real n / real N} \ Q')" apply (simp add: closedin_closed) apply (rule_tac x="{0 .. real n / real N} \ UNIV" in exI) using n_div_N_in apply (auto simp: closed_Times) done show "closedin (top_of_set ({0..(1 + real n) / real N} \ Q')) ({real n / real N..(1 + real n) / real N} \ Q')" apply (simp add: closedin_closed) apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \ UNIV" in exI) apply (auto simp: closed_Times) by (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) show "continuous_on ({0..real n / real N} \ Q') k" apply (rule continuous_on_subset [OF contk]) using Q' by auto have "continuous_on ({real n / real N..(1 + real n) / real N} \ Q') h" proof (rule continuous_on_subset [OF conth]) show "{real n / real N..(1 + real n) / real N} \ Q' \ {0..1} \ U" using \N > 0\ apply auto apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) using Suc.prems order_trans apply fastforce apply (metis IntE cloUQ' closedin_closed) done qed moreover have "continuous_on (h ` ({real n / real N..(1 + real n) / real N} \ Q')) p'" proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']]) have "h ` ({real n / real N..(1 + real n) / real N} \ Q') \ h ` (({0..1} \ K t) \ (U \ NN t))" apply (rule image_mono) using \0 < \\ \N > 0\ Suc.prems apply auto apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans) using Suc.prems order_trans apply fastforce using t Q' apply auto done with him show "h ` ({real n / real N..(1 + real n) / real N} \ Q') \ UU (X t)" using t01 by blast qed ultimately show "continuous_on ({real n / real N..(1 + real n) / real N} \ Q') (p' \ h)" by (rule continuous_on_compose) have "k (real n / real N, b) = p' (h (real n / real N, b))" if "b \ Q'" for b proof - have "k (real n / real N, b) \ W" using that Q' kimw by force then have "k (real n / real N, b) = p' (p (k (real n / real N, b)))" by (simp add: homeomorphism_apply1 [OF hom']) then show ?thesis using Q' that by (force simp: heq) qed then show "\x. x \ {real n / real N..(1 + real n) / real N} \ Q' \ x \ {0..real n / real N} \ Q' \ k x = (p' \ h) x" by auto qed have h_in_UU: "h (x, y) \ UU (X t)" if "y \ Q" "\ x \ real n / real N" "0 \ x" "x \ (1 + real n) / real N" for x y proof - have "x \ 1" using Suc.prems that order_trans by force moreover have "x \ K t" by (meson atLeastAtMost_iff le_less not_le subset_eq t that) moreover have "y \ U" using \y \ Q\ opeUQ openin_imp_subset by blast moreover have "y \ NN t" using Q' \Q \ Q'\ \y \ Q\ by auto ultimately have "(x, y) \ (({0..1} \ K t) \ (U \ NN t))" using that by auto then have "h (x, y) \ h ` (({0..1} \ K t) \ (U \ NN t))" by blast also have "... \ UU (X t)" by (metis him t01) finally show ?thesis . qed let ?k = "(\x. if x \ {0..real n / real N} \ Q' then k x else (p' \ h) x)" show ?case proof (intro exI conjI) show "continuous_on ({0..real (Suc n) / real N} \ Q) ?k" apply (rule continuous_on_subset [OF cont]) using \Q \ Q'\ by auto have "\a b. \a \ real n / real N; b \ Q'; 0 \ a\ \ k (a, b) \ C" using kim Q' by force moreover have "\a b. \b \ Q; 0 \ a; a \ (1 + real n) / real N; \ a \ real n / real N\ \ p' (h (a, b)) \ C" apply (rule \W \ C\ [THEN subsetD]) using homeomorphism_image2 [OF hom', symmetric] h_in_UU Q' \Q \ Q'\ \W \ C\ apply auto done ultimately show "?k ` ({0..real (Suc n) / real N} \ Q) \ C" using Q' \Q \ Q'\ by force show "\z\Q. ?k (0, z) = f z" using Q' keq \Q \ Q'\ by auto show "\z \ {0..real (Suc n) / real N} \ Q. h z = p(?k z)" using \Q \ U \ NN t \ N' \ V\ heq apply clarsimp using h_in_UU Q' \Q \ Q'\ apply (auto simp: homeomorphism_apply2 [OF hom', symmetric]) done qed (auto simp: \y \ Q\ opeUQ) qed show ?thesis using*[OF order_refl] N \0 < \\ by (simp add: split: if_split_asm) qed then obtain V fs where opeV: "\y. y \ U \ openin (top_of_set U) (V y)" and V: "\y. y \ U \ y \ V y" and contfs: "\y. y \ U \ continuous_on ({0..1} \ V y) (fs y)" and *: "\y. y \ U \ (fs y) ` ({0..1} \ V y) \ C \ (\z \ V y. fs y (0, z) = f z) \ (\z \ {0..1} \ V y. h z = p(fs y z))" by (metis (mono_tags)) then have VU: "\y. y \ U \ V y \ U" by (meson openin_imp_subset) obtain k where contk: "continuous_on ({0..1} \ U) k" and k: "\x i. \i \ U; x \ {0..1} \ U \ {0..1} \ V i\ \ k x = fs i x" proof (rule pasting_lemma_exists) let ?X = "top_of_set ({0..1::real} \ U)" show "topspace ?X \ (\i\U. {0..1} \ V i)" using V by force show "\i. i \ U \ openin (top_of_set ({0..1} \ U)) ({0..1} \ V i)" by (simp add: Abstract_Topology.openin_Times opeV) show "\i. i \ U \ continuous_map (subtopology (top_of_set ({0..1} \ U)) ({0..1} \ V i)) euclidean (fs i)" using contfs apply simp by (metis continuous_map_iff_continuous continuous_on_subset openin_imp_subset openin_subtopology_self subtopology_subtopology) show "fs i x = fs j x" if "i \ U" "j \ U" and x: "x \ topspace ?X \ {0..1} \ V i \ {0..1} \ V j" for i j x proof - obtain u y where "x = (u, y)" "y \ V i" "y \ V j" "0 \ u" "u \ 1" using x by auto show ?thesis proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \ {y}" h]) show "fs i (0, y) = fs j (0, y)" using*V by (simp add: \y \ V i\ \y \ V j\ that) show conth_y: "continuous_on ({0..1} \ {y}) h" apply (rule continuous_on_subset [OF conth]) using VU \y \ V j\ that by auto show "h ` ({0..1} \ {y}) \ S" using \y \ V i\ assms(3) VU that by fastforce show "continuous_on ({0..1} \ {y}) (fs i)" using continuous_on_subset [OF contfs] \i \ U\ by (simp add: \y \ V i\ subset_iff) show "fs i ` ({0..1} \ {y}) \ C" using "*" \y \ V i\ \i \ U\ by fastforce show "\x. x \ {0..1} \ {y} \ h x = p (fs i x)" using "*" \y \ V i\ \i \ U\ by blast show "continuous_on ({0..1} \ {y}) (fs j)" using continuous_on_subset [OF contfs] \j \ U\ by (simp add: \y \ V j\ subset_iff) show "fs j ` ({0..1} \ {y}) \ C" using "*" \y \ V j\ \j \ U\ by fastforce show "\x. x \ {0..1} \ {y} \ h x = p (fs j x)" using "*" \y \ V j\ \j \ U\ by blast show "connected ({0..1::real} \ {y})" using connected_Icc connected_Times connected_sing by blast show "(0, y) \ {0..1::real} \ {y}" by force show "x \ {0..1} \ {y}" using \x = (u, y)\ x by blast qed qed qed force show ?thesis proof show "k ` ({0..1} \ U) \ C" using V*k VU by fastforce show "\y. y \ U \ k (0, y) = f y" by (simp add: V*k) show "\z. z \ {0..1} \ U \ h z = p (k z)" using V*k by auto qed (auto simp: contk) qed corollary covering_space_lift_homotopy_alt: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and h :: "'c::real_normed_vector \ real \ 'b" assumes cov: "covering_space C p S" and conth: "continuous_on (U \ {0..1}) h" and him: "h ` (U \ {0..1}) \ S" and heq: "\y. y \ U \ h (y,0) = p(f y)" and contf: "continuous_on U f" and fim: "f ` U \ C" obtains k where "continuous_on (U \ {0..1}) k" "k ` (U \ {0..1}) \ C" "\y. y \ U \ k(y, 0) = f y" "\z. z \ U \ {0..1} \ h z = p(k z)" proof - have "continuous_on ({0..1} \ U) (h \ (\z. (snd z, fst z)))" by (intro continuous_intros continuous_on_subset [OF conth]) auto then obtain k where contk: "continuous_on ({0..1} \ U) k" and kim: "k ` ({0..1} \ U) \ C" and k0: "\y. y \ U \ k(0, y) = f y" and heqp: "\z. z \ {0..1} \ U \ (h \ (\z. Pair (snd z) (fst z))) z = p(k z)" apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim]) using him by (auto simp: contf heq) show ?thesis apply (rule_tac k="k \ (\z. Pair (snd z) (fst z))" in that) apply (intro continuous_intros continuous_on_subset [OF contk]) using kim heqp apply (auto simp: k0) done qed corollary covering_space_lift_homotopic_function: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and g:: "'c::real_normed_vector \ 'a" assumes cov: "covering_space C p S" and contg: "continuous_on U g" and gim: "g ` U \ C" and pgeq: "\y. y \ U \ p(g y) = f y" and hom: "homotopic_with_canon (\x. True) U S f f'" obtains g' where "continuous_on U g'" "image g' U \ C" "\y. y \ U \ p(g' y) = f' y" proof - obtain h where conth: "continuous_on ({0..1::real} \ U) h" and him: "h ` ({0..1} \ U) \ S" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = f' x" using hom by (auto simp: homotopic_with_def) have "\y. y \ U \ h (0, y) = p (g y)" by (simp add: h0 pgeq) then obtain k where contk: "continuous_on ({0..1} \ U) k" and kim: "k ` ({0..1} \ U) \ C" and k0: "\y. y \ U \ k(0, y) = g y" and heq: "\z. z \ {0..1} \ U \ h z = p(k z)" using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis show ?thesis proof show "continuous_on U (k \ Pair 1)" by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one) show "(k \ Pair 1) ` U \ C" using kim by auto show "\y. y \ U \ p ((k \ Pair 1) y) = f' y" by (auto simp: h1 heq [symmetric]) qed qed corollary covering_space_lift_inessential_function: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and U :: "'c::real_normed_vector set" assumes cov: "covering_space C p S" and hom: "homotopic_with_canon (\x. True) U S f (\x. a)" obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" proof (cases "U = {}") case True then show ?thesis using that continuous_on_empty by blast next case False then obtain b where b: "b \ C" "p b = a" using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom] by auto then have gim: "(\y. b) ` U \ C" by blast show ?thesis apply (rule covering_space_lift_homotopic_function [OF cov continuous_on_const gim _ homotopic_with_symD [OF hom]]) using b that apply auto done qed subsection\ Lifting of general functions to covering space\ proposition covering_space_lift_path_strong: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and "a \ C" and "path g" and pag: "path_image g \ S" and pas: "pathstart g = p a" obtains h where "path h" "path_image h \ C" "pathstart h = a" and "\t. t \ {0..1} \ p(h t) = g t" proof - obtain k:: "real \ 'c \ 'a" where contk: "continuous_on ({0..1} \ {undefined}) k" and kim: "k ` ({0..1} \ {undefined}) \ C" and k0: "k (0, undefined) = a" and pk: "\z. z \ {0..1} \ {undefined} \ p(k z) = (g \ fst) z" proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \ fst"]) show "continuous_on ({0..1::real} \ {undefined::'c}) (g \ fst)" apply (intro continuous_intros) using \path g\ by (simp add: path_def) show "(g \ fst) ` ({0..1} \ {undefined}) \ S" using pag by (auto simp: path_image_def) show "(g \ fst) (0, y) = p a" if "y \ {undefined}" for y::'c by (metis comp_def fst_conv pas pathstart_def) qed (use assms in auto) show ?thesis proof show "path (k \ (\t. Pair t undefined))" unfolding path_def by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto show "path_image (k \ (\t. (t, undefined))) \ C" using kim by (auto simp: path_image_def) show "pathstart (k \ (\t. (t, undefined))) = a" by (auto simp: pathstart_def k0) show "\t. t \ {0..1} \ p ((k \ (\t. (t, undefined))) t) = g t" by (auto simp: pk) qed qed corollary covering_space_lift_path: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \ S" obtains h where "path h" "path_image h \ C" "\t. t \ {0..1} \ p(h t) = g t" proof - obtain a where "a \ C" "pathstart g = p a" by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE) show ?thesis using covering_space_lift_path_strong [OF cov \a \ C\ \path g\ pig] by (metis \pathstart g = p a\ that) qed proposition covering_space_lift_homotopic_paths: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g1" and pig1: "path_image g1 \ S" and "path g2" and pig2: "path_image g2 \ S" and hom: "homotopic_paths S g1 g2" and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" and h1h2: "pathstart h1 = pathstart h2" shows "homotopic_paths C h1 h2" proof - obtain h :: "real \ real \ 'b" where conth: "continuous_on ({0..1} \ {0..1}) h" and him: "h ` ({0..1} \ {0..1}) \ S" and h0: "\x. h (0, x) = g1 x" and h1: "\x. h (1, x) = g2 x" and heq0: "\t. t \ {0..1} \ h (t, 0) = g1 0" and heq1: "\t. t \ {0..1} \ h (t, 1) = g1 1" using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def) obtain k where contk: "continuous_on ({0..1} \ {0..1}) k" and kim: "k ` ({0..1} \ {0..1}) \ C" and kh2: "\y. y \ {0..1} \ k (y, 0) = h2 0" and hpk: "\z. z \ {0..1} \ {0..1} \ h z = p (k z)" apply (rule covering_space_lift_homotopy_alt [OF cov conth him, of "\x. h2 0"]) using h1h2 ph1 ph2 apply (force simp: heq0 pathstart_def pathfinish_def) - using path_image_def pih2 continuous_on_const by fastforce+ + using path_image_def pih2 by fastforce+ have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2" using \path g1\ \path g2\ path_def by blast+ have g1im: "g1 ` {0..1} \ S" and g2im: "g2 ` {0..1} \ S" using path_image_def pig1 pig2 by auto have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2" using \path h1\ \path h2\ path_def by blast+ have h1im: "h1 ` {0..1} \ C" and h2im: "h2 ` {0..1} \ C" using path_image_def pih1 pih2 by auto show ?thesis unfolding homotopic_paths pathstart_def pathfinish_def proof (intro exI conjI ballI) show keqh1: "k(0, x) = h1 x" if "x \ {0..1}" for x proof (rule covering_space_lift_unique [OF cov _ contg1 g1im]) show "k (0,0) = h1 0" by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one) show "continuous_on {0..1} (\a. k (0, a))" by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g1 x = p (k (0, x))" by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl) qed (use conth1 h1im kim that in \auto simp: ph1\) show "k(1, x) = h2 x" if "x \ {0..1}" for x proof (rule covering_space_lift_unique [OF cov _ contg2 g2im]) show "k (1,0) = h2 0" by (metis atLeastAtMost_iff kh2 order_refl zero_le_one) show "continuous_on {0..1} (\a. k (1, a))" by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g2 x = p (k (1, x))" by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one) qed (use conth2 h2im kim that in \auto simp: ph2\) show "\t. t \ {0..1} \ (k \ Pair t) 0 = h1 0" by (metis comp_apply h1h2 kh2 pathstart_def) show "(k \ Pair t) 1 = h1 1" if "t \ {0..1}" for t proof (rule covering_space_lift_unique [OF cov, of "\a. (k \ Pair a) 1" 0 "\a. h1 1" "{0..1}" "\x. g1 1"]) show "(k \ Pair 0) 1 = h1 1" using keqh1 by auto show "continuous_on {0..1} (\a. (k \ Pair a) 1)" apply simp by (intro continuous_intros continuous_on_compose2 [OF contk]) auto show "\x. x \ {0..1} \ g1 1 = p ((k \ Pair x) 1)" using heq1 hpk by auto - qed (use contk kim g1im h1im that in \auto simp: ph1 continuous_on_const\) + qed (use contk kim g1im h1im that in \auto simp: ph1\) qed (use contk kim in auto) qed corollary covering_space_monodromy: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and "path g1" and pig1: "path_image g1 \ S" and "path g2" and pig2: "path_image g2 \ S" and hom: "homotopic_paths S g1 g2" and "path h1" and pih1: "path_image h1 \ C" and ph1: "\t. t \ {0..1} \ p(h1 t) = g1 t" and "path h2" and pih2: "path_image h2 \ C" and ph2: "\t. t \ {0..1} \ p(h2 t) = g2 t" and h1h2: "pathstart h1 = pathstart h2" shows "pathfinish h1 = pathfinish h2" using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish by blast corollary covering_space_lift_homotopic_path: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes cov: "covering_space C p S" and hom: "homotopic_paths S f f'" and "path g" and pig: "path_image g \ C" and a: "pathstart g = a" and b: "pathfinish g = b" and pgeq: "\t. t \ {0..1} \ p(g t) = f t" obtains g' where "path g'" "path_image g' \ C" "pathstart g' = a" "pathfinish g' = b" "\t. t \ {0..1} \ p(g' t) = f' t" proof (rule covering_space_lift_path_strong [OF cov, of a f']) show "a \ C" using a pig by auto show "path f'" "path_image f' \ S" using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+ show "pathstart f' = p a" by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one) qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig) proposition covering_space_lift_general: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and "a \ C" "z \ U" and U: "path_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ \ \q. path q \ path_image q \ C \ pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof - have *: "\g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ path h \ path_image h \ C \ pathstart h = a \ (\t \ {0..1}. p(h t) = f(g t))" if "y \ U" for y proof - obtain g where "path g" "path_image g \ U" and pastg: "pathstart g = z" and pafig: "pathfinish g = y" using U \z \ U\ \y \ U\ by (force simp: path_connected_def) obtain h where "path h" "path_image h \ C" "pathstart h = a" and "\t. t \ {0..1} \ p(h t) = (f \ g) t" proof (rule covering_space_lift_path_strong [OF cov \a \ C\]) show "path (f \ g)" using \path g\ \path_image g \ U\ contf continuous_on_subset path_continuous_image by blast show "path_image (f \ g) \ S" by (metis \path_image g \ U\ fim image_mono path_image_compose subset_trans) show "pathstart (f \ g) = p a" by (simp add: feq pastg pathstart_compose) qed auto then show ?thesis by (metis \path g\ \path_image g \ U\ comp_apply pafig pastg) qed have "\l. \g h. path g \ path_image g \ U \ pathstart g = z \ pathfinish g = y \ path h \ path_image h \ C \ pathstart h = a \ (\t \ {0..1}. p(h t) = f(g t)) \ pathfinish h = l" for y proof - have "pathfinish h = pathfinish h'" if g: "path g" "path_image g \ U" "pathstart g = z" "pathfinish g = y" and h: "path h" "path_image h \ C" "pathstart h = a" and phg: "\t. t \ {0..1} \ p(h t) = f(g t)" and g': "path g'" "path_image g' \ U" "pathstart g' = z" "pathfinish g' = y" and h': "path h'" "path_image h' \ C" "pathstart h' = a" and phg': "\t. t \ {0..1} \ p(h' t) = f(g' t)" for g h g' h' proof - obtain q where "path q" and piq: "path_image q \ C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a" and homS: "homotopic_paths S (f \ g +++ reversepath g') (p \ q)" using g g' hom [of "g +++ reversepath g'"] by (auto simp: subset_path_image_join) have papq: "path (p \ q)" using homS homotopic_paths_imp_path by blast have pipq: "path_image (p \ q) \ S" using homS homotopic_paths_imp_subset by blast obtain q' where "path q'" "path_image q' \ C" and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q" and pq'_eq: "\t. t \ {0..1} \ p (q' t) = (f \ g +++ reversepath g') t" using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \path q\ piq refl refl] by auto have "q' t = (h \ (*\<^sub>R) 2) t" if "0 \ t" "t \ 1/2" for t proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \ (*\<^sub>R) 2" "{0..1/2}" "f \ g \ (*\<^sub>R) 2" t]) show "q' 0 = (h \ (*\<^sub>R) 2) 0" - by (metis \pathstart q' = pathstart q\ comp_def g h pastq pathstart_def pth_4(2)) + by (metis \pathstart q' = pathstart q\ comp_def h pastq pathstart_def pth_4(2)) show "continuous_on {0..1/2} (f \ g \ (*\<^sub>R) 2)" apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \path g\] continuous_on_subset [OF contf]) using g(2) path_image_def by fastforce+ show "(f \ g \ (*\<^sub>R) 2) ` {0..1/2} \ S" using g(2) path_image_def fim by fastforce show "(h \ (*\<^sub>R) 2) ` {0..1/2} \ C" using h path_image_def by fastforce show "q' ` {0..1/2} \ C" using \path_image q' \ C\ path_image_def by fastforce show "\x. x \ {0..1/2} \ (f \ g \ (*\<^sub>R) 2) x = p (q' x)" by (auto simp: joinpaths_def pq'_eq) show "\x. x \ {0..1/2} \ (f \ g \ (*\<^sub>R) 2) x = p ((h \ (*\<^sub>R) 2) x)" by (simp add: phg) show "continuous_on {0..1/2} q'" by (simp add: continuous_on_path \path q'\) show "continuous_on {0..1/2} (h \ (*\<^sub>R) 2)" apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \path h\], force) done qed (use that in auto) moreover have "q' t = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \ 1" for t proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \ (\t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)" t]) show "q' 1 = (reversepath h' \ (\t. 2 *\<^sub>R t - 1)) 1" using h' \pathfinish q' = pathfinish q\ pafiq by (simp add: pathstart_def pathfinish_def reversepath_def) show "continuous_on {1/2<..1} (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1))" apply (intro continuous_intros continuous_on_compose continuous_on_path \path g'\ continuous_on_subset [OF contf]) using g' apply simp_all by (auto simp: path_image_def reversepath_def) show "(f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ S" using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def) show "q' ` {1/2<..1} \ C" using \path_image q' \ C\ path_image_def by fastforce show "(reversepath h' \ (\t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \ C" using h' by (simp add: path_image_def reversepath_def subset_eq) show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p (q' x)" by (auto simp: joinpaths_def pq'_eq) show "\x. x \ {1/2<..1} \ (f \ reversepath g' \ (\t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \ (\t. 2 *\<^sub>R t - 1)) x)" by (simp add: phg' reversepath_def) show "continuous_on {1/2<..1} q'" by (auto intro: continuous_on_path [OF \path q'\]) show "continuous_on {1/2<..1} (reversepath h' \ (\t. 2 *\<^sub>R t - 1))" apply (intro continuous_intros continuous_on_compose continuous_on_path \path h'\) using h' apply auto done qed (use that in auto) ultimately have "q' t = (h +++ reversepath h') t" if "0 \ t" "t \ 1" for t using that by (simp add: joinpaths_def) then have "path(h +++ reversepath h')" by (auto intro: path_eq [OF \path q'\]) then show ?thesis by (auto simp: \path h\ \path h'\) qed then show ?thesis by metis qed then obtain l :: "'c \ 'a" where l: "\y g h. \path g; path_image g \ U; pathstart g = z; pathfinish g = y; path h; path_image h \ C; pathstart h = a; \t. t \ {0..1} \ p(h t) = f(g t)\ \ pathfinish h = l y" by metis show ?thesis proof show pleq: "p (l y) = f y" if "y \ U" for y using*[OF \y \ U\] by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one) show "l z = a" using l [of "linepath z z" z "linepath a a"] by (auto simp: assms) show LC: "l ` U \ C" by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE) have "\T. openin (top_of_set U) T \ y \ T \ T \ U \ l -` X" if X: "openin (top_of_set C) X" and "y \ U" "l y \ X" for X y proof - have "X \ C" using X openin_euclidean_subtopology_iff by blast have "f y \ S" using fim \y \ U\ by blast then obtain W \ where WV: "f y \ W \ openin (top_of_set S) W \ (\\ = C \ p -` W \ (\U \ \. openin (top_of_set C) U) \ pairwise disjnt \ \ (\U \ \. \q. homeomorphism U W p q))" using cov by (force simp: covering_space_def) then have "l y \ \\" using \X \ C\ pleq that by auto then obtain W' where "l y \ W'" and "W' \ \" by blast with WV obtain p' where opeCW': "openin (top_of_set C) W'" and homUW': "homeomorphism W' W p p'" by blast then have contp': "continuous_on W p'" and p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+ obtain V where "y \ V" "y \ U" and fimW: "f ` V \ W" "V \ U" and "path_connected V" and opeUV: "openin (top_of_set U) V" proof - have "openin (top_of_set U) (U \ f -` W)" using WV contf continuous_on_open_gen fim by auto then show ?thesis using U WV apply (auto simp: locally_path_connected) apply (drule_tac x="U \ f -` W" in spec) apply (drule_tac x=y in spec) apply (auto simp: \y \ U\ intro: that) done qed have "W' \ C" "W \ S" using opeCW' WV openin_imp_subset by auto have p'im: "p' ` W \ W'" using homUW' homeomorphism_image2 by fastforce show ?thesis proof (intro exI conjI) have "openin (top_of_set S) (W \ p' -` (W' \ X))" proof (rule openin_trans) show "openin (top_of_set W) (W \ p' -` (W' \ X))" apply (rule continuous_openin_preimage [OF contp' p'im]) using X \W' \ C\ apply (auto simp: openin_open) done show "openin (top_of_set S) W" using WV by blast qed then show "openin (top_of_set U) (V \ (U \ (f -` (W \ (p' -` (W' \ X))))))" by (blast intro: opeUV openin_subtopology_self continuous_openin_preimage [OF contf fim]) have "p' (f y) \ X" using \l y \ W'\ homeomorphism_apply1 [OF homUW'] pleq \y \ U\ \l y \ X\ by fastforce then show "y \ V \ (U \ f -` (W \ p' -` (W' \ X)))" using \y \ U\ \y \ V\ WV p'im by auto show "V \ (U \ f -` (W \ p' -` (W' \ X))) \ U \ l -` X" proof (intro subsetI IntI; clarify) fix y' assume y': "y' \ V" "y' \ U" "f y' \ W" "p' (f y') \ W'" "p' (f y') \ X" then obtain \ where "path \" "path_image \ \ V" "pathstart \ = y" "pathfinish \ = y'" by (meson \path_connected V\ \y \ V\ path_connected_def) obtain pp qq where "path pp" "path_image pp \ U" "pathstart pp = z" "pathfinish pp = y" "path qq" "path_image qq \ C" "pathstart qq = a" and pqqeq: "\t. t \ {0..1} \ p(qq t) = f(pp t)" using*[OF \y \ U\] by blast have finW: "\x. \0 \ x; x \ 1\ \ f (\ x) \ W" using \path_image \ \ V\ by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD]) have "pathfinish (qq +++ (p' \ f \ \)) = l y'" proof (rule l [of "pp +++ \" y' "qq +++ (p' \ f \ \)"]) show "path (pp +++ \)" by (simp add: \path \\ \path pp\ \pathfinish pp = y\ \pathstart \ = y\) show "path_image (pp +++ \) \ U" using \V \ U\ \path_image \ \ V\ \path_image pp \ U\ not_in_path_image_join by blast show "pathstart (pp +++ \) = z" by (simp add: \pathstart pp = z\) show "pathfinish (pp +++ \) = y'" by (simp add: \pathfinish \ = y'\) have paqq: "pathfinish qq = pathstart (p' \ f \ \)" apply (simp add: \pathstart \ = y\ pathstart_compose) apply (metis (mono_tags, lifting) \l y \ W'\ \path pp\ \path qq\ \path_image pp \ U\ \path_image qq \ C\ \pathfinish pp = y\ \pathstart pp = z\ \pathstart qq = a\ homeomorphism_apply1 [OF homUW'] l pleq pqqeq \y \ U\) done have "continuous_on (path_image \) (p' \ f)" proof (rule continuous_on_compose) show "continuous_on (path_image \) f" using \path_image \ \ V\ \V \ U\ contf continuous_on_subset by blast show "continuous_on (f ` path_image \) p'" apply (rule continuous_on_subset [OF contp']) apply (auto simp: path_image_def pathfinish_def pathstart_def finW) done qed then show "path (qq +++ (p' \ f \ \))" using \path \\ \path qq\ paqq path_continuous_image path_join_imp by blast show "path_image (qq +++ (p' \ f \ \)) \ C" apply (rule subset_path_image_join) apply (simp add: \path_image qq \ C\) by (metis \W' \ C\ \path_image \ \ V\ dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose) show "pathstart (qq +++ (p' \ f \ \)) = a" by (simp add: \pathstart qq = a\) show "p ((qq +++ (p' \ f \ \)) \) = f ((pp +++ \) \)" if \: "\ \ {0..1}" for \ proof (simp add: joinpaths_def, safe) show "p (qq (2*\)) = f (pp (2*\))" if "\*2 \ 1" using \\ \ {0..1}\ pqqeq that by auto show "p (p' (f (\ (2*\ - 1)))) = f (\ (2*\ - 1))" if "\ \*2 \ 1" apply (rule homeomorphism_apply2 [OF homUW' finW]) using that \ by auto qed qed with \pathfinish \ = y'\ \p' (f y') \ X\ show "y' \ l -` X" unfolding pathfinish_join by (simp add: pathfinish_def) qed qed qed then show "continuous_on U l" by (metis IntD1 IntD2 vimage_eq openin_subopen continuous_on_open_gen [OF LC]) qed qed corollary covering_space_lift_stronger: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" "a \ C" "z \ U" and U: "path_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" and hom: "\r. \path r; path_image r \ U; pathstart r = z; pathfinish r = z\ \ \b. homotopic_paths S (f \ r) (linepath b b)" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof (rule covering_space_lift_general [OF cov U contf fim feq]) fix r assume "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" then obtain b where b: "homotopic_paths S (f \ r) (linepath b b)" using hom by blast then have "f (pathstart r) = b" by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath) then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" by (simp add: b \pathstart r = z\) then have "homotopic_paths S (f \ r) (p \ linepath a a)" by (simp add: o_def feq linepath_def) then show "\q. path q \ path_image q \ C \ pathstart q = a \ pathfinish q = a \ homotopic_paths S (f \ r) (p \ q)" by (force simp: \a \ C\) qed auto corollary covering_space_lift_strong: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" "a \ C" "z \ U" and scU: "simply_connected U" and lpcU: "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" and feq: "f z = p a" obtains g where "continuous_on U g" "g ` U \ C" "g z = a" "\y. y \ U \ p(g y) = f y" proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq]) show "path_connected U" using scU simply_connected_eq_contractible_loop_some by blast fix r assume r: "path r" "path_image r \ U" "pathstart r = z" "pathfinish r = z" have "linepath (f z) (f z) = f \ linepath z z" by (simp add: o_def linepath_def) then have "homotopic_paths S (f \ r) (linepath (f z) (f z))" by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path) then show "\b. homotopic_paths S (f \ r) (linepath b b)" by blast qed blast corollary covering_space_lift: fixes p :: "'a::real_normed_vector \ 'b::real_normed_vector" and f :: "'c::real_normed_vector \ 'b" assumes cov: "covering_space C p S" and U: "simply_connected U" "locally path_connected U" and contf: "continuous_on U f" and fim: "f ` U \ S" obtains g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" proof (cases "U = {}") case True with that show ?thesis by auto next case False then obtain z where "z \ U" by blast then obtain a where "a \ C" "f z = p a" by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff) then show ?thesis by (metis that covering_space_lift_strong [OF cov _ \z \ U\ U contf fim]) qed end diff --git a/src/HOL/Analysis/Homotopy.thy b/src/HOL/Analysis/Homotopy.thy --- a/src/HOL/Analysis/Homotopy.thy +++ b/src/HOL/Analysis/Homotopy.thy @@ -1,5679 +1,5677 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Homotopy of Maps\ theory Homotopy imports Path_Connected Continuum_Not_Denumerable Product_Topology begin definition\<^marker>\tag important\ homotopic_with where "homotopic_with P X Y f g \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ (\x. h(0, x) = f x) \ (\x. h(1, x) = g x) \ (\t \ {0..1}. P(\x. h(t,x))))" text\\p\, \q\ are functions \X \ Y\, and the property \P\ restricts all intermediate maps. We often just want to require that \P\ fixes some subset, but to include the case of a loop homotopy, it is convenient to have a general property \P\.\ abbreviation homotopic_with_canon :: "[('a::topological_space \ 'b::topological_space) \ bool, 'a set, 'b set, 'a \ 'b, 'a \ 'b] \ bool" where "homotopic_with_canon P S T p q \ homotopic_with P (top_of_set S) (top_of_set T) p q" lemma split_01: "{0..1::real} = {0..1/2} \ {1/2..1}" by force lemma split_01_prod: "{0..1::real} \ X = ({0..1/2} \ X) \ ({1/2..1} \ X)" by force lemma image_Pair_const: "(\x. (x, c)) ` A = A \ {c}" by auto lemma fst_o_paired [simp]: "fst \ (\(x,y). (f x y, g x y)) = (\(x,y). f x y)" by auto lemma snd_o_paired [simp]: "snd \ (\(x,y). (f x y, g x y)) = (\(x,y). g x y)" by auto lemma continuous_on_o_Pair: "\continuous_on (T \ X) h; t \ T\ \ continuous_on X (h \ Pair t)" by (fast intro: continuous_intros elim!: continuous_on_subset) lemma continuous_map_o_Pair: assumes h: "continuous_map (prod_topology X Y) Z h" and t: "t \ topspace X" shows "continuous_map Y Z (h \ Pair t)" apply (intro continuous_map_compose [OF _ h] continuous_map_id [unfolded id_def] continuous_intros) apply (simp add: t) done subsection\<^marker>\tag unimportant\\Trivial properties\ text \We often want to just localize the ending function equality or whatever.\ text\<^marker>\tag important\ \%whitespace\ proposition homotopic_with: assumes "\h k. (\x. x \ topspace X \ h x = k x) \ (P h \ P k)" shows "homotopic_with P X Y p q \ (\h. continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h \ (\x \ topspace X. h(0,x) = p x) \ (\x \ topspace X. h(1,x) = q x) \ (\t \ {0..1}. P(\x. h(t, x))))" unfolding homotopic_with_def apply (rule iffI, blast, clarify) apply (rule_tac x="\(u,v). if v \ topspace X then h(u,v) else if u = 0 then p v else q v" in exI) apply auto using continuous_map_eq apply fastforce apply (drule_tac x=t in bspec, force) apply (subst assms; simp) done lemma homotopic_with_mono: assumes hom: "homotopic_with P X Y f g" and Q: "\h. \continuous_map X Y h; P h\ \ Q h" shows "homotopic_with Q X Y f g" using hom apply (simp add: homotopic_with_def) apply (erule ex_forward) apply (force simp: o_def dest: continuous_map_o_Pair intro: Q) done lemma homotopic_with_imp_continuous_maps: assumes "homotopic_with P X Y f g" shows "continuous_map X Y f \ continuous_map X Y g" proof - obtain h where conth: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y h" and h: "\x. h (0, x) = f x" "\x. h (1, x) = g x" using assms by (auto simp: homotopic_with_def) have *: "t \ {0..1} \ continuous_map X Y (h \ (\x. (t,x)))" for t by (rule continuous_map_compose [OF _ conth]) (simp add: o_def continuous_map_pairwise) show ?thesis using h *[of 0] *[of 1] by (simp add: continuous_map_eq) qed lemma homotopic_with_imp_continuous: assumes "homotopic_with_canon P X Y f g" shows "continuous_on X f \ continuous_on X g" by (meson assms continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) lemma homotopic_with_imp_property: assumes "homotopic_with P X Y f g" shows "P f \ P g" proof obtain h where h: "\x. h(0, x) = f x" "\x. h(1, x) = g x" and P: "\t. t \ {0..1::real} \ P(\x. h(t,x))" using assms by (force simp: homotopic_with_def) show "P f" "P g" using P [of 0] P [of 1] by (force simp: h)+ qed lemma homotopic_with_equal: assumes "P f" "P g" and contf: "continuous_map X Y f" and fg: "\x. x \ topspace X \ f x = g x" shows "homotopic_with P X Y f g" unfolding homotopic_with_def proof (intro exI conjI allI ballI) let ?h = "\(t::real,x). if t = 1 then g x else f x" show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y ?h" proof (rule continuous_map_eq) show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X) Y (f \ snd)" by (simp add: contf continuous_map_of_snd) qed (auto simp: fg) show "P (\x. ?h (t, x))" if "t \ {0..1}" for t by (cases "t = 1") (simp_all add: assms) qed auto lemma homotopic_with_imp_subset1: "homotopic_with_canon P X Y f g \ f ` X \ Y" by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) lemma homotopic_with_imp_subset2: "homotopic_with_canon P X Y f g \ g ` X \ Y" by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) lemma homotopic_with_subset_left: "\homotopic_with_canon P X Y f g; Z \ X\ \ homotopic_with_canon P Z Y f g" apply (simp add: homotopic_with_def) apply (fast elim!: continuous_on_subset ex_forward) done lemma homotopic_with_subset_right: "\homotopic_with_canon P X Y f g; Y \ Z\ \ homotopic_with_canon P X Z f g" apply (simp add: homotopic_with_def) apply (fast elim!: continuous_on_subset ex_forward) done subsection\Homotopy with P is an equivalence relation\ text \(on continuous functions mapping X into Y that satisfy P, though this only affects reflexivity)\ lemma homotopic_with_refl [simp]: "homotopic_with P X Y f f \ continuous_map X Y f \ P f" by (auto simp: homotopic_with_imp_continuous_maps intro: homotopic_with_equal dest: homotopic_with_imp_property) lemma homotopic_with_symD: assumes "homotopic_with P X Y f g" shows "homotopic_with P X Y g f" proof - let ?I01 = "subtopology euclideanreal {0..1}" let ?j = "\y. (1 - fst y, snd y)" have 1: "continuous_map (prod_topology ?I01 X) (prod_topology euclideanreal X) ?j" apply (intro continuous_intros) apply (simp_all add: prod_topology_subtopology continuous_map_from_subtopology [OF continuous_map_fst]) done have *: "continuous_map (prod_topology ?I01 X) (prod_topology ?I01 X) ?j" proof - have "continuous_map (prod_topology ?I01 X) (subtopology (prod_topology euclideanreal X) ({0..1} \ topspace X)) ?j" by (simp add: continuous_map_into_subtopology [OF 1] image_subset_iff) then show ?thesis by (simp add: prod_topology_subtopology(1)) qed show ?thesis using assms apply (clarsimp simp add: homotopic_with_def) apply (rename_tac h) apply (rule_tac x="h \ (\y. (1 - fst y, snd y))" in exI) apply (simp add: continuous_map_compose [OF *]) done qed lemma homotopic_with_sym: "homotopic_with P X Y f g \ homotopic_with P X Y g f" by (metis homotopic_with_symD) proposition homotopic_with_trans: assumes "homotopic_with P X Y f g" "homotopic_with P X Y g h" shows "homotopic_with P X Y f h" proof - let ?X01 = "prod_topology (subtopology euclideanreal {0..1}) X" obtain k1 k2 where contk1: "continuous_map ?X01 Y k1" and contk2: "continuous_map ?X01 Y k2" and k12: "\x. k1 (1, x) = g x" "\x. k2 (0, x) = g x" "\x. k1 (0, x) = f x" "\x. k2 (1, x) = h x" and P: "\t\{0..1}. P (\x. k1 (t, x))" "\t\{0..1}. P (\x. k2 (t, x))" using assms by (auto simp: homotopic_with_def) define k where "k \ \y. if fst y \ 1/2 then (k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y else (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" have keq: "k1 (2 * u, v) = k2 (2 * u -1, v)" if "u = 1/2" for u v by (simp add: k12 that) show ?thesis unfolding homotopic_with_def proof (intro exI conjI) show "continuous_map ?X01 Y k" unfolding k_def proof (rule continuous_map_cases_le) show fst: "continuous_map ?X01 euclideanreal fst" using continuous_map_fst continuous_map_in_subtopology by blast show "continuous_map ?X01 euclideanreal (\x. 1/2)" by simp show "continuous_map (subtopology ?X01 {y \ topspace ?X01. fst y \ 1/2}) Y (k1 \ (\x. (2 *\<^sub>R fst x, snd x)))" apply (rule fst continuous_map_compose [OF _ contk1] continuous_intros continuous_map_into_subtopology | simp)+ apply (intro continuous_intros fst continuous_map_from_subtopology) apply (force simp: prod_topology_subtopology) using continuous_map_snd continuous_map_from_subtopology by blast show "continuous_map (subtopology ?X01 {y \ topspace ?X01. 1/2 \ fst y}) Y (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x)))" apply (rule fst continuous_map_compose [OF _ contk2] continuous_intros continuous_map_into_subtopology | simp)+ apply (rule continuous_intros fst continuous_map_from_subtopology | simp)+ - apply (force simp: topspace_subtopology prod_topology_subtopology) + apply (force simp: prod_topology_subtopology) using continuous_map_snd continuous_map_from_subtopology by blast show "(k1 \ (\x. (2 *\<^sub>R fst x, snd x))) y = (k2 \ (\x. (2 *\<^sub>R fst x -1, snd x))) y" if "y \ topspace ?X01" and "fst y = 1/2" for y using that by (simp add: keq) qed show "\x. k (0, x) = f x" by (simp add: k12 k_def) show "\x. k (1, x) = h x" by (simp add: k12 k_def) show "\t\{0..1}. P (\x. k (t, x))" using P apply (clarsimp simp add: k_def) apply (case_tac "t \ 1/2", auto) done qed qed lemma homotopic_with_id2: "(\x. x \ topspace X \ g (f x) = x) \ homotopic_with (\x. True) X X (g \ f) id" by (metis comp_apply continuous_map_id eq_id_iff homotopic_with_equal homotopic_with_symD) subsection\Continuity lemmas\ lemma homotopic_with_compose_continuous_map_left: "\homotopic_with p X1 X2 f g; continuous_map X2 X3 h; \j. p j \ q(h \ j)\ \ homotopic_with q X1 X3 (h \ f) (h \ g)" unfolding homotopic_with_def apply clarify apply (rename_tac k) apply (rule_tac x="h \ k" in exI) apply (rule conjI continuous_map_compose | simp add: o_def)+ done lemma homotopic_compose_continuous_map_left: "\homotopic_with (\k. True) X1 X2 f g; continuous_map X2 X3 h\ \ homotopic_with (\k. True) X1 X3 (h \ f) (h \ g)" by (simp add: homotopic_with_compose_continuous_map_left) lemma homotopic_with_compose_continuous_map_right: assumes hom: "homotopic_with p X2 X3 f g" and conth: "continuous_map X1 X2 h" and q: "\j. p j \ q(j \ h)" shows "homotopic_with q X1 X3 (f \ h) (g \ h)" proof - obtain k where contk: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) X3 k" and k: "\x. k (0, x) = f x" "\x. k (1, x) = g x" and p: "\t. t\{0..1} \ p (\x. k (t, x))" using hom unfolding homotopic_with_def by blast have hsnd: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X2 (h \ snd)" by (rule continuous_map_compose [OF continuous_map_snd conth]) let ?h = "k \ (\(t,x). (t,h x))" show ?thesis unfolding homotopic_with_def proof (intro exI conjI allI ballI) have "continuous_map (prod_topology (top_of_set {0..1}) X1) (prod_topology (top_of_set {0..1::real}) X2) (\(t, x). (t, h x))" by (metis (mono_tags, lifting) case_prod_beta' comp_def continuous_map_eq continuous_map_fst continuous_map_pairedI hsnd) then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) X3 ?h" by (intro conjI continuous_map_compose [OF _ contk]) show "q (\x. ?h (t, x))" if "t \ {0..1}" for t using q [OF p [OF that]] by (simp add: o_def) qed (auto simp: k) qed lemma homotopic_compose_continuous_map_right: "\homotopic_with (\k. True) X2 X3 f g; continuous_map X1 X2 h\ \ homotopic_with (\k. True) X1 X3 (f \ h) (g \ h)" by (meson homotopic_with_compose_continuous_map_right) corollary homotopic_compose: shows "\homotopic_with (\x. True) X Y f f'; homotopic_with (\x. True) Y Z g g'\ \ homotopic_with (\x. True) X Z (g \ f) (g' \ f')" apply (rule homotopic_with_trans [where g = "g \ f'"]) apply (simp add: homotopic_compose_continuous_map_left homotopic_with_imp_continuous_maps) by (simp add: homotopic_compose_continuous_map_right homotopic_with_imp_continuous_maps) proposition homotopic_with_compose_continuous_right: "\homotopic_with_canon (\f. p (f \ h)) X Y f g; continuous_on W h; h ` W \ X\ \ homotopic_with_canon p W Y (f \ h) (g \ h)" apply (clarsimp simp add: homotopic_with_def) apply (rename_tac k) apply (rule_tac x="k \ (\y. (fst y, h (snd y)))" in exI) apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ apply (erule continuous_on_subset) apply (fastforce simp: o_def)+ done proposition homotopic_compose_continuous_right: "\homotopic_with_canon (\f. True) X Y f g; continuous_on W h; h ` W \ X\ \ homotopic_with_canon (\f. True) W Y (f \ h) (g \ h)" using homotopic_with_compose_continuous_right by fastforce proposition homotopic_with_compose_continuous_left: "\homotopic_with_canon (\f. p (h \ f)) X Y f g; continuous_on Y h; h ` Y \ Z\ \ homotopic_with_canon p X Z (h \ f) (h \ g)" apply (clarsimp simp add: homotopic_with_def) apply (rename_tac k) apply (rule_tac x="h \ k" in exI) apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ apply (erule continuous_on_subset) apply (fastforce simp: o_def)+ done proposition homotopic_compose_continuous_left: "\homotopic_with_canon (\_. True) X Y f g; continuous_on Y h; h ` Y \ Z\ \ homotopic_with_canon (\f. True) X Z (h \ f) (h \ g)" using homotopic_with_compose_continuous_left by fastforce lemma homotopic_from_subtopology: "homotopic_with P X X' f g \ homotopic_with P (subtopology X s) X' f g" unfolding homotopic_with_def apply (erule ex_forward) by (simp add: continuous_map_from_subtopology prod_topology_subtopology(2)) lemma homotopic_on_emptyI: assumes "topspace X = {}" "P f" "P g" shows "homotopic_with P X X' f g" unfolding homotopic_with_def proof (intro exI conjI ballI) show "P (\x. (\(t,x). if t = 0 then f x else g x) (t, x))" if "t \ {0..1}" for t::real by (cases "t = 0", auto simp: assms) qed (auto simp: continuous_map_atin assms) lemma homotopic_on_empty: "topspace X = {} \ (homotopic_with P X X' f g \ P f \ P g)" using homotopic_on_emptyI homotopic_with_imp_property by metis lemma homotopic_with_canon_on_empty [simp]: "homotopic_with_canon (\x. True) {} t f g" by (auto intro: homotopic_with_equal) lemma homotopic_constant_maps: "homotopic_with (\x. True) X X' (\x. a) (\x. b) \ topspace X = {} \ path_component_of X' a b" (is "?lhs = ?rhs") proof (cases "topspace X = {}") case False then obtain c where c: "c \ topspace X" by blast have "\g. continuous_map (top_of_set {0..1::real}) X' g \ g 0 = a \ g 1 = b" if "x \ topspace X" and hom: "homotopic_with (\x. True) X X' (\x. a) (\x. b)" for x proof - obtain h :: "real \ 'a \ 'b" where conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X' h" and h: "\x. h (0, x) = a" "\x. h (1, x) = b" using hom by (auto simp: homotopic_with_def) have cont: "continuous_map (top_of_set {0..1}) X' (h \ (\t. (t, c)))" apply (rule continuous_map_compose [OF _ conth]) apply (rule continuous_intros c | simp)+ done then show ?thesis by (force simp: h) qed moreover have "homotopic_with (\x. True) X X' (\x. g 0) (\x. g 1)" if "x \ topspace X" "a = g 0" "b = g 1" "continuous_map (top_of_set {0..1}) X' g" for x and g :: "real \ 'b" unfolding homotopic_with_def by (force intro!: continuous_map_compose continuous_intros c that) ultimately show ?thesis using False by (auto simp: path_component_of_def pathin_def) qed (simp add: homotopic_on_empty) proposition homotopic_with_eq: assumes h: "homotopic_with P X Y f g" and f': "\x. x \ topspace X \ f' x = f x" and g': "\x. x \ topspace X \ g' x = g x" and P: "(\h k. (\x. x \ topspace X \ h x = k x) \ P h \ P k)" shows "homotopic_with P X Y f' g'" using h unfolding homotopic_with_def apply safe apply (rule_tac x="\(u,v). if v \ topspace X then h(u,v) else if u = 0 then f' v else g' v" in exI) apply (simp add: f' g', safe) apply (fastforce intro: continuous_map_eq) apply (subst P; fastforce) done lemma homotopic_with_prod_topology: assumes "homotopic_with p X1 Y1 f f'" and "homotopic_with q X2 Y2 g g'" and r: "\i j. \p i; q j\ \ r(\(x,y). (i x, j y))" shows "homotopic_with r (prod_topology X1 X2) (prod_topology Y1 Y2) (\z. (f(fst z),g(snd z))) (\z. (f'(fst z), g'(snd z)))" proof - obtain h where h: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X1) Y1 h" and h0: "\x. h (0, x) = f x" and h1: "\x. h (1, x) = f' x" and p: "\t. \0 \ t; t \ 1\ \ p (\x. h (t,x))" using assms unfolding homotopic_with_def by auto obtain k where k: "continuous_map (prod_topology (subtopology euclideanreal {0..1}) X2) Y2 k" and k0: "\x. k (0, x) = g x" and k1: "\x. k (1, x) = g' x" and q: "\t. \0 \ t; t \ 1\ \ q (\x. k (t,x))" using assms unfolding homotopic_with_def by auto let ?hk = "\(t,x,y). (h(t,x), k(t,y))" show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (prod_topology X1 X2)) (prod_topology Y1 Y2) ?hk" unfolding continuous_map_pairwise case_prod_unfold by (rule conjI continuous_map_pairedI continuous_intros continuous_map_id [unfolded id_def] continuous_map_fst_of [unfolded o_def] continuous_map_snd_of [unfolded o_def] continuous_map_compose [OF _ h, unfolded o_def] continuous_map_compose [OF _ k, unfolded o_def])+ next fix x show "?hk (0, x) = (f (fst x), g (snd x))" "?hk (1, x) = (f' (fst x), g' (snd x))" by (auto simp: case_prod_beta h0 k0 h1 k1) qed (auto simp: p q r) qed lemma homotopic_with_product_topology: assumes ht: "\i. i \ I \ homotopic_with (p i) (X i) (Y i) (f i) (g i)" and pq: "\h. (\i. i \ I \ p i (h i)) \ q(\x. (\i\I. h i (x i)))" shows "homotopic_with q (product_topology X I) (product_topology Y I) (\z. (\i\I. (f i) (z i))) (\z. (\i\I. (g i) (z i)))" proof - obtain h where h: "\i. i \ I \ continuous_map (prod_topology (subtopology euclideanreal {0..1}) (X i)) (Y i) (h i)" and h0: "\i x. i \ I \ h i (0, x) = f i x" and h1: "\i x. i \ I \ h i (1, x) = g i x" and p: "\i t. \i \ I; t \ {0..1}\ \ p i (\x. h i (t,x))" using ht unfolding homotopic_with_def by metis show ?thesis unfolding homotopic_with_def proof (intro conjI allI exI) let ?h = "\(t,z). \i\I. h i (t,z i)" have "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (Y i) (\x. h i (fst x, snd x i))" if "i \ I" for i unfolding continuous_map_pairwise case_prod_unfold apply (intro conjI that continuous_intros continuous_map_compose [OF _ h, unfolded o_def]) using continuous_map_componentwise continuous_map_snd that apply fastforce done then show "continuous_map (prod_topology (subtopology euclideanreal {0..1}) (product_topology X I)) (product_topology Y I) ?h" by (auto simp: continuous_map_componentwise case_prod_beta) show "?h (0, x) = (\i\I. f i (x i))" "?h (1, x) = (\i\I. g i (x i))" for x by (auto simp: case_prod_beta h0 h1) show "\t\{0..1}. q (\x. ?h (t, x))" by (force intro: p pq) qed qed text\Homotopic triviality implicitly incorporates path-connectedness.\ lemma homotopic_triviality: shows "(\f g. continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T \ homotopic_with_canon (\x. True) S T f g) \ (S = {} \ path_connected T) \ (\f. continuous_on S f \ f ` S \ T \ (\c. homotopic_with_canon (\x. True) S T f (\x. c)))" (is "?lhs = ?rhs") proof (cases "S = {} \ T = {}") case True then show ?thesis by (auto simp: homotopic_on_emptyI) next case False show ?thesis proof assume LHS [rule_format]: ?lhs have pab: "path_component T a b" if "a \ T" "b \ T" for a b proof - have "homotopic_with_canon (\x. True) S T (\x. a) (\x. b)" - by (simp add: LHS continuous_on_const image_subset_iff that) + by (simp add: LHS image_subset_iff that) then show ?thesis using False homotopic_constant_maps [of "top_of_set S" "top_of_set T" a b] by auto qed moreover have "\c. homotopic_with_canon (\x. True) S T f (\x. c)" if "continuous_on S f" "f ` S \ T" for f using False LHS continuous_on_const that by blast ultimately show ?rhs by (simp add: path_connected_component) next assume RHS: ?rhs with False have T: "path_connected T" by blast show ?lhs proof clarify fix f g assume "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" obtain c d where c: "homotopic_with_canon (\x. True) S T f (\x. c)" and d: "homotopic_with_canon (\x. True) S T g (\x. d)" using False \continuous_on S f\ \f ` S \ T\ RHS \continuous_on S g\ \g ` S \ T\ by blast then have "c \ T" "d \ T" using False homotopic_with_imp_continuous_maps by fastforce+ with T have "path_component T c d" using path_connected_component by blast then have "homotopic_with_canon (\x. True) S T (\x. c) (\x. d)" by (simp add: homotopic_constant_maps) with c d show "homotopic_with_canon (\x. True) S T f g" by (meson homotopic_with_symD homotopic_with_trans) qed qed qed subsection\Homotopy of paths, maintaining the same endpoints\ definition\<^marker>\tag important\ homotopic_paths :: "['a set, real \ 'a, real \ 'a::topological_space] \ bool" where "homotopic_paths s p q \ homotopic_with_canon (\r. pathstart r = pathstart p \ pathfinish r = pathfinish p) {0..1} s p q" lemma homotopic_paths: "homotopic_paths s p q \ (\h. continuous_on ({0..1} \ {0..1}) h \ h ` ({0..1} \ {0..1}) \ s \ (\x \ {0..1}. h(0,x) = p x) \ (\x \ {0..1}. h(1,x) = q x) \ (\t \ {0..1::real}. pathstart(h \ Pair t) = pathstart p \ pathfinish(h \ Pair t) = pathfinish p))" by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) proposition homotopic_paths_imp_pathstart: "homotopic_paths s p q \ pathstart p = pathstart q" by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) proposition homotopic_paths_imp_pathfinish: "homotopic_paths s p q \ pathfinish p = pathfinish q" by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) lemma homotopic_paths_imp_path: "homotopic_paths s p q \ path p \ path q" using homotopic_paths_def homotopic_with_imp_continuous_maps path_def continuous_map_subtopology_eu by blast lemma homotopic_paths_imp_subset: "homotopic_paths s p q \ path_image p \ s \ path_image q \ s" by (metis (mono_tags) continuous_map_subtopology_eu homotopic_paths_def homotopic_with_imp_continuous_maps path_image_def) proposition homotopic_paths_refl [simp]: "homotopic_paths s p p \ path p \ path_image p \ s" by (simp add: homotopic_paths_def path_def path_image_def) proposition homotopic_paths_sym: "homotopic_paths s p q \ homotopic_paths s q p" by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) proposition homotopic_paths_sym_eq: "homotopic_paths s p q \ homotopic_paths s q p" by (metis homotopic_paths_sym) proposition homotopic_paths_trans [trans]: assumes "homotopic_paths s p q" "homotopic_paths s q r" shows "homotopic_paths s p r" proof - have "pathstart q = pathstart p" "pathfinish q = pathfinish p" using assms by (simp_all add: homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish) then have "homotopic_with_canon (\f. pathstart f = pathstart p \ pathfinish f = pathfinish p) {0..1} s q r" using \homotopic_paths s q r\ homotopic_paths_def by force then show ?thesis using assms homotopic_paths_def homotopic_with_trans by blast qed proposition homotopic_paths_eq: "\path p; path_image p \ s; \t. t \ {0..1} \ p t = q t\ \ homotopic_paths s p q" apply (simp add: homotopic_paths_def) apply (rule homotopic_with_eq) - apply (auto simp: path_def homotopic_with_refl pathstart_def pathfinish_def path_image_def elim: continuous_on_eq) + apply (auto simp: path_def pathstart_def pathfinish_def path_image_def elim: continuous_on_eq) done proposition homotopic_paths_reparametrize: assumes "path p" and pips: "path_image p \ s" and contf: "continuous_on {0..1} f" and f01:"f ` {0..1} \ {0..1}" and [simp]: "f(0) = 0" "f(1) = 1" and q: "\t. t \ {0..1} \ q(t) = p(f t)" shows "homotopic_paths s p q" proof - have contp: "continuous_on {0..1} p" by (metis \path p\ path_def) then have "continuous_on {0..1} (p \ f)" using contf continuous_on_compose continuous_on_subset f01 by blast then have "path q" by (simp add: path_def) (metis q continuous_on_cong) have piqs: "path_image q \ s" by (metis (no_types, hide_lams) pips f01 image_subset_iff path_image_def q) have fb0: "\a b. \0 \ a; a \ 1; 0 \ b; b \ 1\ \ 0 \ (1 - a) * f b + a * b" using f01 by force have fb1: "\0 \ a; a \ 1; 0 \ b; b \ 1\ \ (1 - a) * f b + a * b \ 1" for a b using f01 [THEN subsetD, of "f b"] by (simp add: convex_bound_le) have "homotopic_paths s q p" proof (rule homotopic_paths_trans) show "homotopic_paths s q (p \ f)" using q by (force intro: homotopic_paths_eq [OF \path q\ piqs]) next show "homotopic_paths s (p \ f) p" apply (simp add: homotopic_paths_def homotopic_with_def) apply (rule_tac x="p \ (\y. (1 - (fst y)) *\<^sub>R ((f \ snd) y) + (fst y) *\<^sub>R snd y)" in exI) apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ using pips [unfolded path_image_def] apply (auto simp: fb0 fb1 pathstart_def pathfinish_def) done qed then show ?thesis by (simp add: homotopic_paths_sym) qed lemma homotopic_paths_subset: "\homotopic_paths s p q; s \ t\ \ homotopic_paths t p q" unfolding homotopic_paths by fast text\ A slightly ad-hoc but useful lemma in constructing homotopies.\ lemma homotopic_join_lemma: fixes q :: "[real,real] \ 'a::topological_space" assumes p: "continuous_on ({0..1} \ {0..1}) (\y. p (fst y) (snd y))" and q: "continuous_on ({0..1} \ {0..1}) (\y. q (fst y) (snd y))" and pf: "\t. t \ {0..1} \ pathfinish(p t) = pathstart(q t)" shows "continuous_on ({0..1} \ {0..1}) (\y. (p(fst y) +++ q(fst y)) (snd y))" proof - have 1: "(\y. p (fst y) (2 * snd y)) = (\y. p (fst y) (snd y)) \ (\y. (fst y, 2 * snd y))" by (rule ext) (simp) have 2: "(\y. q (fst y) (2 * snd y - 1)) = (\y. q (fst y) (snd y)) \ (\y. (fst y, 2 * snd y - 1))" by (rule ext) (simp) show ?thesis apply (simp add: joinpaths_def) apply (rule continuous_on_cases_le) apply (simp_all only: 1 2) apply (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ using pf apply (auto simp: mult.commute pathstart_def pathfinish_def) done qed text\ Congruence properties of homotopy w.r.t. path-combining operations.\ lemma homotopic_paths_reversepath_D: assumes "homotopic_paths s p q" shows "homotopic_paths s (reversepath p) (reversepath q)" using assms apply (simp add: homotopic_paths_def homotopic_with_def, clarify) apply (rule_tac x="h \ (\x. (fst x, 1 - snd x))" in exI) apply (rule conjI continuous_intros)+ apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) done proposition homotopic_paths_reversepath: "homotopic_paths s (reversepath p) (reversepath q) \ homotopic_paths s p q" using homotopic_paths_reversepath_D by force proposition homotopic_paths_join: "\homotopic_paths s p p'; homotopic_paths s q q'; pathfinish p = pathstart q\ \ homotopic_paths s (p +++ q) (p' +++ q')" apply (simp add: homotopic_paths_def homotopic_with_def, clarify) apply (rename_tac k1 k2) apply (rule_tac x="(\y. ((k1 \ Pair (fst y)) +++ (k2 \ Pair (fst y))) (snd y))" in exI) apply (rule conjI continuous_intros homotopic_join_lemma)+ apply (auto simp: joinpaths_def pathstart_def pathfinish_def path_image_def) done proposition homotopic_paths_continuous_image: "\homotopic_paths s f g; continuous_on s h; h ` s \ t\ \ homotopic_paths t (h \ f) (h \ g)" unfolding homotopic_paths_def by (simp add: homotopic_with_compose_continuous_map_left pathfinish_compose pathstart_compose) subsection\Group properties for homotopy of paths\ text\<^marker>\tag important\\So taking equivalence classes under homotopy would give the fundamental group\ proposition homotopic_paths_rid: "\path p; path_image p \ s\ \ homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p" apply (subst homotopic_paths_sym) apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1 / 2 then 2 *\<^sub>R t else 1"]) apply (simp_all del: le_divide_eq_numeral1) apply (subst split_01) apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ done proposition homotopic_paths_lid: "\path p; path_image p \ s\ \ homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p) p" using homotopic_paths_rid [of "reversepath p" s] by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath pathfinish_reversepath reversepath_joinpaths reversepath_linepath) proposition homotopic_paths_assoc: "\path p; path_image p \ s; path q; path_image q \ s; path r; path_image r \ s; pathfinish p = pathstart q; pathfinish q = pathstart r\ \ homotopic_paths s (p +++ (q +++ r)) ((p +++ q) +++ r)" apply (subst homotopic_paths_sym) apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1 / 2 then inverse 2 *\<^sub>R t else if t \ 3 / 4 then t - (1 / 4) else 2 *\<^sub>R t - 1"]) apply (simp_all del: le_divide_eq_numeral1) apply (simp add: subset_path_image_join) apply (rule continuous_on_cases_1 continuous_intros)+ apply (auto simp: joinpaths_def) done proposition homotopic_paths_rinv: assumes "path p" "path_image p \ s" shows "homotopic_paths s (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" proof - have "continuous_on ({0..1} \ {0..1}) (\x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" using assms apply (simp add: joinpaths_def subpath_def reversepath_def path_def del: le_divide_eq_numeral1) apply (rule continuous_on_cases_le) apply (rule_tac [2] continuous_on_compose [of _ _ p, unfolded o_def]) apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) apply (auto intro!: continuous_intros simp del: eq_divide_eq_numeral1) apply (force elim!: continuous_on_subset simp add: mult_le_one)+ done then show ?thesis using assms apply (subst homotopic_paths_sym_eq) unfolding homotopic_paths_def homotopic_with_def apply (rule_tac x="(\y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) apply (simp add: path_defs joinpaths_def subpath_def reversepath_def) apply (force simp: mult_le_one) done qed proposition homotopic_paths_linv: assumes "path p" "path_image p \ s" shows "homotopic_paths s (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" using homotopic_paths_rinv [of "reversepath p" s] assms by simp subsection\Homotopy of loops without requiring preservation of endpoints\ definition\<^marker>\tag important\ homotopic_loops :: "'a::topological_space set \ (real \ 'a) \ (real \ 'a) \ bool" where "homotopic_loops s p q \ homotopic_with_canon (\r. pathfinish r = pathstart r) {0..1} s p q" lemma homotopic_loops: "homotopic_loops s p q \ (\h. continuous_on ({0..1::real} \ {0..1}) h \ image h ({0..1} \ {0..1}) \ s \ (\x \ {0..1}. h(0,x) = p x) \ (\x \ {0..1}. h(1,x) = q x) \ (\t \ {0..1}. pathfinish(h \ Pair t) = pathstart(h \ Pair t)))" by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) proposition homotopic_loops_imp_loop: "homotopic_loops s p q \ pathfinish p = pathstart p \ pathfinish q = pathstart q" using homotopic_with_imp_property homotopic_loops_def by blast proposition homotopic_loops_imp_path: "homotopic_loops s p q \ path p \ path q" unfolding homotopic_loops_def path_def using homotopic_with_imp_continuous_maps continuous_map_subtopology_eu by blast proposition homotopic_loops_imp_subset: "homotopic_loops s p q \ path_image p \ s \ path_image q \ s" unfolding homotopic_loops_def path_image_def by (meson continuous_map_subtopology_eu homotopic_with_imp_continuous_maps) proposition homotopic_loops_refl: "homotopic_loops s p p \ path p \ path_image p \ s \ pathfinish p = pathstart p" by (simp add: homotopic_loops_def path_image_def path_def) proposition homotopic_loops_sym: "homotopic_loops s p q \ homotopic_loops s q p" by (simp add: homotopic_loops_def homotopic_with_sym) proposition homotopic_loops_sym_eq: "homotopic_loops s p q \ homotopic_loops s q p" by (metis homotopic_loops_sym) proposition homotopic_loops_trans: "\homotopic_loops s p q; homotopic_loops s q r\ \ homotopic_loops s p r" unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) proposition homotopic_loops_subset: "\homotopic_loops s p q; s \ t\ \ homotopic_loops t p q" by (fastforce simp add: homotopic_loops) proposition homotopic_loops_eq: "\path p; path_image p \ s; pathfinish p = pathstart p; \t. t \ {0..1} \ p(t) = q(t)\ \ homotopic_loops s p q" unfolding homotopic_loops_def apply (rule homotopic_with_eq) apply (rule homotopic_with_refl [where f = p, THEN iffD2]) apply (simp_all add: path_image_def path_def pathstart_def pathfinish_def) done proposition homotopic_loops_continuous_image: "\homotopic_loops s f g; continuous_on s h; h ` s \ t\ \ homotopic_loops t (h \ f) (h \ g)" unfolding homotopic_loops_def by (simp add: homotopic_with_compose_continuous_map_left pathfinish_def pathstart_def) subsection\Relations between the two variants of homotopy\ proposition homotopic_paths_imp_homotopic_loops: "\homotopic_paths s p q; pathfinish p = pathstart p; pathfinish q = pathstart p\ \ homotopic_loops s p q" by (auto simp: homotopic_with_def homotopic_paths_def homotopic_loops_def) proposition homotopic_loops_imp_homotopic_paths_null: assumes "homotopic_loops s p (linepath a a)" shows "homotopic_paths s p (linepath (pathstart p) (pathstart p))" proof - have "path p" by (metis assms homotopic_loops_imp_path) have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) have pip: "path_image p \ s" by (metis assms homotopic_loops_imp_subset) obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" and hs: "h ` ({0..1} \ {0..1}) \ s" and [simp]: "\x. x \ {0..1} \ h(0,x) = p x" and [simp]: "\x. x \ {0..1} \ h(1,x) = a" and ends: "\t. t \ {0..1} \ pathfinish (h \ Pair t) = pathstart (h \ Pair t)" using assms by (auto simp: homotopic_loops homotopic_with) have conth0: "path (\u. h (u, 0))" unfolding path_def apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) apply (force intro: continuous_intros continuous_on_subset [OF conth])+ done have pih0: "path_image (\u. h (u, 0)) \ s" using hs by (force simp: path_image_def) have c1: "continuous_on ({0..1} \ {0..1}) (\x. h (fst x * snd x, 0))" apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) apply (force simp: mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ done have c2: "continuous_on ({0..1} \ {0..1}) (\x. h (fst x - fst x * snd x, 0))" apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) apply (force simp: mult_left_le mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ apply (rule continuous_on_subset [OF conth]) apply (auto simp: algebra_simps add_increasing2 mult_left_le) done have [simp]: "\t. \0 \ t \ t \ 1\ \ h (t, 1) = h (t, 0)" using ends by (simp add: pathfinish_def pathstart_def) have adhoc_le: "c * 4 \ 1 + c * (d * 4)" if "\ d * 4 \ 3" "0 \ c" "c \ 1" for c d::real proof - have "c * 3 \ c * (d * 4)" using that less_eq_real_def by auto with \c \ 1\ show ?thesis by fastforce qed have *: "\p x. (path p \ path(reversepath p)) \ (path_image p \ s \ path_image(reversepath p) \ s) \ (pathfinish p = pathstart(linepath a a +++ reversepath p) \ pathstart(reversepath p) = a) \ pathstart p = x \ homotopic_paths s (p +++ linepath a a +++ reversepath p) (linepath x x)" by (metis homotopic_paths_lid homotopic_paths_join homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) have 1: "homotopic_paths s p (p +++ linepath (pathfinish p) (pathfinish p))" using \path p\ homotopic_paths_rid homotopic_paths_sym pip by blast moreover have "homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" apply (rule homotopic_paths_sym) using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" s] by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_pathstart homotopic_paths_imp_subset) moreover have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0)))" apply (simp add: homotopic_paths_def homotopic_with_def) apply (rule_tac x="\y. (subpath 0 (fst y) (\u. h (u, 0)) +++ (\u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\u. h (u, 0))) (snd y)" in exI) apply (simp add: subpath_reversepath) apply (intro conjI homotopic_join_lemma) using ploop apply (simp_all add: path_defs joinpaths_def o_def subpath_def conth c1 c2) apply (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le) done moreover have "homotopic_paths s ((\u. h (u, 0)) +++ linepath a a +++ reversepath (\u. h (u, 0))) (linepath (pathstart p) (pathstart p))" apply (rule *) apply (simp add: pih0 pathstart_def pathfinish_def conth0) apply (simp add: reversepath_def joinpaths_def) done ultimately show ?thesis by (blast intro: homotopic_paths_trans) qed proposition homotopic_loops_conjugate: fixes s :: "'a::real_normed_vector set" assumes "path p" "path q" and pip: "path_image p \ s" and piq: "path_image q \ s" and papp: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" shows "homotopic_loops s (p +++ q +++ reversepath p) q" proof - have contp: "continuous_on {0..1} p" using \path p\ [unfolded path_def] by blast have contq: "continuous_on {0..1} q" using \path q\ [unfolded path_def] by blast have c1: "continuous_on ({0..1} \ {0..1}) (\x. p ((1 - fst x) * snd x + fst x))" apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) apply (force simp: mult_le_one intro!: continuous_intros) apply (rule continuous_on_subset [OF contp]) apply (auto simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) done have c2: "continuous_on ({0..1} \ {0..1}) (\x. p ((fst x - 1) * snd x + 1))" apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) apply (force simp: mult_le_one intro!: continuous_intros) apply (rule continuous_on_subset [OF contp]) apply (auto simp: algebra_simps add_increasing2 mult_left_le_one_le) done have ps1: "\a b. \b * 2 \ 1; 0 \ b; 0 \ a; a \ 1\ \ p ((1 - a) * (2 * b) + a) \ s" using sum_le_prod1 by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) have ps2: "\a b. \\ 4 * b \ 3; b \ 1; 0 \ a; a \ 1\ \ p ((a - 1) * (4 * b - 3) + 1) \ s" apply (rule pip [unfolded path_image_def, THEN subsetD]) apply (rule image_eqI, blast) apply (simp add: algebra_simps) by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono not_le add.commute zero_le_numeral) have qs: "\a b. \4 * b \ 3; \ b * 2 \ 1\ \ q (4 * b - 2) \ s" using path_image_def piq by fastforce have "homotopic_loops s (p +++ q +++ reversepath p) (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" apply (simp add: homotopic_loops_def homotopic_with_def) apply (rule_tac x="(\y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" in exI) apply (simp add: subpath_refl subpath_reversepath) apply (intro conjI homotopic_join_lemma) using papp qloop apply (simp_all add: path_defs joinpaths_def o_def subpath_def c1 c2) apply (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) apply (auto simp: ps1 ps2 qs) done moreover have "homotopic_loops s (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" proof - have "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q) q" using \path q\ homotopic_paths_lid qloop piq by auto hence 1: "\f. homotopic_paths s f q \ \ homotopic_paths s f (linepath (pathfinish q) (pathfinish q) +++ q)" using homotopic_paths_trans by blast hence "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" proof - have "homotopic_paths s (q +++ linepath (pathfinish q) (pathfinish q)) q" by (simp add: \path q\ homotopic_paths_rid piq) thus ?thesis by (metis (no_types) 1 \path q\ homotopic_paths_join homotopic_paths_rinv homotopic_paths_sym homotopic_paths_trans qloop pathfinish_linepath piq) qed thus ?thesis by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) qed ultimately show ?thesis by (blast intro: homotopic_loops_trans) qed lemma homotopic_paths_loop_parts: assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" shows "homotopic_paths S p q" proof - have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp then have "path p" using \path q\ homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast show ?thesis proof (cases "pathfinish p = pathfinish q") case True have pipq: "path_image p \ S" "path_image q \ S" by (metis Un_subset_iff paths \path p\ \path q\ homotopic_loops_imp_subset homotopic_paths_imp_path loops path_image_join path_image_reversepath path_imp_reversepath path_join_eq)+ have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" using \path p\ \path_image p \ S\ homotopic_paths_rid homotopic_paths_sym by blast moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" by (simp add: True \path p\ \path q\ pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" by (simp add: True \path p\ \path q\ homotopic_paths_assoc pipq) moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" by (simp add: \path q\ homotopic_paths_join paths pipq) moreover then have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ q) q" by (metis \path q\ homotopic_paths_imp_path homotopic_paths_lid linepath_trivial path_join_path_ends pathfinish_def pipq(2)) ultimately show ?thesis using homotopic_paths_trans by metis next case False then show ?thesis using \path q\ homotopic_loops_imp_path loops path_join_path_ends by fastforce qed qed subsection\<^marker>\tag unimportant\\Homotopy of "nearby" function, paths and loops\ lemma homotopic_with_linear: fixes f g :: "_ \ 'b::real_normed_vector" assumes contf: "continuous_on s f" and contg:"continuous_on s g" and sub: "\x. x \ s \ closed_segment (f x) (g x) \ t" shows "homotopic_with_canon (\z. True) s t f g" apply (simp add: homotopic_with_def) apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) apply (intro conjI) apply (rule subset_refl continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]| simp)+ using sub closed_segment_def apply fastforce+ done lemma homotopic_paths_linear: fixes g h :: "real \ 'a::real_normed_vector" assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" "\t. t \ {0..1} \ closed_segment (g t) (h t) \ s" shows "homotopic_paths s g h" using assms unfolding path_def apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R (g \ snd) y + (fst y) *\<^sub>R (h \ snd) y)" in exI) apply (intro conjI subsetI continuous_intros; force) done lemma homotopic_loops_linear: fixes g h :: "real \ 'a::real_normed_vector" assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" "\t x. t \ {0..1} \ closed_segment (g t) (h t) \ s" shows "homotopic_loops s g h" using assms unfolding path_def apply (simp add: pathstart_def pathfinish_def homotopic_loops_def homotopic_with_def) apply (rule_tac x="\y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) apply (auto intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) apply (force simp: closed_segment_def) done lemma homotopic_paths_nearby_explicit: assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" and no: "\t x. \t \ {0..1}; x \ s\ \ norm(h t - g t) < norm(g t - x)" shows "homotopic_paths s g h" apply (rule homotopic_paths_linear [OF assms(1-4)]) by (metis no segment_bound(1) subsetI norm_minus_commute not_le) lemma homotopic_loops_nearby_explicit: assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" and no: "\t x. \t \ {0..1}; x \ s\ \ norm(h t - g t) < norm(g t - x)" shows "homotopic_loops s g h" apply (rule homotopic_loops_linear [OF assms(1-4)]) by (metis no segment_bound(1) subsetI norm_minus_commute not_le) lemma homotopic_nearby_paths: fixes g h :: "real \ 'a::euclidean_space" assumes "path g" "open s" "path_image g \ s" shows "\e. 0 < e \ (\h. path h \ pathstart h = pathstart g \ pathfinish h = pathfinish g \ (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_paths s g h)" proof - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - s \ e \ dist x y" using separate_compact_closed [of "path_image g" "-s"] assms by force show ?thesis apply (intro exI conjI) using e [unfolded dist_norm] apply (auto simp: intro!: homotopic_paths_nearby_explicit assms \e > 0\) by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) qed lemma homotopic_nearby_loops: fixes g h :: "real \ 'a::euclidean_space" assumes "path g" "open s" "path_image g \ s" "pathfinish g = pathstart g" shows "\e. 0 < e \ (\h. path h \ pathfinish h = pathstart h \ (\t \ {0..1}. norm(h t - g t) < e) \ homotopic_loops s g h)" proof - obtain e where "e > 0" and e: "\x y. x \ path_image g \ y \ - s \ e \ dist x y" using separate_compact_closed [of "path_image g" "-s"] assms by force show ?thesis apply (intro exI conjI) using e [unfolded dist_norm] apply (auto simp: intro!: homotopic_loops_nearby_explicit assms \e > 0\) by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) qed subsection\ Homotopy and subpaths\ lemma homotopic_join_subpaths1: assumes "path g" and pag: "path_image g \ s" and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" "u \ v" "v \ w" shows "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" proof - have 1: "t * 2 \ 1 \ u + t * (v * 2) \ v + t * (u * 2)" for t using affine_ineq \u \ v\ by fastforce have 2: "t * 2 > 1 \ u + (2*t - 1) * v \ v + (2*t - 1) * w" for t by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \u \ v\ \v \ w\) have t2: "\t::real. t*2 = 1 \ t = 1/2" by auto show ?thesis apply (rule homotopic_paths_subset [OF _ pag]) using assms apply (cases "w = u") using homotopic_paths_rinv [of "subpath u v g" "path_image g"] apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl) apply (rule homotopic_paths_sym) apply (rule homotopic_paths_reparametrize [where f = "\t. if t \ 1 / 2 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"]) using \path g\ path_subpath u w apply blast using \path g\ path_image_subpath_subset u w(1) apply blast apply simp_all apply (subst split_01) apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ apply (simp_all add: field_simps not_le) apply (force dest!: t2) apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2) apply (simp add: joinpaths_def subpath_def) apply (force simp: algebra_simps) done qed lemma homotopic_join_subpaths2: assumes "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" shows "homotopic_paths s (subpath w v g +++ subpath v u g) (subpath w u g)" by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) lemma homotopic_join_subpaths3: assumes hom: "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" and "path g" and pag: "path_image g \ s" and u: "u \ {0..1}" and v: "v \ {0..1}" and w: "w \ {0..1}" shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)" proof - have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" apply (rule homotopic_paths_join) using hom homotopic_paths_sym_eq apply blast apply (metis \path g\ homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w, simp) done also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)" apply (rule homotopic_paths_sym [OF homotopic_paths_assoc]) using assms by (simp_all add: path_image_subpath_subset [THEN order_trans]) also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g) (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" apply (rule homotopic_paths_join) apply (metis \path g\ homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v) apply (metis (no_types, lifting) \path g\ homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w) apply simp done also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" apply (rule homotopic_paths_rid) using \path g\ path_subpath u v apply blast apply (meson \path g\ order.trans pag path_image_subpath_subset u v) done finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" . then show ?thesis using homotopic_join_subpaths2 by blast qed proposition homotopic_join_subpaths: "\path g; path_image g \ s; u \ {0..1}; v \ {0..1}; w \ {0..1}\ \ homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" apply (rule le_cases3 [of u v w]) using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+ text\Relating homotopy of trivial loops to path-connectedness.\ lemma path_component_imp_homotopic_points: "path_component S a b \ homotopic_loops S (linepath a a) (linepath b b)" apply (simp add: path_component_def homotopic_loops_def homotopic_with_def pathstart_def pathfinish_def path_image_def path_def, clarify) apply (rule_tac x="g \ fst" in exI) apply (intro conjI continuous_intros continuous_on_compose)+ apply (auto elim!: continuous_on_subset) done lemma homotopic_loops_imp_path_component_value: "\homotopic_loops S p q; 0 \ t; t \ 1\ \ path_component S (p t) (q t)" apply (simp add: path_component_def homotopic_loops_def homotopic_with_def pathstart_def pathfinish_def path_image_def path_def, clarify) apply (rule_tac x="h \ (\u. (u, t))" in exI) apply (intro conjI continuous_intros continuous_on_compose)+ apply (auto elim!: continuous_on_subset) done lemma homotopic_points_eq_path_component: "homotopic_loops S (linepath a a) (linepath b b) \ path_component S a b" by (auto simp: path_component_imp_homotopic_points dest: homotopic_loops_imp_path_component_value [where t=1]) lemma path_connected_eq_homotopic_points: "path_connected S \ (\a b. a \ S \ b \ S \ homotopic_loops S (linepath a a) (linepath b b))" by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) subsection\Simply connected sets\ text\<^marker>\tag important\\defined as "all loops are homotopic (as loops)\ definition\<^marker>\tag important\ simply_connected where "simply_connected S \ \p q. path p \ pathfinish p = pathstart p \ path_image p \ S \ path q \ pathfinish q = pathstart q \ path_image q \ S \ homotopic_loops S p q" lemma simply_connected_empty [iff]: "simply_connected {}" by (simp add: simply_connected_def) lemma simply_connected_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S" by (simp add: simply_connected_def path_connected_eq_homotopic_points) lemma simply_connected_imp_connected: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ connected S" by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) lemma simply_connected_eq_contractible_loop_any: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ (\p a. path p \ path_image p \ S \ pathfinish p = pathstart p \ a \ S \ homotopic_loops S p (linepath a a))" apply (simp add: simply_connected_def) apply (rule iffI, force, clarify) apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans) apply (fastforce simp add:) using homotopic_loops_sym apply blast done lemma simply_connected_eq_contractible_loop_some: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ (\a. a \ S \ homotopic_loops S p (linepath a a)))" apply (rule iffI) apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any) apply (clarsimp simp add: simply_connected_eq_contractible_loop_any) apply (drule_tac x=p in spec) using homotopic_loops_trans path_connected_eq_homotopic_points apply blast done lemma simply_connected_eq_contractible_loop_all: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ S = {} \ (\a \ S. \p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_loops S p (linepath a a))" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by force next case False then obtain a where "a \ S" by blast show ?thesis proof assume "simply_connected S" then show ?rhs using \a \ S\ \simply_connected S\ simply_connected_eq_contractible_loop_any by blast next assume ?rhs then show "simply_connected S" apply (simp add: simply_connected_eq_contractible_loop_any False) by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans path_component_imp_homotopic_points path_component_refl) qed qed lemma simply_connected_eq_contractible_path: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p. path p \ path_image p \ S \ pathfinish p = pathstart p \ homotopic_paths S p (linepath (pathstart p) (pathstart p)))" apply (rule iffI) apply (simp add: simply_connected_imp_path_connected) apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image simply_connected_eq_contractible_loop_some subset_iff) lemma simply_connected_eq_homotopic_paths: fixes S :: "_::real_normed_vector set" shows "simply_connected S \ path_connected S \ (\p q. path p \ path_image p \ S \ path q \ path_image q \ S \ pathstart q = pathstart p \ pathfinish q = pathfinish p \ homotopic_paths S p q)" (is "?lhs = ?rhs") proof assume ?lhs then have pc: "path_connected S" and *: "\p. \path p; path_image p \ S; pathfinish p = pathstart p\ \ homotopic_paths S p (linepath (pathstart p) (pathstart p))" by (auto simp: simply_connected_eq_contractible_path) have "homotopic_paths S p q" if "path p" "path_image p \ S" "path q" "path_image q \ S" "pathstart q = pathstart p" "pathfinish q = pathfinish p" for p q proof - have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" by (simp add: homotopic_paths_rid homotopic_paths_sym that) also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) (p +++ reversepath q +++ q)" using that by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath) also have "homotopic_paths S (p +++ reversepath q +++ q) ((p +++ reversepath q) +++ q)" by (simp add: that homotopic_paths_assoc) also have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart q) (pathstart q) +++ q)" using * [of "p +++ reversepath q"] that by (simp add: homotopic_paths_join path_image_join) also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" using that homotopic_paths_lid by blast finally show ?thesis . qed then show ?rhs by (blast intro: pc *) next assume ?rhs then show ?lhs by (force simp: simply_connected_eq_contractible_path) qed proposition simply_connected_Times: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes S: "simply_connected S" and T: "simply_connected T" shows "simply_connected(S \ T)" proof - have "homotopic_loops (S \ T) p (linepath (a, b) (a, b))" if "path p" "path_image p \ S \ T" "p 1 = p 0" "a \ S" "b \ T" for p a b proof - have "path (fst \ p)" apply (rule Path_Connected.path_continuous_image [OF \path p\]) apply (rule continuous_intros)+ done moreover have "path_image (fst \ p) \ S" using that apply (simp add: path_image_def) by force ultimately have p1: "homotopic_loops S (fst \ p) (linepath a a)" using S that apply (simp add: simply_connected_eq_contractible_loop_any) apply (drule_tac x="fst \ p" in spec) apply (drule_tac x=a in spec) apply (auto simp: pathstart_def pathfinish_def) done have "path (snd \ p)" apply (rule Path_Connected.path_continuous_image [OF \path p\]) apply (rule continuous_intros)+ done moreover have "path_image (snd \ p) \ T" using that apply (simp add: path_image_def) by force ultimately have p2: "homotopic_loops T (snd \ p) (linepath b b)" using T that apply (simp add: simply_connected_eq_contractible_loop_any) apply (drule_tac x="snd \ p" in spec) apply (drule_tac x=b in spec) apply (auto simp: pathstart_def pathfinish_def) done show ?thesis using p1 p2 apply (simp add: homotopic_loops, clarify) apply (rename_tac h k) apply (rule_tac x="\z. Pair (h z) (k z)" in exI) apply (intro conjI continuous_intros | assumption)+ apply (auto simp: pathstart_def pathfinish_def) done qed with assms show ?thesis by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) qed subsection\Contractible sets\ definition\<^marker>\tag important\ contractible where "contractible S \ \a. homotopic_with_canon (\x. True) S S id (\x. a)" proposition contractible_imp_simply_connected: fixes S :: "_::real_normed_vector set" assumes "contractible S" shows "simply_connected S" proof (cases "S = {}") case True then show ?thesis by force next case False obtain a where a: "homotopic_with_canon (\x. True) S S id (\x. a)" using assms by (force simp: contractible_def) then have "a \ S" by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_in_topspace topspace_euclidean_subtopology) show ?thesis apply (simp add: simply_connected_eq_contractible_loop_all False) apply (rule bexI [OF _ \a \ S\]) using a apply (simp add: homotopic_loops_def homotopic_with_def path_def path_image_def pathfinish_def pathstart_def, clarify) apply (rule_tac x="(h \ (\y. (fst y, (p \ snd) y)))" in exI) apply (intro conjI continuous_on_compose continuous_intros) apply (erule continuous_on_subset | force)+ done qed corollary contractible_imp_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ connected S" by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) lemma contractible_imp_path_connected: fixes S :: "_::real_normed_vector set" shows "contractible S \ path_connected S" by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) lemma nullhomotopic_through_contractible: fixes S :: "_::topological_space set" assumes f: "continuous_on S f" "f ` S \ T" and g: "continuous_on T g" "g ` T \ U" and T: "contractible T" obtains c where "homotopic_with_canon (\h. True) S U (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with_canon (\x. True) T T id (\x. b)" using assms by (force simp: contractible_def) have "homotopic_with_canon (\f. True) T U (g \ id) (g \ (\x. b))" by (metis Abstract_Topology.continuous_map_subtopology_eu b g homotopic_compose_continuous_map_left) then have "homotopic_with_canon (\f. True) S U (g \ id \ f) (g \ (\x. b) \ f)" by (simp add: f homotopic_with_compose_continuous_map_right) then show ?thesis by (simp add: comp_def that) qed lemma nullhomotopic_into_contractible: assumes f: "continuous_on S f" "f ` S \ T" and T: "contractible T" obtains c where "homotopic_with_canon (\h. True) S T f (\x. c)" apply (rule nullhomotopic_through_contractible [OF f, of id T]) using assms -apply (auto simp: continuous_on_id) +apply (auto) done lemma nullhomotopic_from_contractible: assumes f: "continuous_on S f" "f ` S \ T" and S: "contractible S" obtains c where "homotopic_with_canon (\h. True) S T f (\x. c)" apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S]) using assms apply (auto simp: comp_def) done lemma homotopic_through_contractible: fixes S :: "_::real_normed_vector set" assumes "continuous_on S f1" "f1 ` S \ T" "continuous_on T g1" "g1 ` T \ U" "continuous_on S f2" "f2 ` S \ T" "continuous_on T g2" "g2 ` T \ U" "contractible T" "path_connected U" shows "homotopic_with_canon (\h. True) S U (g1 \ f1) (g2 \ f2)" proof - obtain c1 where c1: "homotopic_with_canon (\h. True) S U (g1 \ f1) (\x. c1)" apply (rule nullhomotopic_through_contractible [of S f1 T g1 U]) using assms apply auto done obtain c2 where c2: "homotopic_with_canon (\h. True) S U (g2 \ f2) (\x. c2)" apply (rule nullhomotopic_through_contractible [of S f2 T g2 U]) using assms apply auto done have *: "S = {} \ (\t. path_connected t \ t \ U \ c2 \ t \ c1 \ t)" proof (cases "S = {}") case True then show ?thesis by force next case False with c1 c2 have "c1 \ U" "c2 \ U" using homotopic_with_imp_continuous_maps by fastforce+ with \path_connected U\ show ?thesis by blast qed show ?thesis apply (rule homotopic_with_trans [OF c1]) apply (rule homotopic_with_symD) apply (rule homotopic_with_trans [OF c2]) apply (simp add: path_component homotopic_constant_maps *) done qed lemma homotopic_into_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f ` S \ T" and g: "continuous_on S g" "g ` S \ T" and T: "contractible T" shows "homotopic_with_canon (\h. True) S T f g" using homotopic_through_contractible [of S f T id T g id] by (simp add: assms contractible_imp_path_connected) lemma homotopic_from_contractible: fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" assumes f: "continuous_on S f" "f ` S \ T" and g: "continuous_on S g" "g ` S \ T" and "contractible S" "path_connected T" shows "homotopic_with_canon (\h. True) S T f g" using homotopic_through_contractible [of S id S f T id g] by (simp add: assms contractible_imp_path_connected) lemma starlike_imp_contractible_gen: fixes S :: "'a::real_normed_vector set" assumes S: "starlike S" and P: "\a T. \a \ S; 0 \ T; T \ 1\ \ P(\x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" obtains a where "homotopic_with_canon P S S (\x. x) (\x. a)" proof - obtain a where "a \ S" and a: "\x. x \ S \ closed_segment a x \ S" using S by (auto simp: starlike_def) have "(\y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \ S) \ S" apply clarify apply (erule a [unfolded closed_segment_def, THEN subsetD], simp) apply (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) done then show ?thesis apply (rule_tac a=a in that) using \a \ S\ apply (simp add: homotopic_with_def) apply (rule_tac x="\y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) apply (intro conjI ballI continuous_on_compose continuous_intros) apply (simp_all add: P) done qed lemma starlike_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "starlike S \ contractible S" using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" by (simp add: starlike_imp_contractible) lemma starlike_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ simply_connected S" by (simp add: contractible_imp_simply_connected starlike_imp_contractible) lemma convex_imp_simply_connected: fixes S :: "'a::real_normed_vector set" shows "convex S \ simply_connected S" using convex_imp_starlike starlike_imp_simply_connected by blast lemma starlike_imp_path_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ path_connected S" by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) lemma starlike_imp_connected: fixes S :: "'a::real_normed_vector set" shows "starlike S \ connected S" by (simp add: path_connected_imp_connected starlike_imp_path_connected) lemma is_interval_simply_connected_1: fixes S :: "real set" shows "is_interval S \ simply_connected S" using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto lemma contractible_empty [simp]: "contractible {}" by (simp add: contractible_def homotopic_on_emptyI) lemma contractible_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" and TS: "rel_interior S \ T" "T \ closure S" shows "contractible T" proof (cases "S = {}") case True with assms show ?thesis by (simp add: subsetCE) next case False show ?thesis apply (rule starlike_imp_contractible) apply (rule starlike_convex_tweak_boundary_points [OF \convex S\ False TS]) done qed lemma convex_imp_contractible: fixes S :: "'a::real_normed_vector set" shows "convex S \ contractible S" using contractible_empty convex_imp_starlike starlike_imp_contractible by blast lemma contractible_sing [simp]: fixes a :: "'a::real_normed_vector" shows "contractible {a}" by (rule convex_imp_contractible [OF convex_singleton]) lemma is_interval_contractible_1: fixes S :: "real set" shows "is_interval S \ contractible S" using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 is_interval_simply_connected_1 by auto lemma contractible_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "contractible S" and T: "contractible T" shows "contractible (S \ T)" proof - obtain a h where conth: "continuous_on ({0..1} \ S) h" and hsub: "h ` ({0..1} \ S) \ S" and [simp]: "\x. x \ S \ h (0, x) = x" and [simp]: "\x. x \ S \ h (1::real, x) = a" using S by (auto simp: contractible_def homotopic_with) obtain b k where contk: "continuous_on ({0..1} \ T) k" and ksub: "k ` ({0..1} \ T) \ T" and [simp]: "\x. x \ T \ k (0, x) = x" and [simp]: "\x. x \ T \ k (1::real, x) = b" using T by (auto simp: contractible_def homotopic_with) show ?thesis apply (simp add: contractible_def homotopic_with) apply (rule exI [where x=a]) apply (rule exI [where x=b]) apply (rule exI [where x = "\z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) using hsub ksub apply auto done qed subsection\Local versions of topological properties in general\ definition\<^marker>\tag important\ locally :: "('a::topological_space set \ bool) \ 'a set \ bool" where "locally P S \ \w x. openin (top_of_set S) w \ x \ w \ (\u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w)" lemma locallyI: assumes "\w x. \openin (top_of_set S) w; x \ w\ \ \u v. openin (top_of_set S) u \ P v \ x \ u \ u \ v \ v \ w" shows "locally P S" using assms by (force simp: locally_def) lemma locallyE: assumes "locally P S" "openin (top_of_set S) w" "x \ w" obtains u v where "openin (top_of_set S) u" "P v" "x \ u" "u \ v" "v \ w" using assms unfolding locally_def by meson lemma locally_mono: assumes "locally P S" "\t. P t \ Q t" shows "locally Q S" by (metis assms locally_def) lemma locally_open_subset: assumes "locally P S" "openin (top_of_set S) t" shows "locally P t" using assms apply (simp add: locally_def) apply (erule all_forward)+ apply (rule impI) apply (erule impCE) using openin_trans apply blast apply (erule ex_forward) by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset) lemma locally_diff_closed: "\locally P S; closedin (top_of_set S) t\ \ locally P (S - t)" using locally_open_subset closedin_def by fastforce lemma locally_empty [iff]: "locally P {}" by (simp add: locally_def openin_subtopology) lemma locally_singleton [iff]: fixes a :: "'a::metric_space" shows "locally P {a} \ P {a}" apply (simp add: locally_def openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR cong: conj_cong) using zero_less_one by blast lemma locally_iff: "locally P S \ (\T x. open T \ x \ S \ T \ (\U. open U \ (\v. P v \ x \ S \ U \ S \ U \ v \ v \ S \ T)))" apply (simp add: le_inf_iff locally_def openin_open, safe) apply (metis IntE IntI le_inf_iff) apply (metis IntI Int_subset_iff) done lemma locally_Int: assumes S: "locally P S" and t: "locally P t" and P: "\S t. P S \ P t \ P(S \ t)" shows "locally P (S \ t)" using S t unfolding locally_iff apply clarify apply (drule_tac x=T in spec)+ apply (drule_tac x=x in spec)+ apply clarsimp apply (rename_tac U1 U2 V1 V2) apply (rule_tac x="U1 \ U2" in exI) apply (simp add: open_Int) apply (rule_tac x="V1 \ V2" in exI) apply (auto intro: P) done lemma locally_Times: fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" assumes PS: "locally P S" and QT: "locally Q T" and R: "\S T. P S \ Q T \ R(S \ T)" shows "locally R (S \ T)" unfolding locally_def proof (clarify) fix W x y assume W: "openin (top_of_set (S \ T)) W" and xy: "(x, y) \ W" then obtain U V where "openin (top_of_set S) U" "x \ U" "openin (top_of_set T) V" "y \ V" "U \ V \ W" using Times_in_interior_subtopology by metis then obtain U1 U2 V1 V2 where opeS: "openin (top_of_set S) U1 \ P U2 \ x \ U1 \ U1 \ U2 \ U2 \ U" and opeT: "openin (top_of_set T) V1 \ Q V2 \ y \ V1 \ V1 \ V2 \ V2 \ V" by (meson PS QT locallyE) with \U \ V \ W\ show "\u v. openin (top_of_set (S \ T)) u \ R v \ (x,y) \ u \ u \ v \ v \ W" apply (rule_tac x="U1 \ V1" in exI) apply (rule_tac x="U2 \ V2" in exI) apply (auto simp: openin_Times R openin_Times_eq) done qed proposition homeomorphism_locally_imp: fixes S :: "'a::metric_space set" and t :: "'b::t2_space set" assumes S: "locally P S" and hom: "homeomorphism S t f g" and Q: "\S S'. \P S; homeomorphism S S' f g\ \ Q S'" shows "locally Q t" proof (clarsimp simp: locally_def) fix W y assume "y \ W" and "openin (top_of_set t) W" then obtain T where T: "open T" "W = t \ T" by (force simp: openin_open) then have "W \ t" by auto have f: "\x. x \ S \ g(f x) = x" "f ` S = t" "continuous_on S f" and g: "\y. y \ t \ f(g y) = y" "g ` t = S" "continuous_on t g" using hom by (auto simp: homeomorphism_def) have gw: "g ` W = S \ f -` W" using \W \ t\ apply auto using \g ` t = S\ \W \ t\ apply blast using g \W \ t\ apply auto[1] by (simp add: f rev_image_eqI) have \: "openin (top_of_set S) (g ` W)" proof - have "continuous_on S f" using f(3) by blast then show "openin (top_of_set S) (g ` W)" by (simp add: gw Collect_conj_eq \openin (top_of_set t) W\ continuous_on_open f(2)) qed then obtain u v where osu: "openin (top_of_set S) u" and uv: "P v" "g y \ u" "u \ v" "v \ g ` W" using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \y \ W\ by force have "v \ S" using uv by (simp add: gw) have fv: "f ` v = t \ {x. g x \ v}" using \f ` S = t\ f \v \ S\ by auto have "f ` v \ W" using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto have contvf: "continuous_on v f" using \v \ S\ continuous_on_subset f(3) by blast have contvg: "continuous_on (f ` v) g" using \f ` v \ W\ \W \ t\ continuous_on_subset [OF g(3)] by blast have homv: "homeomorphism v (f ` v) f g" using \v \ S\ \W \ t\ f apply (simp add: homeomorphism_def contvf contvg, auto) by (metis f(1) rev_image_eqI rev_subsetD) have 1: "openin (top_of_set t) (t \ g -` u)" apply (rule continuous_on_open [THEN iffD1, rule_format]) apply (rule \continuous_on t g\) using \g ` t = S\ apply (simp add: osu) done have 2: "\V. Q V \ y \ (t \ g -` u) \ (t \ g -` u) \ V \ V \ W" apply (rule_tac x="f ` v" in exI) apply (intro conjI Q [OF \P v\ homv]) using \W \ t\ \y \ W\ \f ` v \ W\ uv apply (auto simp: fv) done show "\U. openin (top_of_set t) U \ (\v. Q v \ y \ U \ U \ v \ v \ W)" by (meson 1 2) qed lemma homeomorphism_locally: fixes f:: "'a::metric_space \ 'b::metric_space" assumes hom: "homeomorphism S t f g" and eq: "\S t. homeomorphism S t f g \ (P S \ Q t)" shows "locally P S \ locally Q t" apply (rule iffI) apply (erule homeomorphism_locally_imp [OF _ hom]) apply (simp add: eq) apply (erule homeomorphism_locally_imp) using eq homeomorphism_sym homeomorphism_symD [OF hom] apply blast+ done lemma homeomorphic_locally: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" assumes hom: "S homeomorphic T" and iff: "\X Y. X homeomorphic Y \ (P X \ Q Y)" shows "locally P S \ locally Q T" proof - obtain f g where hom: "homeomorphism S T f g" using assms by (force simp: homeomorphic_def) then show ?thesis using homeomorphic_def local.iff by (blast intro!: homeomorphism_locally) qed lemma homeomorphic_local_compactness: fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" shows "S homeomorphic T \ locally compact S \ locally compact T" by (simp add: homeomorphic_compactness homeomorphic_locally) lemma locally_translation: fixes P :: "'a :: real_normed_vector set \ bool" shows "(\S. P (image (\x. a + x) S) \ P S) \ locally P (image (\x. a + x) S) \ locally P S" apply (rule homeomorphism_locally [OF homeomorphism_translation]) apply (simp add: homeomorphism_def) by metis lemma locally_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" and iff: "\S. P (f ` S) \ Q S" shows "locally P (f ` S) \ locally Q S" apply (rule linear_homeomorphism_image [OF f]) apply (rule_tac f=g and g = f in homeomorphism_locally, assumption) by (metis iff homeomorphism_def) lemma locally_open_map_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes P: "locally P S" and f: "continuous_on S f" and oo: "\t. openin (top_of_set S) t \ openin (top_of_set (f ` S)) (f ` t)" and Q: "\t. \t \ S; P t\ \ Q(f ` t)" shows "locally Q (f ` S)" proof (clarsimp simp add: locally_def) fix W y assume oiw: "openin (top_of_set (f ` S)) W" and "y \ W" then have "W \ f ` S" by (simp add: openin_euclidean_subtopology_iff) have oivf: "openin (top_of_set S) (S \ f -` W)" by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) then obtain x where "x \ S" "f x = y" using \W \ f ` S\ \y \ W\ by blast then obtain U V where "openin (top_of_set S) U" "P V" "x \ U" "U \ V" "V \ S \ f -` W" using P [unfolded locally_def, rule_format, of "(S \ f -` W)" x] oivf \y \ W\ by auto then show "\X. openin (top_of_set (f ` S)) X \ (\Y. Q Y \ y \ X \ X \ Y \ Y \ W)" apply (rule_tac x="f ` U" in exI) apply (rule conjI, blast intro!: oo) apply (rule_tac x="f ` V" in exI) apply (force simp: \f x = y\ rev_image_eqI intro: Q) done qed subsection\An induction principle for connected sets\ proposition connected_induction: assumes "connected S" and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ Q x \ Q y)" and etc: "a \ S" "b \ S" "P a" "P b" "Q a" shows "Q b" proof - have 1: "openin (top_of_set S) {b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ Q x)}" apply (subst openin_subopen, clarify) apply (rule_tac x=T in exI, auto) done have 2: "openin (top_of_set S) {b. \T. openin (top_of_set S) T \ b \ T \ (\x\T. P x \ \ Q x)}" apply (subst openin_subopen, clarify) apply (rule_tac x=T in exI, auto) done show ?thesis using \connected S\ apply (simp only: connected_openin HOL.not_ex HOL.de_Morgan_conj) apply (elim disjE allE) apply (blast intro: 1) apply (blast intro: 2, simp_all) apply clarify apply (metis opI) using opD apply (blast intro: etc elim: dest:) using opI etc apply meson+ done qed lemma connected_equivalence_relation_gen: assumes "connected S" and etc: "a \ S" "b \ S" "P a" "P b" and trans: "\x y z. \R x y; R y z\ \ R x z" and opD: "\T a. \openin (top_of_set S) T; a \ T\ \ \z. z \ T \ P z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y \ R x y)" shows "R a b" proof - have "\a b c. \a \ S; P a; b \ S; c \ S; P b; P c; R a b\ \ R a c" apply (rule connected_induction [OF \connected S\ opD], simp_all) by (meson trans opI) then show ?thesis by (metis etc opI) qed lemma connected_induction_simple: assumes "connected S" and etc: "a \ S" "b \ S" "P a" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. \y \ T. P x \ P y)" shows "P b" apply (rule connected_induction [OF \connected S\ _, where P = "\x. True"], blast) apply (frule opI) using etc apply simp_all done lemma connected_equivalence_relation: assumes "connected S" and etc: "a \ S" "b \ S" and sym: "\x y. \R x y; x \ S; y \ S\ \ R y x" and trans: "\x y z. \R x y; R y z; x \ S; y \ S; z \ S\ \ R x z" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. R a x)" shows "R a b" proof - have "\a b c. \a \ S; b \ S; c \ S; R a b\ \ R a c" apply (rule connected_induction_simple [OF \connected S\], simp_all) by (meson local.sym local.trans opI openin_imp_subset subsetCE) then show ?thesis by (metis etc opI) qed lemma locally_constant_imp_constant: assumes "connected S" and opI: "\a. a \ S \ \T. openin (top_of_set S) T \ a \ T \ (\x \ T. f x = f a)" shows "f constant_on S" proof - have "\x y. x \ S \ y \ S \ f x = f y" apply (rule connected_equivalence_relation [OF \connected S\], simp_all) by (metis opI) then show ?thesis by (metis constant_on_def) qed lemma locally_constant: "connected S \ locally (\U. f constant_on U) S \ f constant_on S" apply (simp add: locally_def) apply (rule iffI) apply (rule locally_constant_imp_constant, assumption) apply (metis (mono_tags, hide_lams) constant_on_def constant_on_subset openin_subtopology_self) by (meson constant_on_subset openin_imp_subset order_refl) subsection\Basic properties of local compactness\ proposition locally_compact: fixes s :: "'a :: metric_space set" shows "locally compact s \ (\x \ s. \u v. x \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply clarify apply (erule_tac w = "s \ ball x 1" in locallyE) by auto next assume r [rule_format]: ?rhs have *: "\u v. openin (top_of_set s) u \ compact v \ x \ u \ u \ v \ v \ s \ T" if "open T" "x \ s" "x \ T" for x T proof - obtain u v where uv: "x \ u" "u \ v" "v \ s" "compact v" "openin (top_of_set s) u" using r [OF \x \ s\] by auto obtain e where "e>0" and e: "cball x e \ T" using open_contains_cball \open T\ \x \ T\ by blast show ?thesis apply (rule_tac x="(s \ ball x e) \ u" in exI) apply (rule_tac x="cball x e \ v" in exI) using that \e > 0\ e uv apply auto done qed show ?lhs apply (rule locallyI) apply (subst (asm) openin_open) apply (blast intro: *) done qed lemma locally_compactE: fixes s :: "'a :: metric_space set" assumes "locally compact s" obtains u v where "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ openin (top_of_set s) (u x) \ compact (v x)" using assms unfolding locally_compact by metis lemma locally_compact_alt: fixes s :: "'a :: heine_borel set" shows "locally compact s \ (\x \ s. \u. x \ u \ openin (top_of_set s) u \ compact(closure u) \ closure u \ s)" apply (simp add: locally_compact) apply (intro ball_cong ex_cong refl iffI) apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans) by (meson closure_subset compact_closure) lemma locally_compact_Int_cball: fixes s :: "'a :: heine_borel set" shows "locally compact s \ (\x \ s. \e. 0 < e \ closed(cball x e \ s))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: locally_compact openin_contains_cball) apply (clarify | assumption | drule bspec)+ by (metis (no_types, lifting) compact_cball compact_imp_closed compact_Int inf.absorb_iff2 inf.orderE inf_sup_aci(2)) next assume ?rhs then show ?lhs apply (simp add: locally_compact openin_contains_cball) apply (clarify | assumption | drule bspec)+ apply (rule_tac x="ball x e \ s" in exI, simp) apply (rule_tac x="cball x e \ s" in exI) using compact_eq_bounded_closed apply auto apply (metis open_ball le_infI1 mem_ball open_contains_cball_eq) done qed lemma locally_compact_compact: fixes s :: "'a :: heine_borel set" shows "locally compact s \ (\k. k \ s \ compact k \ (\u v. k \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v))" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ s \ x \ u x \ u x \ v x \ v x \ s \ openin (top_of_set s) (u x) \ compact (v x)" by (metis locally_compactE) have *: "\u v. k \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v" if "k \ s" "compact k" for k proof - have "\C. (\c\C. openin (top_of_set k) c) \ k \ \C \ \D\C. finite D \ k \ \D" using that by (simp add: compact_eq_openin_cover) moreover have "\c \ (\x. k \ u x) ` k. openin (top_of_set k) c" using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) moreover have "k \ \((\x. k \ u x) ` k)" using that by clarsimp (meson subsetCE uv) ultimately obtain D where "D \ (\x. k \ u x) ` k" "finite D" "k \ \D" by metis then obtain T where T: "T \ k" "finite T" "k \ \((\x. k \ u x) ` T)" by (metis finite_subset_image) have Tuv: "\(u ` T) \ \(v ` T)" using T that by (force simp: dest!: uv) show ?thesis apply (rule_tac x="\(u ` T)" in exI) apply (rule_tac x="\(v ` T)" in exI) apply (simp add: Tuv) using T that apply (auto simp: dest!: uv) done qed show ?rhs by (blast intro: *) next assume ?rhs then show ?lhs apply (clarsimp simp add: locally_compact) apply (drule_tac x="{x}" in spec, simp) done qed lemma open_imp_locally_compact: fixes s :: "'a :: heine_borel set" assumes "open s" shows "locally compact s" proof - have *: "\u v. x \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v" if "x \ s" for x proof - obtain e where "e>0" and e: "cball x e \ s" using open_contains_cball assms \x \ s\ by blast have ope: "openin (top_of_set s) (ball x e)" by (meson e open_ball ball_subset_cball dual_order.trans open_subset) show ?thesis apply (rule_tac x="ball x e" in exI) apply (rule_tac x="cball x e" in exI) using \e > 0\ e apply (auto simp: ope) done qed show ?thesis unfolding locally_compact by (blast intro: *) qed lemma closed_imp_locally_compact: fixes s :: "'a :: heine_borel set" assumes "closed s" shows "locally compact s" proof - have *: "\u v. x \ u \ u \ v \ v \ s \ openin (top_of_set s) u \ compact v" if "x \ s" for x proof - show ?thesis apply (rule_tac x = "s \ ball x 1" in exI) apply (rule_tac x = "s \ cball x 1" in exI) using \x \ s\ assms apply auto done qed show ?thesis unfolding locally_compact by (blast intro: *) qed lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" by (simp add: closed_imp_locally_compact) lemma locally_compact_Int: fixes s :: "'a :: t2_space set" shows "\locally compact s; locally compact t\ \ locally compact (s \ t)" by (simp add: compact_Int locally_Int) lemma locally_compact_closedin: fixes s :: "'a :: heine_borel set" shows "\closedin (top_of_set s) t; locally compact s\ \ locally compact t" unfolding closedin_closed using closed_imp_locally_compact locally_compact_Int by blast lemma locally_compact_delete: fixes s :: "'a :: t1_space set" shows "locally compact s \ locally compact (s - {a})" by (auto simp: openin_delete locally_open_subset) lemma locally_closed: fixes s :: "'a :: heine_borel set" shows "locally closed s \ locally compact s" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp only: locally_def) apply (erule all_forward imp_forward asm_rl exE)+ apply (rule_tac x = "u \ ball x 1" in exI) apply (rule_tac x = "v \ cball x 1" in exI) apply (force intro: openin_trans) done next assume ?rhs then show ?lhs using compact_eq_bounded_closed locally_mono by blast qed lemma locally_compact_openin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" and opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ S" by (meson \x \ S\ opS openin_contains_cball) then have "cball x e2 \ (S \ T) = cball x e2 \ S" by force ultimately show ?thesis apply (rule_tac x="min e1 e2" in exI) apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) by (metis closed_Int closed_cball inf_left_commute) qed moreover have "\e>0. closed (cball x e \ (S \ T))" if "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \ (S \ T) \ T" by (meson \x \ T\ opT openin_contains_cball) then have "cball x e2 \ (S \ T) = cball x e2 \ T" by force ultimately show ?thesis apply (rule_tac x="min e1 e2" in exI) apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int) by (metis closed_Int closed_cball inf_left_commute) qed ultimately show ?thesis by (force simp: locally_compact_Int_cball) qed lemma locally_compact_closedin_Un: fixes S :: "'a::euclidean_space set" assumes LCS: "locally compact S" and LCT:"locally compact T" and clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" shows "locally compact (S \ T)" proof - have "\e>0. closed (cball x e \ (S \ T))" if "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast ultimately show ?thesis apply (rule_tac x="min e1 e2" in exI) apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) by (metis closed_Int closed_Un closed_cball inf_left_commute) qed moreover have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ S)" using LCS \x \ S\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S - T" using clT x by (fastforce simp: openin_contains_cball closedin_def) then have "closed (cball x e2 \ T)" proof - have "{} = T - (T - cball x e2)" using Diff_subset Int_Diff \cball x e2 \ (S \ T) \ S - T\ by auto then show ?thesis by (simp add: Diff_Diff_Int inf_commute) qed ultimately show ?thesis apply (rule_tac x="min e1 e2" in exI) apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) by (metis closed_Int closed_Un closed_cball inf_left_commute) qed moreover have "\e>0. closed (cball x e \ (S \ T))" if x: "x \ S" "x \ T" for x proof - obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \ T)" using LCT \x \ T\ unfolding locally_compact_Int_cball by blast moreover obtain e2 where "e2>0" and "cball x e2 \ (S \ T) \ S \ T - S" using clS x by (fastforce simp: openin_contains_cball closedin_def) then have "closed (cball x e2 \ S)" by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) ultimately show ?thesis apply (rule_tac x="min e1 e2" in exI) apply (auto simp: cball_min_Int \e2 > 0\ inf_assoc closed_Int Int_Un_distrib) by (metis closed_Int closed_Un closed_cball inf_left_commute) qed ultimately show ?thesis by (auto simp: locally_compact_Int_cball) qed lemma locally_compact_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "\locally compact S; locally compact T\ \ locally compact (S \ T)" by (auto simp: compact_Times locally_Times) lemma locally_compact_compact_subopen: fixes S :: "'a :: heine_borel set" shows "locally compact S \ (\K T. K \ S \ compact K \ open T \ K \ T \ (\U V. K \ U \ U \ V \ U \ T \ V \ S \ openin (top_of_set S) U \ compact V))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof clarify fix K :: "'a set" and T :: "'a set" assume "K \ S" and "compact K" and "open T" and "K \ T" obtain U V where "K \ U" "U \ V" "V \ S" "compact V" and ope: "openin (top_of_set S) U" using L unfolding locally_compact_compact by (meson \K \ S\ \compact K\) show "\U V. K \ U \ U \ V \ U \ T \ V \ S \ openin (top_of_set S) U \ compact V" proof (intro exI conjI) show "K \ U \ T" by (simp add: \K \ T\ \K \ U\) show "U \ T \ closure(U \ T)" by (rule closure_subset) show "closure (U \ T) \ S" by (metis \U \ V\ \V \ S\ \compact V\ closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) show "openin (top_of_set S) (U \ T)" by (simp add: \open T\ ope openin_Int_open) show "compact (closure (U \ T))" by (meson Int_lower1 \U \ V\ \compact V\ bounded_subset compact_closure compact_eq_bounded_closed) qed auto qed next assume ?rhs then show ?lhs unfolding locally_compact_compact by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) qed subsection\Sura-Bura's results about compact components of sets\ proposition Sura_Bura_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" and C: "C \ components S" shows "C = \{T. C \ T \ openin (top_of_set S) T \ closedin (top_of_set S) T}" (is "C = \?\") proof obtain x where x: "C = connected_component_set S x" and "x \ S" using C by (auto simp: components_def) have "C \ S" by (simp add: C in_components_subset) have "\?\ \ connected_component_set S x" proof (rule connected_component_maximal) have "x \ C" by (simp add: \x \ S\ x) then show "x \ \?\" by blast have clo: "closed (\?\)" by (simp add: \compact S\ closed_Inter closedin_compact_eq compact_imp_closed) have False if K1: "closedin (top_of_set (\?\)) K1" and K2: "closedin (top_of_set (\?\)) K2" and K12_Int: "K1 \ K2 = {}" and K12_Un: "K1 \ K2 = \?\" and "K1 \ {}" "K2 \ {}" for K1 K2 proof - have "closed K1" "closed K2" using closedin_closed_trans clo K1 K2 by blast+ then obtain V1 V2 where "open V1" "open V2" "K1 \ V1" "K2 \ V2" and V12: "V1 \ V2 = {}" using separation_normal \K1 \ K2 = {}\ by metis have SV12_ne: "(S - (V1 \ V2)) \ (\?\) \ {}" proof (rule compact_imp_fip) show "compact (S - (V1 \ V2))" by (simp add: \open V1\ \open V2\ \compact S\ compact_diff open_Un) show clo\: "closed T" if "T \ ?\" for T using that \compact S\ by (force intro: closedin_closed_trans simp add: compact_imp_closed) show "(S - (V1 \ V2)) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof assume djo: "(S - (V1 \ V2)) \ \\ = {}" obtain D where opeD: "openin (top_of_set S) D" and cloD: "closedin (top_of_set S) D" and "C \ D" and DV12: "D \ V1 \ V2" proof (cases "\ = {}") case True with \C \ S\ djo that show ?thesis by force next case False show ?thesis proof show ope: "openin (top_of_set S) (\\)" using openin_Inter \finite \\ False \ by blast then show "closedin (top_of_set S) (\\)" by (meson clo\ \ closed_Inter closed_subset openin_imp_subset subset_eq) show "C \ \\" using \ by auto show "\\ \ V1 \ V2" using ope djo openin_imp_subset by fastforce qed qed have "connected C" by (simp add: x) have "closed D" using \compact S\ cloD closedin_closed_trans compact_imp_closed by blast have cloV1: "closedin (top_of_set D) (D \ closure V1)" and cloV2: "closedin (top_of_set D) (D \ closure V2)" by (simp_all add: closedin_closed_Int) moreover have "D \ closure V1 = D \ V1" "D \ closure V2 = D \ V2" apply safe using \D \ V1 \ V2\ \open V1\ \open V2\ V12 apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) done ultimately have cloDV1: "closedin (top_of_set D) (D \ V1)" and cloDV2: "closedin (top_of_set D) (D \ V2)" by metis+ then obtain U1 U2 where "closed U1" "closed U2" and D1: "D \ V1 = D \ U1" and D2: "D \ V2 = D \ U2" by (auto simp: closedin_closed) have "D \ U1 \ C \ {}" proof assume "D \ U1 \ C = {}" then have *: "C \ D \ V2" using D1 DV12 \C \ D\ by auto have "\?\ \ D \ V2" apply (rule Inter_lower) using * apply simp by (meson cloDV2 \open V2\ cloD closedin_trans le_inf_iff opeD openin_Int_open) then show False using K1 V12 \K1 \ {}\ \K1 \ V1\ closedin_imp_subset by blast qed moreover have "D \ U2 \ C \ {}" proof assume "D \ U2 \ C = {}" then have *: "C \ D \ V1" using D2 DV12 \C \ D\ by auto have "\?\ \ D \ V1" apply (rule Inter_lower) using * apply simp by (meson cloDV1 \open V1\ cloD closedin_trans le_inf_iff opeD openin_Int_open) then show False using K2 V12 \K2 \ {}\ \K2 \ V2\ closedin_imp_subset by blast qed ultimately show False using \connected C\ unfolding connected_closed apply (simp only: not_ex) apply (drule_tac x="D \ U1" in spec) apply (drule_tac x="D \ U2" in spec) using \C \ D\ D1 D2 V12 DV12 \closed U1\ \closed U2\ \closed D\ by blast qed qed show False by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \K1 \ V1\ \K2 \ V2\ disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) qed then show "connected (\?\)" by (auto simp: connected_closedin_eq) show "\?\ \ S" by (fastforce simp: C in_components_subset) qed with x show "\?\ \ C" by simp qed auto corollary Sura_Bura_clopen_subset: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and U: "open U" "C \ U" obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof (rule ccontr) assume "\ thesis" with that have neg: "\K. openin (top_of_set S) K \ compact K \ C \ K \ K \ U" by metis obtain V K where "C \ V" "V \ U" "V \ K" "K \ S" "compact K" and opeSV: "openin (top_of_set S) V" using S U \compact C\ apply (simp add: locally_compact_compact_subopen) by (meson C in_components_subset) let ?\ = "{T. C \ T \ openin (top_of_set K) T \ compact T \ T \ K}" have CK: "C \ components K" by (meson C \C \ V\ \K \ S\ \V \ K\ components_intermediate_subset subset_trans) with \compact K\ have "C = \{T. C \ T \ openin (top_of_set K) T \ closedin (top_of_set K) T}" by (simp add: Sura_Bura_compact) then have Ceq: "C = \?\" by (simp add: closedin_compact_eq \compact K\) obtain W where "open W" and W: "V = S \ W" using opeSV by (auto simp: openin_open) have "-(U \ W) \ \?\ \ {}" proof (rule closed_imp_fip_compact) show "- (U \ W) \ \\ \ {}" if "finite \" and \: "\ \ ?\" for \ proof (cases "\ = {}") case True have False if "U = UNIV" "W = UNIV" proof - have "V = S" by (simp add: W \W = UNIV\) with neg show False using \C \ V\ \K \ S\ \V \ K\ \V \ U\ \compact K\ by auto qed with True show ?thesis by auto next case False show ?thesis proof assume "- (U \ W) \ \\ = {}" then have FUW: "\\ \ U \ W" by blast have "C \ \\" using \ by auto moreover have "compact (\\)" by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \) moreover have "\\ \ K" using False that(2) by fastforce moreover have opeKF: "openin (top_of_set K) (\\)" using False \ \finite \\ by blast then have opeVF: "openin (top_of_set V) (\\)" using W \K \ S\ \V \ K\ opeKF \\\ \ K\ FUW openin_subset_trans by fastforce then have "openin (top_of_set S) (\\)" by (metis opeSV openin_trans) moreover have "\\ \ U" by (meson \V \ U\ opeVF dual_order.trans openin_imp_subset) ultimately show False using neg by blast qed qed qed (use \open W\ \open U\ in auto) with W Ceq \C \ V\ \C \ U\ show False by auto qed corollary Sura_Bura_clopen_subset_alt: fixes S :: "'a::euclidean_space set" assumes S: "locally compact S" and C: "C \ components S" and "compact C" and opeSU: "openin (top_of_set S) U" and "C \ U" obtains K where "openin (top_of_set S) K" "compact K" "C \ K" "K \ U" proof - obtain V where "open V" "U = S \ V" using opeSU by (auto simp: openin_open) with \C \ U\ have "C \ V" by auto then show ?thesis using Sura_Bura_clopen_subset [OF S C \compact C\ \open V\] by (metis \U = S \ V\ inf.bounded_iff openin_imp_subset that) qed corollary Sura_Bura: fixes S :: "'a::euclidean_space set" assumes "locally compact S" "C \ components S" "compact C" shows "C = \ {K. C \ K \ compact K \ openin (top_of_set S) K}" (is "C = ?rhs") proof show "?rhs \ C" proof (clarsimp, rule ccontr) fix x assume *: "\X. C \ X \ compact X \ openin (top_of_set S) X \ x \ X" and "x \ C" obtain U V where "open U" "open V" "{x} \ U" "C \ V" "U \ V = {}" using separation_normal [of "{x}" C] by (metis Int_empty_left \x \ C\ \compact C\ closed_empty closed_insert compact_imp_closed insert_disjoint(1)) have "x \ V" using \U \ V = {}\ \{x} \ U\ by blast then show False by (meson "*" Sura_Bura_clopen_subset \C \ V\ \open V\ assms(1) assms(2) assms(3) subsetCE) qed qed blast subsection\Special cases of local connectedness and path connectedness\ lemma locally_connected_1: assumes "\v x. \openin (top_of_set S) v; x \ v\ \ \u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" shows "locally connected S" apply (clarsimp simp add: locally_def) apply (drule assms; blast) done lemma locally_connected_2: assumes "locally connected S" "openin (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (connected_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. connected v \ x \ u \ u \ v \ v \ w))" and "connected_component t x y" then have "y \ t" and y: "y \ connected_component_set t x" using connected_component_subset by blast+ obtain F where "\x y. (\w. openin ?SS w \ (\u. connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. connected u \ x \ F x y \ F x y \ u \ u \ y))" by moura then obtain G where "\a A. (\U. openin ?SS U \ (\V. connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" by moura then have *: "openin ?SS (F y t) \ connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" using 1 \y \ t\ by presburger have "G y t \ connected_component_set t y" by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) then have "\A. openin ?SS A \ y \ A \ A \ connected_component_set t x" by (metis (no_types) * connected_component_eq dual_order.trans y) } then show ?thesis using assms openin_subopen by (force simp: locally_def) qed lemma locally_connected_3: assumes "\t x. \openin (top_of_set S) t; x \ t\ \ openin (top_of_set S) (connected_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v" using assms connected_component_subset by fastforce lemma locally_connected: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ connected u \ x \ u \ u \ v))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_connected_open_connected_component: "locally connected S \ (\t x. openin (top_of_set S) t \ x \ t \ openin (top_of_set S) (connected_component_set t x))" by (metis locally_connected_1 locally_connected_2 locally_connected_3) lemma locally_path_connected_1: assumes "\v x. \openin (top_of_set S) v; x \ v\ \ \u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" shows "locally path_connected S" apply (clarsimp simp add: locally_def) apply (drule assms; blast) done lemma locally_path_connected_2: assumes "locally path_connected S" "openin (top_of_set S) t" "x \ t" shows "openin (top_of_set S) (path_component_set t x)" proof - { fix y :: 'a let ?SS = "top_of_set S" assume 1: "openin ?SS t" "\w x. openin ?SS w \ x \ w \ (\u. openin ?SS u \ (\v. path_connected v \ x \ u \ u \ v \ v \ w))" and "path_component t x y" then have "y \ t" and y: "y \ path_component_set t x" using path_component_mem(2) by blast+ obtain F where "\x y. (\w. openin ?SS w \ (\u. path_connected u \ x \ w \ w \ u \ u \ y)) = (openin ?SS (F x y) \ (\u. path_connected u \ x \ F x y \ F x y \ u \ u \ y))" by moura then obtain G where "\a A. (\U. openin ?SS U \ (\V. path_connected V \ a \ U \ U \ V \ V \ A)) = (openin ?SS (F a A) \ path_connected (G a A) \ a \ F a A \ F a A \ G a A \ G a A \ A)" by moura then have *: "openin ?SS (F y t) \ path_connected (G y t) \ y \ F y t \ F y t \ G y t \ G y t \ t" using 1 \y \ t\ by presburger have "G y t \ path_component_set t y" using * path_component_maximal rev_subsetD by blast then have "\A. openin ?SS A \ y \ A \ A \ path_component_set t x" by (metis "*" \G y t \ path_component_set t y\ dual_order.trans path_component_eq y) } then show ?thesis using assms openin_subopen by (force simp: locally_def) qed lemma locally_path_connected_3: assumes "\t x. \openin (top_of_set S) t; x \ t\ \ openin (top_of_set S) (path_component_set t x)" "openin (top_of_set S) v" "x \ v" shows "\u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v" proof - have "path_component v x x" by (meson assms(3) path_component_refl) then show ?thesis by (metis assms(1) assms(2) assms(3) mem_Collect_eq path_component_subset path_connected_path_component) qed proposition locally_path_connected: "locally path_connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ path_connected u \ x \ u \ u \ v))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) proposition locally_path_connected_open_path_component: "locally path_connected S \ (\t x. openin (top_of_set S) t \ x \ t \ openin (top_of_set S) (path_component_set t x))" by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) lemma locally_connected_open_component: "locally connected S \ (\t c. openin (top_of_set S) t \ c \ components t \ openin (top_of_set S) c)" by (metis components_iff locally_connected_open_connected_component) proposition locally_connected_im_kleinen: "locally connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\c. connected c \ c \ v \ x \ c \ y \ c))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (fastforce simp add: locally_connected) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ c" if "openin (top_of_set S) t" and c: "c \ components t" and "x \ c" for t c x proof - from that \?rhs\ [rule_format, of t x] obtain u where u: "openin (top_of_set S) u \ x \ u \ u \ t \ (\y. y \ u \ (\c. connected c \ c \ t \ x \ c \ y \ c))" using in_components_subset by auto obtain F :: "'a set \ 'a set \ 'a" where "\x y. (\z. z \ x \ y = connected_component_set x z) = (F x y \ x \ y = connected_component_set x (F x y))" by moura then have F: "F t c \ t \ c = connected_component_set t (F t c)" by (meson components_iff c) obtain G :: "'a set \ 'a set \ 'a" where G: "\x y. (\z. z \ y \ z \ x) = (G x y \ y \ G x y \ x)" by moura have "G c u \ u \ G c u \ c" using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) then show ?thesis using G u by auto qed show ?lhs apply (clarsimp simp add: locally_connected_open_component) apply (subst openin_subopen) apply (blast intro: *) done qed proposition locally_path_connected_im_kleinen: "locally path_connected S \ (\v x. openin (top_of_set S) v \ x \ v \ (\u. openin (top_of_set S) u \ x \ u \ u \ v \ (\y. y \ u \ (\p. path p \ path_image p \ v \ pathstart p = x \ pathfinish p = y))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: locally_path_connected path_connected_def) apply (erule all_forward ex_forward imp_forward conjE | simp)+ by (meson dual_order.trans) next assume ?rhs have *: "\T. openin (top_of_set S) T \ x \ T \ T \ path_component_set u z" if "openin (top_of_set S) u" and "z \ u" and c: "path_component u z x" for u z x proof - have "x \ u" by (meson c path_component_mem(2)) with that \?rhs\ [rule_format, of u x] obtain U where U: "openin (top_of_set S) U \ x \ U \ U \ u \ (\y. y \ U \ (\p. path p \ path_image p \ u \ pathstart p = x \ pathfinish p = y))" by blast show ?thesis apply (rule_tac x=U in exI) apply (auto simp: U) apply (metis U c path_component_trans path_component_def) done qed show ?lhs apply (clarsimp simp add: locally_path_connected_open_path_component) apply (subst openin_subopen) apply (blast intro: *) done qed lemma locally_path_connected_imp_locally_connected: "locally path_connected S \ locally connected S" using locally_mono path_connected_imp_connected by blast lemma locally_connected_components: "\locally connected S; c \ components S\ \ locally connected c" by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) lemma locally_path_connected_components: "\locally path_connected S; c \ components S\ \ locally path_connected c" by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) lemma locally_path_connected_connected_component: "locally path_connected S \ locally path_connected (connected_component_set S x)" by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) lemma open_imp_locally_path_connected: fixes S :: "'a :: real_normed_vector set" shows "open S \ locally path_connected S" apply (rule locally_mono [of convex]) apply (simp_all add: locally_def openin_open_eq convex_imp_path_connected) apply (meson open_ball centre_in_ball convex_ball openE order_trans) done lemma open_imp_locally_connected: fixes S :: "'a :: real_normed_vector set" shows "open S \ locally connected S" by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" by (simp add: open_imp_locally_path_connected) lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" by (simp add: open_imp_locally_connected) lemma openin_connected_component_locally_connected: "locally connected S \ openin (top_of_set S) (connected_component_set S x)" apply (simp add: locally_connected_open_connected_component) by (metis connected_component_eq_empty connected_component_subset open_empty open_subset openin_subtopology_self) lemma openin_components_locally_connected: "\locally connected S; c \ components S\ \ openin (top_of_set S) c" using locally_connected_open_component openin_subtopology_self by blast lemma openin_path_component_locally_path_connected: "locally path_connected S \ openin (top_of_set S) (path_component_set S x)" by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) lemma closedin_path_component_locally_path_connected: "locally path_connected S \ closedin (top_of_set S) (path_component_set S x)" apply (simp add: closedin_def path_component_subset complement_path_component_Union) apply (rule openin_Union) using openin_path_component_locally_path_connected by auto lemma convex_imp_locally_path_connected: fixes S :: "'a:: real_normed_vector set" shows "convex S \ locally path_connected S" apply (clarsimp simp add: locally_path_connected) apply (subst (asm) openin_open) apply clarify apply (erule (1) openE) apply (rule_tac x = "S \ ball x e" in exI) apply (force simp: convex_Int convex_imp_path_connected) done lemma convex_imp_locally_connected: fixes S :: "'a:: real_normed_vector set" shows "convex S \ locally connected S" by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) subsection\Relations between components and path components\ lemma path_component_eq_connected_component: assumes "locally path_connected S" shows "(path_component S x = connected_component S x)" proof (cases "x \ S") case True have "openin (top_of_set (connected_component_set S x)) (path_component_set S x)" apply (rule openin_subset_trans [of S]) apply (intro conjI openin_path_component_locally_path_connected [OF assms]) using path_component_subset_connected_component apply (auto simp: connected_component_subset) done moreover have "closedin (top_of_set (connected_component_set S x)) (path_component_set S x)" apply (rule closedin_subset_trans [of S]) apply (intro conjI closedin_path_component_locally_path_connected [OF assms]) using path_component_subset_connected_component apply (auto simp: connected_component_subset) done ultimately have *: "path_component_set S x = connected_component_set S x" by (metis connected_connected_component connected_clopen True path_component_eq_empty) then show ?thesis by blast next case False then show ?thesis by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) qed lemma path_component_eq_connected_component_set: "locally path_connected S \ (path_component_set S x = connected_component_set S x)" by (simp add: path_component_eq_connected_component) lemma locally_path_connected_path_component: "locally path_connected S \ locally path_connected (path_component_set S x)" using locally_path_connected_connected_component path_component_eq_connected_component by fastforce lemma open_path_connected_component: fixes S :: "'a :: real_normed_vector set" shows "open S \ path_component S x = connected_component S x" by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) lemma open_path_connected_component_set: fixes S :: "'a :: real_normed_vector set" shows "open S \ path_component_set S x = connected_component_set S x" by (simp add: open_path_connected_component) proposition locally_connected_quotient_image: assumes lcS: "locally connected S" and oo: "\T. T \ f ` S \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally connected (f ` S)" proof (clarsimp simp: locally_connected_open_component) fix U C assume opefSU: "openin (top_of_set (f ` S)) U" and "C \ components U" then have "C \ U" "U \ f ` S" by (meson in_components_subset openin_imp_subset)+ then have "openin (top_of_set (f ` S)) C \ openin (top_of_set S) (S \ f -` C)" by (auto simp: oo) moreover have "openin (top_of_set S) (S \ f -` C)" proof (subst openin_subopen, clarify) fix x assume "x \ S" "f x \ C" show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` C)" proof (intro conjI exI) show "openin (top_of_set S) (connected_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set S) (connected_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" using \U \ f ` S\ opefSU lcS locally_connected_2 oo by blast with ** show False by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) qed next show "x \ connected_component_set (S \ f -` U) x" using \C \ U\ \f x \ C\ \x \ S\ by auto next have contf: "continuous_on S f" by (simp add: continuous_on_open oo openin_imp_subset) then have "continuous_on (connected_component_set (S \ f -` U) x) f" apply (rule continuous_on_subset) using connected_component_subset apply blast done then have "connected (f ` connected_component_set (S \ f -` U) x)" by (rule connected_continuous_image [OF _ connected_connected_component]) moreover have "f ` connected_component_set (S \ f -` U) x \ U" using connected_component_in by blast moreover have "C \ f ` connected_component_set (S \ f -` U) x \ {}" using \C \ U\ \f x \ C\ \x \ S\ by fastforce ultimately have fC: "f ` (connected_component_set (S \ f -` U) x) \ C" by (rule components_maximal [OF \C \ components U\]) have cUC: "connected_component_set (S \ f -` U) x \ (S \ f -` C)" using connected_component_subset fC by blast have "connected_component_set (S \ f -` U) x \ connected_component_set (S \ f -` C) x" proof - { assume "x \ connected_component_set (S \ f -` U) x" then have ?thesis using cUC connected_component_idemp connected_component_mono by blast } then show ?thesis using connected_component_eq_empty by auto qed also have "\ \ (S \ f -` C)" by (rule connected_component_subset) finally show "connected_component_set (S \ f -` U) x \ (S \ f -` C)" . qed qed ultimately show "openin (top_of_set (f ` S)) C" by metis qed text\The proof resembles that above but is not identical!\ proposition locally_path_connected_quotient_image: assumes lcS: "locally path_connected S" and oo: "\T. T \ f ` S \ openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" shows "locally path_connected (f ` S)" proof (clarsimp simp: locally_path_connected_open_path_component) fix U y assume opefSU: "openin (top_of_set (f ` S)) U" and "y \ U" then have "path_component_set U y \ U" "U \ f ` S" by (meson path_component_subset openin_imp_subset)+ then have "openin (top_of_set (f ` S)) (path_component_set U y) \ openin (top_of_set S) (S \ f -` path_component_set U y)" proof - have "path_component_set U y \ f ` S" using \U \ f ` S\ \path_component_set U y \ U\ by blast then show ?thesis using oo by blast qed moreover have "openin (top_of_set S) (S \ f -` path_component_set U y)" proof (subst openin_subopen, clarify) fix x assume "x \ S" and Uyfx: "path_component U y (f x)" then have "f x \ U" using path_component_mem by blast show "\T. openin (top_of_set S) T \ x \ T \ T \ (S \ f -` path_component_set U y)" proof (intro conjI exI) show "openin (top_of_set S) (path_component_set (S \ f -` U) x)" proof (rule ccontr) assume **: "\ openin (top_of_set S) (path_component_set (S \ f -` U) x)" then have "x \ (S \ f -` U)" by (metis (no_types, lifting) \U \ f ` S\ opefSU lcS oo locally_path_connected_open_path_component) then show False using ** \path_component_set U y \ U\ \x \ S\ \path_component U y (f x)\ by blast qed next show "x \ path_component_set (S \ f -` U) x" by (simp add: \f x \ U\ \x \ S\ path_component_refl) next have contf: "continuous_on S f" by (simp add: continuous_on_open oo openin_imp_subset) then have "continuous_on (path_component_set (S \ f -` U) x) f" apply (rule continuous_on_subset) using path_component_subset apply blast done then have "path_connected (f ` path_component_set (S \ f -` U) x)" by (simp add: path_connected_continuous_image) moreover have "f ` path_component_set (S \ f -` U) x \ U" using path_component_mem by fastforce moreover have "f x \ f ` path_component_set (S \ f -` U) x" by (force simp: \x \ S\ \f x \ U\ path_component_refl_eq) ultimately have "f ` (path_component_set (S \ f -` U) x) \ path_component_set U (f x)" by (meson path_component_maximal) also have "\ \ path_component_set U y" by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) finally have fC: "f ` (path_component_set (S \ f -` U) x) \ path_component_set U y" . have cUC: "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" using path_component_subset fC by blast have "path_component_set (S \ f -` U) x \ path_component_set (S \ f -` path_component_set U y) x" proof - have "\a. path_component_set (path_component_set (S \ f -` U) x) a \ path_component_set (S \ f -` path_component_set U y) a" using cUC path_component_mono by blast then show ?thesis using path_component_path_component by blast qed also have "\ \ (S \ f -` path_component_set U y)" by (rule path_component_subset) finally show "path_component_set (S \ f -` U) x \ (S \ f -` path_component_set U y)" . qed qed ultimately show "openin (top_of_set (f ` S)) (path_component_set U y)" by metis qed subsection\<^marker>\tag unimportant\\Components, continuity, openin, closedin\ lemma continuous_on_components_gen: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "\c. c \ components S \ openin (top_of_set S) c \ continuous_on c f" shows "continuous_on S f" proof (clarsimp simp: continuous_openin_preimage_eq) fix t :: "'b set" assume "open t" have *: "S \ f -` t = (\c \ components S. c \ f -` t)" by auto show "openin (top_of_set S) (S \ f -` t)" unfolding * using \open t\ assms continuous_openin_preimage_gen openin_trans openin_Union by blast qed lemma continuous_on_components: fixes f :: "'a::topological_space \ 'b::topological_space" assumes "locally connected S " "\c. c \ components S \ continuous_on c f" shows "continuous_on S f" apply (rule continuous_on_components_gen) apply (auto simp: assms intro: openin_components_locally_connected) done lemma continuous_on_components_eq: "locally connected S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" by (meson continuous_on_components continuous_on_subset in_components_subset) lemma continuous_on_components_open: fixes S :: "'a::real_normed_vector set" assumes "open S " "\c. c \ components S \ continuous_on c f" shows "continuous_on S f" using continuous_on_components open_imp_locally_connected assms by blast lemma continuous_on_components_open_eq: fixes S :: "'a::real_normed_vector set" shows "open S \ (continuous_on S f \ (\c \ components S. continuous_on c f))" using continuous_on_subset in_components_subset by (blast intro: continuous_on_components_open) lemma closedin_union_complement_components: assumes u: "locally connected u" and S: "closedin (top_of_set u) S" and cuS: "c \ components(u - S)" shows "closedin (top_of_set u) (S \ \c)" proof - have di: "(\S t. S \ c \ t \ c' \ disjnt S t) \ disjnt (\ c) (\ c')" for c' by (simp add: disjnt_def) blast have "S \ u" using S closedin_imp_subset by blast moreover have "u - S = \c \ \(components (u - S) - c)" by (metis Diff_partition Union_components Union_Un_distrib assms(3)) moreover have "disjnt (\c) (\(components (u - S) - c))" apply (rule di) by (metis DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) ultimately have eq: "S \ \c = u - (\(components(u - S) - c))" by (auto simp: disjnt_def) have *: "openin (top_of_set u) (\(components (u - S) - c))" apply (rule openin_Union) apply (rule openin_trans [of "u - S"]) apply (simp add: u S locally_diff_closed openin_components_locally_connected) apply (simp add: openin_diff S) done have "openin (top_of_set u) (u - (u - \(components (u - S) - c)))" apply (rule openin_diff, simp) apply (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) done then show ?thesis by (force simp: eq closedin_def) qed lemma closed_union_complement_components: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and c: "c \ components(- S)" shows "closed(S \ \ c)" proof - have "closedin (top_of_set UNIV) (S \ \c)" apply (rule closedin_union_complement_components [OF locally_connected_UNIV]) using S c apply (simp_all add: Compl_eq_Diff_UNIV) done then show ?thesis by simp qed lemma closedin_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes u: "locally connected u" and S: "closedin (top_of_set u) S" and c: " c \ components(u - S)" shows "closedin (top_of_set u) (S \ c)" proof - have "closedin (top_of_set u) (S \ \{c})" using c by (blast intro: closedin_union_complement_components [OF u S]) then show ?thesis by simp qed lemma closed_Un_complement_component: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and c: " c \ components(-S)" shows "closed (S \ c)" by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component locally_connected_UNIV subtopology_UNIV) subsection\Existence of isometry between subspaces of same dimension\ lemma isometry_subset_subspace: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S \ dim T" obtains f where "linear f" "f ` S \ T" "\x. x \ S \ norm(f x) = norm x" proof - obtain B where "B \ S" and Borth: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" "finite B" "card B = dim S" "span B = S" by (metis orthonormal_basis_subspace [OF S] independent_finite) obtain C where "C \ T" and Corth: "pairwise orthogonal C" and C1:"\x. x \ C \ norm x = 1" and "independent C" "finite C" "card C = dim T" "span C = T" by (metis orthonormal_basis_subspace [OF T] independent_finite) obtain fb where "fb ` B \ C" "inj_on fb B" by (metis \card B = dim S\ \card C = dim T\ \finite B\ \finite C\ card_le_inj d) then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" using Corth apply (auto simp: pairwise_def orthogonal_clauses) by (meson subsetD image_eqI inj_on_def) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce have "span (f ` B) \ span C" by (metis \fb ` B \ C\ ffb image_cong span_mono) then have "f ` S \ T" unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \fb ` B \ C\ by auto have "norm (f x) = norm x" if "x \ S" for x proof - interpret linear f by fact obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce have "norm (f x)^2 = norm (\v\B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" apply (rule norm_sum_Pythagorean [OF \finite B\]) apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) done also have "\ = norm x ^2" by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) finally show ?thesis by (simp add: norm_eq_sqrt_inner) qed then show ?thesis by (rule that [OF \linear f\ \f ` S \ T\]) qed proposition isometries_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" "\x. x \ S \ norm(f x) = norm x" "\x. x \ T \ norm(g x) = norm x" "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" proof - obtain B where "B \ S" and Borth: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" "finite B" "card B = dim S" "span B = S" by (metis orthonormal_basis_subspace [OF S] independent_finite) obtain C where "C \ T" and Corth: "pairwise orthogonal C" and C1:"\x. x \ C \ norm x = 1" and "independent C" "finite C" "card C = dim T" "span C = T" by (metis orthonormal_basis_subspace [OF T] independent_finite) obtain fb where "bij_betw fb B C" by (metis \finite B\ \finite C\ bij_betw_iff_card \card B = dim S\ \card C = dim T\ d) then have pairwise_orth_fb: "pairwise (\v j. orthogonal (fb v) (fb j)) B" using Corth apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) by (meson subsetD image_eqI inj_on_def) obtain f where "linear f" and ffb: "\x. x \ B \ f x = fb x" using linear_independent_extend \independent B\ by fastforce interpret f: linear f by fact define gb where "gb \ inv_into B fb" then have pairwise_orth_gb: "pairwise (\v j. orthogonal (gb v) (gb j)) C" using Borth apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) by (metis \bij_betw fb B C\ bij_betw_imp_surj_on bij_betw_inv_into_right inv_into_into) obtain g where "linear g" and ggb: "\x. x \ C \ g x = gb x" using linear_independent_extend \independent C\ by fastforce interpret g: linear g by fact have "span (f ` B) \ span C" by (metis \bij_betw fb B C\ bij_betw_imp_surj_on eq_iff ffb image_cong) then have "f ` S \ T" unfolding \span B = S\ \span C = T\ span_linear_image[OF \linear f\] . have [simp]: "\x. x \ B \ norm (fb x) = norm x" using B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on by fastforce have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \ S" for x proof - obtain a where x: "x = (\v \ B. a v *\<^sub>R v)" using \finite B\ \span B = S\ \x \ S\ span_finite by fastforce have "f x = (\v \ B. f (a v *\<^sub>R v))" using linear_sum [OF \linear f\] x by auto also have "\ = (\v \ B. a v *\<^sub>R f v)" by (simp add: f.sum f.scale) also have "\ = (\v \ B. a v *\<^sub>R fb v)" by (simp add: ffb cong: sum.cong) finally have *: "f x = (\v\B. a v *\<^sub>R fb v)" . then have "(norm (f x))\<^sup>2 = (norm (\v\B. a v *\<^sub>R fb v))\<^sup>2" by simp also have "\ = (\v\B. norm ((a v *\<^sub>R fb v))^2)" apply (rule norm_sum_Pythagorean [OF \finite B\]) apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) done also have "\ = (norm x)\<^sup>2" by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \finite B\]) finally show "norm (f x) = norm x" by (simp add: norm_eq_sqrt_inner) have "g (f x) = g (\v\B. a v *\<^sub>R fb v)" by (simp add: *) also have "\ = (\v\B. g (a v *\<^sub>R fb v))" by (simp add: g.sum g.scale) also have "\ = (\v\B. a v *\<^sub>R g (fb v))" by (simp add: g.scale) also have "\ = (\v\B. a v *\<^sub>R v)" apply (rule sum.cong [OF refl]) using \bij_betw fb B C\ gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce also have "\ = x" using x by blast finally show "g (f x) = x" . qed have [simp]: "\x. x \ C \ norm (gb x) = norm x" by (metis B1 C1 \bij_betw fb B C\ bij_betw_imp_surj_on gb_def inv_into_into) have g [simp]: "f (g x) = x" if "x \ T" for x proof - obtain a where x: "x = (\v \ C. a v *\<^sub>R v)" using \finite C\ \span C = T\ \x \ T\ span_finite by fastforce have "g x = (\v \ C. g (a v *\<^sub>R v))" by (simp add: x g.sum) also have "\ = (\v \ C. a v *\<^sub>R g v)" by (simp add: g.scale) also have "\ = (\v \ C. a v *\<^sub>R gb v)" by (simp add: ggb cong: sum.cong) finally have "f (g x) = f (\v\C. a v *\<^sub>R gb v)" by simp also have "\ = (\v\C. f (a v *\<^sub>R gb v))" by (simp add: f.scale f.sum) also have "\ = (\v\C. a v *\<^sub>R f (gb v))" by (simp add: f.scale f.sum) also have "\ = (\v\C. a v *\<^sub>R v)" using \bij_betw fb B C\ by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) also have "\ = x" using x by blast finally show "f (g x) = x" . qed have gim: "g ` T = S" by (metis (full_types) S T \f ` S \ T\ d dim_eq_span dim_image_le f(2) g.linear_axioms image_iff linear_subspace_image span_eq_iff subset_iff) have fim: "f ` S = T" using \g ` T = S\ image_iff by fastforce have [simp]: "norm (g x) = norm x" if "x \ T" for x using fim that by auto show ?thesis apply (rule that [OF \linear f\ \linear g\]) apply (simp_all add: fim gim) done qed corollary isometry_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" obtains f where "linear f" "f ` S = T" "\x. x \ S \ norm(f x) = norm x" using isometries_subspaces [OF assms] by metis corollary isomorphisms_UNIV_UNIV: assumes "DIM('M) = DIM('N)" obtains f::"'M::euclidean_space \'N::euclidean_space" and g where "linear f" "linear g" "\x. norm(f x) = norm x" "\y. norm(g y) = norm y" "\x. g (f x) = x" "\y. f(g y) = y" using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) lemma homeomorphic_subspaces: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes S: "subspace S" and T: "subspace T" and d: "dim S = dim T" shows "S homeomorphic T" proof - obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" "\x. x \ S \ g(f x) = x" "\x. x \ T \ f(g x) = x" by (blast intro: isometries_subspaces [OF assms]) then show ?thesis apply (simp add: homeomorphic_def homeomorphism_def) apply (rule_tac x=f in exI) apply (rule_tac x=g in exI) apply (auto simp: linear_continuous_on linear_conv_bounded_linear) done qed lemma homeomorphic_affine_sets: assumes "affine S" "affine T" "aff_dim S = aff_dim T" shows "S homeomorphic T" proof (cases "S = {} \ T = {}") case True with assms aff_dim_empty homeomorphic_empty show ?thesis by metis next case False then obtain a b where ab: "a \ S" "b \ T" by auto then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) have "S homeomorphic ((+) (- a) ` S)" by (fact homeomorphic_translation) also have "\ homeomorphic ((+) (- b) ` T)" by (rule homeomorphic_subspaces [OF ss dd]) also have "\ homeomorphic T" using homeomorphic_translation [of T "- b"] by (simp add: homeomorphic_sym [of T]) finally show ?thesis . qed subsection\Retracts, in a general sense, preserve (co)homotopic triviality)\ locale\<^marker>\tag important\ Retracts = fixes s h t k assumes conth: "continuous_on s h" and imh: "h ` s = t" and contk: "continuous_on t k" and imk: "k ` t \ s" and idhk: "\y. y \ t \ h(k y) = y" begin lemma homotopically_trivial_retraction_gen: assumes P: "\f. \continuous_on u f; f ` u \ t; Q f\ \ P(k \ f)" and Q: "\f. \continuous_on u f; f ` u \ s; P f\ \ Q(h \ f)" and Qeq: "\h k. (\x. x \ u \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on u f; f ` u \ s; P f; continuous_on u g; g ` u \ s; P g\ \ homotopic_with_canon P u s f g" and contf: "continuous_on u f" and imf: "f ` u \ t" and Qf: "Q f" and contg: "continuous_on u g" and img: "g ` u \ t" and Qg: "Q g" shows "homotopic_with_canon Q u t f g" proof - have feq: "\x. x \ u \ (h \ (k \ f)) x = f x" using idhk imf by auto have geq: "\x. x \ u \ (h \ (k \ g)) x = g x" using idhk img by auto have "continuous_on u (k \ f)" using contf continuous_on_compose continuous_on_subset contk imf by blast moreover have "(k \ f) ` u \ s" using imf imk by fastforce moreover have "P (k \ f)" by (simp add: P Qf contf imf) moreover have "continuous_on u (k \ g)" using contg continuous_on_compose continuous_on_subset contk img by blast moreover have "(k \ g) ` u \ s" using img imk by fastforce moreover have "P (k \ g)" by (simp add: P Qg contg img) ultimately have "homotopic_with_canon P u s (k \ f) (k \ g)" by (rule hom) then have "homotopic_with_canon Q u t (h \ (k \ f)) (h \ (k \ g))" apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) using Q by (auto simp: conth imh) then show ?thesis apply (rule homotopic_with_eq) using feq geq apply fastforce+ using Qeq topspace_euclidean_subtopology by blast qed lemma homotopically_trivial_retraction_null_gen: assumes P: "\f. \continuous_on u f; f ` u \ t; Q f\ \ P(k \ f)" and Q: "\f. \continuous_on u f; f ` u \ s; P f\ \ Q(h \ f)" and Qeq: "\h k. (\x. x \ u \ h x = k x) \ Q h = Q k" and hom: "\f. \continuous_on u f; f ` u \ s; P f\ \ \c. homotopic_with_canon P u s f (\x. c)" and contf: "continuous_on u f" and imf:"f ` u \ t" and Qf: "Q f" obtains c where "homotopic_with_canon Q u t f (\x. c)" proof - have feq: "\x. x \ u \ (h \ (k \ f)) x = f x" using idhk imf by auto have "continuous_on u (k \ f)" using contf continuous_on_compose continuous_on_subset contk imf by blast moreover have "(k \ f) ` u \ s" using imf imk by fastforce moreover have "P (k \ f)" by (simp add: P Qf contf imf) ultimately obtain c where "homotopic_with_canon P u s (k \ f) (\x. c)" by (metis hom) then have "homotopic_with_canon Q u t (h \ (k \ f)) (h \ (\x. c))" apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) using Q by (auto simp: conth imh) then show ?thesis apply (rule_tac c = "h c" in that) apply (erule homotopic_with_eq) using feq apply auto[1] apply simp using Qeq topspace_euclidean_subtopology by blast qed lemma cohomotopically_trivial_retraction_gen: assumes P: "\f. \continuous_on t f; f ` t \ u; Q f\ \ P(f \ h)" and Q: "\f. \continuous_on s f; f ` s \ u; P f\ \ Q(f \ k)" and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on s f; f ` s \ u; P f; continuous_on s g; g ` s \ u; P g\ \ homotopic_with_canon P s u f g" and contf: "continuous_on t f" and imf: "f ` t \ u" and Qf: "Q f" and contg: "continuous_on t g" and img: "g ` t \ u" and Qg: "Q g" shows "homotopic_with_canon Q t u f g" proof - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto have geq: "\x. x \ t \ (g \ h \ k) x = g x" using idhk img by auto have "continuous_on s (f \ h)" using contf conth continuous_on_compose imh by blast moreover have "(f \ h) ` s \ u" using imf imh by fastforce moreover have "P (f \ h)" by (simp add: P Qf contf imf) moreover have "continuous_on s (g \ h)" using contg continuous_on_compose continuous_on_subset conth imh by blast moreover have "(g \ h) ` s \ u" using img imh by fastforce moreover have "P (g \ h)" by (simp add: P Qg contg img) ultimately have "homotopic_with_canon P s u (f \ h) (g \ h)" by (rule hom) then have "homotopic_with_canon Q t u (f \ h \ k) (g \ h \ k)" apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) using Q by (auto simp: contk imk) then show ?thesis apply (rule homotopic_with_eq) using feq apply auto[1] using geq apply auto[1] using Qeq topspace_euclidean_subtopology by blast qed lemma cohomotopically_trivial_retraction_null_gen: assumes P: "\f. \continuous_on t f; f ` t \ u; Q f\ \ P(f \ h)" and Q: "\f. \continuous_on s f; f ` s \ u; P f\ \ Q(f \ k)" and Qeq: "\h k. (\x. x \ t \ h x = k x) \ Q h = Q k" and hom: "\f g. \continuous_on s f; f ` s \ u; P f\ \ \c. homotopic_with_canon P s u f (\x. c)" and contf: "continuous_on t f" and imf: "f ` t \ u" and Qf: "Q f" obtains c where "homotopic_with_canon Q t u f (\x. c)" proof - have feq: "\x. x \ t \ (f \ h \ k) x = f x" using idhk imf by auto have "continuous_on s (f \ h)" using contf conth continuous_on_compose imh by blast moreover have "(f \ h) ` s \ u" using imf imh by fastforce moreover have "P (f \ h)" by (simp add: P Qf contf imf) ultimately obtain c where "homotopic_with_canon P s u (f \ h) (\x. c)" by (metis hom) then have "homotopic_with_canon Q t u (f \ h \ k) ((\x. c) \ k)" apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) using Q by (auto simp: contk imk) then show ?thesis apply (rule_tac c = c in that) apply (erule homotopic_with_eq) using feq apply auto[1] apply simp using Qeq topspace_euclidean_subtopology by blast qed end lemma simply_connected_retraction_gen: shows "\simply_connected S; continuous_on S h; h ` S = T; continuous_on T k; k ` T \ S; \y. y \ T \ h(k y) = y\ \ simply_connected T" apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) apply (rule Retracts.homotopically_trivial_retraction_gen [of S h _ k _ "\p. pathfinish p = pathstart p" "\p. pathfinish p = pathstart p"]) apply (simp_all add: Retracts_def pathfinish_def pathstart_def) done lemma homeomorphic_simply_connected: "\S homeomorphic T; simply_connected S\ \ simply_connected T" by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) lemma homeomorphic_simply_connected_eq: "S homeomorphic T \ (simply_connected S \ simply_connected T)" by (metis homeomorphic_simply_connected homeomorphic_sym) subsection\Homotopy equivalence\ subsection\Homotopy equivalence of topological spaces.\ definition\<^marker>\tag important\ homotopy_equivalent_space (infix "homotopy'_equivalent'_space" 50) where "X homotopy_equivalent_space Y \ (\f g. continuous_map X Y f \ continuous_map Y X g \ homotopic_with (\x. True) X X (g \ f) id \ homotopic_with (\x. True) Y Y (f \ g) id)" lemma homeomorphic_imp_homotopy_equivalent_space: "X homeomorphic_space Y \ X homotopy_equivalent_space Y" unfolding homeomorphic_space_def homotopy_equivalent_space_def apply (erule ex_forward)+ by (simp add: homotopic_with_equal homotopic_with_sym homeomorphic_maps_def) lemma homotopy_equivalent_space_refl: "X homotopy_equivalent_space X" by (simp add: homeomorphic_imp_homotopy_equivalent_space homeomorphic_space_refl) lemma homotopy_equivalent_space_sym: "X homotopy_equivalent_space Y \ Y homotopy_equivalent_space X" by (meson homotopy_equivalent_space_def) lemma homotopy_eqv_trans [trans]: assumes 1: "X homotopy_equivalent_space Y" and 2: "Y homotopy_equivalent_space U" shows "X homotopy_equivalent_space U" proof - obtain f1 g1 where f1: "continuous_map X Y f1" and g1: "continuous_map Y X g1" and hom1: "homotopic_with (\x. True) X X (g1 \ f1) id" "homotopic_with (\x. True) Y Y (f1 \ g1) id" using 1 by (auto simp: homotopy_equivalent_space_def) obtain f2 g2 where f2: "continuous_map Y U f2" and g2: "continuous_map U Y g2" and hom2: "homotopic_with (\x. True) Y Y (g2 \ f2) id" "homotopic_with (\x. True) U U (f2 \ g2) id" using 2 by (auto simp: homotopy_equivalent_space_def) have "homotopic_with (\f. True) X Y (g2 \ f2 \ f1) (id \ f1)" using f1 hom2(1) homotopic_compose_continuous_map_right by blast then have "homotopic_with (\f. True) X Y (g2 \ (f2 \ f1)) (id \ f1)" by (simp add: o_assoc) then have "homotopic_with (\x. True) X X (g1 \ (g2 \ (f2 \ f1))) (g1 \ (id \ f1))" by (simp add: g1 homotopic_compose_continuous_map_left) moreover have "homotopic_with (\x. True) X X (g1 \ id \ f1) id" using hom1 by simp ultimately have SS: "homotopic_with (\x. True) X X (g1 \ g2 \ (f2 \ f1)) id" apply (simp add: o_assoc) apply (blast intro: homotopic_with_trans) done have "homotopic_with (\f. True) U Y (f1 \ g1 \ g2) (id \ g2)" using g2 hom1(2) homotopic_with_compose_continuous_map_right by fastforce then have "homotopic_with (\f. True) U Y (f1 \ (g1 \ g2)) (id \ g2)" by (simp add: o_assoc) then have "homotopic_with (\x. True) U U (f2 \ (f1 \ (g1 \ g2))) (f2 \ (id \ g2))" by (simp add: f2 homotopic_with_compose_continuous_map_left) moreover have "homotopic_with (\x. True) U U (f2 \ id \ g2) id" using hom2 by simp ultimately have UU: "homotopic_with (\x. True) U U (f2 \ f1 \ (g1 \ g2)) id" apply (simp add: o_assoc) apply (blast intro: homotopic_with_trans) done show ?thesis unfolding homotopy_equivalent_space_def by (blast intro: f1 f2 g1 g2 continuous_map_compose SS UU) qed lemma deformation_retraction_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X (s \ r) id; retraction_maps X Y r s\ \ X homotopy_equivalent_space Y" unfolding homotopy_equivalent_space_def retraction_maps_def apply (rule_tac x=r in exI) apply (rule_tac x=s in exI) apply (auto simp: homotopic_with_equal continuous_map_compose) done lemma deformation_retract_imp_homotopy_equivalent_space: "\homotopic_with (\x. True) X X r id; retraction_maps X Y r id\ \ X homotopy_equivalent_space Y" using deformation_retraction_imp_homotopy_equivalent_space by force lemma deformation_retract_of_space: "S \ topspace X \ (\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` (topspace X) \ S)" proof (cases "S \ topspace X") case True moreover have "(\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id) \ (S retract_of_space X \ (\f. homotopic_with (\x. True) X X id f \ f ` topspace X \ S))" unfolding retract_of_space_def proof safe fix f r assume f: "homotopic_with (\x. True) X X id f" and fS: "f ` topspace X \ S" and r: "continuous_map X (subtopology X S) r" and req: "\x\S. r x = x" show "\r. homotopic_with (\x. True) X X id r \ retraction_maps X (subtopology X S) r id" proof (intro exI conjI) have "homotopic_with (\x. True) X X f r" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) X X (r \ f) (r \ id)" using homotopic_with_symD continuous_map_into_fulltopology f homotopic_compose_continuous_map_left r by blast show "f x = (r \ f) x" if "x \ topspace X" for x using that fS req by auto qed auto then show "homotopic_with (\x. True) X X id r" by (rule homotopic_with_trans [OF f]) next show "retraction_maps X (subtopology X S) r id" - by (simp add: r req retraction_maps_def topspace_subtopology) + by (simp add: r req retraction_maps_def) qed qed (use True in \auto simp: retraction_maps_def topspace_subtopology_subset continuous_map_in_subtopology\) ultimately show ?thesis by simp qed (auto simp: retract_of_space_def retraction_maps_def) subsection\Contractible spaces\ text\The definition (which agrees with "contractible" on subsets of Euclidean space) is a little cryptic because we don't in fact assume that the constant "a" is in the space. This forces the convention that the empty space / set is contractible, avoiding some special cases. \ definition contractible_space where "contractible_space X \ \a. homotopic_with (\x. True) X X id (\x. a)" lemma contractible_space_top_of_set [simp]:"contractible_space (top_of_set S) \ contractible S" by (auto simp: contractible_space_def contractible_def) lemma contractible_space_empty: "topspace X = {} \ contractible_space X" apply (simp add: contractible_space_def homotopic_with_def) apply (rule_tac x=undefined in exI) apply (rule_tac x="\(t,x). if t = 0 then x else undefined" in exI) apply (auto simp: continuous_map_on_empty) done lemma contractible_space_singleton: "topspace X = {a} \ contractible_space X" apply (simp add: contractible_space_def homotopic_with_def) apply (rule_tac x=a in exI) apply (rule_tac x="\(t,x). if t = 0 then x else a" in exI) apply (auto intro: continuous_map_eq [where f = "\z. a"]) done lemma contractible_space_subset_singleton: "topspace X \ {a} \ contractible_space X" by (meson contractible_space_empty contractible_space_singleton subset_singletonD) lemma contractible_space_subtopology_singleton: "contractible_space(subtopology X {a})" by (meson contractible_space_subset_singleton insert_subset path_connectedin_singleton path_connectedin_subtopology subsetI) lemma contractible_space: "contractible_space X \ topspace X = {} \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" proof (cases "topspace X = {}") case False then show ?thesis apply (auto simp: contractible_space_def) using homotopic_with_imp_continuous_maps by fastforce qed (simp add: contractible_space_empty) lemma contractible_imp_path_connected_space: assumes "contractible_space X" shows "path_connected_space X" proof (cases "topspace X = {}") case False have *: "path_connected_space X" if "a \ topspace X" and conth: "continuous_map (prod_topology (top_of_set {0..1}) X) X h" and h: "\x. h (0, x) = x" "\x. h (1, x) = a" for a and h :: "real \ 'a \ 'a" proof - have "path_component_of X b a" if "b \ topspace X" for b using that unfolding path_component_of_def apply (rule_tac x="h \ (\x. (x,b))" in exI) apply (simp add: h pathin_def) apply (rule continuous_map_compose [OF _ conth]) apply (auto simp: continuous_map_pairwise intro!: continuous_intros continuous_map_compose continuous_map_id [unfolded id_def]) done then show ?thesis by (metis path_component_of_equiv path_connected_space_iff_path_component) qed show ?thesis using assms False by (auto simp: contractible_space homotopic_with_def *) qed (simp add: path_connected_space_topspace_empty) lemma contractible_imp_connected_space: "contractible_space X \ connected_space X" by (simp add: contractible_imp_path_connected_space path_connected_imp_connected_space) lemma contractible_space_alt: "contractible_space X \ (\a \ topspace X. homotopic_with (\x. True) X X id (\x. a))" (is "?lhs = ?rhs") proof assume X: ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) show ?rhs proof show "homotopic_with (\x. True) X X id (\x. b)" if "b \ topspace X" for b apply (rule homotopic_with_trans [OF a]) using homotopic_constant_maps path_connected_space_imp_path_component_of by (metis (full_types) X a continuous_map_const contractible_imp_path_connected_space homotopic_with_imp_continuous_maps that) qed next assume R: ?rhs then show ?lhs apply (simp add: contractible_space_def) by (metis equals0I homotopic_on_emptyI) qed lemma compose_const [simp]: "f \ (\x. a) = (\x. f a)" "(\x. a) \ g = (\x. a)" by (simp_all add: o_def) lemma nullhomotopic_through_contractible_space: assumes f: "continuous_map X Y f" and g: "continuous_map Y Z g" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Z (g \ f) (\x. c)" proof - obtain b where b: "homotopic_with (\x. True) Y Y id (\x. b)" using Y by (auto simp: contractible_space_def) show thesis using homotopic_compose_continuous_map_right [OF homotopic_compose_continuous_map_left [OF b g] f] by (simp add: that) qed lemma nullhomotopic_into_contractible_space: assumes f: "continuous_map X Y f" and Y: "contractible_space Y" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF f _ Y] by (metis continuous_map_id id_comp) lemma nullhomotopic_from_contractible_space: assumes f: "continuous_map X Y f" and X: "contractible_space X" obtains c where "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_through_contractible_space [OF _ f X] by (metis comp_id continuous_map_id) lemma homotopy_dominated_contractibility: assumes f: "continuous_map X Y f" and g: "continuous_map Y X g" and hom: "homotopic_with (\x. True) Y Y (f \ g) id" and X: "contractible_space X" shows "contractible_space Y" proof - obtain c where c: "homotopic_with (\h. True) X Y f (\x. c)" using nullhomotopic_from_contractible_space [OF f X] . have "homotopic_with (\x. True) Y Y (f \ g) (\x. c)" using homotopic_compose_continuous_map_right [OF c g] by fastforce then have "homotopic_with (\x. True) Y Y id (\x. c)" using homotopic_with_trans [OF _ hom] homotopic_with_symD by blast then show ?thesis unfolding contractible_space_def .. qed lemma homotopy_equivalent_space_contractibility: "X homotopy_equivalent_space Y \ (contractible_space X \ contractible_space Y)" unfolding homotopy_equivalent_space_def by (blast intro: homotopy_dominated_contractibility) lemma homeomorphic_space_contractibility: "X homeomorphic_space Y \ (contractible_space X \ contractible_space Y)" by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility) lemma contractible_eq_homotopy_equivalent_singleton_subtopology: "contractible_space X \ topspace X = {} \ (\a \ topspace X. X homotopy_equivalent_space (subtopology X {a}))"(is "?lhs = ?rhs") proof (cases "topspace X = {}") case False show ?thesis proof assume ?lhs then obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" by (auto simp: contractible_space_def) then have "a \ topspace X" by (metis False continuous_map_const homotopic_with_imp_continuous_maps) then have "X homotopy_equivalent_space subtopology X {a}" unfolding homotopy_equivalent_space_def apply (rule_tac x="\x. a" in exI) apply (rule_tac x=id in exI) apply (auto simp: homotopic_with_sym topspace_subtopology_subset a) using connectedin_absolute connectedin_sing contractible_space_alt contractible_space_subtopology_singleton by fastforce with \a \ topspace X\ show ?rhs by blast next assume ?rhs then show ?lhs by (meson False contractible_space_subtopology_singleton homotopy_equivalent_space_contractibility) qed qed (simp add: contractible_space_empty) lemma contractible_space_retraction_map_image: assumes "retraction_map X Y f" and X: "contractible_space X" shows "contractible_space Y" proof - obtain g where f: "continuous_map X Y f" and g: "continuous_map Y X g" and fg: "\y \ topspace Y. f(g y) = y" using assms by (auto simp: retraction_map_def retraction_maps_def) obtain a where a: "homotopic_with (\x. True) X X id (\x. a)" using X by (auto simp: contractible_space_def) have "homotopic_with (\x. True) Y Y id (\x. f a)" proof (rule homotopic_with_eq) show "homotopic_with (\x. True) Y Y (f \ id \ g) (f \ (\x. a) \ g)" using f g a homotopic_compose_continuous_map_left homotopic_compose_continuous_map_right by metis qed (use fg in auto) then show ?thesis unfolding contractible_space_def by blast qed lemma contractible_space_prod_topology: "contractible_space(prod_topology X Y) \ topspace X = {} \ topspace Y = {} \ contractible_space X \ contractible_space Y" proof (cases "topspace X = {} \ topspace Y = {}") case True then have "topspace (prod_topology X Y) = {}" by simp then show ?thesis by (auto simp: contractible_space_empty) next case False have "contractible_space(prod_topology X Y) \ contractible_space X \ contractible_space Y" proof safe assume XY: "contractible_space (prod_topology X Y)" with False have "retraction_map (prod_topology X Y) X fst" by (auto simp: contractible_space False retraction_map_fst) then show "contractible_space X" by (rule contractible_space_retraction_map_image [OF _ XY]) have "retraction_map (prod_topology X Y) Y snd" using False XY by (auto simp: contractible_space False retraction_map_snd) then show "contractible_space Y" by (rule contractible_space_retraction_map_image [OF _ XY]) next assume "contractible_space X" and "contractible_space Y" with False obtain a b where "a \ topspace X" and a: "homotopic_with (\x. True) X X id (\x. a)" and "b \ topspace Y" and b: "homotopic_with (\x. True) Y Y id (\x. b)" by (auto simp: contractible_space) with False show "contractible_space (prod_topology X Y)" apply (simp add: contractible_space) apply (rule_tac x=a in bexI) apply (rule_tac x=b in bexI) using homotopic_with_prod_topology [OF a b] apply (metis (no_types, lifting) case_prod_Pair case_prod_beta' eq_id_iff) apply auto done qed with False show ?thesis by auto qed - - lemma contractible_space_product_topology: "contractible_space(product_topology X I) \ topspace (product_topology X I) = {} \ (\i \ I. contractible_space(X i))" proof (cases "topspace (product_topology X I) = {}") case False have 1: "contractible_space (X i)" if XI: "contractible_space (product_topology X I)" and "i \ I" for i proof (rule contractible_space_retraction_map_image [OF _ XI]) show "retraction_map (product_topology X I) (X i) (\x. x i)" using False by (simp add: retraction_map_product_projection \i \ I\) qed have 2: "contractible_space (product_topology X I)" if "x \ topspace (product_topology X I)" and cs: "\i\I. contractible_space (X i)" for x :: "'a \ 'b" proof - obtain f where f: "\i. i\I \ homotopic_with (\x. True) (X i) (X i) id (\x. f i)" using cs unfolding contractible_space_def by metis have "homotopic_with (\x. True) (product_topology X I) (product_topology X I) id (\x. restrict f I)" - by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto simp: topspace_product_topology) + by (rule homotopic_with_eq [OF homotopic_with_product_topology [OF f]]) (auto) then show ?thesis by (auto simp: contractible_space_def) qed show ?thesis using False 1 2 by blast qed (simp add: contractible_space_empty) lemma contractible_space_subtopology_euclideanreal [simp]: "contractible_space(subtopology euclideanreal S) \ is_interval S" (is "?lhs = ?rhs") proof assume ?lhs then have "path_connectedin (subtopology euclideanreal S) S" using contractible_imp_path_connected_space path_connectedin_topspace path_connectedin_absolute by (simp add: contractible_imp_path_connected) then show ?rhs by (simp add: is_interval_path_connected_1) next assume ?rhs then have "convex S" by (simp add: is_interval_convex_1) show ?lhs proof (cases "S = {}") case False then obtain z where "z \ S" by blast show ?thesis unfolding contractible_space_def homotopic_with_def proof (intro exI conjI allI) show "continuous_map (prod_topology (top_of_set {0..1}) (top_of_set S)) (top_of_set S) (\(t,x). (1 - t) * x + t * z)" apply (auto simp: case_prod_unfold) apply (intro continuous_intros) using \z \ S\ apply (auto intro: convexD [OF \convex S\, simplified]) done qed auto qed (simp add: contractible_space_empty) qed corollary contractible_space_euclideanreal: "contractible_space euclideanreal" proof - have "contractible_space (subtopology euclideanreal UNIV)" using contractible_space_subtopology_euclideanreal by blast then show ?thesis by simp qed abbreviation\<^marker>\tag important\ homotopy_eqv :: "'a::topological_space set \ 'b::topological_space set \ bool" (infix "homotopy'_eqv" 50) where "S homotopy_eqv T \ top_of_set S homotopy_equivalent_space top_of_set T" lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \ S homotopy_eqv T" unfolding homeomorphic_def homeomorphism_def homotopy_equivalent_space_def apply (erule ex_forward)+ apply auto apply (metis homotopic_with_id2 topspace_euclidean_subtopology) by (metis (full_types) homotopic_with_id2 imageE topspace_euclidean_subtopology) lemma homotopy_eqv_inj_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(f ` S) homotopy_eqv S" apply (rule homeomorphic_imp_homotopy_eqv) using assms homeomorphic_sym linear_homeomorphic_image by auto lemma homotopy_eqv_translation: fixes S :: "'a::real_normed_vector set" shows "(+) a ` S homotopy_eqv S" apply (rule homeomorphic_imp_homotopy_eqv) using homeomorphic_translation homeomorphic_sym by blast lemma homotopy_eqv_homotopic_triviality_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on U f" "f ` U \ T" and g: "continuous_on U g" "g ` U \ T" and homUS: "\f g. \continuous_on U f; f ` U \ S; continuous_on U g; g ` U \ S\ \ homotopic_with_canon (\x. True) U S f g" shows "homotopic_with_canon (\x. True) U T f g" proof - obtain h k where h: "continuous_on S h" "h ` S \ T" and k: "continuous_on T k" "k ` T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) have "homotopic_with_canon (\f. True) U S (k \ f) (k \ g)" apply (rule homUS) using f g k apply (safe intro!: continuous_on_compose h k f elim!: continuous_on_subset) apply (force simp: o_def)+ done then have "homotopic_with_canon (\x. True) U T (h \ (k \ f)) (h \ (k \ g))" apply (rule homotopic_with_compose_continuous_left) apply (simp_all add: h) done moreover have "homotopic_with_canon (\x. True) U T (h \ k \ f) (id \ f)" apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) apply (auto simp: hom f) done moreover have "homotopic_with_canon (\x. True) U T (h \ k \ g) (id \ g)" apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) apply (auto simp: hom g) done ultimately show "homotopic_with_canon (\x. True) U T f g" apply (simp add: o_assoc) using homotopic_with_trans homotopic_with_sym by blast qed lemma homotopy_eqv_homotopic_triviality: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f g. continuous_on U f \ f ` U \ S \ continuous_on U g \ g ` U \ S \ homotopic_with_canon (\x. True) U S f g) \ (\f g. continuous_on U f \ f ` U \ T \ continuous_on U g \ g ` U \ T \ homotopic_with_canon (\x. True) U T f g)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis assms homotopy_eqv_homotopic_triviality_imp) next assume ?rhs moreover have "T homotopy_eqv S" using assms homotopy_equivalent_space_sym by blast ultimately show ?lhs by (blast intro: homotopy_eqv_homotopic_triviality_imp) qed lemma homotopy_eqv_cohomotopic_triviality_null_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on T f" "f ` T \ U" and homSU: "\f. \continuous_on S f; f ` S \ U\ \ \c. homotopic_with_canon (\x. True) S U f (\x. c)" obtains c where "homotopic_with_canon (\x. True) T U f (\x. c)" proof - obtain h k where h: "continuous_on S h" "h ` S \ T" and k: "continuous_on T k" "k ` T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) obtain c where "homotopic_with_canon (\x. True) S U (f \ h) (\x. c)" apply (rule exE [OF homSU [of "f \ h"]]) apply (intro continuous_on_compose h) using h f apply (force elim!: continuous_on_subset)+ done then have "homotopic_with_canon (\x. True) T U ((f \ h) \ k) ((\x. c) \ k)" apply (rule homotopic_with_compose_continuous_right [where X=S]) using k by auto moreover have "homotopic_with_canon (\x. True) T U (f \ id) (f \ (h \ k))" apply (rule homotopic_with_compose_continuous_left [where Y=T]) apply (simp add: hom homotopic_with_symD) using f apply auto done ultimately show ?thesis apply (rule_tac c=c in that) apply (simp add: o_def) using homotopic_with_trans by blast qed lemma homotopy_eqv_cohomotopic_triviality_null: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f. continuous_on S f \ f ` S \ U \ (\c. homotopic_with_canon (\x. True) S U f (\x. c))) \ (\f. continuous_on T f \ f ` T \ U \ (\c. homotopic_with_canon (\x. True) T U f (\x. c)))" apply (rule iffI) apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp) by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_equivalent_space_sym) lemma homotopy_eqv_homotopic_triviality_null_imp: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" and f: "continuous_on U f" "f ` U \ T" and homSU: "\f. \continuous_on U f; f ` U \ S\ \ \c. homotopic_with_canon (\x. True) U S f (\x. c)" shows "\c. homotopic_with_canon (\x. True) U T f (\x. c)" proof - obtain h k where h: "continuous_on S h" "h ` S \ T" and k: "continuous_on T k" "k ` T \ S" and hom: "homotopic_with_canon (\x. True) S S (k \ h) id" "homotopic_with_canon (\x. True) T T (h \ k) id" using assms by (auto simp: homotopy_equivalent_space_def) obtain c::'a where "homotopic_with_canon (\x. True) U S (k \ f) (\x. c)" apply (rule exE [OF homSU [of "k \ f"]]) apply (intro continuous_on_compose h) using k f apply (force elim!: continuous_on_subset)+ done then have "homotopic_with_canon (\x. True) U T (h \ (k \ f)) (h \ (\x. c))" apply (rule homotopic_with_compose_continuous_left [where Y=S]) using h by auto moreover have "homotopic_with_canon (\x. True) U T (id \ f) ((h \ k) \ f)" apply (rule homotopic_with_compose_continuous_right [where X=T]) apply (simp add: hom homotopic_with_symD) using f apply auto done ultimately show ?thesis using homotopic_with_trans by (fastforce simp add: o_def) qed lemma homotopy_eqv_homotopic_triviality_null: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" and U :: "'c::real_normed_vector set" assumes "S homotopy_eqv T" shows "(\f. continuous_on U f \ f ` U \ S \ (\c. homotopic_with_canon (\x. True) U S f (\x. c))) \ (\f. continuous_on U f \ f ` U \ T \ (\c. homotopic_with_canon (\x. True) U T f (\x. c)))" apply (rule iffI) apply (metis assms homotopy_eqv_homotopic_triviality_null_imp) by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_equivalent_space_sym) lemma homotopy_eqv_contractible_sets: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "contractible S" "contractible T" "S = {} \ T = {}" shows "S homotopy_eqv T" proof (cases "S = {}") case True with assms show ?thesis by (simp add: homeomorphic_imp_homotopy_eqv) next case False with assms obtain a b where "a \ S" "b \ T" by auto then show ?thesis unfolding homotopy_equivalent_space_def apply (rule_tac x="\x. b" in exI) apply (rule_tac x="\x. a" in exI) apply (intro assms conjI continuous_on_id' homotopic_into_contractible) - apply (auto simp: o_def continuous_on_const) + apply (auto simp: o_def) done qed lemma homotopy_eqv_empty1 [simp]: fixes S :: "'a::real_normed_vector set" shows "S homotopy_eqv ({}::'b::real_normed_vector set) \ S = {}" apply (rule iffI) apply (metis Abstract_Topology.continuous_map_subtopology_eu emptyE equals0I homotopy_equivalent_space_def image_subset_iff) by (simp add: homotopy_eqv_contractible_sets) lemma homotopy_eqv_empty2 [simp]: fixes S :: "'a::real_normed_vector set" shows "({}::'b::real_normed_vector set) homotopy_eqv S \ S = {}" using homotopy_equivalent_space_sym homotopy_eqv_empty1 by blast lemma homotopy_eqv_contractibility: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homotopy_eqv T \ (contractible S \ contractible T)" by (meson contractible_space_top_of_set homotopy_equivalent_space_contractibility) lemma homotopy_eqv_sing: fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" shows "S homotopy_eqv {a} \ S \ {} \ contractible S" proof (cases "S = {}") case True then show ?thesis by simp next case False then show ?thesis by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets) qed lemma homeomorphic_contractible_eq: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "S homeomorphic T \ (contractible S \ contractible T)" by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) lemma homeomorphic_contractible: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" shows "\contractible S; S homeomorphic T\ \ contractible T" by (metis homeomorphic_contractible_eq) subsection\<^marker>\tag unimportant\\Misc other results\ lemma bounded_connected_Compl_real: fixes S :: "real set" assumes "bounded S" and conn: "connected(- S)" shows "S = {}" proof - obtain a b where "S \ box a b" by (meson assms bounded_subset_box_symmetric) then have "a \ S" "b \ S" by auto then have "\x. a \ x \ x \ b \ x \ - S" by (meson Compl_iff conn connected_iff_interval) then show ?thesis using \S \ box a b\ by auto qed corollary bounded_path_connected_Compl_real: fixes S :: "real set" assumes "bounded S" "path_connected(- S)" shows "S = {}" by (simp add: assms bounded_connected_Compl_real path_connected_imp_connected) lemma bounded_connected_Compl_1: fixes S :: "'a::{euclidean_space} set" assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" shows "S = {}" proof - have "DIM('a) = DIM(real)" by (simp add: "1") then obtain f::"'a \ real" and g where "linear f" "\x. norm(f x) = norm x" "\x. g(f x) = x" "\y. f(g y) = y" by (rule isomorphisms_UNIV_UNIV) blast with \bounded S\ have "bounded (f ` S)" using bounded_linear_image linear_linear by blast have "connected (f ` (-S))" using connected_linear_image assms \linear f\ by blast moreover have "f ` (-S) = - (f ` S)" apply (rule bij_image_Compl_eq) apply (auto simp: bij_def) apply (metis \\x. g (f x) = x\ injI) by (metis UNIV_I \\y. f (g y) = y\ image_iff) finally have "connected (- (f ` S))" by simp then have "f ` S = {}" using \bounded (f ` S)\ bounded_connected_Compl_real by blast then show ?thesis by blast qed subsection\<^marker>\tag unimportant\\Some Uncountable Sets\ lemma uncountable_closed_segment: fixes a :: "'a::real_normed_vector" assumes "a \ b" shows "uncountable (closed_segment a b)" unfolding path_image_linepath [symmetric] path_image_def using inj_on_linepath [OF assms] uncountable_closed_interval [of 0 1] countable_image_inj_on by auto lemma uncountable_open_segment: fixes a :: "'a::real_normed_vector" assumes "a \ b" shows "uncountable (open_segment a b)" by (simp add: assms open_segment_def uncountable_closed_segment uncountable_minus_countable) lemma uncountable_convex: fixes a :: "'a::real_normed_vector" assumes "convex S" "a \ S" "b \ S" "a \ b" shows "uncountable S" proof - have "uncountable (closed_segment a b)" by (simp add: uncountable_closed_segment assms) then show ?thesis by (meson assms convex_contains_segment countable_subset) qed lemma uncountable_ball: fixes a :: "'a::euclidean_space" assumes "r > 0" shows "uncountable (ball a r)" proof - have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)))" by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment) moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \ Basis)) \ ball a r" using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis) ultimately show ?thesis by (metis countable_subset) qed lemma ball_minus_countable_nonempty: assumes "countable (A :: 'a :: euclidean_space set)" "r > 0" shows "ball z r - A \ {}" proof assume *: "ball z r - A = {}" have "uncountable (ball z r - A)" by (intro uncountable_minus_countable assms uncountable_ball) thus False by (subst (asm) *) auto qed lemma uncountable_cball: fixes a :: "'a::euclidean_space" assumes "r > 0" shows "uncountable (cball a r)" using assms countable_subset uncountable_ball by auto lemma pairwise_disjnt_countable: fixes \ :: "nat set set" assumes "pairwise disjnt \" shows "countable \" proof - have "inj_on (\X. SOME n. n \ X) (\ - {{}})" apply (clarsimp simp add: inj_on_def) by (metis assms disjnt_insert2 insert_absorb pairwise_def subsetI subset_empty tfl_some) then show ?thesis by (metis countable_Diff_eq countable_def) qed lemma pairwise_disjnt_countable_Union: assumes "countable (\\)" and pwd: "pairwise disjnt \" shows "countable \" proof - obtain f :: "_ \ nat" where f: "inj_on f (\\)" using assms by blast then have "pairwise disjnt (\ X \ \. {f ` X})" using assms by (force simp: pairwise_def disjnt_inj_on_iff [OF f]) then have "countable (\ X \ \. {f ` X})" using pairwise_disjnt_countable by blast then show ?thesis by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable) qed lemma connected_uncountable: fixes S :: "'a::metric_space set" assumes "connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" proof - have "continuous_on S (dist a)" by (intro continuous_intros) then have "connected (dist a ` S)" by (metis connected_continuous_image \connected S\) then have "closed_segment 0 (dist a b) \ (dist a ` S)" by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex) then have "uncountable (dist a ` S)" by (metis \a \ b\ countable_subset dist_eq_0_iff uncountable_closed_segment) then show ?thesis by blast qed lemma path_connected_uncountable: fixes S :: "'a::metric_space set" assumes "path_connected S" "a \ S" "b \ S" "a \ b" shows "uncountable S" using path_connected_imp_connected assms connected_uncountable by metis lemma connected_finite_iff_sing: fixes S :: "'a::metric_space set" assumes "connected S" shows "finite S \ S = {} \ (\a. S = {a})" (is "_ = ?rhs") proof - have "uncountable S" if "\ ?rhs" using connected_uncountable assms that by blast then show ?thesis using uncountable_infinite by auto qed lemma connected_card_eq_iff_nontrivial: fixes S :: "'a::metric_space set" shows "connected S \ uncountable S \ \(\a. S \ {a})" apply (auto simp: countable_finite finite_subset) by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff) lemma simple_path_image_uncountable: fixes g :: "real \ 'a::metric_space" assumes "simple_path g" shows "uncountable (path_image g)" proof - have "g 0 \ path_image g" "g (1/2) \ path_image g" by (simp_all add: path_defs) moreover have "g 0 \ g (1/2)" using assms by (fastforce simp add: simple_path_def) ultimately show ?thesis apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image) by blast qed lemma arc_image_uncountable: fixes g :: "real \ 'a::metric_space" assumes "arc g" shows "uncountable (path_image g)" by (simp add: arc_imp_simple_path assms simple_path_image_uncountable) subsection\<^marker>\tag unimportant\\ Some simple positive connection theorems\ proposition path_connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "path_connected(U - S)" proof (clarsimp simp add: path_connected_def) fix a b assume "a \ U" "a \ S" "b \ U" "b \ S" let ?m = "midpoint a b" show "\g. path g \ path_image g \ U - S \ pathstart g = a \ pathfinish g = b" proof (cases "a = b") case True then show ?thesis by (metis DiffI \a \ U\ \a \ S\ path_component_def path_component_refl) next case False then have "a \ ?m" "b \ ?m" using midpoint_eq_endpoint by fastforce+ have "?m \ U" using \a \ U\ \b \ U\ \convex U\ convex_contains_segment by force obtain c where "c \ U" and nc_abc: "\ collinear {a,b,c}" by (metis False \a \ U\ \b \ U\ \\ collinear U\ collinear_triples insert_absorb) have ncoll_mca: "\ collinear {?m,c,a}" by (metis (full_types) \a \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have ncoll_mcb: "\ collinear {?m,c,b}" by (metis (full_types) \b \ ?m\ collinear_3_trans collinear_midpoint insert_commute nc_abc) have "c \ ?m" by (metis collinear_midpoint insert_commute nc_abc) then have "closed_segment ?m c \ U" by (simp add: \c \ U\ \?m \ U\ \convex U\ closed_segment_subset) then obtain z where z: "z \ closed_segment ?m c" and disjS: "(closed_segment a z \ closed_segment z b) \ S = {}" proof - have False if "closed_segment ?m c \ {z. (closed_segment a z \ closed_segment z b) \ S \ {}}" proof - have closb: "closed_segment ?m c \ {z \ closed_segment ?m c. closed_segment a z \ S \ {}} \ {z \ closed_segment ?m c. closed_segment z b \ S \ {}}" using that by blast have *: "countable {z \ closed_segment ?m c. closed_segment z u \ S \ {}}" if "u \ U" "u \ S" and ncoll: "\ collinear {?m, c, u}" for u proof - have **: False if x1: "x1 \ closed_segment ?m c" and x2: "x2 \ closed_segment ?m c" and "x1 \ x2" "x1 \ u" and w: "w \ closed_segment x1 u" "w \ closed_segment x2 u" and "w \ S" for x1 x2 w proof - have "x1 \ affine hull {?m,c}" "x2 \ affine hull {?m,c}" using segment_as_ball x1 x2 by auto then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) have "\ collinear {x1, u, x2}" proof assume "collinear {x1, u, x2}" then have "collinear {?m, c, u}" by (metis (full_types) \c \ ?m\ coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \x1 \ x2\) with ncoll show False .. qed then have "closed_segment x1 u \ closed_segment u x2 = {u}" by (blast intro!: Int_closed_segment) then have "w = u" using closed_segment_commute w by auto show ?thesis using \u \ S\ \w = u\ that(7) by auto qed then have disj: "disjoint ((\z\closed_segment ?m c. {closed_segment z u \ S}))" by (fastforce simp: pairwise_def disjnt_def) have cou: "countable ((\z \ closed_segment ?m c. {closed_segment z u \ S}) - {{}})" apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) apply (rule countable_subset [OF _ \countable S\], auto) done define f where "f \ \X. (THE z. z \ closed_segment ?m c \ X = closed_segment z u \ S)" show ?thesis proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) fix x assume x: "x \ closed_segment ?m c" "closed_segment x u \ S \ {}" show "x \ f ` ((\z\closed_segment ?m c. {closed_segment z u \ S}) - {{}})" proof (rule_tac x="closed_segment x u \ S" in image_eqI) show "x = f (closed_segment x u \ S)" unfolding f_def apply (rule the_equality [symmetric]) using x apply (auto simp: dest: **) done qed (use x in auto) qed qed have "uncountable (closed_segment ?m c)" by (metis \c \ ?m\ uncountable_closed_segment) then show False using closb * [OF \a \ U\ \a \ S\ ncoll_mca] * [OF \b \ U\ \b \ S\ ncoll_mcb] apply (simp add: closed_segment_commute) by (simp add: countable_subset) qed then show ?thesis by (force intro: that) qed show ?thesis proof (intro exI conjI) have "path_image (linepath a z +++ linepath z b) \ U" by (metis \a \ U\ \b \ U\ \closed_segment ?m c \ U\ z \convex U\ closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) with disjS show "path_image (linepath a z +++ linepath z b) \ U - S" by (force simp: path_image_join) qed auto qed qed corollary connected_convex_diff_countable: fixes U :: "'a::euclidean_space set" assumes "convex U" "\ collinear U" "countable S" shows "connected(U - S)" by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) lemma path_connected_punctured_convex: assumes "convex S" and aff: "aff_dim S \ 1" shows "path_connected(S - {a})" proof - consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \ 2" using assms aff_dim_geq [of S] by linarith then show ?thesis proof cases assume "aff_dim S = -1" then show ?thesis by (metis aff_dim_empty empty_Diff path_connected_empty) next assume "aff_dim S = 0" then show ?thesis by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) next assume ge2: "aff_dim S \ 2" then have "\ collinear S" proof (clarsimp simp add: collinear_affine_hull) fix u v assume "S \ affine hull {u, v}" then have "aff_dim S \ aff_dim {u, v}" by (metis (no_types) aff_dim_affine_hull aff_dim_subset) with ge2 show False by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) qed then show ?thesis apply (rule path_connected_convex_diff_countable [OF \convex S\]) by simp qed qed lemma connected_punctured_convex: shows "\convex S; aff_dim S \ 1\ \ connected(S - {a})" using path_connected_imp_connected path_connected_punctured_convex by blast lemma path_connected_complement_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "countable S" shows "path_connected(- S)" proof - have "path_connected(UNIV - S)" apply (rule path_connected_convex_diff_countable) using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) then show ?thesis by (simp add: Compl_eq_Diff_UNIV) qed proposition path_connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "path_connected(S - T)" proof (clarsimp simp add: path_connected_component) fix x y assume xy: "x \ S" "x \ T" "y \ S" "y \ T" show "path_component (S - T) x y" proof (rule connected_equivalence_relation_gen [OF \connected S\, where P = "\x. x \ T"]) show "\z. z \ U \ z \ T" if opeU: "openin (top_of_set S) U" and "x \ U" for U x proof - have "openin (top_of_set (affine hull S)) U" using opeU ope openin_trans by blast with \x \ U\ obtain r where Usub: "U \ affine hull S" and "r > 0" and subU: "ball x r \ affine hull S \ U" by (auto simp: openin_contains_ball) with \x \ U\ have x: "x \ ball x r \ affine hull S" by auto have "\ S \ {x}" using \\ collinear S\ collinear_subset by blast then obtain x' where "x' \ x" "x' \ S" by blast obtain y where y: "y \ x" "y \ ball x r \ affine hull S" proof show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \ x" using \x' \ x\ \r > 0\ by auto show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \ ball x r \ affine hull S" using \x' \ x\ \r > 0\ \x' \ S\ x by (simp add: dist_norm mem_affine_3_minus hull_inc) qed have "convex (ball x r \ affine hull S)" by (simp add: affine_imp_convex convex_Int) with x y subU have "uncountable U" by (meson countable_subset uncountable_convex) then have "\ U \ T" using \countable T\ countable_subset by blast then show ?thesis by blast qed show "\U. openin (top_of_set S) U \ x \ U \ (\x\U. \y\U. x \ T \ y \ T \ path_component (S - T) x y)" if "x \ S" for x proof - obtain r where Ssub: "S \ affine hull S" and "r > 0" and subS: "ball x r \ affine hull S \ S" using ope \x \ S\ by (auto simp: openin_contains_ball) then have conv: "convex (ball x r \ affine hull S)" by (simp add: affine_imp_convex convex_Int) have "\ aff_dim (affine hull S) \ 1" using \\ collinear S\ collinear_aff_dim by auto then have "\ collinear (ball x r \ affine hull S)" apply (simp add: collinear_aff_dim) by (metis (no_types, hide_lams) aff_dim_convex_Int_open IntI open_ball \0 < r\ aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) then have *: "path_connected ((ball x r \ affine hull S) - T)" by (rule path_connected_convex_diff_countable [OF conv _ \countable T\]) have ST: "ball x r \ affine hull S - T \ S - T" using subS by auto show ?thesis proof (intro exI conjI) show "x \ ball x r \ affine hull S" using \x \ S\ \r > 0\ by (simp add: hull_inc) have "openin (top_of_set (affine hull S)) (ball x r \ affine hull S)" by (subst inf.commute) (simp add: openin_Int_open) then show "openin (top_of_set S) (ball x r \ affine hull S)" by (rule openin_subset_trans [OF _ subS Ssub]) qed (use * path_component_trans in \auto simp: path_connected_component path_component_of_subset [OF ST]\) qed qed (use xy path_component_trans in auto) qed corollary connected_openin_diff_countable: fixes S :: "'a::euclidean_space set" assumes "connected S" and ope: "openin (top_of_set (affine hull S)) S" and "\ collinear S" "countable T" shows "connected(S - T)" by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) corollary path_connected_open_diff_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "open S" "connected S" "countable T" shows "path_connected(S - T)" proof (cases "S = {}") case True then show ?thesis - by (simp add: path_connected_empty) + by (simp) next case False show ?thesis proof (rule path_connected_openin_diff_countable) show "openin (top_of_set (affine hull S)) S" by (simp add: assms hull_subset open_subset) show "\ collinear S" using assms False by (simp add: collinear_aff_dim aff_dim_open) qed (simp_all add: assms) qed corollary connected_open_diff_countable: fixes S :: "'a::euclidean_space set" assumes "2 \ DIM('a)" "open S" "connected S" "countable T" shows "connected(S - T)" by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) subsection\<^marker>\tag unimportant\ \Self-homeomorphisms shuffling points about\ subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_moving_points_exists\\ lemma homeomorphism_moving_point_1: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" "f a = u" "\x. x \ sphere a r \ f x = x" proof - have nou: "norm (u - a) < r" and "u \ T" using u by (auto simp: dist_norm norm_minus_commute) then have "0 < r" by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) define f where "f \ \x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a proof - have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" using eq by (simp add: algebra_simps) then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" by (metis diff_divide_distrib) also have "\ \ norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" using norm_triangle_ineq by blast also have "\ = norm y + (norm x - norm y) * (norm u / r)" using yx \r > 0\ by (simp add: field_split_simps) also have "\ < norm y + (norm x - norm y) * 1" apply (subst add_less_cancel_left) apply (rule mult_strict_left_mono) using nou \0 < r\ yx apply (simp_all add: field_simps) done also have "\ = norm x" by simp finally show False by simp qed have "inj f" unfolding f_def proof (clarsimp simp: inj_on_def) fix x y assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" by (auto simp: algebra_simps) show "x=y" proof (cases "norm (x - a) = norm (y - a)") case True then show ?thesis using eq by auto next case False then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" by linarith then have "False" proof cases case 1 show False using * [OF _ nou 1] eq by simp next case 2 with * [OF eq nou] show False by auto qed then show "x=y" .. qed qed then have inj_onf: "inj_on f (cball a r \ T)" using inj_on_Int by fastforce have contf: "continuous_on (cball a r \ T) f" unfolding f_def using \0 < r\ by (intro continuous_intros) blast have fim: "f ` (cball a r \ T) = cball a r \ T" proof have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \ r" if "norm y \ r" "norm u < r" for y u::'a proof - have "norm (y + (1 - norm y / r) *\<^sub>R u) \ norm y + norm((1 - norm y / r) *\<^sub>R u)" using norm_triangle_ineq by blast also have "\ = norm y + abs(1 - norm y / r) * norm u" by simp also have "\ \ r" proof - have "(r - norm u) * (r - norm y) \ 0" using that by auto then have "r * norm u + r * norm y \ r * r + norm u * norm y" by (simp add: algebra_simps) then show ?thesis using that \0 < r\ by (simp add: abs_if field_simps) qed finally show ?thesis . qed have "f ` (cball a r) \ cball a r" apply (clarsimp simp add: dist_norm norm_minus_commute f_def) using * by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute nou) moreover have "f ` T \ T" unfolding f_def using \affine T\ \a \ T\ \u \ T\ by (force simp: add.commute mem_affine_3_minus) ultimately show "f ` (cball a r \ T) \ cball a r \ T" by blast next show "cball a r \ T \ f ` (cball a r \ T)" proof (clarsimp simp add: dist_norm norm_minus_commute) fix x assume x: "norm (x - a) \ r" and "x \ T" have "\v \ {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \ 1 = 0" by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) then obtain v where "0 \ v" "v \ 1" and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" by auto then have n: "norm (a - (x - v *\<^sub>R (u - a))) = r - r * v" by (simp add: field_simps norm_minus_commute) show "x \ f ` (cball a r \ T)" proof (rule image_eqI) show "x = f (x - v *\<^sub>R (u - a))" using \r > 0\ v by (simp add: f_def) (simp add: field_simps) have "x - v *\<^sub>R (u - a) \ cball a r" using \r > 0\\0 \ v\ by (simp add: dist_norm n) moreover have "x - v *\<^sub>R (u - a) \ T" - by (simp add: f_def \affine T\ \u \ T\ \x \ T\ assms mem_affine_3_minus2) + by (simp add: f_def \u \ T\ \x \ T\ assms mem_affine_3_minus2) ultimately show "x - v *\<^sub>R (u - a) \ cball a r \ T" by blast qed qed qed have "\g. homeomorphism (cball a r \ T) (cball a r \ T) f g" apply (rule homeomorphism_compact [OF _ contf fim inj_onf]) apply (simp add: affine_closed compact_Int_closed \affine T\) done then show ?thesis apply (rule exE) apply (erule_tac f=f in that) using \r > 0\ apply (simp_all add: f_def dist_norm norm_minus_commute) done qed corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_2: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism (cball a r \ T) (cball a r \ T) f g" "f u = v" "\x. \x \ sphere a r; x \ T\ \ f x = x" proof - have "0 < r" by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) obtain f1 g1 where hom1: "homeomorphism (cball a r \ T) (cball a r \ T) f1 g1" and "f1 a = u" and f1: "\x. x \ sphere a r \ f1 x = x" using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ u] by blast obtain f2 g2 where hom2: "homeomorphism (cball a r \ T) (cball a r \ T) f2 g2" and "f2 a = v" and f2: "\x. x \ sphere a r \ f2 x = x" using homeomorphism_moving_point_1 [OF \affine T\ \a \ T\ v] by blast show ?thesis proof show "homeomorphism (cball a r \ T) (cball a r \ T) (f2 \ g1) (f1 \ g2)" by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) have "g1 u = a" using \0 < r\ \f1 a = u\ assms hom1 homeomorphism_apply1 by fastforce then show "(f2 \ g1) u = v" by (simp add: \f2 a = v\) show "\x. \x \ sphere a r; x \ T\ \ (f2 \ g1) x = x" using f1 f2 hom1 homeomorphism_apply1 by fastforce qed qed corollary\<^marker>\tag unimportant\ homeomorphism_moving_point_3: fixes a :: "'a::euclidean_space" assumes "affine T" "a \ T" and ST: "ball a r \ T \ S" "S \ T" and u: "u \ ball a r \ T" and v: "v \ ball a r \ T" obtains f g where "homeomorphism S S f g" "f u = v" "{x. \ (f x = x \ g x = x)} \ ball a r \ T" proof - obtain f g where hom: "homeomorphism (cball a r \ T) (cball a r \ T) f g" and "f u = v" and fid: "\x. \x \ sphere a r; x \ T\ \ f x = x" using homeomorphism_moving_point_2 [OF \affine T\ \a \ T\ u v] by blast have gid: "\x. \x \ sphere a r; x \ T\ \ g x = x" using fid hom homeomorphism_apply1 by fastforce define ff where "ff \ \x. if x \ ball a r \ T then f x else x" define gg where "gg \ \x. if x \ ball a r \ T then g x else x" show ?thesis proof show "homeomorphism S S ff gg" proof (rule homeomorphismI) have "continuous_on ((cball a r \ T) \ (T - ball a r)) ff" apply (simp add: ff_def) apply (rule continuous_on_cases) using homeomorphism_cont1 [OF hom] - apply (auto simp: affine_closed \affine T\ continuous_on_id fid) + apply (auto simp: affine_closed \affine T\ fid) done then show "continuous_on S ff" apply (rule continuous_on_subset) using ST by auto have "continuous_on ((cball a r \ T) \ (T - ball a r)) gg" apply (simp add: gg_def) apply (rule continuous_on_cases) using homeomorphism_cont2 [OF hom] - apply (auto simp: affine_closed \affine T\ continuous_on_id gid) + apply (auto simp: affine_closed \affine T\ gid) done then show "continuous_on S gg" apply (rule continuous_on_subset) using ST by auto show "ff ` S \ S" proof (clarsimp simp add: ff_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "f x \ cball a r \ T" using homeomorphism_image1 [OF hom] by force then show "f x \ S" using ST(1) \x \ T\ gid hom homeomorphism_def x by fastforce qed show "gg ` S \ S" proof (clarsimp simp add: gg_def) fix x assume "x \ S" and x: "dist a x < r" and "x \ T" then have "g x \ cball a r \ T" using homeomorphism_image2 [OF hom] by force then have "g x \ ball a r" using homeomorphism_apply2 [OF hom] by (metis Diff_Diff_Int Diff_iff \x \ T\ cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) then show "g x \ S" using ST(1) \g x \ cball a r \ T\ by force qed show "\x. x \ S \ gg (ff x) = x" apply (simp add: ff_def gg_def) using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] apply auto apply (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) done show "\x. x \ S \ ff (gg x) = x" apply (simp add: ff_def gg_def) using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] apply auto apply (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) done qed show "ff u = v" using u by (auto simp: ff_def \f u = v\) show "{x. \ (ff x = x \ gg x = x)} \ ball a r \ T" by (auto simp: ff_def gg_def) qed qed proposition\<^marker>\tag unimportant\ homeomorphism_moving_point: fixes a :: "'a::euclidean_space" assumes ope: "openin (top_of_set (affine hull S)) S" and "S \ T" and TS: "T \ affine hull S" and S: "connected S" "a \ S" "b \ S" obtains f g where "homeomorphism T T f g" "f a = b" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" proof - have 1: "\h k. homeomorphism T T h k \ h (f d) = d \ {x. \ (h x = x \ k x = x)} \ S \ bounded {x. \ (h x = x \ k x = x)}" if "d \ S" "f d \ S" and homfg: "homeomorphism T T f g" and S: "{x. \ (f x = x \ g x = x)} \ S" and bo: "bounded {x. \ (f x = x \ g x = x)}" for d f g proof (intro exI conjI) show homgf: "homeomorphism T T g f" by (metis homeomorphism_symD homfg) then show "g (f d) = d" by (meson \S \ T\ homeomorphism_def subsetD \d \ S\) show "{x. \ (g x = x \ f x = x)} \ S" using S by blast show "bounded {x. \ (g x = x \ f x = x)}" using bo by (simp add: conj_commute) qed have 2: "\f g. homeomorphism T T f g \ f x = f2 (f1 x) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" if "x \ S" "f1 x \ S" "f2 (f1 x) \ S" and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" and sub: "{x. \ (f1 x = x \ g1 x = x)} \ S" "{x. \ (f2 x = x \ g2 x = x)} \ S" and bo: "bounded {x. \ (f1 x = x \ g1 x = x)}" "bounded {x. \ (f2 x = x \ g2 x = x)}" for x f1 f2 g1 g2 proof (intro exI conjI) show homgf: "homeomorphism T T (f2 \ f1) (g1 \ g2)" by (metis homeomorphism_compose hom) then show "(f2 \ f1) x = f2 (f1 x)" by force show "{x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)} \ S" using sub by force have "bounded ({x. \(f1 x = x \ g1 x = x)} \ {x. \(f2 x = x \ g2 x = x)})" using bo by simp then show "bounded {x. \ ((f2 \ f1) x = x \ (g1 \ g2) x = x)}" by (rule bounded_subset) auto qed have 3: "\U. openin (top_of_set S) U \ d \ U \ (\x\U. \f g. homeomorphism T T f g \ f d = x \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)})" if "d \ S" for d proof - obtain r where "r > 0" and r: "ball d r \ affine hull S \ S" by (metis \d \ S\ ope openin_contains_ball) have *: "\f g. homeomorphism T T f g \ f d = e \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" if "e \ S" "e \ ball d r" for e apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) using r \S \ T\ TS that apply (auto simp: \d \ S\ \0 < r\ hull_inc) using bounded_subset by blast show ?thesis apply (rule_tac x="S \ ball d r" in exI) apply (intro conjI) apply (simp add: openin_open_Int) apply (simp add: \0 < r\ that) apply (blast intro: *) done qed have "\f g. homeomorphism T T f g \ f a = b \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" apply (rule connected_equivalence_relation [OF S], safe) apply (blast intro: 1 2 3)+ done then show ?thesis using that by auto qed lemma homeomorphism_moving_points_exists_gen: assumes K: "finite K" "\i. i \ K \ x i \ S \ y i \ S" "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and "2 \ aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" and "S \ T" "T \ affine hull S" "connected S" shows "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using assms proof (induction K) case empty then show ?case by (force simp: homeomorphism_ident) next case (insert i K) then have xney: "\j. \j \ K; j \ i\ \ x i \ x j \ y i \ y j" and pw: "pairwise (\i j. x i \ x j \ y i \ y j) K" and "x i \ S" "y i \ S" and xyS: "\i. i \ K \ x i \ S \ y i \ S" by (simp_all add: pairwise_insert) obtain f g where homfg: "homeomorphism T T f g" and feq: "\i. i \ K \ f(x i) = y i" and fg_sub: "{x. \ (f x = x \ g x = x)} \ S" and bo_fg: "bounded {x. \ (f x = x \ g x = x)}" using insert.IH [OF xyS pw] insert.prems by (blast intro: that) then have "\f g. homeomorphism T T f g \ (\i \ K. f(x i) = y i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" using insert by blast have aff_eq: "affine hull (S - y ` K) = affine hull S" apply (rule affine_hull_Diff) apply (auto simp: insert) using \y i \ S\ insert.hyps(2) xney xyS by fastforce have f_in_S: "f x \ S" if "x \ S" for x using homfg fg_sub homeomorphism_apply1 \S \ T\ proof - have "(f (f x) \ f x \ g (f x) \ f x) \ f x \ S" by (metis \S \ T\ homfg subsetD homeomorphism_apply1 that) then show ?thesis using fg_sub by force qed obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" and hk_sub: "{x. \ (h x = x \ k x = x)} \ S - y ` K" and bo_hk: "bounded {x. \ (h x = x \ k x = x)}" proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) show "openin (top_of_set (affine hull (S - y ` K))) (S - y ` K)" by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) show "S - y ` K \ T" using \S \ T\ by auto show "T \ affine hull (S - y ` K)" using insert by (simp add: aff_eq) show "connected (S - y ` K)" proof (rule connected_openin_diff_countable [OF \connected S\ ope]) show "\ collinear S" using collinear_aff_dim \2 \ aff_dim S\ by force show "countable (y ` K)" using countable_finite insert.hyps(1) by blast qed show "f (x i) \ S - y ` K" apply (auto simp: f_in_S \x i \ S\) by (metis feq homfg \x i \ S\ homeomorphism_def \S \ T\ \i \ K\ subsetCE xney xyS) show "y i \ S - y ` K" using insert.hyps xney by (auto simp: \y i \ S\) qed blast show ?case proof (intro exI conjI) show "homeomorphism T T (h \ f) (g \ k)" using homfg homhk homeomorphism_compose by blast show "\i \ insert i K. (h \ f) (x i) = y i" using feq hk_sub by (auto simp: heq) show "{x. \ ((h \ f) x = x \ (g \ k) x = x)} \ S" using fg_sub hk_sub by force have "bounded ({x. \(f x = x \ g x = x)} \ {x. \(h x = x \ k x = x)})" using bo_fg bo_hk bounded_Un by blast then show "bounded {x. \ ((h \ f) x = x \ (g \ k) x = x)}" by (rule bounded_subset) auto qed qed proposition\<^marker>\tag unimportant\ homeomorphism_moving_points_exists: fixes S :: "'a::euclidean_space set" assumes 2: "2 \ DIM('a)" "open S" "connected S" "S \ T" "finite K" and KS: "\i. i \ K \ x i \ S \ y i \ S" and pw: "pairwise (\i j. (x i \ x j) \ (y i \ y j)) K" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "\i. i \ K \ f(x i) = y i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof (cases "S = {}") case True then show ?thesis using KS homeomorphism_ident that by fastforce next case False then have affS: "affine hull S = UNIV" by (simp add: affine_hull_open \open S\) then have ope: "openin (top_of_set (affine hull S)) S" using \open S\ open_openin by auto have "2 \ DIM('a)" by (rule 2) also have "\ = aff_dim (UNIV :: 'a set)" by simp also have "\ \ aff_dim S" by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) finally have "2 \ aff_dim S" by linarith then show ?thesis using homeomorphism_moving_points_exists_gen [OF \finite K\ KS pw _ ope S] that by fastforce qed subsubsection\<^marker>\tag unimportant\\The theorem \homeomorphism_grouping_points_exists\\ lemma homeomorphism_grouping_point_1: fixes a::real and c::real assumes "a < b" "c < d" obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" proof - define f where "f \ \x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" have "\g. homeomorphism (cbox a b) (cbox c d) f g" proof (rule homeomorphism_compact) show "continuous_on (cbox a b) f" apply (simp add: f_def) apply (intro continuous_intros) using assms by auto have "f ` {a..b} = {c..d}" unfolding f_def image_affinity_atLeastAtMost - using assms sum_sqs_eq by (auto simp: field_split_simps algebra_simps) + using assms sum_sqs_eq by (auto simp: field_split_simps) then show "f ` cbox a b = cbox c d" by auto show "inj_on f (cbox a b)" unfolding f_def inj_on_def using assms by auto qed auto then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. then show ?thesis proof show "f a = c" by (simp add: f_def) show "f b = d" - using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps algebra_simps) + using assms sum_sqs_eq [of a b] by (auto simp: f_def field_split_simps) qed qed lemma homeomorphism_grouping_point_2: fixes a::real and w::real assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" and "b \ cbox a c" "v \ cbox u w" and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" "\x. x \ cbox a b \ f x = f1 x" "\x. x \ cbox b c \ f x = f2 x" proof - have le: "a \ b" "b \ c" "u \ v" "v \ w" using assms by simp_all then have ac: "cbox a c = cbox a b \ cbox b c" and uw: "cbox u w = cbox u v \ cbox v w" by auto define f where "f \ \x. if x \ b then f1 x else f2 x" have "\g. homeomorphism (cbox a c) (cbox u w) f g" proof (rule homeomorphism_compact) have cf1: "continuous_on (cbox a b) f1" using hom_ab homeomorphism_cont1 by blast have cf2: "continuous_on (cbox b c) f2" using hom_bc homeomorphism_cont1 by blast show "continuous_on (cbox a c) f" apply (simp add: f_def) apply (rule continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) - using le eq apply (force simp: continuous_on_id)+ + using le eq apply (force)+ done have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" unfolding f_def using eq by force+ then show "f ` cbox a c = cbox u w" apply (simp only: ac uw image_Un) by (metis hom_ab hom_bc homeomorphism_def) have neq12: "f1 x \ f2 y" if x: "a \ x" "x \ b" and y: "b < y" "y \ c" for x y proof - have "f1 x \ cbox u v" by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) moreover have "f2 y \ cbox v w" by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) moreover have "f2 y \ f2 b" by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) ultimately show ?thesis using le eq by simp qed have "inj_on f1 (cbox a b)" by (metis (full_types) hom_ab homeomorphism_def inj_onI) moreover have "inj_on f2 (cbox b c)" by (metis (full_types) hom_bc homeomorphism_def inj_onI) ultimately show "inj_on f (cbox a c)" apply (simp (no_asm) add: inj_on_def) apply (simp add: f_def inj_on_eq_iff) using neq12 apply force done qed auto then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. then show ?thesis apply (rule that) using eq le by (auto simp: f_def) qed lemma homeomorphism_grouping_point_3: fixes a::real assumes cbox_sub: "cbox c d \ box a b" "cbox u v \ box a b" and box_ne: "box c d \ {}" "box u v \ {}" obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" "\x. x \ cbox c d \ f x \ cbox u v" proof - have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \ {}" using assms by (simp_all add: cbox_sub subset_eq) obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" and f1_eq: "f1 a = a" "f1 c = u" using homeomorphism_grouping_point_1 [OF \a < c\ \a < u\] . obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" and f2_eq: "f2 c = u" "f2 d = v" using homeomorphism_grouping_point_1 [OF \c < d\ \u < v\] . obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" and f3_eq: "f3 d = v" "f3 b = b" using homeomorphism_grouping_point_1 [OF \d < b\ \v < b\] . obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" and f4_eq: "\x. x \ cbox a c \ f4 x = f1 x" "\x. x \ cbox c d \ f4 x = f2 x" using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" and f_eq: "\x. x \ cbox a d \ f x = f4 x" "\x. x \ cbox d b \ f x = f3 x" using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) show ?thesis apply (rule that [OF fg]) using f4_eq f_eq homeomorphism_image1 [OF 2] apply simp by (metis atLeastAtMost_iff box_real(1) box_real(2) cbox_sub(1) greaterThanLessThan_iff imageI less_eq_real_def subset_eq) qed lemma homeomorphism_grouping_point_4: fixes T :: "real set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "\x. x \ K \ f x \ U" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" proof - obtain c d where "box c d \ {}" "cbox c d \ U" proof - obtain u where "u \ U" using \U \ {}\ by blast then obtain e where "e > 0" "cball u e \ U" using \open U\ open_contains_cball by blast then show ?thesis by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) qed have "compact K" by (simp add: \finite K\ finite_imp_compact) obtain a b where "box a b \ {}" "K \ cbox a b" "cbox a b \ S" proof (cases "K = {}") case True then show ?thesis using \box c d \ {}\ \cbox c d \ U\ \U \ S\ that by blast next case False then obtain a b where "a \ K" "b \ K" and a: "\x. x \ K \ a \ x" and b: "\x. x \ K \ x \ b" using compact_attains_inf compact_attains_sup by (metis \compact K\)+ obtain e where "e > 0" "cball b e \ S" using \open S\ open_contains_cball by (metis \b \ K\ \K \ S\ subsetD) show ?thesis proof show "box a (b + e) \ {}" using \0 < e\ \b \ K\ a by force show "K \ cbox a (b + e)" using \0 < e\ a b by fastforce have "a \ S" using \a \ K\ assms(6) by blast have "b + e \ S" using \0 < e\ \cball b e \ S\ by (force simp: dist_norm) show "cbox a (b + e) \ S" using \a \ S\ \b + e \ S\ \connected S\ connected_contains_Icc by auto qed qed obtain w z where "cbox w z \ S" and sub_wz: "cbox a b \ cbox c d \ box w z" proof - have "a \ S" "b \ S" using \box a b \ {}\ \cbox a b \ S\ by auto moreover have "c \ S" "d \ S" using \box c d \ {}\ \cbox c d \ U\ \U \ S\ by force+ ultimately have "min a c \ S" "max b d \ S" by linarith+ then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \ S" "e2 > 0" "cball (max b d) e2 \ S" using \open S\ open_contains_cball by metis then have *: "min a c - e1 \ S" "max b d + e2 \ S" by (auto simp: dist_norm) show ?thesis proof show "cbox (min a c - e1) (max b d+ e2) \ S" using * \connected S\ connected_contains_Icc by auto show "cbox a b \ cbox c d \ box (min a c - e1) (max b d + e2)" using \0 < e1\ \0 < e2\ by auto qed qed then obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" and "f w = w" "f z = z" and fin: "\x. x \ cbox a b \ f x \ cbox c d" using homeomorphism_grouping_point_3 [of a b w z c d] using \box a b \ {}\ \box c d \ {}\ by blast have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" using hom homeomorphism_def by blast+ define f' where "f' \ \x. if x \ cbox w z then f x else x" define g' where "g' \ \x. if x \ cbox w z then g x else x" show ?thesis proof have T: "cbox w z \ (T - box w z) = T" using \cbox w z \ S\ \S \ T\ by auto show "homeomorphism T T f' g'" proof have clo: "closedin (top_of_set (cbox w z \ (T - box w z))) (T - box w z)" by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) have "continuous_on (cbox w z \ (T - box w z)) f'" "continuous_on (cbox w z \ (T - box w z)) g'" unfolding f'_def g'_def apply (safe intro!: continuous_on_cases_local contfg continuous_on_id clo) apply (simp_all add: closed_subset) using \f w = w\ \f z = z\ apply force by (metis \f w = w\ \f z = z\ hom homeomorphism_def less_eq_real_def mem_box_real(2)) then show "continuous_on T f'" "continuous_on T g'" by (simp_all only: T) show "f' ` T \ T" unfolding f'_def by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) show "g' ` T \ T" unfolding g'_def by clarsimp (metis \cbox w z \ S\ \S \ T\ subsetD hom homeomorphism_def imageI mem_box_real(2)) show "\x. x \ T \ g' (f' x) = x" unfolding f'_def g'_def using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce show "\y. y \ T \ f' (g' y) = y" unfolding f'_def g'_def using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce qed show "\x. x \ K \ f' x \ U" using fin sub_wz \K \ cbox a b\ \cbox c d \ U\ by (force simp: f'_def) show "{x. \ (f' x = x \ g' x = x)} \ S" using \cbox w z \ S\ by (auto simp: f'_def g'_def) show "bounded {x. \ (f' x = x \ g' x = x)}" apply (rule bounded_subset [of "cbox w z"]) using bounded_cbox apply blast apply (auto simp: f'_def g'_def) done qed qed proposition\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists: fixes S :: "'a::euclidean_space set" assumes "open U" "open S" "connected S" "U \ {}" "finite K" "K \ S" "U \ S" "S \ T" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ DIM('a)") case True have TS: "T \ affine hull S" using affine_hull_open assms by blast have "infinite U" using \open U\ \U \ {}\ finite_imp_not_open by blast then obtain P where "P \ U" "finite P" "card K = card P" using infinite_arbitrarily_large by metis then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast obtain f g where "homeomorphism T T f g" "\i. i \ K \ f (id i) = \ i" "{x. \ (f x = x \ g x = x)} \ S" "bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists [OF True \open S\ \connected S\ \S \ T\ \finite K\]) show "\i. i \ K \ id i \ S \ \ i \ S" using \P \ U\ \bij_betw \ K P\ \K \ S\ \U \ S\ bij_betwE by blast show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed (use affine_hull_open assms that in auto) then show ?thesis using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) next case False with DIM_positive have "DIM('a) = 1" by (simp add: dual_order.antisym) then obtain h::"'a \real" and j where "linear h" "linear j" and noh: "\x. norm(h x) = norm x" and noj: "\y. norm(j y) = norm y" and hj: "\x. j(h x) = x" "\y. h(j y) = y" and ranh: "surj h" using isomorphisms_UNIV_UNIV by (metis (mono_tags, hide_lams) DIM_real UNIV_eq_I range_eqI) obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" and f: "\x. x \ h ` K \ f x \ h ` U" and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" and bou: "bounded {x. \ (f x = x \ g x = x)}" apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) by (simp_all add: assms image_mono \linear h\ open_surjective_linear_image connected_linear_image ranh) have jf: "j (f (h x)) = x \ f (h x) = h x" for x by (metis hj) have jg: "j (g (h x)) = x \ g (h x) = h x" for x by (metis hj) have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y by (simp_all add: \linear h\ \linear j\ linear_linear linear_continuous_on) show ?thesis proof show "homeomorphism T T (j \ f \ h) (j \ g \ h)" proof show "continuous_on T (j \ f \ h)" "continuous_on T (j \ g \ h)" using hom homeomorphism_def by (blast intro: continuous_on_compose cont_hj)+ show "(j \ f \ h) ` T \ T" "(j \ g \ h) ` T \ T" by auto (metis (mono_tags, hide_lams) hj(1) hom homeomorphism_def imageE imageI)+ show "\x. x \ T \ (j \ g \ h) ((j \ f \ h) x) = x" using hj hom homeomorphism_apply1 by fastforce show "\y. y \ T \ (j \ f \ h) ((j \ g \ h) y) = y" using hj hom homeomorphism_apply2 by fastforce qed show "{x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)} \ S" apply (clarsimp simp: jf jg hj) using sub hj apply (drule_tac c="h x" in subsetD, force) by (metis imageE) have "bounded (j ` {x. (\ (f x = x \ g x = x))})" by (rule bounded_linear_image [OF bou]) (use \linear j\ linear_conv_bounded_linear in auto) moreover have *: "{x. \((j \ f \ h) x = x \ (j \ g \ h) x = x)} = j ` {x. (\ (f x = x \ g x = x))}" using hj by (auto simp: jf jg image_iff, metis+) ultimately show "bounded {x. \ ((j \ f \ h) x = x \ (j \ g \ h) x = x)}" by metis show "\x. x \ K \ (j \ f \ h) x \ U" using f hj by fastforce qed qed proposition\<^marker>\tag unimportant\ homeomorphism_grouping_points_exists_gen: fixes S :: "'a::euclidean_space set" assumes opeU: "openin (top_of_set S) U" and opeS: "openin (top_of_set (affine hull S)) S" and "U \ {}" "finite K" "K \ S" and S: "S \ T" "T \ affine hull S" "connected S" obtains f g where "homeomorphism T T f g" "{x. (\ (f x = x \ g x = x))} \ S" "bounded {x. (\ (f x = x \ g x = x))}" "\x. x \ K \ f x \ U" proof (cases "2 \ aff_dim S") case True have opeU': "openin (top_of_set (affine hull S)) U" using opeS opeU openin_trans by blast obtain u where "u \ U" "u \ S" using \U \ {}\ opeU openin_imp_subset by fastforce+ have "infinite U" apply (rule infinite_openin [OF opeU \u \ U\]) apply (rule connected_imp_perfect_aff_dim [OF \connected S\ _ \u \ S\]) using True apply simp done then obtain P where "P \ U" "finite P" "card K = card P" using infinite_arbitrarily_large by metis then obtain \ where \: "bij_betw \ K P" using \finite K\ finite_same_card_bij by blast have "\f g. homeomorphism T T f g \ (\i \ K. f(id i) = \ i) \ {x. \ (f x = x \ g x = x)} \ S \ bounded {x. \ (f x = x \ g x = x)}" proof (rule homeomorphism_moving_points_exists_gen [OF \finite K\ _ _ True opeS S]) show "\i. i \ K \ id i \ S \ \ i \ S" by (metis id_apply opeU openin_contains_cball subsetCE \P \ U\ \bij_betw \ K P\ \K \ S\ bij_betwE) show "pairwise (\i j. id i \ id j \ \ i \ \ j) K" using \ by (auto simp: pairwise_def bij_betw_def inj_on_def) qed then show ?thesis using \ \P \ U\ bij_betwE by (fastforce simp add: intro!: that) next case False with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith then show ?thesis proof cases assume "aff_dim S = -1" then have "S = {}" using aff_dim_empty by blast then have "False" using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast then show ?thesis .. next assume "aff_dim S = 0" then obtain a where "S = {a}" using aff_dim_eq_0 by blast then have "K \ U" using \U \ {}\ \K \ S\ openin_imp_subset [OF opeU] by blast show ?thesis apply (rule that [of id id]) - using \K \ U\ by (auto simp: continuous_on_id intro: homeomorphismI) + using \K \ U\ by (auto intro: homeomorphismI) next assume "aff_dim S = 1" then have "affine hull S homeomorphic (UNIV :: real set)" by (auto simp: homeomorphic_affine_sets) then obtain h::"'a\real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" using homeomorphic_def by blast then have h: "\x. x \ affine hull S \ j(h(x)) = x" and j: "\y. j y \ affine hull S \ h(j y) = y" by (auto simp: homeomorphism_def) have connh: "connected (h ` S)" by (meson Topological_Spaces.connected_continuous_image \connected S\ homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) have hUS: "h ` U \ h ` S" by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) have opn: "openin (top_of_set (affine hull S)) U \ open (h ` U)" for U using homeomorphism_imp_open_map [OF homhj] by simp have "open (h ` U)" "open (h ` S)" by (auto intro: opeS opeU openin_trans opn) then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" and f: "\x. x \ h ` K \ f x \ h ` U" and sub: "{x. \ (f x = x \ g x = x)} \ h ` S" and bou: "bounded {x. \ (f x = x \ g x = x)}" apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) using assms by (auto simp: connh hUS) have jf: "\x. x \ affine hull S \ j (f (h x)) = x \ f (h x) = h x" by (metis h j) have jg: "\x. x \ affine hull S \ j (g (h x)) = x \ g (h x) = h x" by (metis h j) have cont_hj: "continuous_on T h" "continuous_on Y j" for Y apply (rule continuous_on_subset [OF _ \T \ affine hull S\]) using homeomorphism_def homhj apply blast by (meson continuous_on_subset homeomorphism_def homhj top_greatest) define f' where "f' \ \x. if x \ affine hull S then (j \ f \ h) x else x" define g' where "g' \ \x. if x \ affine hull S then (j \ g \ h) x else x" show ?thesis proof show "homeomorphism T T f' g'" proof have "continuous_on T (j \ f \ h)" apply (intro continuous_on_compose cont_hj) using hom homeomorphism_def by blast then show "continuous_on T f'" apply (rule continuous_on_eq) using \T \ affine hull S\ f'_def by auto have "continuous_on T (j \ g \ h)" apply (intro continuous_on_compose cont_hj) using hom homeomorphism_def by blast then show "continuous_on T g'" apply (rule continuous_on_eq) using \T \ affine hull S\ g'_def by auto show "f' ` T \ T" proof (clarsimp simp: f'_def) fix x assume "x \ T" then have "f (h x) \ h ` T" by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) then show "j (f (h x)) \ T" using \T \ affine hull S\ h by auto qed show "g' ` T \ T" proof (clarsimp simp: g'_def) fix x assume "x \ T" then have "g (h x) \ h ` T" by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) then show "j (g (h x)) \ T" using \T \ affine hull S\ h by auto qed show "\x. x \ T \ g' (f' x) = x" using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def) show "\y. y \ T \ f' (g' y) = y" using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def) qed next show "{x. \ (f' x = x \ g' x = x)} \ S" apply (clarsimp simp: f'_def g'_def jf jg) apply (rule imageE [OF subsetD [OF sub]], force) by (metis h hull_inc) next have "compact (j ` closure {x. \ (f x = x \ g x = x)})" using bou by (auto simp: compact_continuous_image cont_hj) then have "bounded (j ` {x. \ (f x = x \ g x = x)})" by (rule bounded_closure_image [OF compact_imp_bounded]) moreover have *: "{x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x} = j ` {x. (\ (f x = x \ g x = x))}" using h j by (auto simp: image_iff; metis) ultimately have "bounded {x \ affine hull S. j (f (h x)) \ x \ j (g (h x)) \ x}" by metis then show "bounded {x. \ (f' x = x \ g' x = x)}" by (simp add: f'_def g'_def Collect_mono bounded_subset) next show "f' x \ U" if "x \ K" for x proof - have "U \ S" using opeU openin_imp_subset by blast then have "j (f (h x)) \ U" using f h hull_subset that by fastforce then show "f' x \ U" using \K \ S\ S f'_def that by auto qed qed qed qed subsection\Nullhomotopic mappings\ text\ A mapping out of a sphere is nullhomotopic iff it extends to the ball. This even works out in the degenerate cases when the radius is \\\ 0, and we also don't need to explicitly assume continuity since it's already implicit in both sides of the equivalence.\ lemma nullhomotopic_from_lemma: assumes contg: "continuous_on (cball a r - {a}) g" and fa: "\e. 0 < e \ \d. 0 < d \ (\x. x \ a \ norm(x - a) < d \ norm(g x - f a) < e)" and r: "\x. x \ cball a r \ x \ a \ f x = g x" shows "continuous_on (cball a r) f" proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) fix x assume x: "dist a x \ r" show "continuous (at x within cball a r) f" proof (cases "x=a") case True then show ?thesis by (metis continuous_within_eps_delta fa dist_norm dist_self r) next case False show ?thesis proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) have "\d>0. \x'\cball a r. dist x' x < d \ dist (g x') (g x) < e" if "e>0" for e proof - obtain d where "d > 0" and d: "\x'. \dist x' a \ r; x' \ a; dist x' x < d\ \ dist (g x') (g x) < e" using contg False x \e>0\ unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that) show ?thesis using \d > 0\ \x \ a\ by (rule_tac x="min d (norm(x - a))" in exI) (auto simp: dist_commute dist_norm [symmetric] intro!: d) qed then show "continuous (at x within cball a r) g" using contg False by (auto simp: continuous_within_eps_delta) show "0 < norm (x - a)" using False by force show "x \ cball a r" by (simp add: x) show "\x'. \x' \ cball a r; dist x' x < norm (x - a)\ \ g x' = f x'" by (metis dist_commute dist_norm less_le r) qed qed qed proposition nullhomotopic_from_sphere_extension: fixes f :: "'M::euclidean_space \ 'a::real_normed_vector" shows "(\c. homotopic_with_canon (\x. True) (sphere a r) S f (\x. c)) \ (\g. continuous_on (cball a r) g \ g ` (cball a r) \ S \ (\x \ sphere a r. g x = f x))" (is "?lhs = ?rhs") proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp add: homotopic_on_emptyI) next case equal then show ?thesis apply (auto simp: homotopic_with) apply (rule_tac x="\x. h (0, a)" in exI) apply (fastforce simp add:) using continuous_on_const by blast next case greater let ?P = "continuous_on {x. norm(x - a) = r} f \ f ` {x. norm(x - a) = r} \ S" have ?P if ?lhs using that proof fix c assume c: "homotopic_with_canon (\x. True) (sphere a r) S f (\x. c)" then have contf: "continuous_on (sphere a r) f" by (metis homotopic_with_imp_continuous) moreover have fim: "f ` sphere a r \ S" by (meson continuous_map_subtopology_eu c homotopic_with_imp_continuous_maps) show ?P using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) qed moreover have ?P if ?rhs using that proof fix g assume g: "continuous_on (cball a r) g \ g ` cball a r \ S \ (\xa\sphere a r. g xa = f xa)" then show ?P apply (safe elim!: continuous_on_eq [OF continuous_on_subset]) apply (auto simp: dist_norm norm_minus_commute) by (metis dist_norm image_subset_iff mem_sphere norm_minus_commute sphere_cball subsetCE) qed moreover have ?thesis if ?P proof assume ?lhs then obtain c where "homotopic_with_canon (\x. True) (sphere a r) S (\x. c) f" using homotopic_with_sym by blast then obtain h where conth: "continuous_on ({0..1::real} \ sphere a r) h" and him: "h ` ({0..1} \ sphere a r) \ S" and h: "\x. h(0, x) = c" "\x. h(1, x) = f x" by (auto simp: homotopic_with_def) obtain b1::'M where "b1 \ Basis" using SOME_Basis by auto have "c \ S" apply (rule him [THEN subsetD]) apply (rule_tac x = "(0, a + r *\<^sub>R b1)" in image_eqI) using h greater \b1 \ Basis\ apply (auto simp: dist_norm) done have uconth: "uniformly_continuous_on ({0..1::real} \ (sphere a r)) h" by (force intro: compact_Times conth compact_uniformly_continuous) let ?g = "\x. h (norm (x - a)/r, a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" let ?g' = "\x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" show ?rhs proof (intro exI conjI) have "continuous_on (cball a r - {a}) ?g'" apply (rule continuous_on_compose2 [OF conth]) apply (intro continuous_intros) using greater apply (auto simp: dist_norm norm_minus_commute) done then show "continuous_on (cball a r) ?g" proof (rule nullhomotopic_from_lemma) show "\d>0. \x. x \ a \ norm (x - a) < d \ norm (?g' x - ?g a) < e" if "0 < e" for e proof - obtain d where "0 < d" and d: "\x x'. \x \ {0..1} \ sphere a r; x' \ {0..1} \ sphere a r; dist x' x < d\ \ dist (h x') (h x) < e" using uniformly_continuous_onE [OF uconth \0 < e\] by auto have *: "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" if "x \ a" "norm (x - a) < r" "norm (x - a) < d * r" for x proof - have "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) = norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" by (simp add: h) also have "\ < e" apply (rule d [unfolded dist_norm]) using greater \0 < d\ \b1 \ Basis\ that by (simp_all add: dist_norm) (simp add: field_simps) finally show ?thesis . qed show ?thesis apply (rule_tac x = "min r (d * r)" in exI) using greater \0 < d\ by (auto simp: *) qed show "\x. x \ cball a r \ x \ a \ ?g x = ?g' x" by auto qed next show "?g ` cball a r \ S" using greater him \c \ S\ by (force simp: h dist_norm norm_minus_commute) next show "\x\sphere a r. ?g x = f x" using greater by (auto simp: h dist_norm norm_minus_commute) qed next assume ?rhs then obtain g where contg: "continuous_on (cball a r) g" and gim: "g ` cball a r \ S" and gf: "\x \ sphere a r. g x = f x" by auto let ?h = "\y. g (a + (fst y) *\<^sub>R (snd y - a))" have "continuous_on ({0..1} \ sphere a r) ?h" proof (rule continuous_on_compose2 [OF contg]) show "continuous_on ({0..1} \ sphere a r) (\x. a + fst x *\<^sub>R (snd x - a))" by (intro continuous_intros) qed (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) moreover have "?h ` ({0..1} \ sphere a r) \ S" by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) moreover have "\x\sphere a r. ?h (0, x) = g a" "\x\sphere a r. ?h (1, x) = f x" by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) ultimately have "homotopic_with_canon (\x. True) (sphere a r) S (\x. g a) f" by (auto simp: homotopic_with) then show ?lhs using homotopic_with_symD by blast qed ultimately show ?thesis by meson qed end \ No newline at end of file diff --git a/src/HOL/Analysis/Lebesgue_Measure.thy b/src/HOL/Analysis/Lebesgue_Measure.thy --- a/src/HOL/Analysis/Lebesgue_Measure.thy +++ b/src/HOL/Analysis/Lebesgue_Measure.thy @@ -1,1813 +1,1813 @@ (* Title: HOL/Analysis/Lebesgue_Measure.thy Author: Johannes Hölzl, TU München Author: Robert Himmelmann, TU München Author: Jeremy Avigad Author: Luke Serafin *) section \Lebesgue Measure\ theory Lebesgue_Measure imports Finite_Product_Measure Caratheodory Complete_Measure Summation_Tests Regularity begin lemma measure_eqI_lessThan: fixes M N :: "real measure" assumes sets: "sets M = sets borel" "sets N = sets borel" assumes fin: "\x. emeasure M {x <..} < \" assumes "\x. emeasure M {x <..} = emeasure N {x <..}" shows "M = N" proof (rule measure_eqI_generator_eq_countable) let ?LT = "\a::real. {a <..}" let ?E = "range ?LT" show "Int_stable ?E" by (auto simp: Int_stable_def lessThan_Int_lessThan) show "?E \ Pow UNIV" "sets M = sigma_sets UNIV ?E" "sets N = sigma_sets UNIV ?E" unfolding sets borel_Ioi by auto show "?LT`Rats \ ?E" "(\i\Rats. ?LT i) = UNIV" "\a. a \ ?LT`Rats \ emeasure M a \ \" using fin by (auto intro: Rats_no_bot_less simp: less_top) qed (auto intro: assms countable_rat) subsection \Measures defined by monotonous functions\ text \ Every right-continuous and nondecreasing function gives rise to a measure on the reals: \ definition\<^marker>\tag important\ interval_measure :: "(real \ real) \ real measure" where "interval_measure F = extend_measure UNIV {(a, b). a \ b} (\(a, b). {a<..b}) (\(a, b). ennreal (F b - F a))" lemma\<^marker>\tag important\ emeasure_interval_measure_Ioc: assumes "a \ b" assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\a. continuous (at_right a) F" shows "emeasure (interval_measure F) {a<..b} = F b - F a" proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \a \ b\]) show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \ b}" proof (unfold_locales, safe) fix a b c d :: real assume *: "a \ b" "c \ d" then show "\C\{{a<..b} |a b. a \ b}. finite C \ disjoint C \ {a<..b} - {c<..d} = \C" proof cases let ?C = "{{a<..b}}" assume "b < c \ d \ a \ d \ c" with * have "?C \ {{a<..b} |a b. a \ b} \ finite ?C \ disjoint ?C \ {a<..b} - {c<..d} = \?C" by (auto simp add: disjoint_def) thus ?thesis .. next let ?C = "{{a<..c}, {d<..b}}" assume "\ (b < c \ d \ a \ d \ c)" with * have "?C \ {{a<..b} |a b. a \ b} \ finite ?C \ disjoint ?C \ {a<..b} - {c<..d} = \?C" by (auto simp add: disjoint_def Ioc_inj) (metis linear)+ thus ?thesis .. qed qed (auto simp: Ioc_inj, metis linear) next fix l r :: "nat \ real" and a b :: real assume l_r[simp]: "\n. l n \ r n" and "a \ b" and disj: "disjoint_family (\n. {l n<..r n})" assume lr_eq_ab: "(\i. {l i<..r i}) = {a<..b}" have [intro, simp]: "\a b. a \ b \ F a \ F b" by (auto intro!: l_r mono_F) { fix S :: "nat set" assume "finite S" moreover note \a \ b\ moreover have "\i. i \ S \ {l i <.. r i} \ {a <.. b}" unfolding lr_eq_ab[symmetric] by auto ultimately have "(\i\S. F (r i) - F (l i)) \ F b - F a" proof (induction S arbitrary: a rule: finite_psubset_induct) case (psubset S) show ?case proof cases assume "\i\S. l i < r i" with \finite S\ have "Min (l ` {i\S. l i < r i}) \ l ` {i\S. l i < r i}" by (intro Min_in) auto then obtain m where m: "m \ S" "l m < r m" "l m = Min (l ` {i\S. l i < r i})" by fastforce have "(\i\S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (\i\S - {m}. F (r i) - F (l i))" using m psubset by (intro sum.remove) auto also have "(\i\S - {m}. F (r i) - F (l i)) \ F b - F (r m)" proof (intro psubset.IH) show "S - {m} \ S" using \m\S\ by auto show "r m \ b" using psubset.prems(2)[OF \m\S\] \l m < r m\ by auto next fix i assume "i \ S - {m}" then have i: "i \ S" "i \ m" by auto { assume i': "l i < r i" "l i < r m" with \finite S\ i m have "l m \ l i" by auto with i' have "{l i <.. r i} \ {l m <.. r m} \ {}" by auto then have False using disjoint_family_onD[OF disj, of i m] i by auto } then have "l i \ r i \ r m \ l i" unfolding not_less[symmetric] using l_r[of i] by auto then show "{l i <.. r i} \ {r m <.. b}" using psubset.prems(2)[OF \i\S\] by auto qed also have "F (r m) - F (l m) \ F (r m) - F a" using psubset.prems(2)[OF \m \ S\] \l m < r m\ by (auto simp add: Ioc_subset_iff intro!: mono_F) finally show ?case by (auto intro: add_mono) qed (auto simp add: \a \ b\ less_le) qed } note claim1 = this (* second key induction: a lower bound on the measures of any finite collection of Ai's that cover an interval {u..v} *) { fix S u v and l r :: "nat \ real" assume "finite S" "\i. i\S \ l i < r i" "{u..v} \ (\i\S. {l i<..< r i})" then have "F v - F u \ (\i\S. F (r i) - F (l i))" proof (induction arbitrary: v u rule: finite_psubset_induct) case (psubset S) show ?case proof cases assume "S = {}" then show ?case using psubset by (simp add: mono_F) next assume "S \ {}" then obtain j where "j \ S" by auto let ?R = "r j < u \ l j > v \ (\i\S-{j}. l i \ l j \ r j \ r i)" show ?case proof cases assume "?R" with \j \ S\ psubset.prems have "{u..v} \ (\i\S-{j}. {l i<..< r i})" apply (auto simp: subset_eq Ball_def) apply (metis Diff_iff less_le_trans leD linear singletonD) apply (metis Diff_iff less_le_trans leD linear singletonD) apply (metis order_trans less_le_not_le linear) done with \j \ S\ have "F v - F u \ (\i\S - {j}. F (r i) - F (l i))" by (intro psubset) auto also have "\ \ (\i\S. F (r i) - F (l i))" using psubset.prems by (intro sum_mono2 psubset) (auto intro: less_imp_le) finally show ?thesis . next assume "\ ?R" then have j: "u \ r j" "l j \ v" "\i. i \ S - {j} \ r i < r j \ l i > l j" by (auto simp: not_less) let ?S1 = "{i \ S. l i < l j}" let ?S2 = "{i \ S. r i > r j}" have "(\i\S. F (r i) - F (l i)) \ (\i\?S1 \ ?S2 \ {j}. F (r i) - F (l i))" using \j \ S\ \finite S\ psubset.prems j by (intro sum_mono2) (auto intro: less_imp_le) also have "(\i\?S1 \ ?S2 \ {j}. F (r i) - F (l i)) = (\i\?S1. F (r i) - F (l i)) + (\i\?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))" using psubset(1) psubset.prems(1) j apply (subst sum.union_disjoint) apply simp_all apply (subst sum.union_disjoint) apply auto apply (metis less_le_not_le) done also (xtrans) have "(\i\?S1. F (r i) - F (l i)) \ F (l j) - F u" using \j \ S\ \finite S\ psubset.prems j apply (intro psubset.IH psubset) apply (auto simp: subset_eq Ball_def) apply (metis less_le_trans not_le) done also (xtrans) have "(\i\?S2. F (r i) - F (l i)) \ F v - F (r j)" using \j \ S\ \finite S\ psubset.prems j apply (intro psubset.IH psubset) apply (auto simp: subset_eq Ball_def) apply (metis le_less_trans not_le) done finally (xtrans) show ?case by (auto simp: add_mono) qed qed qed } note claim2 = this (* now prove the inequality going the other way *) have "ennreal (F b - F a) \ (\i. ennreal (F (r i) - F (l i)))" proof (rule ennreal_le_epsilon) fix epsilon :: real assume egt0: "epsilon > 0" have "\i. \d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)" proof fix i note right_cont_F [of "r i"] thus "\d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)" apply - apply (subst (asm) continuous_at_right_real_increasing) apply (rule mono_F, assumption) apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec) apply (erule impE) using egt0 by (auto simp add: field_simps) qed then obtain delta where deltai_gt0: "\i. delta i > 0" and deltai_prop: "\i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)" by metis have "\a' > a. F a' - F a < epsilon / 2" apply (insert right_cont_F [of a]) apply (subst (asm) continuous_at_right_real_increasing) using mono_F apply force apply (drule_tac x = "epsilon / 2" in spec) using egt0 unfolding mult.commute [of 2] by force then obtain a' where a'lea [arith]: "a' > a" and a_prop: "F a' - F a < epsilon / 2" by auto define S' where "S' = {i. l i < r i}" obtain S :: "nat set" where "S \ S'" and finS: "finite S" and Sprop: "{a'..b} \ (\i \ S. {l i<..i. i \ S' \ open ({l i<.. {a <.. b}" by auto also have "{a <.. b} = (\i\S'. {l i<..r i})" unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans) also have "\ \ (\i \ S'. {l i<.. (\i \ S'. {l i<..i. i \ S \ l i < r i" by auto from finS have "\n. \i \ S. i \ n" by (subst finite_nat_set_iff_bounded_le [symmetric]) then obtain n where Sbound [rule_format]: "\i \ S. i \ n" .. have "F b - F a' \ (\i\S. F (r i + delta i) - F (l i))" apply (rule claim2 [rule_format]) using finS Sprop apply auto apply (frule Sprop2) apply (subgoal_tac "delta i > 0") apply arith by (rule deltai_gt0) also have "... \ (\i \ S. F(r i) - F(l i) + epsilon / 2^(i+2))" apply (rule sum_mono) apply simp apply (rule order_trans) apply (rule less_imp_le) apply (rule deltai_prop) by auto also have "... = (\i \ S. F(r i) - F(l i)) + (epsilon / 4) * (\i \ S. (1 / 2)^i)" (is "_ = ?t + _") by (subst sum.distrib) (simp add: field_simps sum_distrib_left) also have "... \ ?t + (epsilon / 4) * (\ i < Suc n. (1 / 2)^i)" apply (rule add_left_mono) apply (rule mult_left_mono) apply (rule sum_mono2) using egt0 apply auto by (frule Sbound, auto) also have "... \ ?t + (epsilon / 2)" apply (rule add_left_mono) apply (subst geometric_sum) apply auto apply (rule mult_left_mono) using egt0 apply auto done finally have aux2: "F b - F a' \ (\i\S. F (r i) - F (l i)) + epsilon / 2" by simp have "F b - F a = (F b - F a') + (F a' - F a)" by auto also have "... \ (F b - F a') + epsilon / 2" using a_prop by (intro add_left_mono) simp also have "... \ (\i\S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2" apply (intro add_right_mono) apply (rule aux2) done also have "... = (\i\S. F (r i) - F (l i)) + epsilon" by auto also have "... \ (\i\n. F (r i) - F (l i)) + epsilon" using finS Sbound Sprop by (auto intro!: add_right_mono sum_mono2) finally have "ennreal (F b - F a) \ (\i\n. ennreal (F (r i) - F (l i))) + epsilon" using egt0 by (simp add: sum_nonneg flip: ennreal_plus) then show "ennreal (F b - F a) \ (\i. ennreal (F (r i) - F (l i))) + (epsilon :: real)" by (rule order_trans) (auto intro!: add_mono sum_le_suminf simp del: sum_ennreal) qed moreover have "(\i. ennreal (F (r i) - F (l i))) \ ennreal (F b - F a)" using \a \ b\ by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1) ultimately show "(\n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)" by (rule antisym[rotated]) qed (auto simp: Ioc_inj mono_F) lemma measure_interval_measure_Ioc: assumes "a \ b" and "\x y. x \ y \ F x \ F y" and "\a. continuous (at_right a) F" shows "measure (interval_measure F) {a <.. b} = F b - F a" unfolding measure_def by (simp add: assms emeasure_interval_measure_Ioc) lemma emeasure_interval_measure_Ioc_eq: "(\x y. x \ y \ F x \ F y) \ (\a. continuous (at_right a) F) \ emeasure (interval_measure F) {a <.. b} = (if a \ b then F b - F a else 0)" using emeasure_interval_measure_Ioc[of a b F] by auto lemma\<^marker>\tag important\ sets_interval_measure [simp, measurable_cong]: "sets (interval_measure F) = sets borel" apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc) apply (rule sigma_sets_eqI) apply auto apply (case_tac "a \ ba") apply (auto intro: sigma_sets.Empty) done lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV" by (simp add: interval_measure_def space_extend_measure) lemma emeasure_interval_measure_Icc: assumes "a \ b" assumes mono_F: "\x y. x \ y \ F x \ F y" assumes cont_F : "continuous_on UNIV F" shows "emeasure (interval_measure F) {a .. b} = F b - F a" proof (rule tendsto_unique) { fix a b :: real assume "a \ b" then have "emeasure (interval_measure F) {a <.. b} = F b - F a" using cont_F by (subst emeasure_interval_measure_Ioc) (auto intro: mono_F continuous_within_subset simp: continuous_on_eq_continuous_within) } note * = this let ?F = "interval_measure F" show "((\a. F b - F a) \ emeasure ?F {a..b}) (at_left a)" proof (rule tendsto_at_left_sequentially) show "a - 1 < a" by simp fix X assume "\n. X n < a" "incseq X" "X \ a" with \a \ b\ have "(\n. emeasure ?F {X n<..b}) \ emeasure ?F (\n. {X n <..b})" apply (intro Lim_emeasure_decseq) apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *) apply force apply (subst (asm ) *) apply (auto intro: less_le_trans less_imp_le) done also have "(\n. {X n <..b}) = {a..b}" using \\n. X n < a\ apply auto apply (rule LIMSEQ_le_const2[OF \X \ a\]) apply (auto intro: less_imp_le) apply (auto intro: less_le_trans) done also have "(\n. emeasure ?F {X n<..b}) = (\n. F b - F (X n))" using \\n. X n < a\ \a \ b\ by (subst *) (auto intro: less_imp_le less_le_trans) finally show "(\n. F b - F (X n)) \ emeasure ?F {a..b}" . qed show "((\a. ennreal (F b - F a)) \ F b - F a) (at_left a)" by (rule continuous_on_tendsto_compose[where g="\x. x" and s=UNIV]) - (auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const) + (auto simp: continuous_on_ennreal continuous_on_diff cont_F) qed (rule trivial_limit_at_left_real) lemma\<^marker>\tag important\ sigma_finite_interval_measure: assumes mono_F: "\x y. x \ y \ F x \ F y" assumes right_cont_F : "\a. continuous (at_right a) F" shows "sigma_finite_measure (interval_measure F)" apply unfold_locales apply (intro exI[of _ "(\(a, b). {a <.. b}) ` (\ \ \)"]) apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms]) done subsection \Lebesgue-Borel measure\ definition\<^marker>\tag important\ lborel :: "('a :: euclidean_space) measure" where "lborel = distr (\\<^sub>M b\Basis. interval_measure (\x. x)) borel (\f. \b\Basis. f b *\<^sub>R b)" abbreviation\<^marker>\tag important\ lebesgue :: "'a::euclidean_space measure" where "lebesgue \ completion lborel" abbreviation\<^marker>\tag important\ lebesgue_on :: "'a set \ 'a::euclidean_space measure" where "lebesgue_on \ \ restrict_space (completion lborel) \" lemma lebesgue_on_mono: assumes major: "AE x in lebesgue_on S. P x" and minor: "\x.\P x; x \ S\ \ Q x" shows "AE x in lebesgue_on S. Q x" proof - have "AE a in lebesgue_on S. P a \ Q a" using minor space_restrict_space by fastforce then show ?thesis using major by auto qed lemma integral_eq_zero_null_sets: assumes "S \ null_sets lebesgue" shows "integral\<^sup>L (lebesgue_on S) f = 0" proof (rule integral_eq_zero_AE) show "AE x in lebesgue_on S. f x = 0" by (metis (no_types, lifting) assms AE_not_in lebesgue_on_mono null_setsD2 null_sets_restrict_space order_refl) qed lemma shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel" and space_lborel[simp]: "space lborel = space borel" and measurable_lborel1[simp]: "measurable M lborel = measurable M borel" and measurable_lborel2[simp]: "measurable lborel M = measurable borel M" by (simp_all add: lborel_def) lemma space_lebesgue_on [simp]: "space (lebesgue_on S) = S" by (simp add: space_restrict_space) lemma sets_lebesgue_on_refl [iff]: "S \ sets (lebesgue_on S)" by (metis inf_top.right_neutral sets.top space_borel space_completion space_lborel space_restrict_space) lemma Compl_in_sets_lebesgue: "-A \ sets lebesgue \ A \ sets lebesgue" by (metis Compl_eq_Diff_UNIV double_compl space_borel space_completion space_lborel Sigma_Algebra.sets.compl_sets) lemma measurable_lebesgue_cong: assumes "\x. x \ S \ f x = g x" shows "f \ measurable (lebesgue_on S) M \ g \ measurable (lebesgue_on S) M" by (metis (mono_tags, lifting) IntD1 assms measurable_cong_simp space_restrict_space) lemma lebesgue_on_UNIV_eq: "lebesgue_on UNIV = lebesgue" proof - have "measure_of UNIV (sets lebesgue) (emeasure lebesgue) = lebesgue" by (metis measure_of_of_measure space_borel space_completion space_lborel) then show ?thesis by (auto simp: restrict_space_def) qed lemma integral_restrict_Int: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "T \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on T) (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on (S \ T)) f" proof - have "(\x. indicat_real T x *\<^sub>R (if x \ S then f x else 0)) = (\x. indicat_real (S \ T) x *\<^sub>R f x)" by (force simp: indicator_def) then show ?thesis by (simp add: assms sets.Int Bochner_Integration.integral_restrict_space) qed lemma integral_restrict: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ T" "S \ sets lebesgue" "T \ sets lebesgue" shows "integral\<^sup>L (lebesgue_on T) (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on S) f" using integral_restrict_Int [of S T f] assms by (simp add: Int_absorb2) lemma integral_restrict_UNIV: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "integral\<^sup>L lebesgue (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on S) f" using integral_restrict_Int [of S UNIV f] assms by (simp add: lebesgue_on_UNIV_eq) lemma integrable_lebesgue_on_empty [iff]: fixes f :: "'a::euclidean_space \ 'b::{second_countable_topology,banach}" shows "integrable (lebesgue_on {}) f" by (simp add: integrable_restrict_space) lemma integral_lebesgue_on_empty [simp]: fixes f :: "'a::euclidean_space \ 'b::{second_countable_topology,banach}" shows "integral\<^sup>L (lebesgue_on {}) f = 0" by (simp add: Bochner_Integration.integral_empty) lemma has_bochner_integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \: "\ \ space M \ sets M" shows "has_bochner_integral (restrict_space M \) f i \ has_bochner_integral M (\x. indicator \ x *\<^sub>R f x) i" by (simp add: integrable_restrict_space [OF assms] integral_restrict_space [OF assms] has_bochner_integral_iff) lemma integrable_restrict_UNIV: fixes f :: "'a::euclidean_space \ 'b::{banach, second_countable_topology}" assumes S: "S \ sets lebesgue" shows "integrable lebesgue (\x. if x \ S then f x else 0) \ integrable (lebesgue_on S) f" using has_bochner_integral_restrict_space [of S lebesgue f] assms by (simp add: integrable.simps indicator_scaleR_eq_if) lemma integral_mono_lebesgue_on_AE: fixes f::"_ \ real" assumes f: "integrable (lebesgue_on T) f" and gf: "AE x in (lebesgue_on S). g x \ f x" and f0: "AE x in (lebesgue_on T). 0 \ f x" and "S \ T" and S: "S \ sets lebesgue" and T: "T \ sets lebesgue" shows "(\x. g x \(lebesgue_on S)) \ (\x. f x \(lebesgue_on T))" proof - have "(\x. g x \(lebesgue_on S)) = (\x. (if x \ S then g x else 0) \lebesgue)" by (simp add: Lebesgue_Measure.integral_restrict_UNIV S) also have "\ \ (\x. (if x \ T then f x else 0) \lebesgue)" proof (rule Bochner_Integration.integral_mono_AE') show "integrable lebesgue (\x. if x \ T then f x else 0)" by (simp add: integrable_restrict_UNIV T f) show "AE x in lebesgue. (if x \ S then g x else 0) \ (if x \ T then f x else 0)" using assms by (auto simp: AE_restrict_space_iff) show "AE x in lebesgue. 0 \ (if x \ T then f x else 0)" using f0 by (simp add: AE_restrict_space_iff T) qed also have "\ = (\x. f x \(lebesgue_on T))" using Lebesgue_Measure.integral_restrict_UNIV T by blast finally show ?thesis . qed subsection \Borel measurability\ lemma borel_measurable_if_I: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" proof - have eq: "{x. x \ S} \ {x. f x \ Y} = {x. x \ S} \ {x. f x \ Y} \ S" for Y by blast show ?thesis using f S apply (simp add: vimage_def in_borel_measurable_borel Ball_def) apply (elim all_forward imp_forward asm_rl) apply (simp only: Collect_conj_eq Collect_disj_eq imp_conv_disj eq) apply (auto simp: Compl_eq [symmetric] Compl_in_sets_lebesgue sets_restrict_space_iff) done qed lemma borel_measurable_if_D: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue" shows "f \ borel_measurable (lebesgue_on S)" using assms apply (simp add: in_borel_measurable_borel Ball_def) apply (elim all_forward imp_forward asm_rl) apply (force simp: space_restrict_space sets_restrict_space image_iff intro: rev_bexI) done lemma borel_measurable_if: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "(\x. if x \ S then f x else 0) \ borel_measurable lebesgue \ f \ borel_measurable (lebesgue_on S)" using assms borel_measurable_if_D borel_measurable_if_I by blast lemma borel_measurable_if_lebesgue_on: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" "T \ sets lebesgue" "S \ T" shows "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on T) \ f \ borel_measurable (lebesgue_on S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using measurable_restrict_mono [OF _ \S \ T\] by (subst measurable_lebesgue_cong [where g = "(\x. if x \ S then f x else 0)"]) auto next assume ?rhs then show ?lhs by (simp add: \S \ sets lebesgue\ borel_measurable_if_I measurable_restrict_space1) qed lemma borel_measurable_vimage_halfspace_component_lt: "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i < a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_less]) apply (fastforce simp add: space_restrict_space) done lemma borel_measurable_vimage_halfspace_component_ge: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i \ a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_ge]) apply (fastforce simp add: space_restrict_space) done lemma borel_measurable_vimage_halfspace_component_gt: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i > a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_greater]) apply (fastforce simp add: space_restrict_space) done lemma borel_measurable_vimage_halfspace_component_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a i. i \ Basis \ {x \ S. f x \ i \ a} \ sets (lebesgue_on S))" apply (rule trans [OF borel_measurable_iff_halfspace_le]) apply (fastforce simp add: space_restrict_space) done lemma fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows borel_measurable_vimage_open_interval: "f \ borel_measurable (lebesgue_on S) \ (\a b. {x \ S. f x \ box a b} \ sets (lebesgue_on S))" (is ?thesis1) and borel_measurable_vimage_open: "f \ borel_measurable (lebesgue_on S) \ (\T. open T \ {x \ S. f x \ T} \ sets (lebesgue_on S))" (is ?thesis2) proof - have "{x \ S. f x \ box a b} \ sets (lebesgue_on S)" if "f \ borel_measurable (lebesgue_on S)" for a b proof - have "S = S \ space lebesgue" by simp then have "S \ (f -` box a b) \ sets (lebesgue_on S)" by (metis (no_types) box_borel in_borel_measurable_borel inf_sup_aci(1) space_restrict_space that) then show ?thesis by (simp add: Collect_conj_eq vimage_def) qed moreover have "{x \ S. f x \ T} \ sets (lebesgue_on S)" if T: "\a b. {x \ S. f x \ box a b} \ sets (lebesgue_on S)" "open T" for T proof - obtain \ where "countable \" and \: "\X. X \ \ \ \a b. X = box a b" "\\ = T" using open_countable_Union_open_box that \open T\ by metis then have eq: "{x \ S. f x \ T} = (\U \ \. {x \ S. f x \ U})" by blast have "{x \ S. f x \ U} \ sets (lebesgue_on S)" if "U \ \" for U using that T \ by blast then show ?thesis by (auto simp: eq intro: Sigma_Algebra.sets.countable_UN' [OF \countable \\]) qed moreover have eq: "{x \ S. f x \ i < a} = {x \ S. f x \ {y. y \ i < a}}" for i a by auto have "f \ borel_measurable (lebesgue_on S)" if "\T. open T \ {x \ S. f x \ T} \ sets (lebesgue_on S)" by (metis (no_types) eq borel_measurable_vimage_halfspace_component_lt open_halfspace_component_lt that) ultimately show "?thesis1" "?thesis2" by blast+ qed lemma borel_measurable_vimage_closed: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\T. closed T \ {x \ S. f x \ T} \ sets (lebesgue_on S))" (is "?lhs = ?rhs") proof - have eq: "{x \ S. f x \ T} = S - {x \ S. f x \ (- T)}" for T by auto show ?thesis apply (simp add: borel_measurable_vimage_open, safe) apply (simp_all (no_asm) add: eq) apply (intro sets.Diff sets_lebesgue_on_refl, force simp: closed_open) apply (intro sets.Diff sets_lebesgue_on_refl, force simp: open_closed) done qed lemma borel_measurable_vimage_closed_interval: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\a b. {x \ S. f x \ cbox a b} \ sets (lebesgue_on S))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using borel_measurable_vimage_closed by blast next assume RHS: ?rhs have "{x \ S. f x \ T} \ sets (lebesgue_on S)" if "open T" for T proof - obtain \ where "countable \" and \: "\ \ Pow T" "\X. X \ \ \ \a b. X = cbox a b" "\\ = T" using open_countable_Union_open_cbox that \open T\ by metis then have eq: "{x \ S. f x \ T} = (\U \ \. {x \ S. f x \ U})" by blast have "{x \ S. f x \ U} \ sets (lebesgue_on S)" if "U \ \" for U using that \ by (metis RHS) then show ?thesis by (auto simp: eq intro: Sigma_Algebra.sets.countable_UN' [OF \countable \\]) qed then show ?lhs by (simp add: borel_measurable_vimage_open) qed lemma borel_measurable_vimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets (lebesgue_on S))" (is "?lhs = ?rhs") proof assume f: ?lhs then show ?rhs using measurable_sets [OF f] by (simp add: Collect_conj_eq inf_sup_aci(1) space_restrict_space vimage_def) qed (simp add: borel_measurable_vimage_open_interval) lemma lebesgue_measurable_vimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "f \ borel_measurable lebesgue" "T \ sets borel" shows "{x. f x \ T} \ sets lebesgue" using assms borel_measurable_vimage_borel [of f UNIV] by auto lemma borel_measurable_lebesgue_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "f \ borel_measurable lebesgue \ (\T. T \ sets borel \ {x. f x \ T} \ sets lebesgue)" apply (intro iffI allI impI lebesgue_measurable_vimage_borel) apply (auto simp: in_borel_measurable_borel vimage_def) done subsection \<^marker>\tag unimportant\ \Measurability of continuous functions\ lemma continuous_imp_measurable_on_sets_lebesgue: assumes f: "continuous_on S f" and S: "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S)" proof - have "sets (restrict_space borel S) \ sets (lebesgue_on S)" by (simp add: mono_restrict_space subsetI) then show ?thesis by (simp add: borel_measurable_continuous_on_restrict [OF f] borel_measurable_subalgebra space_restrict_space) qed lemma id_borel_measurable_lebesgue [iff]: "id \ borel_measurable lebesgue" by (simp add: measurable_completion) lemma id_borel_measurable_lebesgue_on [iff]: "id \ borel_measurable (lebesgue_on S)" by (simp add: measurable_completion measurable_restrict_space1) context begin interpretation sigma_finite_measure "interval_measure (\x. x)" by (rule sigma_finite_interval_measure) auto interpretation finite_product_sigma_finite "\_. interval_measure (\x. x)" Basis proof qed simp lemma lborel_eq_real: "lborel = interval_measure (\x. x)" unfolding lborel_def Basis_real_def using distr_id[of "interval_measure (\x. x)"] by (subst distr_component[symmetric]) (simp_all add: distr_distr comp_def del: distr_id cong: distr_cong) lemma lborel_eq: "lborel = distr (\\<^sub>M b\Basis. lborel) borel (\f. \b\Basis. f b *\<^sub>R b)" by (subst lborel_def) (simp add: lborel_eq_real) lemma nn_integral_lborel_prod: assumes [measurable]: "\b. b \ Basis \ f b \ borel_measurable borel" assumes nn[simp]: "\b x. b \ Basis \ 0 \ f b x" shows "(\\<^sup>+x. (\b\Basis. f b (x \ b)) \lborel) = (\b\Basis. (\\<^sup>+x. f b x \lborel))" by (simp add: lborel_def nn_integral_distr product_nn_integral_prod product_nn_integral_singleton) lemma emeasure_lborel_Icc[simp]: fixes l u :: real assumes [simp]: "l \ u" shows "emeasure lborel {l .. u} = u - l" proof - have "((\f. f 1) -` {l..u} \ space (Pi\<^sub>M {1} (\b. interval_measure (\x. x)))) = {1::real} \\<^sub>E {l..u}" by (auto simp: space_PiM) then show ?thesis - by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc continuous_on_id) + by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc) qed lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \ u then u - l else 0)" by simp lemma\<^marker>\tag important\ emeasure_lborel_cbox[simp]: assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows "emeasure lborel (cbox l u) = (\b\Basis. (u - l) \ b)" proof - have "(\x. \b\Basis. indicator {l\b .. u\b} (x \ b) :: ennreal) = indicator (cbox l u)" by (auto simp: fun_eq_iff cbox_def split: split_indicator) then have "emeasure lborel (cbox l u) = (\\<^sup>+x. (\b\Basis. indicator {l\b .. u\b} (x \ b)) \lborel)" by simp also have "\ = (\b\Basis. (u - l) \ b)" by (subst nn_integral_lborel_prod) (simp_all add: prod_ennreal inner_diff_left) finally show ?thesis . qed lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \ c" using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c] by (auto simp add: power_0_left) lemma emeasure_lborel_Ioo[simp]: assumes [simp]: "l \ u" shows "emeasure lborel {l <..< u} = ennreal (u - l)" proof - have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}" using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto then show ?thesis by simp qed lemma emeasure_lborel_Ioc[simp]: assumes [simp]: "l \ u" shows "emeasure lborel {l <.. u} = ennreal (u - l)" proof - have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}" using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto then show ?thesis by simp qed lemma emeasure_lborel_Ico[simp]: assumes [simp]: "l \ u" shows "emeasure lborel {l ..< u} = ennreal (u - l)" proof - have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}" using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto then show ?thesis by simp qed lemma emeasure_lborel_box[simp]: assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows "emeasure lborel (box l u) = (\b\Basis. (u - l) \ b)" proof - have "(\x. \b\Basis. indicator {l\b <..< u\b} (x \ b) :: ennreal) = indicator (box l u)" by (auto simp: fun_eq_iff box_def split: split_indicator) then have "emeasure lborel (box l u) = (\\<^sup>+x. (\b\Basis. indicator {l\b <..< u\b} (x \ b)) \lborel)" by simp also have "\ = (\b\Basis. (u - l) \ b)" by (subst nn_integral_lborel_prod) (simp_all add: prod_ennreal inner_diff_left) finally show ?thesis . qed lemma emeasure_lborel_cbox_eq: "emeasure lborel (cbox l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le) lemma emeasure_lborel_box_eq: "emeasure lborel (box l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0" using emeasure_lborel_cbox[of x x] nonempty_Basis - by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: prod_constant) + by (auto simp del: emeasure_lborel_cbox nonempty_Basis) lemma fmeasurable_cbox [iff]: "cbox a b \ fmeasurable lborel" and fmeasurable_box [iff]: "box a b \ fmeasurable lborel" by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq) lemma fixes l u :: real assumes [simp]: "l \ u" shows measure_lborel_Icc[simp]: "measure lborel {l .. u} = u - l" and measure_lborel_Ico[simp]: "measure lborel {l ..< u} = u - l" and measure_lborel_Ioc[simp]: "measure lborel {l <.. u} = u - l" and measure_lborel_Ioo[simp]: "measure lborel {l <..< u} = u - l" by (simp_all add: measure_def) lemma assumes [simp]: "\b. b \ Basis \ l \ b \ u \ b" shows measure_lborel_box[simp]: "measure lborel (box l u) = (\b\Basis. (u - l) \ b)" and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\b\Basis. (u - l) \ b)" by (simp_all add: measure_def inner_diff_left prod_nonneg) lemma measure_lborel_cbox_eq: "measure lborel (cbox l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le) lemma measure_lborel_box_eq: "measure lborel (box l u) = (if \b\Basis. l \ b \ u \ b then \b\Basis. (u - l) \ b else 0)" using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force lemma measure_lborel_singleton[simp]: "measure lborel {x} = 0" by (simp add: measure_def) lemma sigma_finite_lborel: "sigma_finite_measure lborel" proof show "\A::'a set set. countable A \ A \ sets lborel \ \A = space lborel \ (\a\A. emeasure lborel a \ \)" by (intro exI[of _ "range (\n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One))"]) (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV) qed end lemma emeasure_lborel_UNIV [simp]: "emeasure lborel (UNIV::'a::euclidean_space set) = \" proof - { fix n::nat let ?Ba = "Basis :: 'a set" have "real n \ (2::real) ^ card ?Ba * real n" by (simp add: mult_le_cancel_right1) also have "... \ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" apply (rule mult_left_mono) apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc) - apply (simp add: DIM_positive) + apply (simp) done finally have "real n \ (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" . } note [intro!] = this show ?thesis unfolding UN_box_eq_UNIV[symmetric] apply (subst SUP_emeasure_incseq[symmetric]) - apply (auto simp: incseq_def subset_box inner_add_left prod_constant + apply (auto simp: incseq_def subset_box inner_add_left simp del: Sup_eq_top_iff SUP_eq_top_iff intro!: ennreal_SUP_eq_top) done qed lemma emeasure_lborel_countable: fixes A :: "'a::euclidean_space set" assumes "countable A" shows "emeasure lborel A = 0" proof - have "A \ (\i. {from_nat_into A i})" using from_nat_into_surj assms by force then have "emeasure lborel A \ emeasure lborel (\i. {from_nat_into A i})" by (intro emeasure_mono) auto also have "emeasure lborel (\i. {from_nat_into A i}) = 0" by (rule emeasure_UN_eq_0) auto finally show ?thesis by (auto simp add: ) qed lemma countable_imp_null_set_lborel: "countable A \ A \ null_sets lborel" by (simp add: null_sets_def emeasure_lborel_countable sets.countable) lemma finite_imp_null_set_lborel: "finite A \ A \ null_sets lborel" by (intro countable_imp_null_set_lborel countable_finite) lemma insert_null_sets_iff [simp]: "insert a N \ null_sets lebesgue \ N \ null_sets lebesgue" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson completion.complete2 subset_insertI) next assume ?rhs then show ?lhs by (simp add: null_sets.insert_in_sets null_setsI) qed lemma insert_null_sets_lebesgue_on_iff [simp]: assumes "a \ S" "S \ sets lebesgue" shows "insert a N \ null_sets (lebesgue_on S) \ N \ null_sets (lebesgue_on S)" by (simp add: assms null_sets_restrict_space) lemma lborel_neq_count_space[simp]: "lborel \ count_space (A::('a::ordered_euclidean_space) set)" proof assume asm: "lborel = count_space A" have "space lborel = UNIV" by simp hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space) have "emeasure lborel {undefined::'a} = 1" by (subst asm, subst emeasure_count_space_finite) auto moreover have "emeasure lborel {undefined} \ 1" by simp ultimately show False by contradiction qed lemma mem_closed_if_AE_lebesgue_open: assumes "open S" "closed C" assumes "AE x \ S in lebesgue. x \ C" assumes "x \ S" shows "x \ C" proof (rule ccontr) assume xC: "x \ C" with openE[of "S - C"] assms obtain e where e: "0 < e" "ball x e \ S - C" by blast then obtain a b where box: "x \ box a b" "box a b \ S - C" by (metis rational_boxes order_trans) then have "0 < emeasure lebesgue (box a b)" by (auto simp: emeasure_lborel_box_eq mem_box algebra_simps intro!: prod_pos) also have "\ \ emeasure lebesgue (S - C)" using assms box by (auto intro!: emeasure_mono) also have "\ = 0" using assms by (auto simp: eventually_ae_filter completion.complete2 set_diff_eq null_setsD1) finally show False by simp qed lemma mem_closed_if_AE_lebesgue: "closed C \ (AE x in lebesgue. x \ C) \ x \ C" using mem_closed_if_AE_lebesgue_open[OF open_UNIV] by simp subsection \Affine transformation on the Lebesgue-Borel\ lemma\<^marker>\tag important\ lborel_eqI: fixes M :: "'a::euclidean_space measure" assumes emeasure_eq: "\l u. (\b. b \ Basis \ l \ b \ u \ b) \ emeasure M (box l u) = (\b\Basis. (u - l) \ b)" assumes sets_eq: "sets M = sets borel" shows "lborel = M" proof (rule measure_eqI_generator_eq) let ?E = "range (\(a, b). box a b::'a set)" show "Int_stable ?E" by (auto simp: Int_stable_def box_Int_box) show "?E \ Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E" by (simp_all add: borel_eq_box sets_eq) let ?A = "\n::nat. box (- (real n *\<^sub>R One)) (real n *\<^sub>R One) :: 'a set" show "range ?A \ ?E" "(\i. ?A i) = UNIV" unfolding UN_box_eq_UNIV by auto { fix i show "emeasure lborel (?A i) \ \" by auto } { fix X assume "X \ ?E" then show "emeasure lborel X = emeasure M X" apply (auto simp: emeasure_eq emeasure_lborel_box_eq) apply (subst box_eq_empty(1)[THEN iffD2]) apply (auto intro: less_imp_le simp: not_le) done } qed lemma\<^marker>\tag important\ lborel_affine_euclidean: fixes c :: "'a::euclidean_space \ real" and t defines "T x \ t + (\j\Basis. (c j * (x \ j)) *\<^sub>R j)" assumes c: "\j. j \ Basis \ c j \ 0" shows "lborel = density (distr lborel borel T) (\_. (\j\Basis. \c j\))" (is "_ = ?D") proof (rule lborel_eqI) let ?B = "Basis :: 'a set" fix l u assume le: "\b. b \ ?B \ l \ b \ u \ b" have [measurable]: "T \ borel \\<^sub>M borel" by (simp add: T_def[abs_def]) have eq: "T -` box l u = box (\j\Basis. (((if 0 < c j then l - t else u - t) \ j) / c j) *\<^sub>R j) (\j\Basis. (((if 0 < c j then u - t else l - t) \ j) / c j) *\<^sub>R j)" using c by (auto simp: box_def T_def field_simps inner_simps divide_less_eq) with le c show "emeasure ?D (box l u) = (\b\?B. (u - l) \ b)" by (auto simp: emeasure_density emeasure_distr nn_integral_multc emeasure_lborel_box_eq inner_simps - field_simps field_split_simps ennreal_mult'[symmetric] prod_nonneg prod.distrib[symmetric] + field_split_simps ennreal_mult'[symmetric] prod_nonneg prod.distrib[symmetric] intro!: prod.cong) qed simp lemma lborel_affine: fixes t :: "'a::euclidean_space" shows "c \ 0 \ lborel = density (distr lborel borel (\x. t + c *\<^sub>R x)) (\_. \c\^DIM('a))" using lborel_affine_euclidean[where c="\_::'a. c" and t=t] unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation prod_constant by simp lemma lborel_real_affine: "c \ 0 \ lborel = density (distr lborel borel (\x. t + c * x)) (\_. ennreal (abs c))" using lborel_affine[of c t] by simp lemma AE_borel_affine: fixes P :: "real \ bool" shows "c \ 0 \ Measurable.pred borel P \ AE x in lborel. P x \ AE x in lborel. P (t + c * x)" by (subst lborel_real_affine[where t="- t / c" and c="1 / c"]) (simp_all add: AE_density AE_distr_iff field_simps) lemma nn_integral_real_affine: fixes c :: real assumes [measurable]: "f \ borel_measurable borel" and c: "c \ 0" shows "(\\<^sup>+x. f x \lborel) = \c\ * (\\<^sup>+x. f (t + c * x) \lborel)" by (subst lborel_real_affine[OF c, of t]) (simp add: nn_integral_density nn_integral_distr nn_integral_cmult) lemma lborel_integrable_real_affine: fixes f :: "real \ 'a :: {banach, second_countable_topology}" assumes f: "integrable lborel f" shows "c \ 0 \ integrable lborel (\x. f (t + c * x))" using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top) lemma lborel_integrable_real_affine_iff: fixes f :: "real \ 'a :: {banach, second_countable_topology}" shows "c \ 0 \ integrable lborel (\x. f (t + c * x)) \ integrable lborel f" using lborel_integrable_real_affine[of f c t] lborel_integrable_real_affine[of "\x. f (t + c * x)" "1/c" "-t/c"] by (auto simp add: field_simps) lemma\<^marker>\tag important\ lborel_integral_real_affine: fixes f :: "real \ 'a :: {banach, second_countable_topology}" and c :: real assumes c: "c \ 0" shows "(\x. f x \ lborel) = \c\ *\<^sub>R (\x. f (t + c * x) \lborel)" proof cases assume f[measurable]: "integrable lborel f" then show ?thesis using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t] by (subst lborel_real_affine[OF c, of t]) (simp add: integral_density integral_distr) next assume "\ integrable lborel f" with c show ?thesis by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq) qed lemma fixes c :: "'a::euclidean_space \ real" and t assumes c: "\j. j \ Basis \ c j \ 0" defines "T == (\x. t + (\j\Basis. (c j * (x \ j)) *\<^sub>R j))" shows lebesgue_affine_euclidean: "lebesgue = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" (is "_ = ?D") and lebesgue_affine_measurable: "T \ lebesgue \\<^sub>M lebesgue" proof - have T_borel[measurable]: "T \ borel \\<^sub>M borel" by (auto simp: T_def[abs_def]) { fix A :: "'a set" assume A: "A \ sets borel" then have "emeasure lborel A = 0 \ emeasure (density (distr lborel borel T) (\_. (\j\Basis. \c j\))) A = 0" unfolding T_def using c by (subst lborel_affine_euclidean[symmetric]) auto also have "\ \ emeasure (distr lebesgue lborel T) A = 0" using A c by (simp add: distr_completion emeasure_density nn_integral_cmult prod_nonneg cong: distr_cong) finally have "emeasure lborel A = 0 \ emeasure (distr lebesgue lborel T) A = 0" . } then have eq: "null_sets lborel = null_sets (distr lebesgue lborel T)" by (auto simp: null_sets_def) show "T \ lebesgue \\<^sub>M lebesgue" by (rule completion.measurable_completion2) (auto simp: eq measurable_completion) have "lebesgue = completion (density (distr lborel borel T) (\_. (\j\Basis. \c j\)))" using c by (subst lborel_affine_euclidean[of c t]) (simp_all add: T_def[abs_def]) also have "\ = density (completion (distr lebesgue lborel T)) (\_. (\j\Basis. \c j\))" using c by (auto intro!: always_eventually prod_pos completion_density_eq simp: distr_completion cong: distr_cong) also have "\ = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" by (subst completion.completion_distr_eq) (auto simp: eq measurable_completion) finally show "lebesgue = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" . qed corollary lebesgue_real_affine: "c \ 0 \ lebesgue = density (distr lebesgue lebesgue (\x. t + c * x)) (\_. ennreal (abs c))" using lebesgue_affine_euclidean [where c= "\x::real. c"] by simp lemma nn_integral_real_affine_lebesgue: fixes c :: real assumes f[measurable]: "f \ borel_measurable lebesgue" and c: "c \ 0" shows "(\\<^sup>+x. f x \lebesgue) = ennreal\c\ * (\\<^sup>+x. f(t + c * x) \lebesgue)" proof - have "(\\<^sup>+x. f x \lebesgue) = (\\<^sup>+x. f x \density (distr lebesgue lebesgue (\x. t + c * x)) (\x. ennreal \c\))" using lebesgue_real_affine c by auto also have "\ = \\<^sup>+ x. ennreal \c\ * f x \distr lebesgue lebesgue (\x. t + c * x)" by (subst nn_integral_density) auto also have "\ = ennreal \c\ * integral\<^sup>N (distr lebesgue lebesgue (\x. t + c * x)) f" using f measurable_distr_eq1 nn_integral_cmult by blast also have "\ = \c\ * (\\<^sup>+x. f(t + c * x) \lebesgue)" using lebesgue_affine_measurable[where c= "\x::real. c"] by (subst nn_integral_distr) (force+) finally show ?thesis . qed lemma lebesgue_measurable_scaling[measurable]: "(*\<^sub>R) x \ lebesgue \\<^sub>M lebesgue" proof cases assume "x = 0" then have "(*\<^sub>R) x = (\x. 0::'a)" by (auto simp: fun_eq_iff) then show ?thesis by auto next assume "x \ 0" then show ?thesis using lebesgue_affine_measurable[of "\_. x" 0] unfolding scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] euclidean_representation by (auto simp add: ac_simps) qed lemma fixes m :: real and \ :: "'a::euclidean_space" defines "T r d x \ r *\<^sub>R x + d" shows emeasure_lebesgue_affine: "emeasure lebesgue (T m \ ` S) = \m\ ^ DIM('a) * emeasure lebesgue S" (is ?e) and measure_lebesgue_affine: "measure lebesgue (T m \ ` S) = \m\ ^ DIM('a) * measure lebesgue S" (is ?m) proof - show ?e proof cases assume "m = 0" then show ?thesis by (simp add: image_constant_conv T_def[abs_def]) next let ?T = "T m \" and ?T' = "T (1 / m) (- ((1/m) *\<^sub>R \))" assume "m \ 0" then have s_comp_s: "?T' \ ?T = id" "?T \ ?T' = id" by (auto simp: T_def[abs_def] fun_eq_iff scaleR_add_right scaleR_diff_right) then have "inv ?T' = ?T" "bij ?T'" by (auto intro: inv_unique_comp o_bij) then have eq: "T m \ ` S = T (1 / m) ((-1/m) *\<^sub>R \) -` S \ space lebesgue" using bij_vimage_eq_inv_image[OF \bij ?T'\, of S] by auto have trans_eq_T: "(\x. \ + (\j\Basis. (m * (x \ j)) *\<^sub>R j)) = T m \" for m \ unfolding T_def[abs_def] scaleR_scaleR[symmetric] scaleR_sum_right[symmetric] by (auto simp add: euclidean_representation ac_simps) have T[measurable]: "T r d \ lebesgue \\<^sub>M lebesgue" for r d using lebesgue_affine_measurable[of "\_. r" d] by (cases "r = 0") (auto simp: trans_eq_T T_def[abs_def]) show ?thesis proof cases assume "S \ sets lebesgue" with \m \ 0\ show ?thesis unfolding eq apply (subst lebesgue_affine_euclidean[of "\_. m" \]) apply (simp_all add: emeasure_density trans_eq_T nn_integral_cmult emeasure_distr del: space_completion emeasure_completion) - apply (simp add: vimage_comp s_comp_s prod_constant) + apply (simp add: vimage_comp s_comp_s) done next assume "S \ sets lebesgue" moreover have "?T ` S \ sets lebesgue" proof assume "?T ` S \ sets lebesgue" then have "?T -` (?T ` S) \ space lebesgue \ sets lebesgue" by (rule measurable_sets[OF T]) also have "?T -` (?T ` S) \ space lebesgue = S" by (simp add: vimage_comp s_comp_s eq) finally show False using \S \ sets lebesgue\ by auto qed ultimately show ?thesis by (simp add: emeasure_notin_sets) qed qed show ?m unfolding measure_def \?e\ by (simp add: enn2real_mult prod_nonneg) qed lemma lebesgue_real_scale: assumes "c \ 0" shows "lebesgue = density (distr lebesgue lebesgue (\x. c * x)) (\x. ennreal \c\)" using assms by (subst lebesgue_affine_euclidean[of "\_. c" 0]) simp_all lemma divideR_right: fixes x y :: "'a::real_normed_vector" shows "r \ 0 \ y = x /\<^sub>R r \ r *\<^sub>R y = x" using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp lemma lborel_has_bochner_integral_real_affine_iff: fixes x :: "'a :: {banach, second_countable_topology}" shows "c \ 0 \ has_bochner_integral lborel f x \ has_bochner_integral lborel (\x. f (t + c * x)) (x /\<^sub>R \c\)" unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong) lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)" by (subst lborel_real_affine[of "-1" 0]) (auto simp: density_1 one_ennreal_def[symmetric]) lemma lborel_distr_mult: assumes "(c::real) \ 0" shows "distr lborel borel ((*) c) = density lborel (\_. inverse \c\)" proof- have "distr lborel borel ((*) c) = distr lborel lborel ((*) c)" by (simp cong: distr_cong) also from assms have "... = density lborel (\_. inverse \c\)" by (subst lborel_real_affine[of "inverse c" 0]) (auto simp: o_def distr_density_distr) finally show ?thesis . qed lemma lborel_distr_mult': assumes "(c::real) \ 0" shows "lborel = density (distr lborel borel ((*) c)) (\_. \c\)" proof- have "lborel = density lborel (\_. 1)" by (rule density_1[symmetric]) also from assms have "(\_. 1 :: ennreal) = (\_. inverse \c\ * \c\)" by (intro ext) simp also have "density lborel ... = density (density lborel (\_. inverse \c\)) (\_. \c\)" by (subst density_density_eq) (auto simp: ennreal_mult) also from assms have "density lborel (\_. inverse \c\) = distr lborel borel ((*) c)" by (rule lborel_distr_mult[symmetric]) finally show ?thesis . qed lemma lborel_distr_plus: fixes c :: "'a::euclidean_space" shows "distr lborel borel ((+) c) = lborel" by (subst lborel_affine[of 1 c], auto simp: density_1) interpretation lborel: sigma_finite_measure lborel by (rule sigma_finite_lborel) interpretation lborel_pair: pair_sigma_finite lborel lborel .. lemma\<^marker>\tag important\ lborel_prod: "lborel \\<^sub>M lborel = (lborel :: ('a::euclidean_space \ 'b::euclidean_space) measure)" proof (rule lborel_eqI[symmetric], clarify) fix la ua :: 'a and lb ub :: 'b assume lu: "\a b. (a, b) \ Basis \ (la, lb) \ (a, b) \ (ua, ub) \ (a, b)" have [simp]: "\b. b \ Basis \ la \ b \ ua \ b" "\b. b \ Basis \ lb \ b \ ub \ b" "inj_on (\u. (u, 0)) Basis" "inj_on (\u. (0, u)) Basis" "(\u. (u, 0)) ` Basis \ (\u. (0, u)) ` Basis = {}" "box (la, lb) (ua, ub) = box la ua \ box lb ub" using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def) show "emeasure (lborel \\<^sub>M lborel) (box (la, lb) (ua, ub)) = ennreal (prod ((\) ((ua, ub) - (la, lb))) Basis)" by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def prod.union_disjoint prod.reindex ennreal_mult inner_diff_left prod_nonneg) qed (simp add: borel_prod[symmetric]) (* FIXME: conversion in measurable prover *) lemma lborelD_Collect[measurable (raw)]: "{x\space borel. P x} \ sets borel \ {x\space lborel. P x} \ sets lborel" by simp lemma lborelD[measurable (raw)]: "A \ sets borel \ A \ sets lborel" by simp lemma emeasure_bounded_finite: assumes "bounded A" shows "emeasure lborel A < \" proof - obtain a b where "A \ cbox a b" by (meson bounded_subset_cbox_symmetric \bounded A\) then have "emeasure lborel A \ emeasure lborel (cbox a b)" by (intro emeasure_mono) auto then show ?thesis by (auto simp: emeasure_lborel_cbox_eq prod_nonneg less_top[symmetric] top_unique split: if_split_asm) qed lemma emeasure_compact_finite: "compact A \ emeasure lborel A < \" using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded) lemma borel_integrable_compact: fixes f :: "'a::euclidean_space \ 'b::{banach, second_countable_topology}" assumes "compact S" "continuous_on S f" shows "integrable lborel (\x. indicator S x *\<^sub>R f x)" proof cases assume "S \ {}" have "continuous_on S (\x. norm (f x))" using assms by (intro continuous_intros) from continuous_attains_sup[OF \compact S\ \S \ {}\ this] obtain M where M: "\x. x \ S \ norm (f x) \ M" by auto show ?thesis proof (rule integrable_bound) show "integrable lborel (\x. indicator S x * M)" using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left) show "(\x. indicator S x *\<^sub>R f x) \ borel_measurable lborel" using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact) show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \ norm (indicator S x * M)" by (auto split: split_indicator simp: abs_real_def dest!: M) qed qed simp lemma borel_integrable_atLeastAtMost: fixes f :: "real \ real" assumes f: "\x. a \ x \ x \ b \ isCont f x" shows "integrable lborel (\x. f x * indicator {a .. b} x)" (is "integrable _ ?f") proof - have "integrable lborel (\x. indicator {a .. b} x *\<^sub>R f x)" proof (rule borel_integrable_compact) from f show "continuous_on {a..b} f" by (auto intro: continuous_at_imp_continuous_on) qed simp then show ?thesis by (auto simp: mult.commute) qed subsection \Lebesgue measurable sets\ abbreviation\<^marker>\tag important\ lmeasurable :: "'a::euclidean_space set set" where "lmeasurable \ fmeasurable lebesgue" lemma not_measurable_UNIV [simp]: "UNIV \ lmeasurable" by (simp add: fmeasurable_def) lemma\<^marker>\tag important\ lmeasurable_iff_integrable: "S \ lmeasurable \ integrable lebesgue (indicator S :: 'a::euclidean_space \ real)" by (auto simp: fmeasurable_def integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator) lemma lmeasurable_cbox [iff]: "cbox a b \ lmeasurable" and lmeasurable_box [iff]: "box a b \ lmeasurable" by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq) lemma fixes a::real shows lmeasurable_interval [iff]: "{a..b} \ lmeasurable" "{a<.. lmeasurable" apply (metis box_real(2) lmeasurable_cbox) by (metis box_real(1) lmeasurable_box) lemma fmeasurable_compact: "compact S \ S \ fmeasurable lborel" using emeasure_compact_finite[of S] by (intro fmeasurableI) (auto simp: borel_compact) lemma lmeasurable_compact: "compact S \ S \ lmeasurable" using fmeasurable_compact by (force simp: fmeasurable_def) lemma measure_frontier: "bounded S \ measure lebesgue (frontier S) = measure lebesgue (closure S) - measure lebesgue (interior S)" using closure_subset interior_subset by (auto simp: frontier_def fmeasurable_compact intro!: measurable_measure_Diff) lemma lmeasurable_closure: "bounded S \ closure S \ lmeasurable" by (simp add: lmeasurable_compact) lemma lmeasurable_frontier: "bounded S \ frontier S \ lmeasurable" by (simp add: compact_frontier_bounded lmeasurable_compact) lemma lmeasurable_open: "bounded S \ open S \ S \ lmeasurable" using emeasure_bounded_finite[of S] by (intro fmeasurableI) (auto simp: borel_open) lemma lmeasurable_ball [simp]: "ball a r \ lmeasurable" by (simp add: lmeasurable_open) lemma lmeasurable_cball [simp]: "cball a r \ lmeasurable" by (simp add: lmeasurable_compact) lemma lmeasurable_interior: "bounded S \ interior S \ lmeasurable" by (simp add: bounded_interior lmeasurable_open) lemma null_sets_cbox_Diff_box: "cbox a b - box a b \ null_sets lborel" proof - have "emeasure lborel (cbox a b - box a b) = 0" by (subst emeasure_Diff) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq box_subset_cbox) then have "cbox a b - box a b \ null_sets lborel" by (auto simp: null_sets_def) then show ?thesis by (auto dest!: AE_not_in) qed lemma bounded_set_imp_lmeasurable: assumes "bounded S" "S \ sets lebesgue" shows "S \ lmeasurable" by (metis assms bounded_Un emeasure_bounded_finite emeasure_completion fmeasurableI main_part_null_part_Un) lemma finite_measure_lebesgue_on: "S \ lmeasurable \ finite_measure (lebesgue_on S)" by (auto simp: finite_measureI fmeasurable_def emeasure_restrict_space) lemma integrable_const_ivl [iff]: fixes a::"'a::ordered_euclidean_space" shows "integrable (lebesgue_on {a..b}) (\x. c)" by (metis cbox_interval finite_measure.integrable_const finite_measure_lebesgue_on lmeasurable_cbox) subsection\<^marker>\tag unimportant\\Translation preserves Lebesgue measure\ lemma sigma_sets_image: assumes S: "S \ sigma_sets \ M" and "M \ Pow \" "f ` \ = \" "inj_on f \" and M: "\y. y \ M \ f ` y \ M" shows "(f ` S) \ sigma_sets \ M" using S proof (induct S rule: sigma_sets.induct) case (Basic a) then show ?case by (simp add: M) next case Empty then show ?case by (simp add: sigma_sets.Empty) next case (Compl a) then have "\ - a \ \" "a \ \" by (auto simp: sigma_sets_into_sp [OF \M \ Pow \\]) then show ?case by (auto simp: inj_on_image_set_diff [OF \inj_on f \\] assms intro: Compl sigma_sets.Compl) next case (Union a) then show ?case by (metis image_UN sigma_sets.simps) qed lemma null_sets_translation: assumes "N \ null_sets lborel" shows "{x. x - a \ N} \ null_sets lborel" proof - have [simp]: "(\x. x + a) ` N = {x. x - a \ N}" by force show ?thesis using assms emeasure_lebesgue_affine [of 1 a N] by (auto simp: null_sets_def) qed lemma lebesgue_sets_translation: fixes f :: "'a \ 'a::euclidean_space" assumes S: "S \ sets lebesgue" shows "((\x. a + x) ` S) \ sets lebesgue" proof - have im_eq: "(+) a ` A = {x. x - a \ A}" for A by force have "((\x. a + x) ` S) = ((\x. -a + x) -` S) \ (space lebesgue)" using image_iff by fastforce also have "\ \ sets lebesgue" proof (rule measurable_sets [OF measurableI assms]) fix A :: "'b set" assume A: "A \ sets lebesgue" have vim_eq: "(\x. x - a) -` A = (+) a ` A" for A by force have "\s n N'. (+) a ` (S \ N) = s \ n \ s \ sets borel \ N' \ null_sets lborel \ n \ N'" if "S \ sets borel" and "N' \ null_sets lborel" and "N \ N'" for S N N' proof (intro exI conjI) show "(+) a ` (S \ N) = (\x. a + x) ` S \ (\x. a + x) ` N" by auto show "(\x. a + x) ` N' \ null_sets lborel" using that by (auto simp: null_sets_translation im_eq) qed (use that im_eq in auto) with A have "(\x. x - a) -` A \ sets lebesgue" by (force simp: vim_eq completion_def intro!: sigma_sets_image) then show "(+) (- a) -` A \ space lebesgue \ sets lebesgue" by (auto simp: vimage_def im_eq) qed auto finally show ?thesis . qed lemma measurable_translation: "S \ lmeasurable \ ((+) a ` S) \ lmeasurable" using emeasure_lebesgue_affine [of 1 a S] apply (auto intro: lebesgue_sets_translation simp add: fmeasurable_def cong: image_cong_simp) apply (simp add: ac_simps) done lemma measurable_translation_subtract: "S \ lmeasurable \ ((\x. x - a) ` S) \ lmeasurable" using measurable_translation [of S "- a"] by (simp cong: image_cong_simp) lemma measure_translation: "measure lebesgue ((+) a ` S) = measure lebesgue S" using measure_lebesgue_affine [of 1 a S] by (simp add: ac_simps cong: image_cong_simp) lemma measure_translation_subtract: "measure lebesgue ((\x. x - a) ` S) = measure lebesgue S" using measure_translation [of "- a"] by (simp cong: image_cong_simp) subsection \A nice lemma for negligibility proofs\ lemma summable_iff_suminf_neq_top: "(\n. f n \ 0) \ \ summable f \ (\i. ennreal (f i)) = top" by (metis summable_suminf_not_top) proposition\<^marker>\tag important\ starlike_negligible_bounded_gmeasurable: fixes S :: "'a :: euclidean_space set" assumes S: "S \ sets lebesgue" and "bounded S" and eq1: "\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1" shows "S \ null_sets lebesgue" proof - obtain M where "0 < M" "S \ ball 0 M" using \bounded S\ by (auto dest: bounded_subset_ballD) let ?f = "\n. root DIM('a) (Suc n)" have vimage_eq_image: "(*\<^sub>R) (?f n) -` S = (*\<^sub>R) (1 / ?f n) ` S" for n apply safe subgoal for x by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto subgoal by auto done have eq: "(1 / ?f n) ^ DIM('a) = 1 / Suc n" for n by (simp add: field_simps) { fix n x assume x: "root DIM('a) (1 + real n) *\<^sub>R x \ S" have "1 * norm x \ root DIM('a) (1 + real n) * norm x" by (rule mult_mono) auto also have "\ < M" using x \S \ ball 0 M\ by auto finally have "norm x < M" by simp } note less_M = this have "(\n. ennreal (1 / Suc n)) = top" using not_summable_harmonic[where 'a=real] summable_Suc_iff[where f="\n. 1 / (real n)"] by (intro summable_iff_suminf_neq_top) (auto simp add: inverse_eq_divide) then have "top * emeasure lebesgue S = (\n. (1 / ?f n)^DIM('a) * emeasure lebesgue S)" unfolding ennreal_suminf_multc eq by simp also have "\ = (\n. emeasure lebesgue ((*\<^sub>R) (?f n) -` S))" unfolding vimage_eq_image using emeasure_lebesgue_affine[of "1 / ?f n" 0 S for n] by simp also have "\ = emeasure lebesgue (\n. (*\<^sub>R) (?f n) -` S)" proof (intro suminf_emeasure) show "disjoint_family (\n. (*\<^sub>R) (?f n) -` S)" unfolding disjoint_family_on_def proof safe fix m n :: nat and x assume "m \ n" "?f m *\<^sub>R x \ S" "?f n *\<^sub>R x \ S" with eq1[of "?f m / ?f n" "?f n *\<^sub>R x"] show "x \ {}" by auto qed have "(*\<^sub>R) (?f i) -` S \ sets lebesgue" for i using measurable_sets[OF lebesgue_measurable_scaling[of "?f i"] S] by auto then show "range (\i. (*\<^sub>R) (?f i) -` S) \ sets lebesgue" by auto qed also have "\ \ emeasure lebesgue (ball 0 M :: 'a set)" using less_M by (intro emeasure_mono) auto also have "\ < top" using lmeasurable_ball by (auto simp: fmeasurable_def) finally have "emeasure lebesgue S = 0" by (simp add: ennreal_top_mult split: if_split_asm) then show "S \ null_sets lebesgue" unfolding null_sets_def using \S \ sets lebesgue\ by auto qed corollary starlike_negligible_compact: "compact S \ (\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1) \ S \ null_sets lebesgue" using starlike_negligible_bounded_gmeasurable[of S] by (auto simp: compact_eq_bounded_closed) proposition outer_regular_lborel_le: assumes B[measurable]: "B \ sets borel" and "0 < (e::real)" obtains U where "open U" "B \ U" and "emeasure lborel (U - B) \ e" proof - let ?\ = "emeasure lborel" let ?B = "\n::nat. ball 0 n :: 'a set" let ?e = "\n. e*((1/2)^Suc n)" have "\n. \U. open U \ ?B n \ B \ U \ ?\ (U - B) < ?e n" proof fix n :: nat let ?A = "density lborel (indicator (?B n))" have emeasure_A: "X \ sets borel \ emeasure ?A X = ?\ (?B n \ X)" for X by (auto simp: emeasure_density borel_measurable_indicator indicator_inter_arith[symmetric]) have finite_A: "emeasure ?A (space ?A) \ \" using emeasure_bounded_finite[of "?B n"] by (auto simp: emeasure_A) interpret A: finite_measure ?A by rule fact have "emeasure ?A B + ?e n > (INF U\{U. B \ U \ open U}. emeasure ?A U)" using \0 by (auto simp: outer_regular[OF _ finite_A B, symmetric]) then obtain U where U: "B \ U" "open U" and muU: "?\ (?B n \ B) + ?e n > ?\ (?B n \ U)" unfolding INF_less_iff by (auto simp: emeasure_A) moreover { have "?\ ((?B n \ U) - B) = ?\ ((?B n \ U) - (?B n \ B))" using U by (intro arg_cong[where f="?\"]) auto also have "\ = ?\ (?B n \ U) - ?\ (?B n \ B)" using U A.emeasure_finite[of B] by (intro emeasure_Diff) (auto simp del: A.emeasure_finite simp: emeasure_A) also have "\ < ?e n" using U muU A.emeasure_finite[of B] by (subst minus_less_iff_ennreal) (auto simp del: A.emeasure_finite simp: emeasure_A less_top ac_simps intro!: emeasure_mono) finally have "?\ ((?B n \ U) - B) < ?e n" . } ultimately show "\U. open U \ ?B n \ B \ U \ ?\ (U - B) < ?e n" by (intro exI[of _ "?B n \ U"]) auto qed then obtain U where U: "\n. open (U n)" "\n. ?B n \ B \ U n" "\n. ?\ (U n - B) < ?e n" by metis show ?thesis proof { fix x assume "x \ B" moreover obtain n where "norm x < real n" using reals_Archimedean2 by blast ultimately have "x \ (\n. U n)" using U(2)[of n] by auto } note * = this then show "open (\n. U n)" "B \ (\n. U n)" using U by auto have "?\ (\n. U n - B) \ (\n. ?\ (U n - B))" using U(1) by (intro emeasure_subadditive_countably) auto also have "\ \ (\n. ennreal (?e n))" using U(3) by (intro suminf_le) (auto intro: less_imp_le) also have "\ = ennreal (e * 1)" using \0 by (intro suminf_ennreal_eq sums_mult power_half_series) auto finally show "emeasure lborel ((\n. U n) - B) \ ennreal e" by simp qed qed lemma\<^marker>\tag important\ outer_regular_lborel: assumes B: "B \ sets borel" and "0 < (e::real)" obtains U where "open U" "B \ U" "emeasure lborel (U - B) < e" proof - obtain U where U: "open U" "B \ U" and "emeasure lborel (U-B) \ e/2" using outer_regular_lborel_le [OF B, of "e/2"] \e > 0\ by force moreover have "ennreal (e/2) < ennreal e" using \e > 0\ by (simp add: ennreal_lessI) ultimately have "emeasure lborel (U-B) < e" by auto with U show ?thesis using that by auto qed lemma completion_upper: assumes A: "A \ sets (completion M)" obtains A' where "A \ A'" "A' \ sets M" "A' - A \ null_sets (completion M)" "emeasure (completion M) A = emeasure M A'" proof - from AE_notin_null_part[OF A] obtain N where N: "N \ null_sets M" "null_part M A \ N" unfolding eventually_ae_filter using null_part_null_sets[OF A, THEN null_setsD2, THEN sets.sets_into_space] by auto let ?A' = "main_part M A \ N" show ?thesis proof show "A \ ?A'" using \null_part M A \ N\ by (subst main_part_null_part_Un[symmetric, OF A]) auto have "main_part M A \ A" using assms main_part_null_part_Un by auto then have "?A' - A \ N" by blast with N show "?A' - A \ null_sets (completion M)" by (blast intro: null_sets_completionI completion.complete_measure_axioms complete_measure.complete2) show "emeasure (completion M) A = emeasure M (main_part M A \ N)" using A \N \ null_sets M\ by (simp add: emeasure_Un_null_set) qed (use A N in auto) qed lemma sets_lebesgue_outer_open: fixes e::real assumes S: "S \ sets lebesgue" and "e > 0" obtains T where "open T" "S \ T" "(T - S) \ lmeasurable" "emeasure lebesgue (T - S) < ennreal e" proof - obtain S' where S': "S \ S'" "S' \ sets borel" and null: "S' - S \ null_sets lebesgue" and em: "emeasure lebesgue S = emeasure lborel S'" using completion_upper[of S lborel] S by auto then have f_S': "S' \ sets borel" using S by (auto simp: fmeasurable_def) with outer_regular_lborel[OF _ \0] obtain U where U: "open U" "S' \ U" "emeasure lborel (U - S') < e" by blast show thesis proof show "open U" "S \ U" using f_S' U S' by auto have "(U - S) = (U - S') \ (S' - S)" using S' U by auto then have eq: "emeasure lebesgue (U - S) = emeasure lborel (U - S')" using null by (simp add: U(1) emeasure_Un_null_set f_S' sets.Diff) have "(U - S) \ sets lebesgue" by (simp add: S U(1) sets.Diff) then show "(U - S) \ lmeasurable" unfolding fmeasurable_def using U(3) eq less_le_trans by fastforce with eq U show "emeasure lebesgue (U - S) < e" by (simp add: eq) qed qed lemma sets_lebesgue_inner_closed: fixes e::real assumes "S \ sets lebesgue" "e > 0" obtains T where "closed T" "T \ S" "S-T \ lmeasurable" "emeasure lebesgue (S - T) < ennreal e" proof - have "-S \ sets lebesgue" using assms by (simp add: Compl_in_sets_lebesgue) then obtain T where "open T" "-S \ T" and T: "(T - -S) \ lmeasurable" "emeasure lebesgue (T - -S) < e" using sets_lebesgue_outer_open assms by blast show thesis proof show "closed (-T)" using \open T\ by blast show "-T \ S" using \- S \ T\ by auto show "S - ( -T) \ lmeasurable" "emeasure lebesgue (S - (- T)) < e" using T by (auto simp: Int_commute) qed qed lemma lebesgue_openin: "\openin (top_of_set S) T; S \ sets lebesgue\ \ T \ sets lebesgue" by (metis borel_open openin_open sets.Int sets_completionI_sets sets_lborel) lemma lebesgue_closedin: "\closedin (top_of_set S) T; S \ sets lebesgue\ \ T \ sets lebesgue" by (metis borel_closed closedin_closed sets.Int sets_completionI_sets sets_lborel) subsection\\F_sigma\ and \G_delta\ sets.\ \ \\<^url>\https://en.wikipedia.org/wiki/F-sigma_set\\ inductive\<^marker>\tag important\ fsigma :: "'a::topological_space set \ bool" where "(\n::nat. closed (F n)) \ fsigma (\(F ` UNIV))" inductive\<^marker>\tag important\ gdelta :: "'a::topological_space set \ bool" where "(\n::nat. open (F n)) \ gdelta (\(F ` UNIV))" lemma fsigma_Union_compact: fixes S :: "'a::{real_normed_vector,heine_borel} set" shows "fsigma S \ (\F::nat \ 'a set. range F \ Collect compact \ S = \(F ` UNIV))" proof safe assume "fsigma S" then obtain F :: "nat \ 'a set" where F: "range F \ Collect closed" "S = \(F ` UNIV)" by (meson fsigma.cases image_subsetI mem_Collect_eq) then have "\D::nat \ 'a set. range D \ Collect compact \ \(D ` UNIV) = F i" for i using closed_Union_compact_subsets [of "F i"] by (metis image_subsetI mem_Collect_eq range_subsetD) then obtain D :: "nat \ nat \ 'a set" where D: "\i. range (D i) \ Collect compact \ \((D i) ` UNIV) = F i" by metis let ?DD = "\n. (\(i,j). D i j) (prod_decode n)" show "\F::nat \ 'a set. range F \ Collect compact \ S = \(F ` UNIV)" proof (intro exI conjI) show "range ?DD \ Collect compact" using D by clarsimp (metis mem_Collect_eq rangeI split_conv subsetCE surj_pair) show "S = \ (range ?DD)" proof show "S \ \ (range ?DD)" using D F by clarsimp (metis UN_iff old.prod.case prod_decode_inverse prod_encode_eq) show "\ (range ?DD) \ S" using D F by fastforce qed qed next fix F :: "nat \ 'a set" assume "range F \ Collect compact" and "S = \(F ` UNIV)" then show "fsigma (\(F ` UNIV))" by (simp add: compact_imp_closed fsigma.intros image_subset_iff) qed lemma gdelta_imp_fsigma: "gdelta S \ fsigma (- S)" proof (induction rule: gdelta.induct) case (1 F) have "- \(F ` UNIV) = (\i. -(F i))" by auto then show ?case by (simp add: fsigma.intros closed_Compl 1) qed lemma fsigma_imp_gdelta: "fsigma S \ gdelta (- S)" proof (induction rule: fsigma.induct) case (1 F) have "- \(F ` UNIV) = (\i. -(F i))" by auto then show ?case by (simp add: 1 gdelta.intros open_closed) qed lemma gdelta_complement: "gdelta(- S) \ fsigma S" using fsigma_imp_gdelta gdelta_imp_fsigma by force lemma lebesgue_set_almost_fsigma: assumes "S \ sets lebesgue" obtains C T where "fsigma C" "T \ null_sets lebesgue" "C \ T = S" "disjnt C T" proof - { fix n::nat obtain T where "closed T" "T \ S" "S-T \ lmeasurable" "emeasure lebesgue (S - T) < ennreal (1 / Suc n)" using sets_lebesgue_inner_closed [OF assms] by (metis of_nat_0_less_iff zero_less_Suc zero_less_divide_1_iff) then have "\T. closed T \ T \ S \ S - T \ lmeasurable \ measure lebesgue (S-T) < 1 / Suc n" by (metis emeasure_eq_measure2 ennreal_leI not_le) } then obtain F where F: "\n::nat. closed (F n) \ F n \ S \ S - F n \ lmeasurable \ measure lebesgue (S - F n) < 1 / Suc n" by metis let ?C = "\(F ` UNIV)" show thesis proof show "fsigma ?C" using F by (simp add: fsigma.intros) show "(S - ?C) \ null_sets lebesgue" proof (clarsimp simp add: completion.null_sets_outer_le) fix e :: "real" assume "0 < e" then obtain n where n: "1 / Suc n < e" using nat_approx_posE by metis show "\T \ lmeasurable. S - (\x. F x) \ T \ measure lebesgue T \ e" proof (intro bexI conjI) show "measure lebesgue (S - F n) \ e" by (meson F n less_trans not_le order.asym) qed (use F in auto) qed show "?C \ (S - ?C) = S" using F by blast show "disjnt ?C (S - ?C)" by (auto simp: disjnt_def) qed qed lemma lebesgue_set_almost_gdelta: assumes "S \ sets lebesgue" obtains C T where "gdelta C" "T \ null_sets lebesgue" "S \ T = C" "disjnt S T" proof - have "-S \ sets lebesgue" using assms Compl_in_sets_lebesgue by blast then obtain C T where C: "fsigma C" "T \ null_sets lebesgue" "C \ T = -S" "disjnt C T" using lebesgue_set_almost_fsigma by metis show thesis proof show "gdelta (-C)" by (simp add: \fsigma C\ fsigma_imp_gdelta) show "S \ T = -C" "disjnt S T" using C by (auto simp: disjnt_def) qed (use C in auto) qed end diff --git a/src/HOL/Analysis/Locally.thy b/src/HOL/Analysis/Locally.thy --- a/src/HOL/Analysis/Locally.thy +++ b/src/HOL/Analysis/Locally.thy @@ -1,451 +1,451 @@ section \Neighbourhood bases and Locally path-connected spaces\ theory Locally imports Path_Connected Function_Topology begin subsection\Neighbourhood Bases\ text \Useful for "local" properties of various kinds\ definition neighbourhood_base_at where "neighbourhood_base_at x P X \ \W. openin X W \ x \ W \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W)" definition neighbourhood_base_of where "neighbourhood_base_of P X \ \x \ topspace X. neighbourhood_base_at x P X" lemma neighbourhood_base_of: "neighbourhood_base_of P X \ (\W x. openin X W \ x \ W \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W))" unfolding neighbourhood_base_at_def neighbourhood_base_of_def using openin_subset by blast lemma neighbourhood_base_at_mono: "\neighbourhood_base_at x P X; \S. \P S; x \ S\ \ Q S\ \ neighbourhood_base_at x Q X" unfolding neighbourhood_base_at_def by (meson subset_eq) lemma neighbourhood_base_of_mono: "\neighbourhood_base_of P X; \S. P S \ Q S\ \ neighbourhood_base_of Q X" unfolding neighbourhood_base_of_def using neighbourhood_base_at_mono by force lemma open_neighbourhood_base_at: "(\S. \P S; x \ S\ \ openin X S) \ neighbourhood_base_at x P X \ (\W. openin X W \ x \ W \ (\U. P U \ x \ U \ U \ W))" unfolding neighbourhood_base_at_def by (meson subsetD) lemma open_neighbourhood_base_of: "(\S. P S \ openin X S) \ neighbourhood_base_of P X \ (\W x. openin X W \ x \ W \ (\U. P U \ x \ U \ U \ W))" apply (simp add: neighbourhood_base_of, safe, blast) by meson lemma neighbourhood_base_of_open_subset: "\neighbourhood_base_of P X; openin X S\ \ neighbourhood_base_of P (subtopology X S)" apply (clarsimp simp add: neighbourhood_base_of openin_subtopology_alt image_def) apply (rename_tac x V) apply (drule_tac x="S \ V" in spec) apply (simp add: Int_ac) by (metis IntI le_infI1 openin_Int) lemma neighbourhood_base_of_topology_base: "openin X = arbitrary union_of (\W. W \ \) \ neighbourhood_base_of P X \ (\W x. W \ \ \ x \ W \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W))" apply (auto simp: openin_topology_base_unique neighbourhood_base_of) by (meson subset_trans) lemma neighbourhood_base_at_unlocalized: assumes "\S T. \P S; openin X T; x \ T; T \ S\ \ P T" shows "neighbourhood_base_at x P X \ (x \ topspace X \ (\U V. openin X U \ P V \ x \ U \ U \ V \ V \ topspace X))" (is "?lhs = ?rhs") proof assume R: ?rhs show ?lhs unfolding neighbourhood_base_at_def proof clarify fix W assume "openin X W" "x \ W" then have "x \ topspace X" using openin_subset by blast with R obtain U V where "openin X U" "P V" "x \ U" "U \ V" "V \ topspace X" by blast then show "\U V. openin X U \ P V \ x \ U \ U \ V \ V \ W" by (metis IntI \openin X W\ \x \ W\ assms inf_le1 inf_le2 openin_Int) qed qed (simp add: neighbourhood_base_at_def) lemma neighbourhood_base_at_with_subset: "\openin X U; x \ U\ \ (neighbourhood_base_at x P X \ neighbourhood_base_at x (\T. T \ U \ P T) X)" apply (simp add: neighbourhood_base_at_def) apply (metis IntI Int_subset_iff openin_Int) done lemma neighbourhood_base_of_with_subset: "neighbourhood_base_of P X \ neighbourhood_base_of (\t. t \ topspace X \ P t) X" using neighbourhood_base_at_with_subset by (fastforce simp add: neighbourhood_base_of_def) subsection\Locally path-connected spaces\ definition weakly_locally_path_connected_at where "weakly_locally_path_connected_at x X \ neighbourhood_base_at x (path_connectedin X) X" definition locally_path_connected_at where "locally_path_connected_at x X \ neighbourhood_base_at x (\U. openin X U \ path_connectedin X U) X" definition locally_path_connected_space where "locally_path_connected_space X \ neighbourhood_base_of (path_connectedin X) X" lemma locally_path_connected_space_alt: "locally_path_connected_space X \ neighbourhood_base_of (\U. openin X U \ path_connectedin X U) X" (is "?P = ?Q") and locally_path_connected_space_eq_open_path_component_of: "locally_path_connected_space X \ (\U x. openin X U \ x \ U \ openin X (Collect (path_component_of (subtopology X U) x)))" (is "?P = ?R") proof - have ?P if ?Q using locally_path_connected_space_def neighbourhood_base_of_mono that by auto moreover have ?R if P: ?P proof - show ?thesis proof clarify fix U y assume "openin X U" "y \ U" have "\T. openin X T \ x \ T \ T \ Collect (path_component_of (subtopology X U) y)" if "path_component_of (subtopology X U) y x" for x proof - have "x \ U" using path_component_of_equiv that topspace_subtopology by fastforce then have "\Ua. openin X Ua \ (\V. path_connectedin X V \ x \ Ua \ Ua \ V \ V \ U)" using P by (simp add: \openin X U\ locally_path_connected_space_def neighbourhood_base_of) then show ?thesis by (metis dual_order.trans path_component_of_equiv path_component_of_maximal path_connectedin_subtopology subset_iff that) qed then show "openin X (Collect (path_component_of (subtopology X U) y))" using openin_subopen by force qed qed moreover have ?Q if ?R using that apply (simp add: open_neighbourhood_base_of) by (metis mem_Collect_eq openin_subset path_component_of_refl path_connectedin_path_component_of path_connectedin_subtopology that topspace_subtopology_subset) ultimately show "?P = ?Q" "?P = ?R" by blast+ qed lemma locally_path_connected_space: "locally_path_connected_space X \ (\V x. openin X V \ x \ V \ (\U. openin X U \ path_connectedin X U \ x \ U \ U \ V))" by (simp add: locally_path_connected_space_alt open_neighbourhood_base_of) lemma locally_path_connected_space_open_path_components: "locally_path_connected_space X \ (\U c. openin X U \ c \ path_components_of(subtopology X U) \ openin X c)" - apply (auto simp: locally_path_connected_space_eq_open_path_component_of path_components_of_def topspace_subtopology) + apply (auto simp: locally_path_connected_space_eq_open_path_component_of path_components_of_def) by (metis imageI inf.absorb_iff2 openin_closedin_eq) lemma openin_path_component_of_locally_path_connected_space: "locally_path_connected_space X \ openin X (Collect (path_component_of X x))" apply (auto simp: locally_path_connected_space_eq_open_path_component_of) by (metis openin_empty openin_topspace path_component_of_eq_empty subtopology_topspace) lemma openin_path_components_of_locally_path_connected_space: "\locally_path_connected_space X; c \ path_components_of X\ \ openin X c" apply (auto simp: locally_path_connected_space_eq_open_path_component_of) by (metis (no_types, lifting) imageE openin_topspace path_components_of_def subtopology_topspace) lemma closedin_path_components_of_locally_path_connected_space: "\locally_path_connected_space X; c \ path_components_of X\ \ closedin X c" by (simp add: closedin_def complement_path_components_of_Union openin_path_components_of_locally_path_connected_space openin_clauses(3) path_components_of_subset) lemma closedin_path_component_of_locally_path_connected_space: assumes "locally_path_connected_space X" shows "closedin X (Collect (path_component_of X x))" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: assms closedin_path_components_of_locally_path_connected_space path_component_in_path_components_of) next case False then show ?thesis by (metis closedin_empty path_component_of_eq_empty) qed lemma weakly_locally_path_connected_at: "weakly_locally_path_connected_at x X \ (\V. openin X V \ x \ V \ (\U. openin X U \ x \ U \ U \ V \ (\y \ U. \C. path_connectedin X C \ C \ V \ x \ C \ y \ C)))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: weakly_locally_path_connected_at_def neighbourhood_base_at_def) by (meson order_trans subsetD) next have *: "\V. path_connectedin X V \ U \ V \ V \ W" if "(\y\U. \C. path_connectedin X C \ C \ W \ x \ C \ y \ C)" for W U proof (intro exI conjI) let ?V = "(Collect (path_component_of (subtopology X W) x))" show "path_connectedin X (Collect (path_component_of (subtopology X W) x))" by (meson path_connectedin_path_component_of path_connectedin_subtopology) show "U \ ?V" by (auto simp: path_component_of path_connectedin_subtopology that) show "?V \ W" by (meson path_connectedin_path_component_of path_connectedin_subtopology) qed assume ?rhs then show ?lhs unfolding weakly_locally_path_connected_at_def neighbourhood_base_at_def by (metis "*") qed lemma locally_path_connected_space_im_kleinen: "locally_path_connected_space X \ (\V x. openin X V \ x \ V \ (\U. openin X U \ x \ U \ U \ V \ (\y \ U. \c. path_connectedin X c \ c \ V \ x \ c \ y \ c)))" apply (simp add: locally_path_connected_space_def neighbourhood_base_of_def) apply (simp add: weakly_locally_path_connected_at flip: weakly_locally_path_connected_at_def) using openin_subset apply force done lemma locally_path_connected_space_open_subset: "\locally_path_connected_space X; openin X s\ \ locally_path_connected_space (subtopology X s)" apply (simp add: locally_path_connected_space_def neighbourhood_base_of openin_open_subtopology path_connectedin_subtopology) by (meson order_trans) lemma locally_path_connected_space_quotient_map_image: assumes f: "quotient_map X Y f" and X: "locally_path_connected_space X" shows "locally_path_connected_space Y" unfolding locally_path_connected_space_open_path_components proof clarify fix V C assume V: "openin Y V" and c: "C \ path_components_of (subtopology Y V)" then have sub: "C \ topspace Y" using path_connectedin_path_components_of path_connectedin_subset_topspace path_connectedin_subtopology by blast have "openin X {x \ topspace X. f x \ C}" proof (subst openin_subopen, clarify) fix x assume x: "x \ topspace X" and "f x \ C" let ?T = "Collect (path_component_of (subtopology X {z \ topspace X. f z \ V}) x)" show "\T. openin X T \ x \ T \ T \ {x \ topspace X. f x \ C}" proof (intro exI conjI) have "\U. openin X U \ ?T \ path_components_of (subtopology X U)" proof (intro exI conjI) show "openin X ({z \ topspace X. f z \ V})" using V f openin_subset quotient_map_def by fastforce show "Collect (path_component_of (subtopology X {z \ topspace X. f z \ V}) x) \ path_components_of (subtopology X {z \ topspace X. f z \ V})" by (metis (no_types, lifting) Int_iff \f x \ C\ c mem_Collect_eq path_component_in_path_components_of path_components_of_subset topspace_subtopology topspace_subtopology_subset x) qed with X show "openin X ?T" using locally_path_connected_space_open_path_components by blast show x: "x \ ?T" using V \f x \ C\ c openin_subset path_component_of_equiv path_components_of_subset topspace_subtopology topspace_subtopology_subset x by fastforce have "f ` ?T \ C" proof (rule path_components_of_maximal) show "C \ path_components_of (subtopology Y V)" by (simp add: c) show "path_connectedin (subtopology Y V) (f ` ?T)" proof - have "quotient_map (subtopology X {a \ topspace X. f a \ V}) (subtopology Y V) f" by (simp add: V f quotient_map_restriction) then show ?thesis by (meson path_connectedin_continuous_map_image path_connectedin_path_component_of quotient_imp_continuous_map) qed show "\ disjnt C (f ` ?T)" by (metis (no_types, lifting) \f x \ C\ x disjnt_iff image_eqI) qed then show "?T \ {x \ topspace X. f x \ C}" - by (force simp: path_component_of_equiv topspace_subtopology) + by (force simp: path_component_of_equiv) qed qed then show "openin Y C" using f sub by (simp add: quotient_map_def) qed lemma homeomorphic_locally_path_connected_space: assumes "X homeomorphic_space Y" shows "locally_path_connected_space X \ locally_path_connected_space Y" proof - obtain f g where "homeomorphic_maps X Y f g" using assms homeomorphic_space_def by fastforce then show ?thesis by (metis (no_types) homeomorphic_map_def homeomorphic_maps_map locally_path_connected_space_quotient_map_image) qed lemma locally_path_connected_space_retraction_map_image: "\retraction_map X Y r; locally_path_connected_space X\ \ locally_path_connected_space Y" using Abstract_Topology.retraction_imp_quotient_map locally_path_connected_space_quotient_map_image by blast lemma locally_path_connected_space_euclideanreal: "locally_path_connected_space euclideanreal" unfolding locally_path_connected_space_def neighbourhood_base_of proof clarsimp fix W and x :: "real" assume "open W" "x \ W" then obtain e where "e > 0" and e: "\x'. \x' - x\ < e \ x' \ W" by (auto simp: open_real) then show "\U. open U \ (\V. path_connected V \ x \ U \ U \ V \ V \ W)" by (force intro!: convex_imp_path_connected exI [where x = "{x-e<.. topspace X") case True then show ?thesis using connectedin_connected_component_of [of X x] apply (clarsimp simp add: connectedin_def connected_space_clopen_in topspace_subtopology_subset cong: conj_cong) apply (drule_tac x="path_component_of_set X x" in spec) by (metis assms closedin_closed_subtopology closedin_connected_component_of closedin_path_component_of_locally_path_connected_space inf.absorb_iff2 inf.orderE openin_path_component_of_locally_path_connected_space openin_subtopology path_component_of_eq_empty path_component_subset_connected_component_of) next case False then show ?thesis using connected_component_of_eq_empty path_component_of_eq_empty by fastforce qed lemma path_components_eq_connected_components_of: "locally_path_connected_space X \ (path_components_of X = connected_components_of X)" by (simp add: path_components_of_def connected_components_of_def image_def path_component_eq_connected_component_of) lemma path_connected_eq_connected_space: "locally_path_connected_space X \ path_connected_space X \ connected_space X" by (metis connected_components_of_subset_sing path_components_eq_connected_components_of path_components_of_subset_singleton) lemma locally_path_connected_space_product_topology: "locally_path_connected_space(product_topology X I) \ topspace(product_topology X I) = {} \ finite {i. i \ I \ ~path_connected_space(X i)} \ (\i \ I. locally_path_connected_space(X i))" (is "?lhs \ ?empty \ ?rhs") proof (cases ?empty) case True then show ?thesis by (simp add: locally_path_connected_space_def neighbourhood_base_of openin_closedin_eq) next case False then obtain z where z: "z \ (\\<^sub>E i\I. topspace (X i))" by auto have ?rhs if L: ?lhs proof - obtain U C where U: "openin (product_topology X I) U" and V: "path_connectedin (product_topology X I) C" and "z \ U" "U \ C" and Csub: "C \ (\\<^sub>E i\I. topspace (X i))" using L apply (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of) by (metis openin_topspace topspace_product_topology z) then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" and XV: "\i. i\I \ openin (X i) (V i)" and "z \ Pi\<^sub>E I V" and subU: "Pi\<^sub>E I V \ U" by (force simp: openin_product_topology_alt) show ?thesis proof (intro conjI ballI) have "path_connected_space (X i)" if "i \ I" "V i = topspace (X i)" for i proof - have pc: "path_connectedin (X i) ((\x. x i) ` C)" apply (rule path_connectedin_continuous_map_image [OF _ V]) by (simp add: continuous_map_product_projection \i \ I\) moreover have "((\x. x i) ` C) = topspace (X i)" proof show "(\x. x i) ` C \ topspace (X i)" by (simp add: pc path_connectedin_subset_topspace) have "V i \ (\x. x i) ` (\\<^sub>E i\I. V i)" by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl that(1)) also have "\ \ (\x. x i) ` U" using subU by blast finally show "topspace (X i) \ (\x. x i) ` C" using \U \ C\ that by blast qed ultimately show ?thesis by (simp add: path_connectedin_topspace) qed then have "{i \ I. \ path_connected_space (X i)} \ {i \ I. V i \ topspace (X i)}" by blast with finV show "finite {i \ I. \ path_connected_space (X i)}" using finite_subset by blast next show "locally_path_connected_space (X i)" if "i \ I" for i apply (rule locally_path_connected_space_quotient_map_image [OF _ L, where f = "\x. x i"]) by (metis False Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection that) qed qed moreover have ?lhs if R: ?rhs proof (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of) fix F z assume "openin (product_topology X I) F" and "z \ F" then obtain W where finW: "finite {i \ I. W i \ topspace (X i)}" and opeW: "\i. i \ I \ openin (X i) (W i)" and "z \ Pi\<^sub>E I W" "Pi\<^sub>E I W \ F" by (auto simp: openin_product_topology_alt) have "\i \ I. \U C. openin (X i) U \ path_connectedin (X i) C \ z i \ U \ U \ C \ C \ W i \ (W i = topspace (X i) \ path_connected_space (X i) \ U = topspace (X i) \ C = topspace (X i))" (is "\i \ I. ?\ i") proof fix i assume "i \ I" have "locally_path_connected_space (X i)" by (simp add: R \i \ I\) moreover have "openin (X i) (W i) " "z i \ W i" using \z \ Pi\<^sub>E I W\ opeW \i \ I\ by auto ultimately obtain U C where UC: "openin (X i) U" "path_connectedin (X i) C" "z i \ U" "U \ C" "C \ W i" using \i \ I\ by (force simp: locally_path_connected_space_def neighbourhood_base_of) show "?\ i" proof (cases "W i = topspace (X i) \ path_connected_space(X i)") case True then show ?thesis using \z i \ W i\ path_connectedin_topspace by blast next case False then show ?thesis by (meson UC) qed qed then obtain U C where *: "\i. i \ I \ openin (X i) (U i) \ path_connectedin (X i) (C i) \ z i \ (U i) \ (U i) \ (C i) \ (C i) \ W i \ (W i = topspace (X i) \ path_connected_space (X i) \ (U i) = topspace (X i) \ (C i) = topspace (X i))" by metis let ?A = "{i \ I. \ path_connected_space (X i)} \ {i \ I. W i \ topspace (X i)}" have "{i \ I. U i \ topspace (X i)} \ ?A" by (clarsimp simp add: "*") moreover have "finite ?A" by (simp add: that finW) ultimately have "finite {i \ I. U i \ topspace (X i)}" using finite_subset by auto then have "openin (product_topology X I) (Pi\<^sub>E I U)" using * by (simp add: openin_PiE_gen) then show "\U. openin (product_topology X I) U \ (\V. path_connectedin (product_topology X I) V \ z \ U \ U \ V \ V \ F)" apply (rule_tac x="PiE I U" in exI, simp) apply (rule_tac x="PiE I C" in exI) using \z \ Pi\<^sub>E I W\ \Pi\<^sub>E I W \ F\ * apply (simp add: path_connectedin_PiE subset_PiE PiE_iff PiE_mono dual_order.trans) done qed ultimately show ?thesis using False by blast qed end diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy +++ b/src/HOL/Analysis/Path_Connected.thy @@ -1,4006 +1,4006 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" by (rule ext) simp show ?thesis unfolding path_def apply safe apply (subst g) apply (rule continuous_on_compose) apply (auto intro: continuous_intros) done qed lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast then have g: "g = h \ (f \ g)" by (metis comp_assoc id_comp) show ?thesis unfolding path_def using h assms by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def path_translation_eq) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms inj_on_eq_iff [of f] by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) = arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" using continuous_on_subset path_def by blast lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def by force lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" proof - have *: "\g. path_image (reversepath g) \ path_image g" unfolding path_image_def subset_eq reversepath_def Ball_def image_iff by force show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed lemma path_reversepath [simp]: "path (reversepath g) \ path g" proof - have *: "\g. path g \ path (reversepath g)" unfolding path_def reversepath_def apply (rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"]) done show ?thesis using *[of "reversepath g"] *[of g] unfolding reversepath_reversepath by (rule iffI) qed lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" apply (simp add: simple_path_def) apply (force simp: reversepath_def) done lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" and "path_image g2 \ s" shows "path_image (g1 +++ g2) \ s" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: "pathfinish g1 = pathstart g2 \ path_image(g1 +++ g2) = path_image g1 \ path_image g2" apply (rule subset_antisym [OF path_image_join_subset]) apply (auto simp: pathfinish_def pathstart_def path_image_def joinpaths_def image_def) apply (drule sym) apply (rule_tac x="xa/2" in bexI, auto) apply (rule ccontr) apply (drule_tac x="(xa+1)/2" in bspec) apply (auto simp: field_simps) apply (drule_tac x="1/2" in bspec, auto) done lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def Bex_def image_def) apply (metis eq_iff le_less_linear) apply (metis leD linear) using less_eq_real_def zero_le_one apply blast using less_eq_real_def zero_le_one apply blast done lemma connected_simple_path_endless: "simple_path c \ connected(path_image c - {pathstart c,pathfinish c})" apply (simp add: simple_path_endless) apply (rule connected_continuous_image) apply (meson continuous_on_subset greaterThanLessThan_subseteq_atLeastAtMost_iff le_numeral_extra(3) le_numeral_extra(4) path_def simple_path_imp_path) by auto lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by (auto simp: path_image_def reversepath_def) lemma path_imp_reversepath: "path g \ path(reversepath g)" apply (auto simp: path_def reversepath_def) using continuous_on_compose [of "{0..1}" "\x. 1 - x" g] apply (auto simp: continuous_on_op_minus) done lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" proof - have *: "{0..1::real} = {0..1/2} \ {1/2..1}" by auto have gg: "g2 0 = g1 1" by (metis assms(3) pathfinish_def pathstart_def) have 1: "continuous_on {0..1/2} (g1 +++ g2)" apply (rule continuous_on_eq [of _ "g1 \ (\x. 2*x)"]) apply (rule continuous_intros | simp add: joinpaths_def assms)+ done have "continuous_on {1/2..1} (g2 \ (\x. 2*x-1))" apply (rule continuous_on_subset [of "{1/2..1}"]) apply (rule continuous_intros | simp add: image_affinity_atLeastAtMost_diff assms)+ done then have 2: "continuous_on {1/2..1} (g1 +++ g2)" apply (rule continuous_on_eq [of "{1/2..1}" "g2 \ (\x. 2*x-1)"]) apply (rule assms continuous_intros | simp add: joinpaths_def mult.commute half_bounded_equal gg)+ done show ?thesis apply (subst *) apply (rule continuous_on_closed_Un) using 1 2 apply auto done qed lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" - by (simp add: path_join) + by (simp) lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xyI: "x = 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy apply auto apply (drule inj_onD [OF injg1]) using g21 [symmetric] xyI apply (auto dest: inj_onD [OF injg2]) done } note * = this { fix x and y::real assume xy: "y \ 1" "0 \ x" "\ y * 2 \ 1" "x * 2 \ 1" "g1 (2 * x) = g2 (2 * y - 1)" have g1im: "g1 (2 * x) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x" in image_eqI, auto) done have "x = 0 \ y = 1" using subsetD [OF sb g1im] xy apply auto apply (force dest: inj_onD [OF injg1]) using g21 [symmetric] apply (auto dest: inj_onD [OF injg2]) done } note ** = this show ?thesis using assms apply (simp add: arc_def simple_path_def path_join, clarify) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis **) apply (force dest: inj_onD [OF injg2]) done qed lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy by (auto dest: inj_onD [OF injg2]) } note * = this show ?thesis apply (simp add: arc_def inj_on_def) - apply (clarsimp simp add: arc_imp_path assms path_join) + apply (clarsimp simp add: arc_imp_path assms) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis *) apply (force dest: inj_onD [OF injg2]) done qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" using assms(2) unfolding path_def continuous_on_iff apply (drule_tac x=0 in bspec, simp) by (metis half_gt_zero_iff norm_conv_dist) obtain d2 where "d2 > 0" and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) apply (drule_tac x="e/2" in spec) apply (force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast then show False by (simp add: e_def [symmetric]) qed lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x + 1) / 2" "(y + 1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") proof assume ?lhs then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ apply (rule simple_path_joinE [OF arc_imp_simple_path assms]) using n1 by force next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2})" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp lemma simple_path_assoc: assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" by auto lemma simple_path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ simple_path(p +++ q) \ simple_path(q +++ p)" by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) lemma path_image_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path_image(p +++ q) = path_image(q +++ p)" by (simp add: path_image_join sup_commute) subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" apply (rule continuous_intros | simp)+ apply (simp add: image_affinity_atLeastAtMost [where c=u]) using assms apply (auto simp: path_def continuous_on_subset) done then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" -by (metis add.commute affine_ineq less_eq_real_def mult.right_neutral) +by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y)" by (auto simp: arc_def inj_on_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (cases "v = u") (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, simp add: field_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs unfolding inj_on_def by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) show ?lhs using psp ne unfolding arc_def subpath_def inj_on_def by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms apply (simp add: simple_path_subpath_eq simple_path_imp_path) apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) done lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" apply (rule arc_simple_path_subpath) apply (force simp: simple_path_def)+ done lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost path_image_subpath) apply (auto simp: path_image_def) done lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) then have com: "compact ({0..1} \ {u. g u \ closure (- S)})" apply (simp add: Int_commute [of "{0..1}"] compact_eq_bounded_closed closed_vimage_Int [unfolded vimage_def]) using compact_eq_bounded_closed apply fastforce done have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce { assume "u \ 0" then have "u > 0" using \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then have "g u \ closure S" by (simp add: frontier_def closure_approachable) } then show ?thesis apply (rule_tac u=u in that) apply (auto simp: \0 \ u\ \u \ 1\ gu interior_closure umin) using \_ \ 1\ interior_closure umin apply fastforce done qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (simp add: gunot) using \0 \ u\ u0 by (force simp: subpath_def gxin) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" using subpath_to_frontier_strong [OF g g1] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (metis DiffI disj frontier_def g0 notin pathstart_def) using \0 \ u\ g0 disj apply (simp add: path_image_subpath_gen) apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def) apply (rename_tac y) apply (drule_tac x="y/u" in spec) apply (auto split: if_split_asm) done qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis apply (rule that [of "subpath 0 u g"]) using assms u apply (simp_all add: path_image_subpath) apply (simp add: pathstart_def) apply (force simp: closed_segment_eq_real_ivl path_image_def) done qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" proof - obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis apply (rule that [OF \path h\]) using assms h apply auto apply (metis Diff_single_insert frontier_subset_eq insert_iff interior_subset subset_iff) done qed subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" lemma pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" using assms(2)[unfolded pathfinish_def pathstart_def] by auto show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a apply (force simp: field_simps)+ done next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" unfolding simple_path_def proof (intro conjI impI ballI) show "path (shiftpath a g)" by (simp add: assms path_shiftpath simple_path_imp_path) have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a unfolding shiftpath_def by (force split: if_split_asm dest!: *) qed subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" proof (clarsimp simp: inj_on_def linepath_def) fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" by (auto simp: algebra_simps) then show "x=y" using assms by auto qed lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" proof - have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \ convex hull s" by (rule convexD_alt) (use assms in auto) then show ?thesis by (simp add: midpoint_def algebra_simps) qed lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto show ?thesis proof (cases "c \ open_segment a b \ d \ open_segment a b") case True then show ?thesis using c d that by blast next case False then have "(c = a \ c = b) \ (d = a \ d = b)" by (simp add: \c \ closed_segment a b\ \d \ closed_segment a b\ open_segment_def) with \a \ b\ \f a = f b\ c d show ?thesis by (rule_tac z = "midpoint a b" in that) (fastforce+) qed qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" proof - obtain e where "ball z e \ path_image g = {}" "e > 0" using not_on_path_ball[OF assms] by auto moreover have "cball z (e/2) \ ball z e" using \e > 0\ by auto ultimately show ?thesis by (rule_tac x="e/2" in exI) auto qed subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set s x \ Collect (path_component s x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component s x y" shows "x \ s" and "y \ s" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ s" shows "path_component s x x" unfolding path_defs apply (rule_tac x="\u. x" in exI) using assms apply (auto intro!: continuous_intros) done lemma path_component_refl_eq: "path_component s x x \ x \ s" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component s x y \ path_component s y x" unfolding path_component_def apply (erule exE) apply (rule_tac x="reversepath g" in exI, auto) done lemma path_component_trans: assumes "path_component s x y" and "path_component s y z" shows "path_component s x z" using assms unfolding path_component_def apply (elim exE) apply (rule_tac x="g +++ ga" in exI) apply (auto simp: path_image_join) done lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" unfolding path_component_def by auto lemma path_component_linepath: fixes s :: "'a::real_normed_vector set" shows "closed_segment a b \ s \ path_component s a b" unfolding path_component_def by (rule_tac x="linepath a b" in exI, auto) subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set s x = {y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set s x \ s" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set s x = {} \ x \ s" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "s \ t \ (path_component_set s x) \ (path_component_set t x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set s x \ path_component_set s y = path_component_set s x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected s \ (\x\s. \y\s. \g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs) lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected s \ (\x\s. path_component_set s x = s)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ t; path_connected t; t \ s\ \ t \ (path_component_set s x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes s :: "'a::real_normed_vector set" assumes "convex s" shows "path_connected s" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" - by (auto intro!: convex_connected convex_real_interval) + by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y assume as: "y \ path_component_set S x" then have "y \ S" by (simp add: path_component_mem(2)) then obtain e where e: "e > 0" "ball y e \ S" using assms[unfolded open_contains_ball] by auto have "\u. dist y u < e \ path_component S x u" by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) then show "\e > 0. ball y e \ path_component_set S x" using \e>0\ by auto qed lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z proof - have "y \ path_component_set S z" by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) then have "y \ path_component_set S x" using path_component_eq that(2) by blast then show False using y by blast qed qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' assume "x' \ f ` S" "y' \ f ` S" then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' apply (rule_tac x="f \ g" in exI) unfolding path_defs apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)]) apply auto done qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def apply clarify apply (rule_tac x="\x. a" in exI) apply (simp add: image_constant_conv) - apply (simp add: path_def continuous_on_const) + apply (simp add: path_def) done lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto show "path_component (S \ T) x y" using x y proof safe assume "x \ S" "y \ S" then show "path_component (S \ T) x y" by (meson Un_upper1 \path_connected S\ path_component_of_subset path_connected_component) next assume "x \ S" "y \ T" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ S" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ T" then show "path_component (S \ T) x y" by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) qed qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (auto elim!: path_component_of_subset [rotated]) then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on {0..1} (p \ ((*) y))" apply (rule continuous_intros)+ using p [unfolded path_def] y apply (auto simp: mult_le_one intro: continuous_on_subset [of _ p]) done then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z assume pa: "path_component s x y" "path_component s x z" then have pae: "path_component_set s x = path_component_set s y" using path_component_eq by auto have yz: "path_component s y z" using pa path_component_sym path_component_trans by blast then have "\g. path g \ path_image g \ path_component_set s x \ pathstart g = y \ pathfinish g = z" apply (simp add: path_component_def, clarify) apply (rule_tac x=g in exI) by (simp add: pae path_component_maximal path_connected_path_image pathstart_in_path_image) } then show ?thesis by (simp add: path_connected_def) qed lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" apply (intro iffI) apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) using path_component_of_subset path_connected_component by blast lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" proof (cases "x \ S") case True show ?thesis apply (rule subset_antisym) apply (simp add: path_component_subset) - by (simp add: True path_component_maximal path_component_refl path_connected_path_component) + by (simp add: True path_component_maximal path_component_refl) next case False then show ?thesis by (metis False empty_iff path_component_eq_empty) qed lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis apply (rule connected_component_maximal) apply (auto simp: True path_component_subset path_component_refl path_connected_imp_connected) done next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" unfolding path_connected_space_def Ball_def apply (intro all_cong1 imp_cong refl, safe) using path_connectedin_path_image apply fastforce by (meson path_connectedin) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" using linear_injective_left_inverse [OF assms] apply clarify apply (rule_tac g=g in that) using assms apply (auto simp: homeomorphism_def eq_id_iff [symmetric] image_comp comp_def linear_conv_bounded_linear linear_continuous_on) done lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = h]) apply (rule continuous_intros | force)+ done moreover have "path (\z. (g z, y2))" using \path g\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = g]) apply (rule continuous_intros | force)+ done ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" apply (intro exI conjI) apply (rule 1) apply (rule 2) apply (metis hs pathstart_def pathstart_join) by (metis gf pathfinish_def pathfinish_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" apply (rule subset_antisym) using path_component_subset apply force using path_component_refl by auto lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" apply (auto simp: disjnt_def) using path_component_eq apply fastforce using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" apply (rule iffI, metis (no_types) path_component_mem(1) path_component_refl) apply (erule disjE, metis Collect_empty_eq_bot path_component_eq_empty) apply (rule ext) apply (metis path_component_trans path_component_sym) done lemma path_component_unique: assumes "x \ c" "c \ S" "path_connected c" "\c'. \x \ c'; c' \ S; path_connected c'\ \ c' \ c" shows "path_component_set S x = c" apply (rule subset_antisym) using assms apply (metis mem_Collect_eq subsetCE path_component_eq_eq path_component_subset path_connected_path_component) by (simp add: assms path_component_maximal) lemma path_component_intermediate_subset: "path_component_set u a \ t \ t \ u \ path_component_set t a = path_component_set u a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g ` {0..1} \ T" by (simp add: path_def pathin_def) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" apply (auto simp: path_component_in_topspace) apply (force simp: path_component_of_def pathin_const) done lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose) apply (force simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" apply (auto simp: path_component_of_def) using path_connectedin_path_image apply fastforce apply (metis path_connectedin) done lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: fun_eq_iff path_component_in_topspace) apply (meson path_component_of_sym path_component_of_trans) done qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "\y. path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) then show ?thesis apply (simp add: path_connectedin_def path_component_of_subset_topspace path_connected_space_iff_path_component) by (metis Int_absorb1 mem_Collect_eq path_component_of_equiv path_component_of_subset_topspace topspace_subtopology) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" apply (auto simp: path_components_of_def path_component_of_equiv) using path_component_of_maximal path_connectedin_def apply fastforce by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: "C \ path_components_of X \ (C \ {})" apply (clarsimp simp: path_components_of_def path_component_of_eq_empty) by (meson path_component_of_refl) lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" apply (rule iffI) using nonempty_path_components_of path_component_of_eq_empty apply fastforce by (simp add: path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" - apply (simp add: topspace_subtopology) + apply (simp) by (meson Union_upper \ path_component_of path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ topspace X = {}" using Union_path_components_of nonempty_path_components_of by fastforce lemma path_components_of_empty_space: "topspace X = {} \ path_components_of X = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" proof (intro iffI conjI) assume L: "path_components_of X = {S}" then show "path_connected_space X" by (simp add: path_connected_space_iff_components_eq) show "topspace X = S" by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) next assume R: "path_connected_space X \ topspace X = S" then show "path_components_of X = {S}" using ccpo_Sup_singleton [of S] by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) qed with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" proof (cases "S = {}") case True then show ?thesis using ne path_components_of_eq_empty that by fastforce next case False then obtain a where "a \ S" by blast show ?thesis proof show "Collect (path_component_of X a) \ path_components_of X" by (meson \a \ S\ S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) show "S \ Collect (path_component_of X a)" by (simp add: S \a \ S\ path_component_of_maximal) qed qed lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using g x unfolding homeomorphic_maps_def by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] apply safe using assms homeomorphic_map_path_component_of apply fastforce by (metis assms homeomorphic_map_path_component_of imageI) subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using ex_card[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis - by (simp add: path_connected_empty) + by (simp) next case equal then show ?thesis - by (simp add: path_connected_singleton) + by (simp) next case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected (sphere (0::'a) r)" by auto then have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" and 2: "2 \ DIM('a)" shows "path_connected (- s)" proof (cases "s = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ s" by auto { fix x y assume "x \ s" "y \ s" then have "x \ a" "y \ a" using \a \ s\ by auto then have bxy: "bounded(insert x (insert y s))" by (simp add: \bounded s\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "s \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \ s" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (C - 1) * u) *\<^sub>R (x - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (C - 1) * u)" in spec) using \x \ a\ \x \ s\ \0 \ u\ CC apply (auto simp: xeq) done } then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \ s" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (D - 1) * u) *\<^sub>R (y - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (D - 1) * u)" in spec) using \y \ a\ \y \ s\ \0 \ u\ DD apply (auto simp: xeq) done } then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))" apply (rule path_component_of_subset [of "sphere a B"]) using \s \ ball a B\ apply (force simp: ball_def dist_norm norm_minus_commute) apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format]) using \x \ a\ using \y \ a\ B apply (auto simp: dist_norm C_def D_def) done have "path_component (- s) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" "2 \ DIM('a)" shows "connected (- s)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes s :: "'a :: euclidean_space set" assumes "connected s" "cball a r \ s" "2 \ DIM('a)" shows "connected (s - ball a r)" apply (rule connected_diff_open_from_closed [OF ball_subset_cball]) using assms connected_sphere apply (auto simp: cball_diff_eq_sphere dist_norm) done proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast have "dist a (a + \ *\<^sub>R (SOME i. i \ Basis)) = \" by (simp add: dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "\{S - ball a r |r. 0 < r \ r < \} \ {} \ False" apply (drule_tac c="a + scaleR (\) ((SOME i. i \ Basis))" in subsetD) by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x apply (rule UnionI [of "S - ball a (min \ (dist a x) / 2)"]) using that \0 < \\ apply simp_all apply (rule_tac x="min \ (dist a x) / 2" in exI) apply auto done then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" shows "connected(S - T)" using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x F) then have "connected (S-F)" by auto moreover have "open (S - F)" using finite_imp_closed[OF \finite F\] \open S\ by auto ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis apply (simp add: abs_if split: if_split_asm) apply (metis add.inverse_inverse real_vector.scale_minus_left xb) done qed with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" by (force simp: sphere_def dist_norm) have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" by (simp add: dist_norm) also have "\ = norm ((2*r) *\<^sub>R b)" by (metis mult_2 scaleR_add_left) also have "\ = 2*r" using \r > 0\ b norm_Basis by fastforce finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . qed qed lemma sphere_1D_doubleton: fixes a :: "'a :: euclidean_space" assumes "DIM('a) = 1" and "r > 0" obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" proof - have "sphere a r = (+) a ` sphere 0 r" by (metis add.right_neutral sphere_translation) then show ?thesis using sphere_1D_doubleton_zero [OF assms] by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) qed lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" apply (rule connected_intermediate_closure [of "ball a r"]) using assms by auto moreover have "connected {x. r \ dist a x \ x \ S}" apply (rule connected_intermediate_closure [of "- cball a r"]) using assms apply (auto intro: connected_complement_bounded_convex) apply (metis ComplI interior_cball interior_closure mem_ball not_less) done ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) apply (rule_tac x=0 in exI, simp) using vector_choose_size zero_le_one by blast next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis apply (subst *) apply (rule path_connected_continuous_image, auto) done qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes s :: "'a::real_normed_vector set" shows "open s \ open (connected_component_set s x)" apply (simp add: open_contains_ball, clarify) apply (rename_tac y) apply (drule_tac x=y in bspec) apply (simp add: connected_component_in, clarify) apply (rule_tac x=e in exI) by (metis mem_Collect_eq connected_component_eq connected_component_maximal centre_in_ball connected_ball) corollary open_components: fixes s :: "'a::real_normed_vector set" shows "\open u; s \ components u\ \ open s" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes s :: "'a::real_normed_vector set" assumes x: "x \ s" and s: "open s" shows "x \ closure (connected_component_set s y) \ x \ connected_component_set s y" proof - { assume "x \ closure (connected_component_set s y)" moreover have "x \ connected_component_set s x" using x by simp ultimately have "x \ connected_component_set s y" using s by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) } then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto from connectedD [OF \connected S\ 1 2 3 4] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ have "S \ T" using "3" by auto show ?thesis using \S \ \(B - {T}) = {}\ \S \ T\ \T \ B\ that by auto qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" apply (rule connected_disjoint_Union_open_pick [OF B, of A]) using SA SB \S \ A\ by auto moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" apply (rule connected_disjoint_Union_open_pick [OF A, of B]) using SA SB \T \ B\ by auto ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis apply (rule connected_disjoint_Union_open_unique) apply (simp add: components_eq disjnt_def pairwise_def) using \open S\ apply (simp_all add: assms open_components in_components_connected in_components_nonempty) done qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes "bounded (-s)" shows "\x. x \ s \ \ bounded (connected_component_set s x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" apply (auto simp: bounded_def dist_norm) apply (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) apply simp using i apply (auto simp: algebra_simps) done have **: "{x. B \ i \ x} \ connected_component_set s (B *\<^sub>R i)" apply (rule connected_component_maximal) apply (auto simp: i intro: convex_connected convex_halfspace_ge [of B]) apply (rule *) apply (rule order_trans [OF _ Basis_le_norm [OF i]]) by (simp add: inner_commute) have "B *\<^sub>R i \ s" by (rule *) (simp add: norm_Basis [OF i]) then show ?thesis apply (rule_tac x="B *\<^sub>R i" in exI, clarify) apply (frule bounded_subset [of _ "{x. B \ i \ x}", OF _ **]) using unbounded_inner apply blast done qed lemma cobounded_unique_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes bs: "bounded (-s)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set s x)" "\ bounded(connected_component_set s y)" shows "connected_component_set s x = connected_component_set s y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have ccb: "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) obtain x' where x': "connected_component s x x'" "norm x' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) obtain y' where y': "connected_component s y y'" "norm y' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) have x'y': "connected_component s x' y'" apply (simp add: connected_component_def) apply (rule_tac x="- ball 0 B" in exI) using x' y' apply (auto simp: ccb dist_norm *) done show ?thesis apply (rule connected_component_eq) using x' y' x'y' by (metis (no_types, lifting) connected_component_eq_empty connected_component_eq_eq connected_component_idemp connected_component_in) qed lemma cobounded_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "bounded (-s) \ \c. c \ components s \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "\bounded (- s); c \ components s; \ bounded c; c' \ components s; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac s="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) apply clarify apply (metis connected_component_eq_eq connected_component_in) done lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" apply (auto simp: outside bounded_def dist_norm) apply (metis diff_0 norm_minus_cancel not_less) by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt) apply (simp add: Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto { fix y::'a and z::'a assume yz: "B < norm z" "B < norm y" have "connected_component (- cball 0 B) y z" apply (rule connected_componentI [OF _ subset_refl]) apply (rule connected_complement_bounded_convex) using assms yz by (auto simp: dist_norm) then have "connected_component (- S) y z" apply (rule connected_component_of_subset) apply (metis Bno Compl_anti_mono mem_cball_0 subset_iff) done } note cyz = this show ?thesis apply (auto simp: outside) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) apply (simp add: bounded_pos) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" apply (rule connected_component_maximal) using that interior_subset mem_ball apply auto done then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have 1: "closed_segment x y \ T \ {}" using \y \ T\ by blast have 2: "closed_segment x y - T \ {}" using False by blast obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment 1 2] by auto then show ?thesis proof - have "norm (y - x) < e" by (metis dist_norm \dist y x < e\) moreover have "norm (c - x) \ norm (y - x)" by (simp add: \c \ closed_segment x y\ segment_bound(1)) ultimately have "norm (c - x) < e" by linarith then show ?thesis by (metis (no_types) \c \ frontier T\ dist_norm that(1)) qed qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "interior s \ inside (frontier s)" proof - { fix x y assume x: "x \ interior s" and y: "y \ s" and cc: "connected_component (- frontier s) x y" have "connected_component_set (- frontier s) x \ frontier s \ {}" apply (rule connected_Int_frontier, simp) apply (metis IntI cc connected_component_in connected_component_refl empty_iff interiorE mem_Collect_eq rev_subsetD x) using y cc by blast then have "bounded (connected_component_set (- frontier s) x)" using connected_component_in by auto } then show ?thesis apply (auto simp: inside_def frontier_def) apply (rule classical) apply (rule bounded_subset [OF assms], blast) done qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" - by (simp add: inside_def connected_component_UNIV) + by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- s) x y; x \ inside s\ \ y \ inside s" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- s) x y; x \ outside s\ \ y \ outside s" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "convex s" and z: "z \ s" shows "z \ outside s" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain a where "a \ s" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- s) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- s) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ s" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF s \a \ s\ ins, of "1/(u*C + 1)"] \C>0\ \z \ s\ Cpos u - by (simp add: * field_split_simps algebra_simps) + by (simp add: * field_split_simps) } note contra = this have "connected_component (- s) z (z + C *\<^sub>R (z-a))" apply (rule connected_componentI [OF connected_segment [of z "z + C *\<^sub>R (z-a)"]]) apply (simp add: closed_segment_def) using contra apply auto done then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex s" shows "outside s = - s" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "convex s \ inside s = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex t; s \ t\ \ - t \ outside s" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "outside(frontier s) \ - closure s" unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff proof - { assume "interior s \ inside (frontier s)" hence "interior s \ inside (frontier s) = inside (frontier s)" by (simp add: subset_Un_eq) then have "closure s \ frontier s \ inside (frontier s)" using frontier_def by auto } then show "closure s \ frontier s \ inside (frontier s)" using interior_inside_frontier [OF assms] by blast qed lemma outside_frontier_eq_complement_closure: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" "convex s" shows "outside(frontier s) = - closure s" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded s; convex s\ \ inside(frontier s) = interior s" apply (simp add: inside_outside outside_frontier_eq_complement_closure) using closure_subset interior_subset apply (auto simp: frontier_def) done lemma open_inside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (inside s)" proof - { fix x assume x: "x \ inside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside s" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (outside s)" proof - { fix x assume x: "x \ outside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis Int_iff outside_def connected_component_refl_eq x) then have "\e>0. ball x e \ outside s" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(inside s) \ s \ inside s" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(inside s) \ s" proof - have "closure (inside s) \ - inside s = closure (inside s) - interior (inside s)" by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) moreover have "- inside s \ - outside s = s" by (metis (no_types) compl_sup double_compl inside_Un_outside) moreover have "closure (inside s) \ - outside s" by (metis (no_types) assms closure_inside_subset union_with_inside) ultimately have "closure (inside s) - interior (inside s) \ s" by blast then show ?thesis by (simp add: frontier_def open_inside interior_open) qed lemma closure_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(outside s) \ s \ outside s" apply (rule closure_minimal, simp) by (metis assms closed_open inside_outside open_inside) lemma frontier_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(outside s) \ s" apply (simp add: frontier_def open_outside interior_open) by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup.commute) lemma inside_complement_unbounded_connected_empty: "\connected (- s); \ bounded (- s)\ \ inside s = {}" apply (simp add: inside_def) by (meson Compl_iff bounded_subset connected_component_maximal order_refl) lemma inside_bounded_complement_connected_empty: fixes s :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- s); bounded s\ \ inside s = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "s \ inside t" shows "inside s - t \ inside t" unfolding inside_def proof clarify fix x assume x: "x \ t" "x \ s" and bo: "bounded (connected_component_set (- s) x)" show "bounded (connected_component_set (- t) x)" proof (cases "s \ connected_component_set (- t) x = {}") case True show ?thesis apply (rule bounded_subset [OF bo]) apply (rule connected_component_maximal) using x True apply auto done next case False then show ?thesis using assms [unfolded inside_def] x apply (simp add: disjoint_iff_not_equal, clarify) apply (drule subsetD, assumption, auto) by (metis (no_types, hide_lams) ComplI connected_component_eq_eq) qed qed lemma inside_inside_subset: "inside(inside s) \ s" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected t; inside s \ t \ {}; outside s \ t \ {}\ \ s \ t \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis (no_types, hide_lams) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) lemma outside_bounded_nonempty: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" shows "outside s \ {}" by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball double_complement order_refl outside_convex outside_def) lemma outside_compact_in_open: fixes s :: "'a :: {real_normed_vector,perfect_space} set" assumes s: "compact s" and t: "open t" and "s \ t" "t \ {}" shows "outside s \ t \ {}" proof - have "outside s \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty s) with assms obtain a b where a: "a \ outside s" and b: "b \ t" by auto show ?thesis proof (cases "a \ t") case True with a show ?thesis by blast next case False have front: "frontier t \ - s" using \s \ t\ frontier_disjoint_eq t by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- t)" and pf: "pathfinish \ \ frontier t" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -s" unfolding c_def using front pf by blast moreover have "open (-s)" using s compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -s" using open_contains_cball[of "-s"] s by blast then obtain d where "d \ t" and d: "dist d c < \" using closure_approachable [of c t] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -s" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -s" using pimg_sbs apply (auto simp: path_image_def) apply (drule subsetD) using \c \ - s\ \s \ t\ interior_subset apply (auto simp: c_def) done have "closed_segment c d \ cball c \" apply (simp add: segment_convex_hull) apply (rule hull_minimal) using \\ > 0\ d apply (auto simp: dist_commute) done with \ have "closed_segment c d \ -s" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- s) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside s \ t \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ t\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- t)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes s :: "'a :: euclidean_space set" assumes s: "closed s" and t: "compact t" and "connected t" "s \ inside t" shows "inside s \ inside t" proof (cases "inside t = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have coms: "compact s" using assms apply (simp add: s compact_eq_bounded_closed) by (meson bounded_inside bounded_subset compact_imp_bounded) then have bst: "bounded (s \ t)" by (simp add: compact_imp_bounded t) then obtain r where "0 < r" and r: "s \ t \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside s \ outside t \ {}" proof - have "- ball 0 r \ outside s" apply (rule outside_subset_convex) using r by auto moreover have "- ball 0 r \ outside t" apply (rule outside_subset_convex) using r by auto ultimately show ?thesis by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) qed have "s \ t = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside s \ inside t \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t) ultimately have "inside s \ t = {}" using inside_outside_intersect_connected [OF \connected t\, of s] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \s \ inside t\] by blast qed qed lemma connected_with_inside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ inside s)" proof (cases "s \ inside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ inside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ inside s)" if "a \ (s \ inside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ inside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_inside_subset s apply fastforce by (metis (no_types, lifting) frontier_inside_subset insertE insert_Diff interior_eq open_inside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ outside s)" proof (cases "s \ outside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ outside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ outside s)" if "a \ (s \ outside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ outside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_outside_subset s apply fastforce by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "closed s" and cons: "connected s" shows "inside (inside s) = {}" by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) lemma inside_in_components: "inside s \ components (- s) \ connected(inside s) \ inside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: inside_same_component connected_componentI) apply (metis IntI empty_iff inside_no_overlap) done text\The proof is virtually the same as that above.\ lemma outside_in_components: "outside s \ components (- s) \ connected(outside s) \ outside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: outside_same_component connected_componentI) apply (metis IntI empty_iff outside_no_overlap) done lemma bounded_unique_outside: fixes s :: "'a :: euclidean_space set" shows "\bounded s; DIM('a) \ 2\ \ (c \ components (- s) \ \ bounded c \ c = outside s)" apply (rule iffI) apply (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) by (simp add: connected_outside outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" apply (rule distance_attains_inf [of "frontier(f ` S)" "f a"]) using \a \ S\ by (auto simp: frontier_eq_empty dist_norm False) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ apply (simp add: compact_closure [symmetric] compact_def) apply (drule_tac x=z in spec) using closure_subset apply force done then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have rle: "r \ norm (f y - f a)" apply (rule le_no) using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast show ?thesis apply (rule **) (*such a horrible mess*) apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball]) using \a \ S\ \0 < r\ apply (auto simp: disjoint_iff_not_equal dist_norm) by (metis dw_le norm_minus_commute not_less order_trans rle wy) qed subsubsection\Special characterizations of classes of functions into and out of R.\ proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: continuous_map_def gf) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast end diff --git a/src/HOL/Analysis/Retracts.thy b/src/HOL/Analysis/Retracts.thy --- a/src/HOL/Analysis/Retracts.thy +++ b/src/HOL/Analysis/Retracts.thy @@ -1,2593 +1,2593 @@ section \Absolute Retracts, Absolute Neighbourhood Retracts and Euclidean Neighbourhood Retracts\ theory Retracts imports Brouwer_Fixpoint Continuous_Extension Ordered_Euclidean_Space begin text \Absolute retracts (AR), absolute neighbourhood retracts (ANR) and also Euclidean neighbourhood retracts (ENR). We define AR and ANR by specializing the standard definitions for a set to embedding in spaces of higher dimension. John Harrison writes: "This turns out to be sufficient (since any set in \\\<^sup>n\ can be embedded as a closed subset of a convex subset of \\\<^sup>n\<^sup>+\<^sup>1\) to derive the usual definitions, but we need to split them into two implications because of the lack of type quantifiers. Then ENR turns out to be equivalent to ANR plus local compactness."\ definition\<^marker>\tag important\ AR :: "'a::topological_space set \ bool" where "AR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ S' retract_of U" definition\<^marker>\tag important\ ANR :: "'a::topological_space set \ bool" where "ANR S \ \U. \S'::('a * real) set. S homeomorphic S' \ closedin (top_of_set U) S' \ (\T. openin (top_of_set U) T \ S' retract_of T)" definition\<^marker>\tag important\ ENR :: "'a::topological_space set \ bool" where "ENR S \ \U. open U \ S retract_of U" text \First, show that we do indeed get the "usual" properties of ARs and ANRs.\ lemma AR_imp_absolute_extensor: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "AR S" and contf: "continuous_on T f" and "f ` T \ S" and cloUT: "closedin (top_of_set U) T" obtains g where "continuous_on U g" "g ` U \ S" "\x. x \ T \ g x = f x" proof - have "aff_dim S < int (DIM('b \ real))" using aff_dim_le_DIM [of S] by simp then obtain C and S' :: "('b * real) set" where C: "convex C" "C \ {}" and cloCS: "closedin (top_of_set C) S'" and hom: "S homeomorphic S'" by (metis that homeomorphic_closedin_convex) then have "S' retract_of C" using \AR S\ by (simp add: AR_def) then obtain r where "S' \ C" and contr: "continuous_on C r" and "r ` C \ S'" and rid: "\x. x\S' \ r x = x" by (auto simp: retraction_def retract_of_def) obtain g h where "homeomorphism S S' g h" using hom by (force simp: homeomorphic_def) then have "continuous_on (f ` T) g" by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def) then have contgf: "continuous_on T (g \ f)" by (metis continuous_on_compose contf) have gfTC: "(g \ f) ` T \ C" proof - have "g ` S = S'" by (metis (no_types) \homeomorphism S S' g h\ homeomorphism_def) with \S' \ C\ \f ` T \ S\ show ?thesis by force qed obtain f' where f': "continuous_on U f'" "f' ` U \ C" "\x. x \ T \ f' x = (g \ f) x" by (metis Dugundji [OF C cloUT contgf gfTC]) show ?thesis proof (rule_tac g = "h \ r \ f'" in that) show "continuous_on U (h \ r \ f')" apply (intro continuous_on_compose f') using continuous_on_subset contr f' apply blast by (meson \homeomorphism S S' g h\ \r ` C \ S'\ continuous_on_subset \f' ` U \ C\ homeomorphism_def image_mono) show "(h \ r \ f') ` U \ S" using \homeomorphism S S' g h\ \r ` C \ S'\ \f' ` U \ C\ by (fastforce simp: homeomorphism_def) show "\x. x \ T \ (h \ r \ f') x = f x" using \homeomorphism S S' g h\ \f ` T \ S\ f' by (auto simp: rid homeomorphism_def) qed qed lemma AR_imp_absolute_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "AR S" "S homeomorphic S'" and clo: "closedin (top_of_set U) S'" shows "S' retract_of U" proof - obtain g h where hom: "homeomorphism S S' g h" using assms by (force simp: homeomorphic_def) have h: "continuous_on S' h" " h ` S' \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done obtain h' where h': "continuous_on U h'" "h' ` U \ S" and h'h: "\x. x \ S' \ h' x = h x" by (blast intro: AR_imp_absolute_extensor [OF \AR S\ h clo]) have [simp]: "S' \ U" using clo closedin_limpt by blast show ?thesis proof (simp add: retraction_def retract_of_def, intro exI conjI) show "continuous_on U (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` U \ S'" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\S'. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed qed lemma AR_imp_absolute_retract_UNIV: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "AR S" and hom: "S homeomorphic S'" and clo: "closed S'" shows "S' retract_of UNIV" apply (rule AR_imp_absolute_retract [OF \AR S\ hom]) using clo closed_closedin by auto lemma absolute_extensor_imp_AR: fixes S :: "'a::euclidean_space set" assumes "\f :: 'a * real \ 'a. \U T. \continuous_on T f; f ` T \ S; closedin (top_of_set U) T\ \ \g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)" shows "AR S" proof (clarsimp simp: AR_def) fix U and T :: "('a * real) set" assume "S homeomorphic T" and clo: "closedin (top_of_set U) T" then obtain g h where hom: "homeomorphism S T g h" by (force simp: homeomorphic_def) have h: "continuous_on T h" " h ` T \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done obtain h' where h': "continuous_on U h'" "h' ` U \ S" and h'h: "\x\T. h' x = h x" using assms [OF h clo] by blast have [simp]: "T \ U" using clo closedin_imp_subset by auto show "T retract_of U" proof (simp add: retraction_def retract_of_def, intro exI conjI) show "continuous_on U (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` U \ T" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\T. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed qed lemma AR_eq_absolute_extensor: fixes S :: "'a::euclidean_space set" shows "AR S \ (\f :: 'a * real \ 'a. \U T. continuous_on T f \ f ` T \ S \ closedin (top_of_set U) T \ (\g. continuous_on U g \ g ` U \ S \ (\x \ T. g x = f x)))" apply (rule iffI) apply (metis AR_imp_absolute_extensor) apply (simp add: absolute_extensor_imp_AR) done lemma AR_imp_retract: fixes S :: "'a::euclidean_space set" assumes "AR S \ closedin (top_of_set U) S" shows "S retract_of U" using AR_imp_absolute_retract assms homeomorphic_refl by blast lemma AR_homeomorphic_AR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "AR T" "S homeomorphic T" shows "AR S" unfolding AR_def by (metis assms AR_imp_absolute_retract homeomorphic_trans [of _ S] homeomorphic_sym) lemma homeomorphic_AR_iff_AR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "S homeomorphic T \ AR S \ AR T" by (metis AR_homeomorphic_AR homeomorphic_sym) lemma ANR_imp_absolute_neighbourhood_extensor: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "ANR S" and contf: "continuous_on T f" and "f ` T \ S" and cloUT: "closedin (top_of_set U) T" obtains V g where "T \ V" "openin (top_of_set U) V" "continuous_on V g" "g ` V \ S" "\x. x \ T \ g x = f x" proof - have "aff_dim S < int (DIM('b \ real))" using aff_dim_le_DIM [of S] by simp then obtain C and S' :: "('b * real) set" where C: "convex C" "C \ {}" and cloCS: "closedin (top_of_set C) S'" and hom: "S homeomorphic S'" by (metis that homeomorphic_closedin_convex) then obtain D where opD: "openin (top_of_set C) D" and "S' retract_of D" using \ANR S\ by (auto simp: ANR_def) then obtain r where "S' \ D" and contr: "continuous_on D r" and "r ` D \ S'" and rid: "\x. x \ S' \ r x = x" by (auto simp: retraction_def retract_of_def) obtain g h where homgh: "homeomorphism S S' g h" using hom by (force simp: homeomorphic_def) have "continuous_on (f ` T) g" by (meson \f ` T \ S\ continuous_on_subset homeomorphism_def homgh) then have contgf: "continuous_on T (g \ f)" by (intro continuous_on_compose contf) have gfTC: "(g \ f) ` T \ C" proof - have "g ` S = S'" by (metis (no_types) homeomorphism_def homgh) then show ?thesis by (metis (no_types) assms(3) cloCS closedin_def image_comp image_mono order.trans topspace_euclidean_subtopology) qed obtain f' where contf': "continuous_on U f'" and "f' ` U \ C" and eq: "\x. x \ T \ f' x = (g \ f) x" by (metis Dugundji [OF C cloUT contgf gfTC]) show ?thesis proof (rule_tac V = "U \ f' -` D" and g = "h \ r \ f'" in that) show "T \ U \ f' -` D" using cloUT closedin_imp_subset \S' \ D\ \f ` T \ S\ eq homeomorphism_image1 homgh by fastforce show ope: "openin (top_of_set U) (U \ f' -` D)" using \f' ` U \ C\ by (auto simp: opD contf' continuous_openin_preimage) have conth: "continuous_on (r ` f' ` (U \ f' -` D)) h" apply (rule continuous_on_subset [of S']) using homeomorphism_def homgh apply blast using \r ` D \ S'\ by blast show "continuous_on (U \ f' -` D) (h \ r \ f')" apply (intro continuous_on_compose conth continuous_on_subset [OF contr] continuous_on_subset [OF contf'], auto) done show "(h \ r \ f') ` (U \ f' -` D) \ S" using \homeomorphism S S' g h\ \f' ` U \ C\ \r ` D \ S'\ by (auto simp: homeomorphism_def) show "\x. x \ T \ (h \ r \ f') x = f x" using \homeomorphism S S' g h\ \f ` T \ S\ eq by (auto simp: rid homeomorphism_def) qed qed corollary ANR_imp_absolute_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" "S homeomorphic S'" and clo: "closedin (top_of_set U) S'" obtains V where "openin (top_of_set U) V" "S' retract_of V" proof - obtain g h where hom: "homeomorphism S S' g h" using assms by (force simp: homeomorphic_def) have h: "continuous_on S' h" " h ` S' \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done from ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo] obtain V h' where "S' \ V" and opUV: "openin (top_of_set U) V" and h': "continuous_on V h'" "h' ` V \ S" and h'h:"\x. x \ S' \ h' x = h x" by (blast intro: ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ h clo]) have "S' retract_of V" proof (simp add: retraction_def retract_of_def, intro exI conjI \S' \ V\) show "continuous_on V (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` V \ S'" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\S'. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed then show ?thesis by (rule that [OF opUV]) qed corollary ANR_imp_absolute_neighbourhood_retract_UNIV: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" and hom: "S homeomorphic S'" and clo: "closed S'" obtains V where "open V" "S' retract_of V" using ANR_imp_absolute_neighbourhood_retract [OF \ANR S\ hom] by (metis clo closed_closedin open_openin subtopology_UNIV) corollary neighbourhood_extension_into_ANR: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ T" and "ANR T" "closed S" obtains V g where "S \ V" "open V" "continuous_on V g" "g ` V \ T" "\x. x \ S \ g x = f x" using ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf fim] by (metis \closed S\ closed_closedin open_openin subtopology_UNIV) lemma absolute_neighbourhood_extensor_imp_ANR: fixes S :: "'a::euclidean_space set" assumes "\f :: 'a * real \ 'a. \U T. \continuous_on T f; f ` T \ S; closedin (top_of_set U) T\ \ \V g. T \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)" shows "ANR S" proof (clarsimp simp: ANR_def) fix U and T :: "('a * real) set" assume "S homeomorphic T" and clo: "closedin (top_of_set U) T" then obtain g h where hom: "homeomorphism S T g h" by (force simp: homeomorphic_def) have h: "continuous_on T h" " h ` T \ S" using hom homeomorphism_def apply blast apply (metis hom equalityE homeomorphism_def) done obtain V h' where "T \ V" and opV: "openin (top_of_set U) V" and h': "continuous_on V h'" "h' ` V \ S" and h'h: "\x\T. h' x = h x" using assms [OF h clo] by blast have [simp]: "T \ U" using clo closedin_imp_subset by auto have "T retract_of V" proof (simp add: retraction_def retract_of_def, intro exI conjI \T \ V\) show "continuous_on V (g \ h')" apply (intro continuous_on_compose h') apply (meson hom continuous_on_subset h' homeomorphism_cont1) done show "(g \ h') ` V \ T" using h' by clarsimp (metis hom subsetD homeomorphism_def imageI) show "\x\T. (g \ h') x = x" by clarsimp (metis h'h hom homeomorphism_def) qed then show "\V. openin (top_of_set U) V \ T retract_of V" using opV by blast qed lemma ANR_eq_absolute_neighbourhood_extensor: fixes S :: "'a::euclidean_space set" shows "ANR S \ (\f :: 'a * real \ 'a. \U T. continuous_on T f \ f ` T \ S \ closedin (top_of_set U) T \ (\V g. T \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x \ T. g x = f x)))" apply (rule iffI) apply (metis ANR_imp_absolute_neighbourhood_extensor) apply (simp add: absolute_neighbourhood_extensor_imp_ANR) done lemma ANR_imp_neighbourhood_retract: fixes S :: "'a::euclidean_space set" assumes "ANR S" "closedin (top_of_set U) S" obtains V where "openin (top_of_set U) V" "S retract_of V" using ANR_imp_absolute_neighbourhood_retract assms homeomorphic_refl by blast lemma ANR_imp_absolute_closed_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ANR S" "S homeomorphic S'" and US': "closedin (top_of_set U) S'" obtains V W where "openin (top_of_set U) V" "closedin (top_of_set U) W" "S' \ V" "V \ W" "S' retract_of W" proof - obtain Z where "openin (top_of_set U) Z" and S'Z: "S' retract_of Z" by (blast intro: assms ANR_imp_absolute_neighbourhood_retract) then have UUZ: "closedin (top_of_set U) (U - Z)" by auto have "S' \ (U - Z) = {}" using \S' retract_of Z\ closedin_retract closedin_subtopology by fastforce then obtain V W where "openin (top_of_set U) V" and "openin (top_of_set U) W" and "S' \ V" "U - Z \ W" "V \ W = {}" using separation_normal_local [OF US' UUZ] by auto moreover have "S' retract_of U - W" apply (rule retract_of_subset [OF S'Z]) using US' \S' \ V\ \V \ W = {}\ closedin_subset apply fastforce using Diff_subset_conv \U - Z \ W\ by blast ultimately show ?thesis apply (rule_tac V=V and W = "U-W" in that) using openin_imp_subset apply force+ done qed lemma ANR_imp_closed_neighbourhood_retract: fixes S :: "'a::euclidean_space set" assumes "ANR S" "closedin (top_of_set U) S" obtains V W where "openin (top_of_set U) V" "closedin (top_of_set U) W" "S \ V" "V \ W" "S retract_of W" by (meson ANR_imp_absolute_closed_neighbourhood_retract assms homeomorphic_refl) lemma ANR_homeomorphic_ANR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ANR T" "S homeomorphic T" shows "ANR S" unfolding ANR_def by (metis assms ANR_imp_absolute_neighbourhood_retract homeomorphic_trans [of _ S] homeomorphic_sym) lemma homeomorphic_ANR_iff_ANR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" shows "S homeomorphic T \ ANR S \ ANR T" by (metis ANR_homeomorphic_ANR homeomorphic_sym) subsection \Analogous properties of ENRs\ lemma ENR_imp_absolute_neighbourhood_retract: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ENR S" and hom: "S homeomorphic S'" and "S' \ U" obtains V where "openin (top_of_set U) V" "S' retract_of V" proof - obtain X where "open X" "S retract_of X" using \ENR S\ by (auto simp: ENR_def) then obtain r where "retraction X S r" by (auto simp: retract_of_def) have "locally compact S'" using retract_of_locally_compact open_imp_locally_compact homeomorphic_local_compactness \S retract_of X\ \open X\ hom by blast then obtain W where UW: "openin (top_of_set U) W" and WS': "closedin (top_of_set W) S'" apply (rule locally_compact_closedin_open) apply (rename_tac W) apply (rule_tac W = "U \ W" in that, blast) by (simp add: \S' \ U\ closedin_limpt) obtain f g where hom: "homeomorphism S S' f g" using assms by (force simp: homeomorphic_def) have contg: "continuous_on S' g" using hom homeomorphism_def by blast moreover have "g ` S' \ S" by (metis hom equalityE homeomorphism_def) ultimately obtain h where conth: "continuous_on W h" and hg: "\x. x \ S' \ h x = g x" using Tietze_unbounded [of S' g W] WS' by blast have "W \ U" using UW openin_open by auto have "S' \ W" using WS' closedin_closed by auto have him: "\x. x \ S' \ h x \ X" by (metis (no_types) \S retract_of X\ hg hom homeomorphism_def image_insert insert_absorb insert_iff retract_of_imp_subset subset_eq) have "S' retract_of (W \ h -` X)" proof (simp add: retraction_def retract_of_def, intro exI conjI) show "S' \ W" "S' \ h -` X" using him WS' closedin_imp_subset by blast+ show "continuous_on (W \ h -` X) (f \ r \ h)" proof (intro continuous_on_compose) show "continuous_on (W \ h -` X) h" by (meson conth continuous_on_subset inf_le1) show "continuous_on (h ` (W \ h -` X)) r" proof - have "h ` (W \ h -` X) \ X" by blast then show "continuous_on (h ` (W \ h -` X)) r" by (meson \retraction X S r\ continuous_on_subset retraction) qed show "continuous_on (r ` h ` (W \ h -` X)) f" apply (rule continuous_on_subset [of S]) using hom homeomorphism_def apply blast apply clarify apply (meson \retraction X S r\ subsetD imageI retraction_def) done qed show "(f \ r \ h) ` (W \ h -` X) \ S'" using \retraction X S r\ hom by (auto simp: retraction_def homeomorphism_def) show "\x\S'. (f \ r \ h) x = x" using \retraction X S r\ hom by (auto simp: retraction_def homeomorphism_def hg) qed then show ?thesis apply (rule_tac V = "W \ h -` X" in that) apply (rule openin_trans [OF _ UW]) using \continuous_on W h\ \open X\ continuous_openin_preimage_eq apply blast+ done qed corollary ENR_imp_absolute_neighbourhood_retract_UNIV: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "ENR S" "S homeomorphic S'" obtains T' where "open T'" "S' retract_of T'" by (metis ENR_imp_absolute_neighbourhood_retract UNIV_I assms(1) assms(2) open_openin subsetI subtopology_UNIV) lemma ENR_homeomorphic_ENR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ENR T" "S homeomorphic T" shows "ENR S" unfolding ENR_def by (meson ENR_imp_absolute_neighbourhood_retract_UNIV assms homeomorphic_sym) lemma homeomorphic_ENR_iff_ENR: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" shows "ENR S \ ENR T" by (meson ENR_homeomorphic_ENR assms homeomorphic_sym) lemma ENR_translation: fixes S :: "'a::euclidean_space set" shows "ENR(image (\x. a + x) S) \ ENR S" by (meson homeomorphic_sym homeomorphic_translation homeomorphic_ENR_iff_ENR) lemma ENR_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "ENR (image f S) \ ENR S" apply (rule homeomorphic_ENR_iff_ENR) using assms homeomorphic_sym linear_homeomorphic_image by auto text \Some relations among the concepts. We also relate AR to being a retract of UNIV, which is often a more convenient proxy in the closed case.\ lemma AR_imp_ANR: "AR S \ ANR S" using ANR_def AR_def by fastforce lemma ENR_imp_ANR: fixes S :: "'a::euclidean_space set" shows "ENR S \ ANR S" apply (simp add: ANR_def) by (metis ENR_imp_absolute_neighbourhood_retract closedin_imp_subset) lemma ENR_ANR: fixes S :: "'a::euclidean_space set" shows "ENR S \ ANR S \ locally compact S" proof assume "ENR S" then have "locally compact S" using ENR_def open_imp_locally_compact retract_of_locally_compact by auto then show "ANR S \ locally compact S" using ENR_imp_ANR \ENR S\ by blast next assume "ANR S \ locally compact S" then have "ANR S" "locally compact S" by auto then obtain T :: "('a * real) set" where "closed T" "S homeomorphic T" using locally_compact_homeomorphic_closed by (metis DIM_prod DIM_real Suc_eq_plus1 lessI) then show "ENR S" using \ANR S\ apply (simp add: ANR_def) apply (drule_tac x=UNIV in spec) apply (drule_tac x=T in spec, clarsimp) apply (meson ENR_def ENR_homeomorphic_ENR open_openin) done qed lemma AR_ANR: fixes S :: "'a::euclidean_space set" shows "AR S \ ANR S \ contractible S \ S \ {}" (is "?lhs = ?rhs") proof assume ?lhs obtain C and S' :: "('a * real) set" where "convex C" "C \ {}" "closedin (top_of_set C) S'" "S homeomorphic S'" apply (rule homeomorphic_closedin_convex [of S, where 'n = "'a * real"]) using aff_dim_le_DIM [of S] by auto with \AR S\ have "contractible S" apply (simp add: AR_def) apply (drule_tac x=C in spec) apply (drule_tac x="S'" in spec, simp) using convex_imp_contractible homeomorphic_contractible_eq retract_of_contractible by fastforce with \AR S\ show ?rhs apply (auto simp: AR_imp_ANR) apply (force simp: AR_def) done next assume ?rhs then obtain a and h:: "real \ 'a \ 'a" where conth: "continuous_on ({0..1} \ S) h" and hS: "h ` ({0..1} \ S) \ S" and [simp]: "\x. h(0, x) = x" and [simp]: "\x. h(1, x) = a" and "ANR S" "S \ {}" by (auto simp: contractible_def homotopic_with_def) then have "a \ S" by (metis all_not_in_conv atLeastAtMost_iff image_subset_iff mem_Sigma_iff order_refl zero_le_one) have "\g. continuous_on W g \ g ` W \ S \ (\x\T. g x = f x)" if f: "continuous_on T f" "f ` T \ S" and WT: "closedin (top_of_set W) T" for W T and f :: "'a \ real \ 'a" proof - obtain U g where "T \ U" and WU: "openin (top_of_set W) U" and contg: "continuous_on U g" and "g ` U \ S" and gf: "\x. x \ T \ g x = f x" using iffD1 [OF ANR_eq_absolute_neighbourhood_extensor \ANR S\, rule_format, OF f WT] by auto have WWU: "closedin (top_of_set W) (W - U)" using WU closedin_diff by fastforce moreover have "(W - U) \ T = {}" using \T \ U\ by auto ultimately obtain V V' where WV': "openin (top_of_set W) V'" and WV: "openin (top_of_set W) V" and "W - U \ V'" "T \ V" "V' \ V = {}" using separation_normal_local [of W "W-U" T] WT by blast then have WVT: "T \ (W - V) = {}" by auto have WWV: "closedin (top_of_set W) (W - V)" using WV closedin_diff by fastforce obtain j :: " 'a \ real \ real" where contj: "continuous_on W j" and j: "\x. x \ W \ j x \ {0..1}" and j0: "\x. x \ W - V \ j x = 1" and j1: "\x. x \ T \ j x = 0" by (rule Urysohn_local [OF WT WWV WVT, of 0 "1::real"]) (auto simp: in_segment) have Weq: "W = (W - V) \ (W - V')" using \V' \ V = {}\ by force show ?thesis proof (intro conjI exI) have *: "continuous_on (W - V') (\x. h (j x, g x))" apply (rule continuous_on_compose2 [OF conth continuous_on_Pair]) apply (rule continuous_on_subset [OF contj Diff_subset]) apply (rule continuous_on_subset [OF contg]) apply (metis Diff_subset_conv Un_commute \W - U \ V'\) using j \g ` U \ S\ \W - U \ V'\ apply fastforce done show "continuous_on W (\x. if x \ W - V then a else h (j x, g x))" apply (subst Weq) apply (rule continuous_on_cases_local) - apply (simp_all add: Weq [symmetric] WWV continuous_on_const *) + apply (simp_all add: Weq [symmetric] WWV *) using WV' closedin_diff apply fastforce apply (auto simp: j0 j1) done next have "h (j (x, y), g (x, y)) \ S" if "(x, y) \ W" "(x, y) \ V" for x y proof - have "j(x, y) \ {0..1}" using j that by blast moreover have "g(x, y) \ S" using \V' \ V = {}\ \W - U \ V'\ \g ` U \ S\ that by fastforce ultimately show ?thesis using hS by blast qed with \a \ S\ \g ` U \ S\ show "(\x. if x \ W - V then a else h (j x, g x)) ` W \ S" by auto next show "\x\T. (if x \ W - V then a else h (j x, g x)) = f x" using \T \ V\ by (auto simp: j0 j1 gf) qed qed then show ?lhs by (simp add: AR_eq_absolute_extensor) qed lemma ANR_retract_of_ANR: fixes S :: "'a::euclidean_space set" assumes "ANR T" "S retract_of T" shows "ANR S" using assms apply (simp add: ANR_eq_absolute_neighbourhood_extensor retract_of_def retraction_def) apply (clarsimp elim!: all_forward) apply (erule impCE, metis subset_trans) apply (clarsimp elim!: ex_forward) apply (rule_tac x="r \ g" in exI) by (metis comp_apply continuous_on_compose continuous_on_subset subsetD imageI image_comp image_mono subset_trans) lemma AR_retract_of_AR: fixes S :: "'a::euclidean_space set" shows "\AR T; S retract_of T\ \ AR S" using ANR_retract_of_ANR AR_ANR retract_of_contractible by fastforce lemma ENR_retract_of_ENR: "\ENR T; S retract_of T\ \ ENR S" by (meson ENR_def retract_of_trans) lemma retract_of_UNIV: fixes S :: "'a::euclidean_space set" shows "S retract_of UNIV \ AR S \ closed S" by (metis AR_ANR AR_imp_retract ENR_def ENR_imp_ANR closed_UNIV closed_closedin contractible_UNIV empty_not_UNIV open_UNIV retract_of_closed retract_of_contractible retract_of_empty(1) subtopology_UNIV) lemma compact_AR: fixes S :: "'a::euclidean_space set" shows "compact S \ AR S \ compact S \ S retract_of UNIV" using compact_imp_closed retract_of_UNIV by blast text \More properties of ARs, ANRs and ENRs\ lemma not_AR_empty [simp]: "\ AR({})" by (auto simp: AR_def) lemma ENR_empty [simp]: "ENR {}" by (simp add: ENR_def) lemma ANR_empty [simp]: "ANR ({} :: 'a::euclidean_space set)" by (simp add: ENR_imp_ANR) lemma convex_imp_AR: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ {}\ \ AR S" apply (rule absolute_extensor_imp_AR) apply (rule Dugundji, assumption+) by blast lemma convex_imp_ANR: fixes S :: "'a::euclidean_space set" shows "convex S \ ANR S" using ANR_empty AR_imp_ANR convex_imp_AR by blast lemma ENR_convex_closed: fixes S :: "'a::euclidean_space set" shows "\closed S; convex S\ \ ENR S" using ENR_def ENR_empty convex_imp_AR retract_of_UNIV by blast lemma AR_UNIV [simp]: "AR (UNIV :: 'a::euclidean_space set)" using retract_of_UNIV by auto lemma ANR_UNIV [simp]: "ANR (UNIV :: 'a::euclidean_space set)" by (simp add: AR_imp_ANR) lemma ENR_UNIV [simp]:"ENR UNIV" using ENR_def by blast lemma AR_singleton: fixes a :: "'a::euclidean_space" shows "AR {a}" using retract_of_UNIV by blast lemma ANR_singleton: fixes a :: "'a::euclidean_space" shows "ANR {a}" by (simp add: AR_imp_ANR AR_singleton) lemma ENR_singleton: "ENR {a}" using ENR_def by blast text \ARs closed under union\ lemma AR_closed_Un_local_aux: fixes U :: "'a::euclidean_space set" assumes "closedin (top_of_set U) S" "closedin (top_of_set U) T" "AR S" "AR T" "AR(S \ T)" shows "(S \ T) retract_of U" proof - have "S \ T \ {}" using assms AR_def by fastforce have "S \ U" "T \ U" using assms by (auto simp: closedin_imp_subset) define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" have US': "closedin (top_of_set U) S'" using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have UT': "closedin (top_of_set U) T'" using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have "S \ S'" using S'_def \S \ U\ setdist_sing_in_set by fastforce have "T \ T'" using T'_def \T \ U\ setdist_sing_in_set by fastforce have "S \ T \ W" "W \ U" using \S \ U\ by (auto simp: W_def setdist_sing_in_set) have "(S \ T) retract_of W" apply (rule AR_imp_absolute_retract [OF \AR(S \ T)\]) apply (simp add: homeomorphic_refl) apply (rule closedin_subset_trans [of U]) apply (simp_all add: assms closedin_Int \S \ T \ W\ \W \ U\) done then obtain r0 where "S \ T \ W" and contr0: "continuous_on W r0" and "r0 ` W \ S \ T" and r0 [simp]: "\x. x \ S \ T \ r0 x = x" by (auto simp: retract_of_def retraction_def) have ST: "x \ W \ x \ S \ x \ T" for x using setdist_eq_0_closedin \S \ T \ {}\ assms by (force simp: W_def setdist_sing_in_set) have "S' \ T' = W" by (auto simp: S'_def T'_def W_def) then have cloUW: "closedin (top_of_set U) W" using closedin_Int US' UT' by blast define r where "r \ \x. if x \ W then r0 x else x" have "r ` (W \ S) \ S" "r ` (W \ T) \ T" using \r0 ` W \ S \ T\ r_def by auto have contr: "continuous_on (W \ (S \ T)) r" unfolding r_def proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) show "closedin (top_of_set (W \ (S \ T))) W" using \S \ U\ \T \ U\ \W \ U\ \closedin (top_of_set U) W\ closedin_subset_trans by fastforce show "closedin (top_of_set (W \ (S \ T))) (S \ T)" by (meson \S \ U\ \T \ U\ \W \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) show "\x. x \ W \ x \ W \ x \ S \ T \ x \ W \ r0 x = x" by (auto simp: ST) qed have cloUWS: "closedin (top_of_set U) (W \ S)" by (simp add: cloUW assms closedin_Un) obtain g where contg: "continuous_on U g" and "g ` U \ S" and geqr: "\x. x \ W \ S \ g x = r x" apply (rule AR_imp_absolute_extensor [OF \AR S\ _ _ cloUWS]) apply (rule continuous_on_subset [OF contr]) using \r ` (W \ S) \ S\ apply auto done have cloUWT: "closedin (top_of_set U) (W \ T)" by (simp add: cloUW assms closedin_Un) obtain h where conth: "continuous_on U h" and "h ` U \ T" and heqr: "\x. x \ W \ T \ h x = r x" apply (rule AR_imp_absolute_extensor [OF \AR T\ _ _ cloUWT]) apply (rule continuous_on_subset [OF contr]) using \r ` (W \ T) \ T\ apply auto done have "U = S' \ T'" by (force simp: S'_def T'_def) then have cont: "continuous_on U (\x. if x \ S' then g x else h x)" apply (rule ssubst) apply (rule continuous_on_cases_local) using US' UT' \S' \ T' = W\ \U = S' \ T'\ contg conth continuous_on_subset geqr heqr apply auto done have UST: "(\x. if x \ S' then g x else h x) ` U \ S \ T" using \g ` U \ S\ \h ` U \ T\ by auto show ?thesis apply (simp add: retract_of_def retraction_def \S \ U\ \T \ U\) apply (rule_tac x="\x. if x \ S' then g x else h x" in exI) apply (intro conjI cont UST) by (metis IntI ST Un_iff \S \ S'\ \S' \ T' = W\ \T \ T'\ subsetD geqr heqr r0 r_def) qed lemma AR_closed_Un_local: fixes S :: "'a::euclidean_space set" assumes STS: "closedin (top_of_set (S \ T)) S" and STT: "closedin (top_of_set (S \ T)) T" and "AR S" "AR T" "AR(S \ T)" shows "AR(S \ T)" proof - have "C retract_of U" if hom: "S \ T homeomorphic C" and UC: "closedin (top_of_set U) C" for U and C :: "('a * real) set" proof - obtain f g where hom: "homeomorphism (S \ T) C f g" using hom by (force simp: homeomorphic_def) have US: "closedin (top_of_set U) (C \ g -` S)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) using hom homeomorphism_def apply blast apply (metis hom homeomorphism_def set_eq_subset) done have UT: "closedin (top_of_set U) (C \ g -` T)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) using hom homeomorphism_def apply blast apply (metis hom homeomorphism_def set_eq_subset) done have ARS: "AR (C \ g -` S)" apply (rule AR_homeomorphic_AR [OF \AR S\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ART: "AR (C \ g -` T)" apply (rule AR_homeomorphic_AR [OF \AR T\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ARI: "AR ((C \ g -` S) \ (C \ g -` T))" apply (rule AR_homeomorphic_AR [OF \AR (S \ T)\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have "C = (C \ g -` S) \ (C \ g -` T)" using hom by (auto simp: homeomorphism_def) then show ?thesis by (metis AR_closed_Un_local_aux [OF US UT ARS ART ARI]) qed then show ?thesis by (force simp: AR_def) qed corollary AR_closed_Un: fixes S :: "'a::euclidean_space set" shows "\closed S; closed T; AR S; AR T; AR (S \ T)\ \ AR (S \ T)" by (metis AR_closed_Un_local_aux closed_closedin retract_of_UNIV subtopology_UNIV) text \ANRs closed under union\ lemma ANR_closed_Un_local_aux: fixes U :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "ANR S" "ANR T" "ANR(S \ T)" obtains V where "openin (top_of_set U) V" "(S \ T) retract_of V" proof (cases "S = {} \ T = {}") case True with assms that show ?thesis by (metis ANR_imp_neighbourhood_retract Un_commute inf_bot_right sup_inf_absorb) next case False then have [simp]: "S \ {}" "T \ {}" by auto have "S \ U" "T \ U" using assms by (auto simp: closedin_imp_subset) define S' where "S' \ {x \ U. setdist {x} S \ setdist {x} T}" define T' where "T' \ {x \ U. setdist {x} T \ setdist {x} S}" define W where "W \ {x \ U. setdist {x} S = setdist {x} T}" have cloUS': "closedin (top_of_set U) S'" using continuous_closedin_preimage [of U "\x. setdist {x} S - setdist {x} T" "{..0}"] by (simp add: S'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have cloUT': "closedin (top_of_set U) T'" using continuous_closedin_preimage [of U "\x. setdist {x} T - setdist {x} S" "{..0}"] by (simp add: T'_def vimage_def Collect_conj_eq continuous_on_diff continuous_on_setdist) have "S \ S'" using S'_def \S \ U\ setdist_sing_in_set by fastforce have "T \ T'" using T'_def \T \ U\ setdist_sing_in_set by fastforce have "S' \ T' = U" by (auto simp: S'_def T'_def) have "W \ S'" by (simp add: Collect_mono S'_def W_def) have "W \ T'" by (simp add: Collect_mono T'_def W_def) have ST_W: "S \ T \ W" and "W \ U" using \S \ U\ by (force simp: W_def setdist_sing_in_set)+ have "S' \ T' = W" by (auto simp: S'_def T'_def W_def) then have cloUW: "closedin (top_of_set U) W" using closedin_Int cloUS' cloUT' by blast obtain W' W0 where "openin (top_of_set W) W'" and cloWW0: "closedin (top_of_set W) W0" and "S \ T \ W'" "W' \ W0" and ret: "(S \ T) retract_of W0" apply (rule ANR_imp_closed_neighbourhood_retract [OF \ANR(S \ T)\]) apply (rule closedin_subset_trans [of U, OF _ ST_W \W \ U\]) apply (blast intro: assms)+ done then obtain U0 where opeUU0: "openin (top_of_set U) U0" and U0: "S \ T \ U0" "U0 \ W \ W0" unfolding openin_open using \W \ U\ by blast have "W0 \ U" using \W \ U\ cloWW0 closedin_subset by fastforce obtain r0 where "S \ T \ W0" and contr0: "continuous_on W0 r0" and "r0 ` W0 \ S \ T" and r0 [simp]: "\x. x \ S \ T \ r0 x = x" using ret by (force simp: retract_of_def retraction_def) have ST: "x \ W \ x \ S \ x \ T" for x using assms by (auto simp: W_def setdist_sing_in_set dest!: setdist_eq_0_closedin) define r where "r \ \x. if x \ W0 then r0 x else x" have "r ` (W0 \ S) \ S" "r ` (W0 \ T) \ T" using \r0 ` W0 \ S \ T\ r_def by auto have contr: "continuous_on (W0 \ (S \ T)) r" unfolding r_def proof (rule continuous_on_cases_local [OF _ _ contr0 continuous_on_id]) show "closedin (top_of_set (W0 \ (S \ T))) W0" apply (rule closedin_subset_trans [of U]) using cloWW0 cloUW closedin_trans \W0 \ U\ \S \ U\ \T \ U\ apply blast+ done show "closedin (top_of_set (W0 \ (S \ T))) (S \ T)" by (meson \S \ U\ \T \ U\ \W0 \ U\ assms closedin_Un closedin_subset_trans sup.bounded_iff sup.cobounded2) show "\x. x \ W0 \ x \ W0 \ x \ S \ T \ x \ W0 \ r0 x = x" using ST cloWW0 closedin_subset by fastforce qed have cloS'WS: "closedin (top_of_set S') (W0 \ S)" by (meson closedin_subset_trans US cloUS' \S \ S'\ \W \ S'\ cloUW cloWW0 closedin_Un closedin_imp_subset closedin_trans) obtain W1 g where "W0 \ S \ W1" and contg: "continuous_on W1 g" and opeSW1: "openin (top_of_set S') W1" and "g ` W1 \ S" and geqr: "\x. x \ W0 \ S \ g x = r x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ _ \r ` (W0 \ S) \ S\ cloS'WS]) apply (rule continuous_on_subset [OF contr], blast+) done have cloT'WT: "closedin (top_of_set T') (W0 \ T)" by (meson closedin_subset_trans UT cloUT' \T \ T'\ \W \ T'\ cloUW cloWW0 closedin_Un closedin_imp_subset closedin_trans) obtain W2 h where "W0 \ T \ W2" and conth: "continuous_on W2 h" and opeSW2: "openin (top_of_set T') W2" and "h ` W2 \ T" and heqr: "\x. x \ W0 \ T \ h x = r x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ _ \r ` (W0 \ T) \ T\ cloT'WT]) apply (rule continuous_on_subset [OF contr], blast+) done have "S' \ T' = W" by (force simp: S'_def T'_def W_def) obtain O1 O2 where "open O1" "W1 = S' \ O1" "open O2" "W2 = T' \ O2" using opeSW1 opeSW2 by (force simp: openin_open) show ?thesis proof have eq: "W1 - (W - U0) \ (W2 - (W - U0)) = ((U - T') \ O1 \ (U - S') \ O2 \ U \ O1 \ O2) - (W - U0)" using \U0 \ W \ W0\ \W0 \ S \ W1\ \W0 \ T \ W2\ by (auto simp: \S' \ T' = U\ [symmetric] \S' \ T' = W\ [symmetric] \W1 = S' \ O1\ \W2 = T' \ O2\) show "openin (top_of_set U) (W1 - (W - U0) \ (W2 - (W - U0)))" apply (subst eq) apply (intro openin_Un openin_Int_open openin_diff closedin_diff cloUW opeUU0 cloUS' cloUT' \open O1\ \open O2\, simp_all) done have cloW1: "closedin (top_of_set (W1 - (W - U0) \ (W2 - (W - U0)))) (W1 - (W - U0))" using cloUS' apply (simp add: closedin_closed) apply (erule ex_forward) using U0 \W0 \ S \ W1\ apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) done have cloW2: "closedin (top_of_set (W1 - (W - U0) \ (W2 - (W - U0)))) (W2 - (W - U0))" using cloUT' apply (simp add: closedin_closed) apply (erule ex_forward) using U0 \W0 \ T \ W2\ apply (auto simp: \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = U\ [symmetric]\S' \ T' = W\ [symmetric]) done have *: "\x\S \ T. (if x \ S' then g x else h x) = x" using ST \S' \ T' = W\ cloT'WT closedin_subset geqr heqr apply (auto simp: r_def, fastforce) using \S \ S'\ \T \ T'\ \W0 \ S \ W1\ \W1 = S' \ O1\ by auto have "\r. continuous_on (W1 - (W - U0) \ (W2 - (W - U0))) r \ r ` (W1 - (W - U0) \ (W2 - (W - U0))) \ S \ T \ (\x\S \ T. r x = x)" apply (rule_tac x = "\x. if x \ S' then g x else h x" in exI) apply (intro conjI *) apply (rule continuous_on_cases_local [OF cloW1 cloW2 continuous_on_subset [OF contg] continuous_on_subset [OF conth]]) using \W1 = S' \ O1\ \W2 = T' \ O2\ \S' \ T' = W\ \g ` W1 \ S\ \h ` W2 \ T\ apply auto using \U0 \ W \ W0\ \W0 \ S \ W1\ apply (fastforce simp add: geqr heqr)+ done then show "S \ T retract_of W1 - (W - U0) \ (W2 - (W - U0))" using \W0 \ S \ W1\ \W0 \ T \ W2\ ST opeUU0 U0 by (auto simp: retract_of_def retraction_def) qed qed lemma ANR_closed_Un_local: fixes S :: "'a::euclidean_space set" assumes STS: "closedin (top_of_set (S \ T)) S" and STT: "closedin (top_of_set (S \ T)) T" and "ANR S" "ANR T" "ANR(S \ T)" shows "ANR(S \ T)" proof - have "\T. openin (top_of_set U) T \ C retract_of T" if hom: "S \ T homeomorphic C" and UC: "closedin (top_of_set U) C" for U and C :: "('a * real) set" proof - obtain f g where hom: "homeomorphism (S \ T) C f g" using hom by (force simp: homeomorphic_def) have US: "closedin (top_of_set U) (C \ g -` S)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STS]) using hom [unfolded homeomorphism_def] apply blast apply (metis hom homeomorphism_def set_eq_subset) done have UT: "closedin (top_of_set U) (C \ g -` T)" apply (rule closedin_trans [OF _ UC]) apply (rule continuous_closedin_preimage_gen [OF _ _ STT]) using hom [unfolded homeomorphism_def] apply blast apply (metis hom homeomorphism_def set_eq_subset) done have ANRS: "ANR (C \ g -` S)" apply (rule ANR_homeomorphic_ANR [OF \ANR S\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ANRT: "ANR (C \ g -` T)" apply (rule ANR_homeomorphic_ANR [OF \ANR T\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have ANRI: "ANR ((C \ g -` S) \ (C \ g -` T))" apply (rule ANR_homeomorphic_ANR [OF \ANR (S \ T)\]) apply (simp add: homeomorphic_def) apply (rule_tac x=g in exI) apply (rule_tac x=f in exI) using hom apply (auto simp: homeomorphism_def elim!: continuous_on_subset) apply (rule_tac x="f x" in image_eqI, auto) done have "C = (C \ g -` S) \ (C \ g -` T)" using hom by (auto simp: homeomorphism_def) then show ?thesis by (metis ANR_closed_Un_local_aux [OF US UT ANRS ANRT ANRI]) qed then show ?thesis by (auto simp: ANR_def) qed corollary ANR_closed_Un: fixes S :: "'a::euclidean_space set" shows "\closed S; closed T; ANR S; ANR T; ANR (S \ T)\ \ ANR (S \ T)" by (simp add: ANR_closed_Un_local closedin_def diff_eq open_Compl openin_open_Int) lemma ANR_openin: fixes S :: "'a::euclidean_space set" assumes "ANR T" and opeTS: "openin (top_of_set T) S" shows "ANR S" proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) fix f :: "'a \ real \ 'a" and U C assume contf: "continuous_on C f" and fim: "f ` C \ S" and cloUC: "closedin (top_of_set U) C" have "f ` C \ T" using fim opeTS openin_imp_subset by blast obtain W g where "C \ W" and UW: "openin (top_of_set U) W" and contg: "continuous_on W g" and gim: "g ` W \ T" and geq: "\x. x \ C \ g x = f x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf \f ` C \ T\ cloUC]) using fim by auto show "\V g. C \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ (\x\C. g x = f x)" proof (intro exI conjI) show "C \ W \ g -` S" using \C \ W\ fim geq by blast show "openin (top_of_set U) (W \ g -` S)" by (metis (mono_tags, lifting) UW contg continuous_openin_preimage gim opeTS openin_trans) show "continuous_on (W \ g -` S) g" by (blast intro: continuous_on_subset [OF contg]) show "g ` (W \ g -` S) \ S" using gim by blast show "\x\C. g x = f x" using geq by blast qed qed lemma ENR_openin: fixes S :: "'a::euclidean_space set" assumes "ENR T" and opeTS: "openin (top_of_set T) S" shows "ENR S" using assms apply (simp add: ENR_ANR) using ANR_openin locally_open_subset by blast lemma ANR_neighborhood_retract: fixes S :: "'a::euclidean_space set" assumes "ANR U" "S retract_of T" "openin (top_of_set U) T" shows "ANR S" using ANR_openin ANR_retract_of_ANR assms by blast lemma ENR_neighborhood_retract: fixes S :: "'a::euclidean_space set" assumes "ENR U" "S retract_of T" "openin (top_of_set U) T" shows "ENR S" using ENR_openin ENR_retract_of_ENR assms by blast lemma ANR_rel_interior: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(rel_interior S)" by (blast intro: ANR_openin openin_set_rel_interior) lemma ANR_delete: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(S - {a})" by (blast intro: ANR_openin openin_delete openin_subtopology_self) lemma ENR_rel_interior: fixes S :: "'a::euclidean_space set" shows "ENR S \ ENR(rel_interior S)" by (blast intro: ENR_openin openin_set_rel_interior) lemma ENR_delete: fixes S :: "'a::euclidean_space set" shows "ENR S \ ENR(S - {a})" by (blast intro: ENR_openin openin_delete openin_subtopology_self) lemma open_imp_ENR: "open S \ ENR S" using ENR_def by blast lemma open_imp_ANR: fixes S :: "'a::euclidean_space set" shows "open S \ ANR S" by (simp add: ENR_imp_ANR open_imp_ENR) lemma ANR_ball [iff]: fixes a :: "'a::euclidean_space" shows "ANR(ball a r)" by (simp add: convex_imp_ANR) lemma ENR_ball [iff]: "ENR(ball a r)" by (simp add: open_imp_ENR) lemma AR_ball [simp]: fixes a :: "'a::euclidean_space" shows "AR(ball a r) \ 0 < r" by (auto simp: AR_ANR convex_imp_contractible) lemma ANR_cball [iff]: fixes a :: "'a::euclidean_space" shows "ANR(cball a r)" by (simp add: convex_imp_ANR) lemma ENR_cball: fixes a :: "'a::euclidean_space" shows "ENR(cball a r)" using ENR_convex_closed by blast lemma AR_cball [simp]: fixes a :: "'a::euclidean_space" shows "AR(cball a r) \ 0 \ r" by (auto simp: AR_ANR convex_imp_contractible) lemma ANR_box [iff]: fixes a :: "'a::euclidean_space" shows "ANR(cbox a b)" "ANR(box a b)" by (auto simp: convex_imp_ANR open_imp_ANR) lemma ENR_box [iff]: fixes a :: "'a::euclidean_space" shows "ENR(cbox a b)" "ENR(box a b)" apply (simp add: ENR_convex_closed closed_cbox) by (simp add: open_box open_imp_ENR) lemma AR_box [simp]: "AR(cbox a b) \ cbox a b \ {}" "AR(box a b) \ box a b \ {}" by (auto simp: AR_ANR convex_imp_contractible) lemma ANR_interior: fixes S :: "'a::euclidean_space set" shows "ANR(interior S)" by (simp add: open_imp_ANR) lemma ENR_interior: fixes S :: "'a::euclidean_space set" shows "ENR(interior S)" by (simp add: open_imp_ENR) lemma AR_imp_contractible: fixes S :: "'a::euclidean_space set" shows "AR S \ contractible S" by (simp add: AR_ANR) lemma ENR_imp_locally_compact: fixes S :: "'a::euclidean_space set" shows "ENR S \ locally compact S" by (simp add: ENR_ANR) lemma ANR_imp_locally_path_connected: fixes S :: "'a::euclidean_space set" assumes "ANR S" shows "locally path_connected S" proof - obtain U and T :: "('a \ real) set" where "convex U" "U \ {}" and UT: "closedin (top_of_set U) T" and "S homeomorphic T" apply (rule homeomorphic_closedin_convex [of S]) using aff_dim_le_DIM [of S] apply auto done then have "locally path_connected T" by (meson ANR_imp_absolute_neighbourhood_retract assms convex_imp_locally_path_connected locally_open_subset retract_of_locally_path_connected) then have S: "locally path_connected S" if "openin (top_of_set U) V" "T retract_of V" "U \ {}" for V using \S homeomorphic T\ homeomorphic_locally homeomorphic_path_connectedness by blast show ?thesis using assms apply (clarsimp simp: ANR_def) apply (drule_tac x=U in spec) apply (drule_tac x=T in spec) using \S homeomorphic T\ \U \ {}\ UT apply (blast intro: S) done qed lemma ANR_imp_locally_connected: fixes S :: "'a::euclidean_space set" assumes "ANR S" shows "locally connected S" using locally_path_connected_imp_locally_connected ANR_imp_locally_path_connected assms by auto lemma AR_imp_locally_path_connected: fixes S :: "'a::euclidean_space set" assumes "AR S" shows "locally path_connected S" by (simp add: ANR_imp_locally_path_connected AR_imp_ANR assms) lemma AR_imp_locally_connected: fixes S :: "'a::euclidean_space set" assumes "AR S" shows "locally connected S" using ANR_imp_locally_connected AR_ANR assms by blast lemma ENR_imp_locally_path_connected: fixes S :: "'a::euclidean_space set" assumes "ENR S" shows "locally path_connected S" by (simp add: ANR_imp_locally_path_connected ENR_imp_ANR assms) lemma ENR_imp_locally_connected: fixes S :: "'a::euclidean_space set" assumes "ENR S" shows "locally connected S" using ANR_imp_locally_connected ENR_ANR assms by blast lemma ANR_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ANR S" "ANR T" shows "ANR(S \ T)" proof (clarsimp simp only: ANR_eq_absolute_neighbourhood_extensor) fix f :: " ('a \ 'b) \ real \ 'a \ 'b" and U C assume "continuous_on C f" and fim: "f ` C \ S \ T" and cloUC: "closedin (top_of_set U) C" have contf1: "continuous_on C (fst \ f)" by (simp add: \continuous_on C f\ continuous_on_fst) obtain W1 g where "C \ W1" and UW1: "openin (top_of_set U) W1" and contg: "continuous_on W1 g" and gim: "g ` W1 \ S" and geq: "\x. x \ C \ g x = (fst \ f) x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR S\ contf1 _ cloUC]) using fim apply auto done have contf2: "continuous_on C (snd \ f)" by (simp add: \continuous_on C f\ continuous_on_snd) obtain W2 h where "C \ W2" and UW2: "openin (top_of_set U) W2" and conth: "continuous_on W2 h" and him: "h ` W2 \ T" and heq: "\x. x \ C \ h x = (snd \ f) x" apply (rule ANR_imp_absolute_neighbourhood_extensor [OF \ANR T\ contf2 _ cloUC]) using fim apply auto done show "\V g. C \ V \ openin (top_of_set U) V \ continuous_on V g \ g ` V \ S \ T \ (\x\C. g x = f x)" proof (intro exI conjI) show "C \ W1 \ W2" by (simp add: \C \ W1\ \C \ W2\) show "openin (top_of_set U) (W1 \ W2)" by (simp add: UW1 UW2 openin_Int) show "continuous_on (W1 \ W2) (\x. (g x, h x))" by (metis (no_types) contg conth continuous_on_Pair continuous_on_subset inf_commute inf_le1) show "(\x. (g x, h x)) ` (W1 \ W2) \ S \ T" using gim him by blast show "(\x\C. (g x, h x) = f x)" using geq heq by auto qed qed lemma AR_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "AR S" "AR T" shows "AR(S \ T)" using assms by (simp add: AR_ANR ANR_Times contractible_Times) subsection\<^marker>\tag unimportant\\Retracts and intervals in ordered euclidean space\ lemma ANR_interval [iff]: fixes a :: "'a::ordered_euclidean_space" shows "ANR{a..b}" by (simp add: interval_cbox) lemma ENR_interval [iff]: fixes a :: "'a::ordered_euclidean_space" shows "ENR{a..b}" by (auto simp: interval_cbox) subsection \More advanced properties of ANRs and ENRs\ lemma ENR_rel_frontier_convex: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" shows "ENR(rel_frontier S)" proof (cases "S = {}") case True then show ?thesis by simp next case False with assms have "rel_interior S \ {}" by (simp add: rel_interior_eq_empty) then obtain a where a: "a \ rel_interior S" by auto have ahS: "affine hull S - {a} \ {x. closest_point (affine hull S) x \ a}" by (auto simp: closest_point_self) have "rel_frontier S retract_of affine hull S - {a}" by (simp add: assms a rel_frontier_retract_of_punctured_affine_hull) also have "\ retract_of {x. closest_point (affine hull S) x \ a}" apply (simp add: retract_of_def retraction_def ahS) apply (rule_tac x="closest_point (affine hull S)" in exI) apply (auto simp: False closest_point_self affine_imp_convex closest_point_in_set continuous_on_closest_point) done finally have "rel_frontier S retract_of {x. closest_point (affine hull S) x \ a}" . moreover have "openin (top_of_set UNIV) (UNIV \ closest_point (affine hull S) -` (- {a}))" apply (rule continuous_openin_preimage_gen) apply (auto simp: False affine_imp_convex continuous_on_closest_point) done ultimately show ?thesis unfolding ENR_def apply (rule_tac x = "closest_point (affine hull S) -` (- {a})" in exI) apply (simp add: vimage_def) done qed lemma ANR_rel_frontier_convex: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" shows "ANR(rel_frontier S)" by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms) lemma ENR_closedin_Un_local: fixes S :: "'a::euclidean_space set" shows "\ENR S; ENR T; ENR(S \ T); closedin (top_of_set (S \ T)) S; closedin (top_of_set (S \ T)) T\ \ ENR(S \ T)" by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un) lemma ENR_closed_Un: fixes S :: "'a::euclidean_space set" shows "\closed S; closed T; ENR S; ENR T; ENR(S \ T)\ \ ENR(S \ T)" by (auto simp: closed_subset ENR_closedin_Un_local) lemma absolute_retract_Un: fixes S :: "'a::euclidean_space set" shows "\S retract_of UNIV; T retract_of UNIV; (S \ T) retract_of UNIV\ \ (S \ T) retract_of UNIV" by (meson AR_closed_Un_local_aux closed_subset retract_of_UNIV retract_of_imp_subset) lemma retract_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" and Un: "(S \ T) retract_of U" and Int: "(S \ T) retract_of T" shows "S retract_of U" proof - obtain r where r: "continuous_on T r" "r ` T \ S \ T" "\x\S \ T. r x = x" using Int by (auto simp: retraction_def retract_of_def) have "S retract_of S \ T" unfolding retraction_def retract_of_def proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then x else r x)" apply (rule continuous_on_cases_local [OF clS clT]) - using r by (auto simp: continuous_on_id) + using r by (auto) qed (use r in auto) also have "\ retract_of U" by (rule Un) finally show ?thesis . qed lemma AR_from_Un_Int_local: fixes S :: "'a::euclidean_space set" assumes clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" and Un: "AR(S \ T)" and Int: "AR(S \ T)" shows "AR S" apply (rule AR_retract_of_AR [OF Un]) by (meson AR_imp_retract clS clT closedin_closed_subset local.Int retract_from_Un_Int retract_of_refl sup_ge2) lemma AR_from_Un_Int_local': fixes S :: "'a::euclidean_space set" assumes "closedin (top_of_set (S \ T)) S" and "closedin (top_of_set (S \ T)) T" and "AR(S \ T)" "AR(S \ T)" shows "AR T" using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute) lemma AR_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes clo: "closed S" "closed T" and Un: "AR(S \ T)" and Int: "AR(S \ T)" shows "AR S" by (metis AR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) lemma ANR_from_Un_Int_local: fixes S :: "'a::euclidean_space set" assumes clS: "closedin (top_of_set (S \ T)) S" and clT: "closedin (top_of_set (S \ T)) T" and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" shows "ANR S" proof - obtain V where clo: "closedin (top_of_set (S \ T)) (S \ T)" and ope: "openin (top_of_set (S \ T)) V" and ret: "S \ T retract_of V" using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int) then obtain r where r: "continuous_on V r" and rim: "r ` V \ S \ T" and req: "\x\S \ T. r x = x" by (auto simp: retraction_def retract_of_def) have Vsub: "V \ S \ T" by (meson ope openin_contains_cball) have Vsup: "S \ T \ V" by (simp add: retract_of_imp_subset ret) then have eq: "S \ V = ((S \ T) - T) \ V" by auto have eq': "S \ V = S \ (V \ T)" using Vsub by blast have "continuous_on (S \ V \ T) (\x. if x \ S then x else r x)" proof (rule continuous_on_cases_local) show "closedin (top_of_set (S \ V \ T)) S" using clS closedin_subset_trans inf.boundedE by blast show "closedin (top_of_set (S \ V \ T)) (V \ T)" using clT Vsup by (auto simp: closedin_closed) show "continuous_on (V \ T) r" by (meson Int_lower1 continuous_on_subset r) qed (use req continuous_on_id in auto) with rim have "S retract_of S \ V" unfolding retraction_def retract_of_def apply (rule_tac x="\x. if x \ S then x else r x" in exI) apply (auto simp: eq') done then show ?thesis using ANR_neighborhood_retract [OF Un] using \S \ V = S \ T - T \ V\ clT ope by fastforce qed lemma ANR_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes clo: "closed S" "closed T" and Un: "ANR(S \ T)" and Int: "ANR(S \ T)" shows "ANR S" by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest) lemma ANR_finite_Union_convex_closed: fixes \ :: "'a::euclidean_space set set" assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" shows "ANR(\\)" proof - have "ANR(\\)" if "card \ < n" for n using assms that proof (induction n arbitrary: \) case 0 then show ?case by simp next case (Suc n) have "ANR(\\)" if "finite \" "\ \ \" for \ using that proof (induction \) case empty then show ?case by simp next case (insert C \) have "ANR (C \ \\)" proof (rule ANR_closed_Un) show "ANR (C \ \\)" unfolding Int_Union proof (rule Suc) show "finite ((\) C ` \)" by (simp add: insert.hyps(1)) show "\Ca. Ca \ (\) C ` \ \ closed Ca" by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2) show "\Ca. Ca \ (\) C ` \ \ convex Ca" by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE) show "card ((\) C ` \) < n" proof - have "card \ \ n" by (meson Suc.prems(4) not_less not_less_eq) then show ?thesis by (metis Suc.prems(1) card_image_le card_seteq insert.hyps insert.prems insert_subset le_trans not_less) qed qed show "closed (\\)" using Suc.prems(2) insert.hyps(1) insert.prems by blast qed (use Suc.prems convex_imp_ANR insert.prems insert.IH in auto) then show ?case by simp qed then show ?case using Suc.prems(1) by blast qed then show ?thesis by blast qed lemma finite_imp_ANR: fixes S :: "'a::euclidean_space set" assumes "finite S" shows "ANR S" proof - have "ANR(\x \ S. {x})" by (blast intro: ANR_finite_Union_convex_closed assms) then show ?thesis by simp qed lemma ANR_insert: fixes S :: "'a::euclidean_space set" assumes "ANR S" "closed S" shows "ANR(insert a S)" by (metis ANR_closed_Un ANR_empty ANR_singleton Diff_disjoint Diff_insert_absorb assms closed_singleton insert_absorb insert_is_Un) lemma ANR_path_component_ANR: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(path_component_set S x)" using ANR_imp_locally_path_connected ANR_openin openin_path_component_locally_path_connected by blast lemma ANR_connected_component_ANR: fixes S :: "'a::euclidean_space set" shows "ANR S \ ANR(connected_component_set S x)" by (metis ANR_openin openin_connected_component_locally_connected ANR_imp_locally_connected) lemma ANR_component_ANR: fixes S :: "'a::euclidean_space set" assumes "ANR S" "c \ components S" shows "ANR c" by (metis ANR_connected_component_ANR assms componentsE) subsection\Original ANR material, now for ENRs\ lemma ENR_bounded: fixes S :: "'a::euclidean_space set" assumes "bounded S" shows "ENR S \ (\U. open U \ bounded U \ S retract_of U)" (is "?lhs = ?rhs") proof obtain r where "0 < r" and r: "S \ ball 0 r" using bounded_subset_ballD assms by blast assume ?lhs then show ?rhs apply (clarsimp simp: ENR_def) apply (rule_tac x="ball 0 r \ U" in exI, auto) using r retract_of_imp_subset retract_of_subset by fastforce next assume ?rhs then show ?lhs using ENR_def by blast qed lemma absolute_retract_imp_AR_gen: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "S retract_of T" "convex T" "T \ {}" "S homeomorphic S'" "closedin (top_of_set U) S'" shows "S' retract_of U" proof - have "AR T" by (simp add: assms convex_imp_AR) then have "AR S" using AR_retract_of_AR assms by auto then show ?thesis using assms AR_imp_absolute_retract by metis qed lemma absolute_retract_imp_AR: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "S retract_of UNIV" "S homeomorphic S'" "closed S'" shows "S' retract_of UNIV" using AR_imp_absolute_retract_UNIV assms retract_of_UNIV by blast lemma homeomorphic_compact_arness: fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set" assumes "S homeomorphic S'" shows "compact S \ S retract_of UNIV \ compact S' \ S' retract_of UNIV" using assms homeomorphic_compactness apply auto apply (meson assms compact_AR homeomorphic_AR_iff_AR homeomorphic_compactness)+ done lemma absolute_retract_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes "(S \ T) retract_of UNIV" "(S \ T) retract_of UNIV" "closed S" "closed T" shows "S retract_of UNIV" using AR_from_Un_Int assms retract_of_UNIV by auto lemma ENR_from_Un_Int_gen: fixes S :: "'a::euclidean_space set" assumes "closedin (top_of_set (S \ T)) S" "closedin (top_of_set (S \ T)) T" "ENR(S \ T)" "ENR(S \ T)" shows "ENR S" apply (simp add: ENR_ANR) using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast lemma ENR_from_Un_Int: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "ENR(S \ T)" "ENR(S \ T)" shows "ENR S" by (meson ENR_from_Un_Int_gen assms closed_subset sup_ge1 sup_ge2) lemma ENR_finite_Union_convex_closed: fixes \ :: "'a::euclidean_space set set" assumes \: "finite \" and clo: "\C. C \ \ \ closed C" and con: "\C. C \ \ \ convex C" shows "ENR(\ \)" by (simp add: ENR_ANR ANR_finite_Union_convex_closed \ clo closed_Union closed_imp_locally_compact con) lemma finite_imp_ENR: fixes S :: "'a::euclidean_space set" shows "finite S \ ENR S" by (simp add: ENR_ANR finite_imp_ANR finite_imp_closed closed_imp_locally_compact) lemma ENR_insert: fixes S :: "'a::euclidean_space set" assumes "closed S" "ENR S" shows "ENR(insert a S)" proof - have "ENR ({a} \ S)" by (metis ANR_insert ENR_ANR Un_commute Un_insert_right assms closed_imp_locally_compact closed_insert sup_bot_right) then show ?thesis by auto qed lemma ENR_path_component_ENR: fixes S :: "'a::euclidean_space set" assumes "ENR S" shows "ENR(path_component_set S x)" by (metis ANR_imp_locally_path_connected ENR_empty ENR_imp_ANR ENR_openin assms locally_path_connected_2 openin_subtopology_self path_component_eq_empty) (*UNUSED lemma ENR_Times: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "ENR S" "ENR T" shows "ENR(S \ T)" using assms apply (simp add: ENR_ANR ANR_Times) thm locally_compact_Times oops SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);; *) subsection\Finally, spheres are ANRs and ENRs\ lemma absolute_retract_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set" assumes "S homeomorphic U" "S \ {}" "S \ T" "convex U" "compact U" shows "S retract_of T" by (metis UNIV_I assms compact_AR convex_imp_AR homeomorphic_AR_iff_AR homeomorphic_compactness homeomorphic_empty(1) retract_of_subset subsetI) lemma frontier_retract_of_punctured_universe: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ interior S" shows "(frontier S) retract_of (- {a})" using rel_frontier_retract_of_punctured_affine_hull by (metis Compl_eq_Diff_UNIV affine_hull_nonempty_interior assms empty_iff rel_frontier_frontier rel_interior_nonempty_interior) lemma sphere_retract_of_punctured_universe_gen: fixes a :: "'a::euclidean_space" assumes "b \ ball a r" shows "sphere a r retract_of (- {b})" proof - have "frontier (cball a r) retract_of (- {b})" apply (rule frontier_retract_of_punctured_universe) using assms by auto then show ?thesis by simp qed lemma sphere_retract_of_punctured_universe: fixes a :: "'a::euclidean_space" assumes "0 < r" shows "sphere a r retract_of (- {a})" by (simp add: assms sphere_retract_of_punctured_universe_gen) lemma ENR_sphere: fixes a :: "'a::euclidean_space" shows "ENR(sphere a r)" proof (cases "0 < r") case True then have "sphere a r retract_of -{a}" by (simp add: sphere_retract_of_punctured_universe) with open_delete show ?thesis by (auto simp: ENR_def) next case False then show ?thesis using finite_imp_ENR by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial) qed corollary\<^marker>\tag unimportant\ ANR_sphere: fixes a :: "'a::euclidean_space" shows "ANR(sphere a r)" by (simp add: ENR_imp_ANR ENR_sphere) subsection\Spheres are connected, etc\ lemma locally_path_connected_sphere_gen: fixes S :: "'a::euclidean_space set" assumes "bounded S" and "convex S" shows "locally path_connected (rel_frontier S)" proof (cases "rel_interior S = {}") case True with assms show ?thesis by (simp add: rel_interior_eq_empty) next case False then obtain a where a: "a \ rel_interior S" by blast show ?thesis proof (rule retract_of_locally_path_connected) show "locally path_connected (affine hull S - {a})" by (meson convex_affine_hull convex_imp_locally_path_connected locally_open_subset openin_delete openin_subtopology_self) show "rel_frontier S retract_of affine hull S - {a}" using a assms rel_frontier_retract_of_punctured_affine_hull by blast qed qed lemma locally_connected_sphere_gen: fixes S :: "'a::euclidean_space set" assumes "bounded S" and "convex S" shows "locally connected (rel_frontier S)" by (simp add: ANR_imp_locally_connected ANR_rel_frontier_convex assms) lemma locally_path_connected_sphere: fixes a :: "'a::euclidean_space" shows "locally path_connected (sphere a r)" using ENR_imp_locally_path_connected ENR_sphere by blast lemma locally_connected_sphere: fixes a :: "'a::euclidean_space" shows "locally connected(sphere a r)" using ANR_imp_locally_connected ANR_sphere by blast subsection\Borsuk homotopy extension theorem\ text\It's only this late so we can use the concept of retraction, saying that the domain sets or range set are ENRs.\ theorem Borsuk_homotopy_extension_homotopic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes cloTS: "closedin (top_of_set T) S" and anr: "(ANR S \ ANR T) \ ANR U" and contf: "continuous_on T f" and "f ` T \ U" and "homotopic_with_canon (\x. True) S U f g" obtains g' where "homotopic_with_canon (\x. True) T U f g'" "continuous_on T g'" "image g' T \ U" "\x. x \ S \ g' x = g x" proof - have "S \ T" using assms closedin_imp_subset by blast obtain h where conth: "continuous_on ({0..1} \ S) h" and him: "h ` ({0..1} \ S) \ U" and [simp]: "\x. h(0, x) = f x" "\x. h(1::real, x) = g x" using assms by (auto simp: homotopic_with_def) define h' where "h' \ \z. if snd z \ S then h z else (f \ snd) z" define B where "B \ {0::real} \ T \ {0..1} \ S" have clo0T: "closedin (top_of_set ({0..1} \ T)) ({0::real} \ T)" by (simp add: Abstract_Topology.closedin_Times) moreover have cloT1S: "closedin (top_of_set ({0..1} \ T)) ({0..1} \ S)" by (simp add: Abstract_Topology.closedin_Times assms) ultimately have clo0TB:"closedin (top_of_set ({0..1} \ T)) B" by (auto simp: B_def) have cloBS: "closedin (top_of_set B) ({0..1} \ S)" by (metis (no_types) Un_subset_iff B_def closedin_subset_trans [OF cloT1S] clo0TB closedin_imp_subset closedin_self) moreover have cloBT: "closedin (top_of_set B) ({0} \ T)" using \S \ T\ closedin_subset_trans [OF clo0T] by (metis B_def Un_upper1 clo0TB closedin_closed inf_le1) moreover have "continuous_on ({0} \ T) (f \ snd)" apply (rule continuous_intros)+ apply (simp add: contf) done ultimately have conth': "continuous_on B h'" apply (simp add: h'_def B_def Un_commute [of "{0} \ T"]) apply (auto intro!: continuous_on_cases_local conth) done have "image h' B \ U" using \f ` T \ U\ him by (auto simp: h'_def B_def) obtain V k where "B \ V" and opeTV: "openin (top_of_set ({0..1} \ T)) V" and contk: "continuous_on V k" and kim: "k ` V \ U" and keq: "\x. x \ B \ k x = h' x" using anr proof assume ST: "ANR S \ ANR T" have eq: "({0} \ T \ {0..1} \ S) = {0::real} \ S" using \S \ T\ by auto have "ANR B" apply (simp add: B_def) apply (rule ANR_closed_Un_local) apply (metis cloBT B_def) apply (metis Un_commute cloBS B_def) apply (simp_all add: ANR_Times convex_imp_ANR ANR_singleton ST eq) done note Vk = that have *: thesis if "openin (top_of_set ({0..1::real} \ T)) V" "retraction V B r" for V r using that apply (clarsimp simp add: retraction_def) apply (rule Vk [of V "h' \ r"], assumption+) apply (metis continuous_on_compose conth' continuous_on_subset) using \h' ` B \ U\ apply force+ done show thesis apply (rule ANR_imp_neighbourhood_retract [OF \ANR B\ clo0TB]) apply (auto simp: ANR_Times ANR_singleton ST retract_of_def *) done next assume "ANR U" with ANR_imp_absolute_neighbourhood_extensor \h' ` B \ U\ clo0TB conth' that show ?thesis by blast qed define S' where "S' \ {x. \u::real. u \ {0..1} \ (u, x::'a) \ {0..1} \ T - V}" have "closedin (top_of_set T) S'" unfolding S'_def apply (rule closedin_compact_projection, blast) using closedin_self opeTV by blast have S'_def: "S' = {x. \u::real. (u, x::'a) \ {0..1} \ T - V}" by (auto simp: S'_def) have cloTS': "closedin (top_of_set T) S'" using S'_def \closedin (top_of_set T) S'\ by blast have "S \ S' = {}" using S'_def B_def \B \ V\ by force obtain a :: "'a \ real" where conta: "continuous_on T a" and "\x. x \ T \ a x \ closed_segment 1 0" and a1: "\x. x \ S \ a x = 1" and a0: "\x. x \ S' \ a x = 0" apply (rule Urysohn_local [OF cloTS cloTS' \S \ S' = {}\, of 1 0], blast) done then have ain: "\x. x \ T \ a x \ {0..1}" using closed_segment_eq_real_ivl by auto have inV: "(u * a t, t) \ V" if "t \ T" "0 \ u" "u \ 1" for t u proof (rule ccontr) assume "(u * a t, t) \ V" with ain [OF \t \ T\] have "a t = 0" apply simp apply (rule a0) by (metis (no_types, lifting) Diff_iff S'_def SigmaI atLeastAtMost_iff mem_Collect_eq mult_le_one mult_nonneg_nonneg that) show False using B_def \(u * a t, t) \ V\ \B \ V\ \a t = 0\ that by auto qed show ?thesis proof show hom: "homotopic_with_canon (\x. True) T U f (\x. k (a x, x))" proof (simp add: homotopic_with, intro exI conjI) show "continuous_on ({0..1} \ T) (k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z)))" apply (intro continuous_on_compose continuous_intros) apply (rule continuous_on_subset [OF conta], force) apply (rule continuous_on_subset [OF contk]) apply (force intro: inV) done show "(k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z))) ` ({0..1} \ T) \ U" using inV kim by auto show "\x\T. (k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z))) (0, x) = f x" by (simp add: B_def h'_def keq) show "\x\T. (k \ (\z. (fst z *\<^sub>R (a \ snd) z, snd z))) (1, x) = k (a x, x)" by auto qed show "continuous_on T (\x. k (a x, x))" using homotopic_with_imp_continuous_maps [OF hom] by auto show "(\x. k (a x, x)) ` T \ U" proof clarify fix t assume "t \ T" show "k (a t, t) \ U" by (metis \t \ T\ image_subset_iff inV kim not_one_le_zero linear mult_cancel_right1) qed show "\x. x \ S \ k (a x, x) = g x" by (simp add: B_def a1 h'_def keq) qed qed corollary\<^marker>\tag unimportant\ nullhomotopic_into_ANR_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "ANR T" and fim: "f ` S \ T" and "S \ {}" shows "(\c. homotopic_with_canon (\x. True) S T f (\x. c)) \ (\g. continuous_on UNIV g \ range g \ T \ (\x \ S. g x = f x))" (is "?lhs = ?rhs") proof assume ?lhs then obtain c where c: "homotopic_with_canon (\x. True) S T (\x. c) f" by (blast intro: homotopic_with_symD) have "closedin (top_of_set UNIV) S" using \closed S\ closed_closedin by fastforce then obtain g where "continuous_on UNIV g" "range g \ T" "\x. x \ S \ g x = f x" apply (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ c, where T=UNIV]) using \ANR T\ \S \ {}\ c homotopic_with_imp_subset1 apply fastforce+ done then show ?rhs by blast next assume ?rhs then obtain g where "continuous_on UNIV g" "range g \ T" "\x. x\S \ g x = f x" by blast then obtain c where "homotopic_with_canon (\h. True) UNIV T g (\x. c)" using nullhomotopic_from_contractible [of UNIV g T] contractible_UNIV by blast then have "homotopic_with_canon (\x. True) S T g (\x. c)" by (simp add: homotopic_from_subtopology) then show ?lhs by (force elim: homotopic_with_eq [of _ _ _ g "\x. c"] simp: \\x. x \ S \ g x = f x\) qed corollary\<^marker>\tag unimportant\ nullhomotopic_into_rel_frontier_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "convex T" "bounded T" and fim: "f ` S \ rel_frontier T" and "S \ {}" shows "(\c. homotopic_with_canon (\x. True) S (rel_frontier T) f (\x. c)) \ (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x \ S. g x = f x))" by (simp add: nullhomotopic_into_ANR_extension assms ANR_rel_frontier_convex) corollary\<^marker>\tag unimportant\ nullhomotopic_into_sphere_extension: fixes f :: "'a::euclidean_space \ 'b :: euclidean_space" assumes "closed S" and contf: "continuous_on S f" and "S \ {}" and fim: "f ` S \ sphere a r" shows "((\c. homotopic_with_canon (\x. True) S (sphere a r) f (\x. c)) \ (\g. continuous_on UNIV g \ range g \ sphere a r \ (\x \ S. g x = f x)))" (is "?lhs = ?rhs") proof (cases "r = 0") case True with fim show ?thesis apply auto using fim continuous_on_const apply fastforce by (metis contf contractible_sing nullhomotopic_into_contractible) next case False then have eq: "sphere a r = rel_frontier (cball a r)" by simp show ?thesis using fim unfolding eq apply (rule nullhomotopic_into_rel_frontier_extension [OF \closed S\ contf convex_cball bounded_cball]) apply (rule \S \ {}\) done qed proposition\<^marker>\tag unimportant\ Borsuk_map_essential_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S" shows "bounded (connected_component_set (- S) a) \ \(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. inverse(norm(x - a)) *\<^sub>R (x - a)) (\x. c))" (is "?lhs = ?rhs") proof (cases "S = {}") case True then show ?thesis by simp next case False have "closed S" "bounded S" using \compact S\ compact_eq_bounded_closed by auto have s01: "(\x. (x - a) /\<^sub>R norm (x - a)) ` S \ sphere 0 1" using \a \ S\ by clarsimp (metis dist_eq_0_iff dist_norm mult.commute right_inverse) have aincc: "a \ connected_component_set (- S) a" by (simp add: \a \ S\) obtain r where "r>0" and r: "S \ ball 0 r" using bounded_subset_ballD \bounded S\ by blast have "\ ?rhs \ \ ?lhs" proof assume notr: "\ ?rhs" have nog: "\g. continuous_on (S \ connected_component_set (- S) a) g \ g ` (S \ connected_component_set (- S) a) \ sphere 0 1 \ (\x\S. g x = (x - a) /\<^sub>R norm (x - a))" if "bounded (connected_component_set (- S) a)" apply (rule non_extensible_Borsuk_map [OF \compact S\ componentsI _ aincc]) using \a \ S\ that by auto obtain g where "range g \ sphere 0 1" "continuous_on UNIV g" "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" using notr by (auto simp: nullhomotopic_into_sphere_extension [OF \closed S\ continuous_on_Borsuk_map [OF \a \ S\] False s01]) with \a \ S\ show "\ ?lhs" apply (clarsimp simp: Borsuk_map_into_sphere [of a S, symmetric] dest!: nog) apply (drule_tac x=g in spec) using continuous_on_subset by fastforce next assume "\ ?lhs" then obtain b where b: "b \ connected_component_set (- S) a" and "r \ norm b" using bounded_iff linear by blast then have bnot: "b \ ball 0 r" by simp have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. (x - b) /\<^sub>R norm (x - b))" apply (rule Borsuk_maps_homotopic_in_path_component) using \closed S\ b open_Compl open_path_connected_component apply fastforce done moreover obtain c where "homotopic_with_canon (\x. True) (ball 0 r) (sphere 0 1) (\x. inverse (norm (x - b)) *\<^sub>R (x - b)) (\x. c)" proof (rule nullhomotopic_from_contractible) show "contractible (ball (0::'a) r)" by (metis convex_imp_contractible convex_ball) show "continuous_on (ball 0 r) (\x. inverse(norm (x - b)) *\<^sub>R (x - b))" by (rule continuous_on_Borsuk_map [OF bnot]) show "(\x. (x - b) /\<^sub>R norm (x - b)) ` ball 0 r \ sphere 0 1" using bnot Borsuk_map_into_sphere by blast qed blast ultimately have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. c)" by (meson homotopic_with_subset_left homotopic_with_trans r) then show "\ ?rhs" by blast qed then show ?thesis by blast qed lemma homotopic_Borsuk_maps_in_bounded_component: fixes a :: "'a :: euclidean_space" assumes "compact S" and "a \ S"and "b \ S" and boc: "bounded (connected_component_set (- S) a)" and hom: "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. (x - b) /\<^sub>R norm (x - b))" shows "connected_component (- S) a b" proof (rule ccontr) assume notcc: "\ connected_component (- S) a b" let ?T = "S \ connected_component_set (- S) a" have "\g. continuous_on (S \ connected_component_set (- S) a) g \ g ` (S \ connected_component_set (- S) a) \ sphere 0 1 \ (\x\S. g x = (x - a) /\<^sub>R norm (x - a))" by (simp add: \a \ S\ componentsI non_extensible_Borsuk_map [OF \compact S\ _ boc]) moreover obtain g where "continuous_on (S \ connected_component_set (- S) a) g" "g ` (S \ connected_component_set (- S) a) \ sphere 0 1" "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" proof (rule Borsuk_homotopy_extension_homotopic) show "closedin (top_of_set ?T) S" by (simp add: \compact S\ closed_subset compact_imp_closed) show "continuous_on ?T (\x. (x - b) /\<^sub>R norm (x - b))" by (simp add: \b \ S\ notcc continuous_on_Borsuk_map) show "(\x. (x - b) /\<^sub>R norm (x - b)) ` ?T \ sphere 0 1" by (simp add: \b \ S\ notcc Borsuk_map_into_sphere) show "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - b) /\<^sub>R norm (x - b)) (\x. (x - a) /\<^sub>R norm (x - a))" by (simp add: hom homotopic_with_symD) qed (auto simp: ANR_sphere intro: that) ultimately show False by blast qed lemma Borsuk_maps_homotopic_in_connected_component_eq: fixes a :: "'a :: euclidean_space" assumes S: "compact S" "a \ S" "b \ S" and 2: "2 \ DIM('a)" shows "(homotopic_with_canon (\x. True) S (sphere 0 1) (\x. (x - a) /\<^sub>R norm (x - a)) (\x. (x - b) /\<^sub>R norm (x - b)) \ connected_component (- S) a b)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "bounded(connected_component_set (- S) a)") case True show ?thesis by (rule homotopic_Borsuk_maps_in_bounded_component [OF S True L]) next case not_bo_a: False show ?thesis proof (cases "bounded(connected_component_set (- S) b)") case True show ?thesis using homotopic_Borsuk_maps_in_bounded_component [OF S] by (simp add: L True assms connected_component_sym homotopic_Borsuk_maps_in_bounded_component homotopic_with_sym) next case False then show ?thesis using cobounded_unique_unbounded_component [of "-S" a b] \compact S\ not_bo_a by (auto simp: compact_eq_bounded_closed assms connected_component_eq_eq) qed qed next assume R: ?rhs then have "path_component (- S) a b" using assms(1) compact_eq_bounded_closed open_Compl open_path_connected_component_set by fastforce then show ?lhs by (simp add: Borsuk_maps_homotopic_in_path_component) qed subsection\More extension theorems\ lemma extension_from_clopen: assumes ope: "openin (top_of_set S) T" and clo: "closedin (top_of_set S) T" and contf: "continuous_on T f" and fim: "f ` T \ U" and null: "U = {} \ S = {}" obtains g where "continuous_on S g" "g ` S \ U" "\x. x \ T \ g x = f x" proof (cases "U = {}") case True then show ?thesis by (simp add: null that) next case False then obtain a where "a \ U" by auto let ?g = "\x. if x \ T then f x else a" have Seq: "S = T \ (S - T)" using clo closedin_imp_subset by fastforce show ?thesis proof have "continuous_on (T \ (S - T)) ?g" apply (rule continuous_on_cases_local) - using Seq clo ope by (auto simp: contf continuous_on_const intro: continuous_on_cases_local) + using Seq clo ope by (auto simp: contf intro: continuous_on_cases_local) with Seq show "continuous_on S ?g" by metis show "?g ` S \ U" using \a \ U\ fim by auto show "\x. x \ T \ ?g x = f x" by auto qed qed lemma extension_from_component: fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes S: "locally connected S \ compact S" and "ANR U" and C: "C \ components S" and contf: "continuous_on C f" and fim: "f ` C \ U" obtains g where "continuous_on S g" "g ` S \ U" "\x. x \ C \ g x = f x" proof - obtain T g where ope: "openin (top_of_set S) T" and clo: "closedin (top_of_set S) T" and "C \ T" and contg: "continuous_on T g" and gim: "g ` T \ U" and gf: "\x. x \ C \ g x = f x" using S proof assume "locally connected S" show ?thesis by (metis C \locally connected S\ openin_components_locally_connected closedin_component contf fim order_refl that) next assume "compact S" then obtain W g where "C \ W" and opeW: "openin (top_of_set S) W" and contg: "continuous_on W g" and gim: "g ` W \ U" and gf: "\x. x \ C \ g x = f x" using ANR_imp_absolute_neighbourhood_extensor [of U C f S] C \ANR U\ closedin_component contf fim by blast then obtain V where "open V" and V: "W = S \ V" by (auto simp: openin_open) moreover have "locally compact S" by (simp add: \compact S\ closed_imp_locally_compact compact_imp_closed) ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \ K" "K \ V" by (metis C Int_subset_iff \C \ W\ \compact S\ compact_components Sura_Bura_clopen_subset) show ?thesis proof show "closedin (top_of_set S) K" by (meson \compact K\ \compact S\ closedin_compact_eq opeK openin_imp_subset) show "continuous_on K g" by (metis Int_subset_iff V \K \ V\ contg continuous_on_subset opeK openin_subtopology subset_eq) show "g ` K \ U" using V \K \ V\ gim opeK openin_imp_subset by fastforce qed (use opeK gf \C \ K\ in auto) qed obtain h where "continuous_on S h" "h ` S \ U" "\x. x \ T \ h x = g x" using extension_from_clopen by (metis C bot.extremum_uniqueI clo contg gim fim image_is_empty in_components_nonempty ope) then show ?thesis by (metis \C \ T\ gf subset_eq that) qed lemma tube_lemma: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and S: "S \ {}" "(\x. (x,a)) ` S \ U" and ope: "openin (top_of_set (S \ T)) U" obtains V where "openin (top_of_set T) V" "a \ V" "S \ V \ U" proof - let ?W = "{y. \x. x \ S \ (x, y) \ (S \ T - U)}" have "U \ S \ T" "closedin (top_of_set (S \ T)) (S \ T - U)" using ope by (auto simp: openin_closedin_eq) then have "closedin (top_of_set T) ?W" using \compact S\ closedin_compact_projection by blast moreover have "a \ T - ?W" using \U \ S \ T\ S by auto moreover have "S \ (T - ?W) \ U" by auto ultimately show ?thesis by (metis (no_types, lifting) Sigma_cong closedin_def that topspace_euclidean_subtopology) qed lemma tube_lemma_gen: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "S \ {}" "T \ T'" "S \ T \ U" and ope: "openin (top_of_set (S \ T')) U" obtains V where "openin (top_of_set T') V" "T \ V" "S \ V \ U" proof - have "\x. x \ T \ \V. openin (top_of_set T') V \ x \ V \ S \ V \ U" using assms by (auto intro: tube_lemma [OF \compact S\]) then obtain F where F: "\x. x \ T \ openin (top_of_set T') (F x) \ x \ F x \ S \ F x \ U" by metis show ?thesis proof show "openin (top_of_set T') (\(F ` T))" using F by blast show "T \ \(F ` T)" using F by blast show "S \ \(F ` T) \ U" using F by auto qed qed proposition\<^marker>\tag unimportant\ homotopic_neighbourhood_extension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and fim: "f ` S \ U" and contg: "continuous_on S g" and gim: "g ` S \ U" and clo: "closedin (top_of_set S) T" and "ANR U" and hom: "homotopic_with_canon (\x. True) T U f g" obtains V where "T \ V" "openin (top_of_set S) V" "homotopic_with_canon (\x. True) V U f g" proof - have "T \ S" using clo closedin_imp_subset by blast obtain h where conth: "continuous_on ({0..1::real} \ T) h" and him: "h ` ({0..1} \ T) \ U" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = g x" using hom by (auto simp: homotopic_with_def) define h' where "h' \ \z. if fst z \ {0} then f(snd z) else if fst z \ {1} then g(snd z) else h z" let ?S0 = "{0::real} \ S" and ?S1 = "{1::real} \ S" have "continuous_on(?S0 \ (?S1 \ {0..1} \ T)) h'" unfolding h'_def proof (intro continuous_on_cases_local) show "closedin (top_of_set (?S0 \ (?S1 \ {0..1} \ T))) ?S0" "closedin (top_of_set (?S1 \ {0..1} \ T)) ?S1" using \T \ S\ by (force intro: closedin_Times closedin_subset_trans [of "{0..1} \ S"])+ show "closedin (top_of_set (?S0 \ (?S1 \ {0..1} \ T))) (?S1 \ {0..1} \ T)" "closedin (top_of_set (?S1 \ {0..1} \ T)) ({0..1} \ T)" using \T \ S\ by (force intro: clo closedin_Times closedin_subset_trans [of "{0..1} \ S"])+ show "continuous_on (?S0) (\x. f (snd x))" by (intro continuous_intros continuous_on_compose2 [OF contf]) auto show "continuous_on (?S1) (\x. g (snd x))" by (intro continuous_intros continuous_on_compose2 [OF contg]) auto qed (use h0 h1 conth in auto) then have "continuous_on ({0,1} \ S \ ({0..1} \ T)) h'" by (metis Sigma_Un_distrib1 Un_assoc insert_is_Un) moreover have "h' ` ({0,1} \ S \ {0..1} \ T) \ U" using fim gim him \T \ S\ unfolding h'_def by force moreover have "closedin (top_of_set ({0..1::real} \ S)) ({0,1} \ S \ {0..1::real} \ T)" by (intro closedin_Times closedin_Un clo) (simp_all add: closed_subset) ultimately obtain W k where W: "({0,1} \ S) \ ({0..1} \ T) \ W" and opeW: "openin (top_of_set ({0..1} \ S)) W" and contk: "continuous_on W k" and kim: "k ` W \ U" and kh': "\x. x \ ({0,1} \ S) \ ({0..1} \ T) \ k x = h' x" by (metis ANR_imp_absolute_neighbourhood_extensor [OF \ANR U\, of "({0,1} \ S) \ ({0..1} \ T)" h' "{0..1} \ S"]) obtain T' where opeT': "openin (top_of_set S) T'" and "T \ T'" and TW: "{0..1} \ T' \ W" using tube_lemma_gen [of "{0..1::real}" T S W] W \T \ S\ opeW by auto moreover have "homotopic_with_canon (\x. True) T' U f g" proof (simp add: homotopic_with, intro exI conjI) show "continuous_on ({0..1} \ T') k" using TW continuous_on_subset contk by auto show "k ` ({0..1} \ T') \ U" using TW kim by fastforce have "T' \ S" by (meson opeT' subsetD openin_imp_subset) then show "\x\T'. k (0, x) = f x" "\x\T'. k (1, x) = g x" by (auto simp: kh' h'_def) qed ultimately show ?thesis by (blast intro: that) qed text\ Homotopy on a union of closed-open sets.\ proposition\<^marker>\tag unimportant\ homotopic_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and "\S. S \ \ \ homotopic_with_canon (\x. True) S T f g" shows "homotopic_with_canon (\x. True) (\\) T f g" proof - obtain \ where "\ \ \" "countable \" and eqU: "\\ = \\" using Lindelof_openin assms by blast show ?thesis proof (cases "\ = {}") case True then show ?thesis by (metis Union_empty eqU homotopic_with_canon_on_empty) next case False then obtain V :: "nat \ 'a set" where V: "range V = \" using range_from_nat_into \countable \\ by metis with \\ \ \\ have clo: "\n. closedin (top_of_set (\\)) (V n)" and ope: "\n. openin (top_of_set (\\)) (V n)" and hom: "\n. homotopic_with_canon (\x. True) (V n) T f g" using assms by auto then obtain h where conth: "\n. continuous_on ({0..1::real} \ V n) (h n)" and him: "\n. h n ` ({0..1} \ V n) \ T" and h0: "\n. \x. x \ V n \ h n (0, x) = f x" and h1: "\n. \x. x \ V n \ h n (1, x) = g x" by (simp add: homotopic_with) metis have wop: "b \ V x \ \k. b \ V k \ (\j V j)" for b x using nat_less_induct [where P = "\i. b \ V i"] by meson obtain \ where cont: "continuous_on ({0..1} \ \(V ` UNIV)) \" and eq: "\x i. \x \ {0..1} \ \(V ` UNIV) \ {0..1} \ (V i - (\m \ \ x = h i x" proof (rule pasting_lemma_exists) let ?X = "top_of_set ({0..1::real} \ \(range V))" show "topspace ?X \ (\i. {0..1::real} \ (V i - (\m \(V ` UNIV))) ({0..1::real} \ (V i - (\m(V ` UNIV))) (V i)" using ope V eqU by auto show "closedin (top_of_set (\(V ` UNIV))) (\m (V i - \ (V ` {..i j x. x \ topspace ?X \ {0..1} \ (V i - (\m {0..1} \ (V j - (\m h i x = h j x" by clarsimp (metis lessThan_iff linorder_neqE_nat) qed auto show ?thesis proof (simp add: homotopic_with eqU [symmetric], intro exI conjI ballI) show "continuous_on ({0..1} \ \\) \" using V eqU by (blast intro!: continuous_on_subset [OF cont]) show "\` ({0..1} \ \\) \ T" proof clarsimp fix t :: real and y :: "'a" and X :: "'a set" assume "y \ X" "X \ \" and t: "0 \ t" "t \ 1" then obtain k where "y \ V k" and j: "\j V j" by (metis image_iff V wop) with him t show "\(t, y) \ T" by (subst eq) force+ qed fix X y assume "X \ \" "y \ X" then obtain k where "y \ V k" and j: "\j V j" by (metis image_iff V wop) then show "\(0, y) = f y" and "\(1, y) = g y" by (subst eq [where i=k]; force simp: h0 h1)+ qed qed qed lemma homotopic_on_components_eq: fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" assumes S: "locally connected S \ compact S" and "ANR T" shows "homotopic_with_canon (\x. True) S T f g \ (continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T) \ (\C \ components S. homotopic_with_canon (\x. True) C T f g)" (is "?lhs \ ?C \ ?rhs") proof - have "continuous_on S f" "f ` S \ T" "continuous_on S g" "g ` S \ T" if ?lhs using homotopic_with_imp_continuous homotopic_with_imp_subset1 homotopic_with_imp_subset2 that by blast+ moreover have "?lhs \ ?rhs" if contf: "continuous_on S f" and fim: "f ` S \ T" and contg: "continuous_on S g" and gim: "g ` S \ T" proof assume ?lhs with that show ?rhs by (simp add: homotopic_with_subset_left in_components_subset) next assume R: ?rhs have "\U. C \ U \ closedin (top_of_set S) U \ openin (top_of_set S) U \ homotopic_with_canon (\x. True) U T f g" if C: "C \ components S" for C proof - have "C \ S" by (simp add: in_components_subset that) show ?thesis using S proof assume "locally connected S" show ?thesis proof (intro exI conjI) show "closedin (top_of_set S) C" by (simp add: closedin_component that) show "openin (top_of_set S) C" by (simp add: \locally connected S\ openin_components_locally_connected that) show "homotopic_with_canon (\x. True) C T f g" by (simp add: R that) qed auto next assume "compact S" have hom: "homotopic_with_canon (\x. True) C T f g" using R that by blast obtain U where "C \ U" and opeU: "openin (top_of_set S) U" and hom: "homotopic_with_canon (\x. True) U T f g" using homotopic_neighbourhood_extension [OF contf fim contg gim _ \ANR T\ hom] \C \ components S\ closedin_component by blast then obtain V where "open V" and V: "U = S \ V" by (auto simp: openin_open) moreover have "locally compact S" by (simp add: \compact S\ closed_imp_locally_compact compact_imp_closed) ultimately obtain K where opeK: "openin (top_of_set S) K" and "compact K" "C \ K" "K \ V" by (metis C Int_subset_iff Sura_Bura_clopen_subset \C \ U\ \compact S\ compact_components) show ?thesis proof (intro exI conjI) show "closedin (top_of_set S) K" by (meson \compact K\ \compact S\ closedin_compact_eq opeK openin_imp_subset) show "homotopic_with_canon (\x. True) K T f g" using V \K \ V\ hom homotopic_with_subset_left opeK openin_imp_subset by fastforce qed (use opeK \C \ K\ in auto) qed qed then obtain \ where \: "\C. C \ components S \ C \ \ C" and clo\: "\C. C \ components S \ closedin (top_of_set S) (\ C)" and ope\: "\C. C \ components S \ openin (top_of_set S) (\ C)" and hom\: "\C. C \ components S \ homotopic_with_canon (\x. True) (\ C) T f g" by metis have Seq: "S = \ (\ ` components S)" proof show "S \ \ (\ ` components S)" by (metis Sup_mono Union_components \ imageI) show "\ (\ ` components S) \ S" using ope\ openin_imp_subset by fastforce qed show ?lhs apply (subst Seq) apply (rule homotopic_on_clopen_Union) using Seq clo\ ope\ hom\ by auto qed ultimately show ?thesis by blast qed lemma cohomotopically_trivial_on_components: fixes S :: "'a :: euclidean_space set" and T :: "'b :: euclidean_space set" assumes S: "locally connected S \ compact S" and "ANR T" shows "(\f g. continuous_on S f \ f ` S \ T \ continuous_on S g \ g ` S \ T \ homotopic_with_canon (\x. True) S T f g) \ (\C\components S. \f g. continuous_on C f \ f ` C \ T \ continuous_on C g \ g ` C \ T \ homotopic_with_canon (\x. True) C T f g)" (is "?lhs = ?rhs") proof assume L[rule_format]: ?lhs show ?rhs proof clarify fix C f g assume contf: "continuous_on C f" and fim: "f ` C \ T" and contg: "continuous_on C g" and gim: "g ` C \ T" and C: "C \ components S" obtain f' where contf': "continuous_on S f'" and f'im: "f' ` S \ T" and f'f: "\x. x \ C \ f' x = f x" using extension_from_component [OF S \ANR T\ C contf fim] by metis obtain g' where contg': "continuous_on S g'" and g'im: "g' ` S \ T" and g'g: "\x. x \ C \ g' x = g x" using extension_from_component [OF S \ANR T\ C contg gim] by metis have "homotopic_with_canon (\x. True) C T f' g'" using L [OF contf' f'im contg' g'im] homotopic_with_subset_left C in_components_subset by fastforce then show "homotopic_with_canon (\x. True) C T f g" using f'f g'g homotopic_with_eq by force qed next assume R [rule_format]: ?rhs show ?lhs proof clarify fix f g assume contf: "continuous_on S f" and fim: "f ` S \ T" and contg: "continuous_on S g" and gim: "g ` S \ T" moreover have "homotopic_with_canon (\x. True) C T f g" if "C \ components S" for C using R [OF that] by (meson contf contg continuous_on_subset fim gim image_mono in_components_subset order.trans that) ultimately show "homotopic_with_canon (\x. True) S T f g" by (subst homotopic_on_components_eq [OF S \ANR T\]) auto qed qed subsection\The complement of a set and path-connectedness\ text\Complement in dimension N > 1 of set homeomorphic to any interval in any dimension is (path-)connected. This naively generalizes the argument in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer fixed point theorem", American Mathematical Monthly 1984.\ lemma unbounded_components_complement_absolute_retract: fixes S :: "'a::euclidean_space set" assumes C: "C \ components(- S)" and S: "compact S" "AR S" shows "\ bounded C" proof - obtain y where y: "C = connected_component_set (- S) y" and "y \ S" using C by (auto simp: components_def) have "open(- S)" using S by (simp add: closed_open compact_eq_bounded_closed) have "S retract_of UNIV" using S compact_AR by blast then obtain r where contr: "continuous_on UNIV r" and ontor: "range r \ S" and r: "\x. x \ S \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof assume "bounded C" have "connected_component_set (- S) y \ S" proof (rule frontier_subset_retraction) show "bounded (connected_component_set (- S) y)" using \bounded C\ y by blast show "frontier (connected_component_set (- S) y) \ S" using C \compact S\ compact_eq_bounded_closed frontier_of_components_closed_complement y by blast show "continuous_on (closure (connected_component_set (- S) y)) r" by (blast intro: continuous_on_subset [OF contr]) qed (use ontor r in auto) with \y \ S\ show False by force qed qed lemma connected_complement_absolute_retract: fixes S :: "'a::euclidean_space set" assumes S: "compact S" "AR S" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S retract_of UNIV" using S compact_AR by blast show ?thesis apply (clarsimp simp: connected_iff_connected_component_eq) apply (rule cobounded_unique_unbounded_component [OF _ 2]) apply (simp add: \compact S\ compact_imp_bounded) apply (meson ComplI S componentsI unbounded_components_complement_absolute_retract)+ done qed lemma path_connected_complement_absolute_retract: fixes S :: "'a::euclidean_space set" assumes "compact S" "AR S" "2 \ DIM('a)" shows "path_connected(- S)" using connected_complement_absolute_retract [OF assms] using \compact S\ compact_eq_bounded_closed connected_open_path_connected by blast theorem connected_complement_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes hom: "S homeomorphic T" and T: "convex T" "compact T" and 2: "2 \ DIM('a)" shows "connected(- S)" proof (cases "S = {}") case True then show ?thesis by (simp add: connected_UNIV) next case False show ?thesis proof (rule connected_complement_absolute_retract) show "compact S" using \compact T\ hom homeomorphic_compactness by auto show "AR S" by (meson AR_ANR False \convex T\ convex_imp_ANR convex_imp_contractible hom homeomorphic_ANR_iff_ANR homeomorphic_contractible_eq) qed (rule 2) qed corollary path_connected_complement_homeomorphic_convex_compact: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes hom: "S homeomorphic T" "convex T" "compact T" "2 \ DIM('a)" shows "path_connected(- S)" using connected_complement_homeomorphic_convex_compact [OF assms] using \compact T\ compact_eq_bounded_closed connected_open_path_connected hom homeomorphic_compactness by blast lemma path_connected_complement_homeomorphic_interval: fixes S :: "'a::euclidean_space set" assumes "S homeomorphic cbox a b" "2 \ DIM('a)" shows "path_connected(-S)" using assms compact_cbox convex_box(1) path_connected_complement_homeomorphic_convex_compact by blast lemma connected_complement_homeomorphic_interval: fixes S :: "'a::euclidean_space set" assumes "S homeomorphic cbox a b" "2 \ DIM('a)" shows "connected(-S)" using assms path_connected_complement_homeomorphic_interval path_connected_imp_connected by blast end diff --git a/src/HOL/Analysis/Starlike.thy b/src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy +++ b/src/HOL/Analysis/Starlike.thy @@ -1,6924 +1,6924 @@ (* Title: HOL/Analysis/Starlike.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) chapter \Unsorted\ theory Starlike imports Convex_Euclidean_Space Abstract_Limits Line_Segment begin subsection\Starlike sets\ definition\<^marker>\tag important\ "starlike S \ (\a\S. \x\S. closed_segment a x \ S)" lemma starlike_UNIV [simp]: "starlike UNIV" by (simp add: starlike_def) lemma convex_imp_starlike: "convex S \ S \ {} \ starlike S" unfolding convex_contains_segment starlike_def by auto lemma affine_hull_closed_segment [simp]: "affine hull (closed_segment a b) = affine hull {a,b}" by (simp add: segment_convex_hull) lemma affine_hull_open_segment [simp]: fixes a :: "'a::euclidean_space" shows "affine hull (open_segment a b) = (if a = b then {} else affine hull {a,b})" by (metis affine_hull_convex_hull affine_hull_empty closure_open_segment closure_same_affine_hull segment_convex_hull) lemma rel_interior_closure_convex_segment: fixes S :: "_::euclidean_space set" assumes "convex S" "a \ rel_interior S" "b \ closure S" shows "open_segment a b \ rel_interior S" proof fix x have [simp]: "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" for u by (simp add: algebra_simps) assume "x \ open_segment a b" then show "x \ rel_interior S" unfolding closed_segment_def open_segment_def using assms by (auto intro: rel_interior_closure_convex_shrink) qed lemma convex_hull_insert_segments: "convex hull (insert a S) = (if S = {} then {a} else \x \ convex hull S. closed_segment a x)" by (force simp add: convex_hull_insert_alt in_segment) lemma Int_convex_hull_insert_rel_exterior: fixes z :: "'a::euclidean_space" assumes "convex C" "T \ C" and z: "z \ rel_interior C" and dis: "disjnt S (rel_interior C)" shows "S \ (convex hull (insert z T)) = S \ (convex hull T)" (is "?lhs = ?rhs") proof have "T = {} \ z \ S" using dis z by (auto simp add: disjnt_def) then show "?lhs \ ?rhs" proof (clarsimp simp add: convex_hull_insert_segments) fix x y assume "x \ S" and y: "y \ convex hull T" and "x \ closed_segment z y" have "y \ closure C" by (metis y \convex C\ \T \ C\ closure_subset contra_subsetD convex_hull_eq hull_mono) moreover have "x \ rel_interior C" by (meson \x \ S\ dis disjnt_iff) moreover have "x \ open_segment z y \ {z, y}" using \x \ closed_segment z y\ closed_segment_eq_open by blast ultimately show "x \ convex hull T" using rel_interior_closure_convex_segment [OF \convex C\ z] using y z by blast qed show "?rhs \ ?lhs" by (meson hull_mono inf_mono subset_insertI subset_refl) qed subsection\<^marker>\tag unimportant\\More results about segments\ lemma dist_half_times2: fixes a :: "'a :: real_normed_vector" shows "dist ((1 / 2) *\<^sub>R (a + b)) x * 2 = dist (a+b) (2 *\<^sub>R x)" proof - have "norm ((1 / 2) *\<^sub>R (a + b) - x) * 2 = norm (2 *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x))" by simp also have "... = norm ((a + b) - 2 *\<^sub>R x)" by (simp add: real_vector.scale_right_diff_distrib) finally show ?thesis by (simp only: dist_norm) qed lemma closed_segment_as_ball: "closed_segment a b = affine hull {a,b} \ cball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)" proof (cases "b = a") case True then show ?thesis by (auto simp: hull_inc) next case False then have *: "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a)) = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1)" for x proof - have "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a)) = ((\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 \ norm (b - a))" unfolding eq_diff_eq [symmetric] by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R x)) \ norm (b - a))" by (simp add: dist_half_times2) (simp add: dist_norm) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) \ norm (b - a))" by auto also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((1 - u * 2) *\<^sub>R (b - a)) \ norm (b - a))" by (simp add: algebra_simps scaleR_2) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ * norm (b - a) \ norm (b - a))" by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ \ 1)" by (simp add: mult_le_cancel_right2 False) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1)" by auto finally show ?thesis . qed show ?thesis by (simp add: affine_hull_2 Set.set_eq_iff closed_segment_def *) qed lemma open_segment_as_ball: "open_segment a b = affine hull {a,b} \ ball(inverse 2 *\<^sub>R (a + b))(norm(b - a) / 2)" proof (cases "b = a") case True then show ?thesis by (auto simp: hull_inc) next case False then have *: "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 < u \ u < 1)" for x proof - have "((\u v. x = u *\<^sub>R a + v *\<^sub>R b \ u + v = 1) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a)) = ((\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b) \ dist ((1 / 2) *\<^sub>R (a + b)) x * 2 < norm (b - a))" unfolding eq_diff_eq [symmetric] by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R x)) < norm (b - a))" by (simp add: dist_half_times2) (simp add: dist_norm) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((a+b) - (2 *\<^sub>R ((1 - u) *\<^sub>R a + u *\<^sub>R b))) < norm (b - a))" by auto also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ norm ((1 - u * 2) *\<^sub>R (b - a)) < norm (b - a))" by (simp add: algebra_simps scaleR_2) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ * norm (b - a) < norm (b - a))" by simp also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ \1 - u * 2\ < 1)" by (simp add: mult_le_cancel_right2 False) also have "... = (\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 < u \ u < 1)" by auto finally show ?thesis . qed show ?thesis using False by (force simp: affine_hull_2 Set.set_eq_iff open_segment_image_interval *) qed lemmas segment_as_ball = closed_segment_as_ball open_segment_as_ball lemma closed_segment_neq_empty [simp]: "closed_segment a b \ {}" by auto lemma open_segment_eq_empty [simp]: "open_segment a b = {} \ a = b" proof - { assume a1: "open_segment a b = {}" have "{} \ {0::real<..<1}" by simp then have "a = b" using a1 open_segment_image_interval by fastforce } then show ?thesis by auto qed lemma open_segment_eq_empty' [simp]: "{} = open_segment a b \ a = b" using open_segment_eq_empty by blast lemmas segment_eq_empty = closed_segment_neq_empty open_segment_eq_empty lemma inj_segment: fixes a :: "'a :: real_vector" assumes "a \ b" shows "inj_on (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) I" proof fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" then have "x *\<^sub>R (b - a) = y *\<^sub>R (b - a)" by (simp add: algebra_simps) with assms show "x = y" by (simp add: real_vector.scale_right_imp_eq) qed lemma finite_closed_segment [simp]: "finite(closed_segment a b) \ a = b" apply auto apply (rule ccontr) apply (simp add: segment_image_interval) using infinite_Icc [OF zero_less_one] finite_imageD [OF _ inj_segment] apply blast done lemma finite_open_segment [simp]: "finite(open_segment a b) \ a = b" by (auto simp: open_segment_def) lemmas finite_segment = finite_closed_segment finite_open_segment lemma closed_segment_eq_sing: "closed_segment a b = {c} \ a = c \ b = c" by auto lemma open_segment_eq_sing: "open_segment a b \ {c}" by (metis finite_insert finite_open_segment insert_not_empty open_segment_image_interval) lemmas segment_eq_sing = closed_segment_eq_sing open_segment_eq_sing lemma subset_closed_segment: "closed_segment a b \ closed_segment c d \ a \ closed_segment c d \ b \ closed_segment c d" by auto (meson contra_subsetD convex_closed_segment convex_contains_segment) lemma subset_co_segment: "closed_segment a b \ open_segment c d \ a \ open_segment c d \ b \ open_segment c d" using closed_segment_subset by blast lemma subset_open_segment: fixes a :: "'a::euclidean_space" shows "open_segment a b \ open_segment c d \ a = b \ a \ closed_segment c d \ b \ closed_segment c d" (is "?lhs = ?rhs") proof (cases "a = b") case True then show ?thesis by simp next case False show ?thesis proof assume rhs: ?rhs with \a \ b\ have "c \ d" using closed_segment_idem singleton_iff by auto have "\uc. (1 - u) *\<^sub>R ((1 - ua) *\<^sub>R c + ua *\<^sub>R d) + u *\<^sub>R ((1 - ub) *\<^sub>R c + ub *\<^sub>R d) = (1 - uc) *\<^sub>R c + uc *\<^sub>R d \ 0 < uc \ uc < 1" if neq: "(1 - ua) *\<^sub>R c + ua *\<^sub>R d \ (1 - ub) *\<^sub>R c + ub *\<^sub>R d" "c \ d" and "a = (1 - ua) *\<^sub>R c + ua *\<^sub>R d" "b = (1 - ub) *\<^sub>R c + ub *\<^sub>R d" and u: "0 < u" "u < 1" and uab: "0 \ ua" "ua \ 1" "0 \ ub" "ub \ 1" for u ua ub proof - have "ua \ ub" using neq by auto moreover have "(u - 1) * ua \ 0" using u uab by (simp add: mult_nonpos_nonneg) ultimately have lt: "(u - 1) * ua < u * ub" using u uab by (metis antisym_conv diff_ge_0_iff_ge le_less_trans mult_eq_0_iff mult_le_0_iff not_less) have "p * ua + q * ub < p+q" if p: "0 < p" and q: "0 < q" for p q proof - have "\ p \ 0" "\ q \ 0" using p q not_less by blast+ then show ?thesis by (metis \ua \ ub\ add_less_cancel_left add_less_cancel_right add_mono_thms_linordered_field(5) less_eq_real_def mult_cancel_left1 mult_less_cancel_left2 uab(2) uab(4)) qed then have "(1 - u) * ua + u * ub < 1" using u \ua \ ub\ by (metis diff_add_cancel diff_gt_0_iff_gt) with lt show ?thesis by (rule_tac x="ua + u*(ub-ua)" in exI) (simp add: algebra_simps) qed with rhs \a \ b\ \c \ d\ show ?lhs unfolding open_segment_image_interval closed_segment_def by (fastforce simp add:) next assume lhs: ?lhs with \a \ b\ have "c \ d" by (meson finite_open_segment rev_finite_subset) have "closure (open_segment a b) \ closure (open_segment c d)" using lhs closure_mono by blast then have "closed_segment a b \ closed_segment c d" by (simp add: \a \ b\ \c \ d\) then show ?rhs by (force simp: \a \ b\) qed qed lemma subset_oc_segment: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d \ a = b \ a \ closed_segment c d \ b \ closed_segment c d" apply (simp add: subset_open_segment [symmetric]) apply (rule iffI) apply (metis closure_closed_segment closure_mono closure_open_segment subset_closed_segment subset_open_segment) apply (meson dual_order.trans segment_open_subset_closed) done lemmas subset_segment = subset_closed_segment subset_co_segment subset_oc_segment subset_open_segment subsection\<^marker>\tag unimportant\ \Shrinking towards the interior of a convex set\ lemma mem_interior_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto show ?thesis unfolding mem_interior proof (intro exI subsetI conjI) fix y assume "y \ ball (x - e *\<^sub>R (x - c)) (e*d)" then have as: "dist (x - e *\<^sub>R (x - c)) y < e * d" by simp have *: "y = (1 - (1 - e)) *\<^sub>R ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) + (1 - e) *\<^sub>R x" using \e > 0\ by (auto simp add: scaleR_left_diff_distrib scaleR_right_diff_distrib) have "dist c ((1 / e) *\<^sub>R y - ((1 - e) / e) *\<^sub>R x) = \1/e\ * norm (e *\<^sub>R c - y + (1 - e) *\<^sub>R x)" unfolding dist_norm unfolding norm_scaleR[symmetric] apply (rule arg_cong[where f=norm]) using \e > 0\ by (auto simp add: euclidean_eq_iff[where 'a='a] field_simps inner_simps) also have "\ = \1/e\ * norm (x - e *\<^sub>R (x - c) - y)" by (auto intro!:arg_cong[where f=norm] simp add: algebra_simps) also have "\ < d" using as[unfolded dist_norm] and \e > 0\ by (auto simp add:pos_divide_less_eq[OF \e > 0\] mult.commute) finally show "y \ S" apply (subst *) apply (rule assms(1)[unfolded convex_alt,rule_format]) apply (rule d[unfolded subset_eq,rule_format]) unfolding mem_ball using assms(3-5) apply auto done qed (insert \e>0\ \d>0\, auto) qed lemma mem_interior_closure_convex_shrink: fixes S :: "'a::euclidean_space set" assumes "convex S" and "c \ interior S" and "x \ closure S" and "0 < e" and "e \ 1" shows "x - e *\<^sub>R (x - c) \ interior S" proof - obtain d where "d > 0" and d: "ball c d \ S" using assms(2) unfolding mem_interior by auto have "\y\S. norm (y - x) * (1 - e) < e * d" proof (cases "x \ S") case True then show ?thesis using \e > 0\ \d > 0\ apply (rule_tac bexI[where x=x]) apply (auto) done next case False then have x: "x islimpt S" using assms(3)[unfolded closure_def] by auto show ?thesis proof (cases "e = 1") case True obtain y where "y \ S" "y \ x" "dist y x < 1" using x[unfolded islimpt_approachable,THEN spec[where x=1]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding True using \d > 0\ apply auto done next case False then have "0 < e * d / (1 - e)" and *: "1 - e > 0" using \e \ 1\ \e > 0\ \d > 0\ by auto then obtain y where "y \ S" "y \ x" "dist y x < e * d / (1 - e)" using x[unfolded islimpt_approachable,THEN spec[where x="e*d / (1 - e)"]] by auto then show ?thesis apply (rule_tac x=y in bexI) unfolding dist_norm using pos_less_divide_eq[OF *] apply auto done qed qed then obtain y where "y \ S" and y: "norm (y - x) * (1 - e) < e * d" by auto define z where "z = c + ((1 - e) / e) *\<^sub>R (x - y)" have *: "x - e *\<^sub>R (x - c) = y - e *\<^sub>R (y - z)" unfolding z_def using \e > 0\ by (auto simp add: scaleR_right_diff_distrib scaleR_right_distrib scaleR_left_diff_distrib) have "z \ interior S" apply (rule interior_mono[OF d,unfolded subset_eq,rule_format]) unfolding interior_open[OF open_ball] mem_ball z_def dist_norm using y and assms(4,5) by simp (simp add: field_simps norm_minus_commute) then show ?thesis unfolding * using mem_interior_convex_shrink \y \ S\ assms by blast qed lemma in_interior_closure_convex_segment: fixes S :: "'a::euclidean_space set" assumes "convex S" and a: "a \ interior S" and b: "b \ closure S" shows "open_segment a b \ interior S" proof (clarsimp simp: in_segment) fix u::real assume u: "0 < u" "u < 1" have "(1 - u) *\<^sub>R a + u *\<^sub>R b = b - (1 - u) *\<^sub>R (b - a)" by (simp add: algebra_simps) also have "... \ interior S" using mem_interior_closure_convex_shrink [OF assms] u by simp finally show "(1 - u) *\<^sub>R a + u *\<^sub>R b \ interior S" . qed lemma closure_open_Int_superset: assumes "open S" "S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(S \ T)" by (metis assms closed_closure closure_minimal inf.orderE open_Int_closure_subset) then show ?thesis by (simp add: closure_mono dual_order.antisym) qed lemma convex_closure_interior: fixes S :: "'a::euclidean_space set" assumes "convex S" and int: "interior S \ {}" shows "closure(interior S) = closure S" proof - obtain a where a: "a \ interior S" using int by auto have "closure S \ closure(interior S)" proof fix x assume x: "x \ closure S" show "x \ closure (interior S)" proof (cases "x=a") case True then show ?thesis using \a \ interior S\ closure_subset by blast next case False show ?thesis proof (clarsimp simp add: closure_def islimpt_approachable) fix e::real assume xnotS: "x \ interior S" and "0 < e" show "\x'\interior S. x' \ x \ dist x' x < e" proof (intro bexI conjI) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ x" using False \0 < e\ by (auto simp: algebra_simps min_def) show "dist (x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a)) x < e" using \0 < e\ by (auto simp: dist_norm min_def) show "x - min (e/2 / norm (x - a)) 1 *\<^sub>R (x - a) \ interior S" apply (clarsimp simp add: min_def a) apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a x]) using \0 < e\ False apply (auto simp: field_split_simps) done qed qed qed qed then show ?thesis by (simp add: closure_mono interior_subset subset_antisym) qed lemma closure_convex_Int_superset: fixes S :: "'a::euclidean_space set" assumes "convex S" "interior S \ {}" "interior S \ closure T" shows "closure(S \ T) = closure S" proof - have "closure S \ closure(interior S)" by (simp add: convex_closure_interior assms) also have "... \ closure (S \ T)" using interior_subset [of S] assms by (metis (no_types, lifting) Int_assoc Int_lower2 closure_mono closure_open_Int_superset inf.orderE open_interior) finally show ?thesis by (simp add: closure_mono dual_order.antisym) qed subsection\<^marker>\tag unimportant\ \Some obvious but surprisingly hard simplex lemmas\ lemma simplex: assumes "finite S" and "0 \ S" shows "convex hull (insert 0 S) = {y. \u. (\x\S. 0 \ u x) \ sum u S \ 1 \ sum (\x. u x *\<^sub>R x) S = y}" proof (simp add: convex_hull_finite set_eq_iff assms, safe) fix x and u :: "'a \ real" assume "0 \ u 0" "\x\S. 0 \ u x" "u 0 + sum u S = 1" then show "\v. (\x\S. 0 \ v x) \ sum v S \ 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by force next fix x and u :: "'a \ real" assume "\x\S. 0 \ u x" "sum u S \ 1" then show "\v. 0 \ v 0 \ (\x\S. 0 \ v x) \ v 0 + sum v S = 1 \ (\x\S. v x *\<^sub>R x) = (\x\S. u x *\<^sub>R x)" by (rule_tac x="\x. if x = 0 then 1 - sum u S else u x" in exI) (auto simp: sum_delta_notmem assms if_smult) qed lemma substd_simplex: assumes d: "d \ Basis" shows "convex hull (insert 0 d) = {x. (\i\Basis. 0 \ x\i) \ (\i\d. x\i) \ 1 \ (\i\Basis. i \ d \ x\i = 0)}" (is "convex hull (insert 0 ?p) = ?s") proof - let ?D = d have "0 \ ?p" using assms by (auto simp: image_def) from d have "finite d" by (blast intro: finite_subset finite_Basis) show ?thesis unfolding simplex[OF \finite d\ \0 \ ?p\] proof (intro set_eqI; safe) fix u :: "'a \ real" assume as: "\x\?D. 0 \ u x" "sum u ?D \ 1" let ?x = "(\x\?D. u x *\<^sub>R x)" have ind: "\i\Basis. i \ d \ u i = ?x \ i" and notind: "(\i\Basis. i \ d \ ?x \ i = 0)" using substdbasis_expansion_unique[OF assms] by blast+ then have **: "sum u ?D = sum ((\) ?x) ?D" using assms by (auto intro!: sum.cong) show "0 \ ?x \ i" if "i \ Basis" for i using as(1) ind notind that by fastforce show "sum ((\) ?x) ?D \ 1" using "**" as(2) by linarith show "?x \ i = 0" if "i \ Basis" "i \ d" for i using notind that by blast next fix x assume "\i\Basis. 0 \ x \ i" "sum ((\) x) ?D \ 1" "(\i\Basis. i \ d \ x \ i = 0)" with d show "\u. (\x\?D. 0 \ u x) \ sum u ?D \ 1 \ (\x\?D. u x *\<^sub>R x) = x" unfolding substdbasis_expansion_unique[OF assms] by (rule_tac x="inner x" in exI) auto qed qed lemma std_simplex: "convex hull (insert 0 Basis) = {x::'a::euclidean_space. (\i\Basis. 0 \ x\i) \ sum (\i. x\i) Basis \ 1}" using substd_simplex[of Basis] by auto lemma interior_std_simplex: "interior (convex hull (insert 0 Basis)) = {x::'a::euclidean_space. (\i\Basis. 0 < x\i) \ sum (\i. x\i) Basis < 1}" unfolding set_eq_iff mem_interior std_simplex proof (intro allI iffI CollectI; clarify) fix x :: 'a fix e assume "e > 0" and as: "ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" show "(\i\Basis. 0 < x \ i) \ sum ((\) x) Basis < 1" proof safe fix i :: 'a assume i: "i \ Basis" then show "0 < x \ i" using as[THEN subsetD[where c="x - (e / 2) *\<^sub>R i"]] and \e > 0\ by (force simp add: inner_simps) next have **: "dist x (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) < e" using \e > 0\ unfolding dist_norm by (auto intro!: mult_strict_left_mono simp: SOME_Basis) have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R (SOME i. i\Basis)) \ i = x\i + (if i = (SOME i. i\Basis) then e/2 else 0)" by (auto simp: SOME_Basis inner_Basis inner_simps) then have *: "sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis = sum (\i. x\i + (if (SOME i. i\Basis) = i then e/2 else 0)) Basis" by (auto simp: intro!: sum.cong) have "sum ((\) x) Basis < sum ((\) (x + (e / 2) *\<^sub>R (SOME i. i\Basis))) Basis" using \e > 0\ DIM_positive by (auto simp: SOME_Basis sum.distrib *) also have "\ \ 1" using ** as by force finally show "sum ((\) x) Basis < 1" by auto qed next fix x :: 'a assume as: "\i\Basis. 0 < x \ i" "sum ((\) x) Basis < 1" obtain a :: 'b where "a \ UNIV" using UNIV_witness .. let ?d = "(1 - sum ((\) x) Basis) / real (DIM('a))" show "\e>0. ball x e \ {x. (\i\Basis. 0 \ x \ i) \ sum ((\) x) Basis \ 1}" proof (rule_tac x="min (Min (((\) x) ` Basis)) D" for D in exI, intro conjI subsetI CollectI) fix y assume y: "y \ ball x (min (Min ((\) x ` Basis)) ?d)" have "sum ((\) y) Basis \ sum (\i. x\i + ?d) Basis" proof (rule sum_mono) fix i :: 'a assume i: "i \ Basis" have "\y\i - x\i\ \ norm (y - x)" by (metis Basis_le_norm i inner_commute inner_diff_right) also have "... < ?d" using y by (simp add: dist_norm norm_minus_commute) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant by (auto simp add: Suc_le_eq) finally show "sum ((\) y) Basis \ 1" . show "(\i\Basis. 0 \ y \ i)" proof safe fix i :: 'a assume i: "i \ Basis" have "norm (x - y) < Min (((\) x) ` Basis)" using y by (auto simp: dist_norm less_eq_real_def) also have "... \ x\i" using i by auto finally have "norm (x - y) < x\i" . then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format, OF i] by (auto simp: inner_simps) qed next have "Min (((\) x) ` Basis) > 0" using as by simp moreover have "?d > 0" using as by (auto simp: Suc_le_eq) ultimately show "0 < min (Min ((\) x ` Basis)) ((1 - sum ((\) x) Basis) / real DIM('a))" by linarith qed qed lemma interior_std_simplex_nonempty: obtains a :: "'a::euclidean_space" where "a \ interior(convex hull (insert 0 Basis))" proof - let ?D = "Basis :: 'a set" let ?a = "sum (\b::'a. inverse (2 * real DIM('a)) *\<^sub>R b) Basis" { fix i :: 'a assume i: "i \ Basis" have "?a \ i = inverse (2 * real DIM('a))" by (rule trans[of _ "sum (\j. if i = j then inverse (2 * real DIM('a)) else 0) ?D"]) (simp_all add: sum.If_cases i) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding interior_std_simplex mem_Collect_eq proof safe fix i :: 'a assume i: "i \ Basis" show "0 < ?a \ i" - unfolding **[OF i] by (auto simp add: Suc_le_eq DIM_positive) + unfolding **[OF i] by (auto simp add: Suc_le_eq) next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real DIM('a))) ?D" apply (rule sum.cong) apply rule apply auto done also have "\ < 1" unfolding sum_constant divide_inverse[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto qed qed lemma rel_interior_substd_simplex: assumes D: "D \ Basis" shows "rel_interior (convex hull (insert 0 D)) = {x::'a::euclidean_space. (\i\D. 0 < x\i) \ (\i\D. x\i) < 1 \ (\i\Basis. i \ D \ x\i = 0)}" (is "rel_interior (convex hull (insert 0 ?p)) = ?s") proof - have "finite D" using D finite_Basis finite_subset by blast show ?thesis proof (cases "D = {}") case True then show ?thesis using rel_interior_sing using euclidean_eq_iff[of _ 0] by auto next case False have h0: "affine hull (convex hull (insert 0 ?p)) = {x::'a::euclidean_space. (\i\Basis. i \ D \ x\i = 0)}" using affine_hull_convex_hull affine_hull_substd_basis assms by auto have aux: "\x::'a. \i\Basis. (\i\D. 0 \ x\i) \ (\i\Basis. i \ D \ x\i = 0) \ 0 \ x\i" by auto { fix x :: "'a::euclidean_space" assume x: "x \ rel_interior (convex hull (insert 0 ?p))" then obtain e where "e > 0" and "ball x e \ {xa. (\i\Basis. i \ D \ xa\i = 0)} \ convex hull (insert 0 ?p)" using mem_rel_interior_ball[of x "convex hull (insert 0 ?p)"] h0 by auto then have as [rule_format]: "\y. dist x y < e \ (\i\Basis. i \ D \ y\i = 0) \ (\i\D. 0 \ y \ i) \ sum ((\) y) D \ 1" unfolding ball_def unfolding substd_simplex[OF assms] using assms by auto have x0: "(\i\Basis. i \ D \ x\i = 0)" using x rel_interior_subset substd_simplex[OF assms] by auto have "(\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x\i = 0)" proof (intro conjI ballI) fix i :: 'a assume "i \ D" then have "\j\D. 0 \ (x - (e / 2) *\<^sub>R i) \ j" apply - apply (rule as[THEN conjunct1]) using D \e > 0\ x0 apply (auto simp: dist_norm inner_simps inner_Basis) done then show "0 < x \ i" using \e > 0\ \i \ D\ D by (force simp: inner_simps inner_Basis) next obtain a where a: "a \ D" using \D \ {}\ by auto then have **: "dist x (x + (e / 2) *\<^sub>R a) < e" using \e > 0\ norm_Basis[of a] D unfolding dist_norm by auto have "\i. i \ Basis \ (x + (e / 2) *\<^sub>R a) \ i = x\i + (if i = a then e/2 else 0)" using a D by (auto simp: inner_simps inner_Basis) then have *: "sum ((\) (x + (e / 2) *\<^sub>R a)) D = sum (\i. x\i + (if a = i then e/2 else 0)) D" using D by (intro sum.cong) auto have "a \ Basis" using \a \ D\ D by auto then have h1: "(\i\Basis. i \ D \ (x + (e / 2) *\<^sub>R a) \ i = 0)" using x0 D \a\D\ by (auto simp add: inner_add_left inner_Basis) have "sum ((\) x) D < sum ((\) (x + (e / 2) *\<^sub>R a)) D" using \e > 0\ \a \ D\ \finite D\ by (auto simp add: * sum.distrib) also have "\ \ 1" using ** h1 as[rule_format, of "x + (e / 2) *\<^sub>R a"] by auto finally show "sum ((\) x) D < 1" "\i. i\Basis \ i \ D \ x\i = 0" using x0 by auto qed } moreover { fix x :: "'a::euclidean_space" assume as: "x \ ?s" have "\i. 0 < x\i \ 0 = x\i \ 0 \ x\i" by auto moreover have "\i. i \ D \ i \ D" by auto ultimately have "\i. (\i\D. 0 < x\i) \ (\i. i \ D \ x\i = 0) \ 0 \ x\i" by metis then have h2: "x \ convex hull (insert 0 ?p)" using as assms unfolding substd_simplex[OF assms] by fastforce obtain a where a: "a \ D" using \D \ {}\ by auto let ?d = "(1 - sum ((\) x) D) / real (card D)" have "0 < card D" using \D \ {}\ \finite D\ by (simp add: card_gt_0_iff) have "Min (((\) x) ` D) > 0" using as \D \ {}\ \finite D\ by (simp add: Min_gr_iff) moreover have "?d > 0" using as using \0 < card D\ by auto ultimately have h3: "min (Min (((\) x) ` D)) ?d > 0" by auto have "x \ rel_interior (convex hull (insert 0 ?p))" unfolding rel_interior_ball mem_Collect_eq h0 apply (rule,rule h2) unfolding substd_simplex[OF assms] apply (rule_tac x="min (Min (((\) x) ` D)) ?d" in exI) apply (rule, rule h3) apply safe unfolding mem_ball proof - fix y :: 'a assume y: "dist x y < min (Min ((\) x ` D)) ?d" assume y2: "\i\Basis. i \ D \ y\i = 0" have "sum ((\) y) D \ sum (\i. x\i + ?d) D" proof (rule sum_mono) fix i assume "i \ D" with D have i: "i \ Basis" by auto have "\y\i - x\i\ \ norm (y - x)" by (metis i inner_commute inner_diff_right norm_bound_Basis_le order_refl) also have "... < ?d" by (metis dist_norm min_less_iff_conj norm_minus_commute y) finally have "\y\i - x\i\ < ?d" . then show "y \ i \ x \ i + ?d" by auto qed also have "\ \ 1" unfolding sum.distrib sum_constant using \0 < card D\ by auto finally show "sum ((\) y) D \ 1" . fix i :: 'a assume i: "i \ Basis" then show "0 \ y\i" proof (cases "i\D") case True have "norm (x - y) < x\i" using y[unfolded min_less_iff_conj dist_norm, THEN conjunct1] using Min_gr_iff[of "(\) x ` D" "norm (x - y)"] \0 < card D\ \i \ D\ by (simp add: card_gt_0_iff) then show "0 \ y\i" using Basis_le_norm[OF i, of "x - y"] and as(1)[rule_format] by (auto simp: inner_simps) qed (insert y2, auto) qed } ultimately have "\x. x \ rel_interior (convex hull insert 0 D) \ x \ {x. (\i\D. 0 < x \ i) \ sum ((\) x) D < 1 \ (\i\Basis. i \ D \ x \ i = 0)}" by blast then show ?thesis by (rule set_eqI) qed qed lemma rel_interior_substd_simplex_nonempty: assumes "D \ {}" and "D \ Basis" obtains a :: "'a::euclidean_space" where "a \ rel_interior (convex hull (insert 0 D))" proof - let ?D = D let ?a = "sum (\b::'a::euclidean_space. inverse (2 * real (card D)) *\<^sub>R b) ?D" have "finite D" apply (rule finite_subset) using assms(2) apply auto done then have d1: "0 < real (card D)" using \D \ {}\ by auto { fix i assume "i \ D" have "?a \ i = inverse (2 * real (card D))" apply (rule trans[of _ "sum (\j. if i = j then inverse (2 * real (card D)) else 0) ?D"]) unfolding inner_sum_left apply (rule sum.cong) using \i \ D\ \finite D\ sum.delta'[of D i "(\k. inverse (2 * real (card D)))"] d1 assms(2) by (auto simp: inner_Basis rev_subsetD[OF _ assms(2)]) } note ** = this show ?thesis apply (rule that[of ?a]) unfolding rel_interior_substd_simplex[OF assms(2)] mem_Collect_eq proof safe fix i assume "i \ D" have "0 < inverse (2 * real (card D))" using d1 by auto also have "\ = ?a \ i" using **[of i] \i \ D\ by auto finally show "0 < ?a \ i" by auto next have "sum ((\) ?a) ?D = sum (\i. inverse (2 * real (card D))) ?D" by (rule sum.cong) (rule refl, rule **) also have "\ < 1" unfolding sum_constant divide_real_def[symmetric] by (auto simp add: field_simps) finally show "sum ((\) ?a) ?D < 1" by auto next fix i assume "i \ Basis" and "i \ D" have "?a \ span D" proof (rule span_sum[of D "(\b. b /\<^sub>R (2 * real (card D)))" D]) { fix x :: "'a::euclidean_space" assume "x \ D" then have "x \ span D" using span_base[of _ "D"] by auto then have "x /\<^sub>R (2 * real (card D)) \ span D" using span_mul[of x "D" "(inverse (real (card D)) / 2)"] by auto } then show "\x. x\D \ x /\<^sub>R (2 * real (card D)) \ span D" by auto qed then show "?a \ i = 0 " using \i \ D\ unfolding span_substd_basis[OF assms(2)] using \i \ Basis\ by auto qed qed subsection\<^marker>\tag unimportant\ \Relative interior of convex set\ lemma rel_interior_convex_nonempty_aux: fixes S :: "'n::euclidean_space set" assumes "convex S" and "0 \ S" shows "rel_interior S \ {}" proof (cases "S = {0}") case True then show ?thesis using rel_interior_sing by auto next case False obtain B where B: "independent B \ B \ S \ S \ span B \ card B = dim S" using basis_exists[of S] by metis then have "B \ {}" using B assms \S \ {0}\ span_empty by auto have "insert 0 B \ span B" using subspace_span[of B] subspace_0[of "span B"] span_superset by auto then have "span (insert 0 B) \ span B" using span_span[of B] span_mono[of "insert 0 B" "span B"] by blast then have "convex hull insert 0 B \ span B" using convex_hull_subset_span[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) \ span B" using span_span[of B] span_mono[of "convex hull insert 0 B" "span B"] by blast then have *: "span (convex hull insert 0 B) = span B" using span_mono[of B "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto then have "span (convex hull insert 0 B) = span S" using B span_mono[of B S] span_mono[of S "span B"] span_span[of B] by auto moreover have "0 \ affine hull (convex hull insert 0 B)" using hull_subset[of "convex hull insert 0 B"] hull_subset[of "insert 0 B"] by auto ultimately have **: "affine hull (convex hull insert 0 B) = affine hull S" using affine_hull_span_0[of "convex hull insert 0 B"] affine_hull_span_0[of "S"] assms hull_subset[of S] by auto obtain d and f :: "'n \ 'n" where fd: "card d = card B" "linear f" "f ` B = d" "f ` span B = {x. \i\Basis. i \ d \ x \ i = (0::real)} \ inj_on f (span B)" and d: "d \ Basis" using basis_to_substdbasis_subspace_isomorphism[of B,OF _ ] B by auto then have "bounded_linear f" using linear_conv_bounded_linear by auto have "d \ {}" using fd B \B \ {}\ by auto have "insert 0 d = f ` (insert 0 B)" using fd linear_0 by auto then have "(convex hull (insert 0 d)) = f ` (convex hull (insert 0 B))" using convex_hull_linear_image[of f "(insert 0 d)"] convex_hull_linear_image[of f "(insert 0 B)"] \linear f\ by auto moreover have "rel_interior (f ` (convex hull insert 0 B)) = f ` rel_interior (convex hull insert 0 B)" apply (rule rel_interior_injective_on_span_linear_image[of f "(convex hull insert 0 B)"]) using \bounded_linear f\ fd * apply auto done ultimately have "rel_interior (convex hull insert 0 B) \ {}" using rel_interior_substd_simplex_nonempty[OF \d \ {}\ d] apply auto apply blast done moreover have "convex hull (insert 0 B) \ S" using B assms hull_mono[of "insert 0 B" "S" "convex"] convex_hull_eq by auto ultimately show ?thesis using subset_rel_interior[of "convex hull insert 0 B" S] ** by auto qed lemma rel_interior_eq_empty: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior S = {} \ S = {}" proof - { assume "S \ {}" then obtain a where "a \ S" by auto then have "0 \ (+) (-a) ` S" using assms exI[of "(\x. x \ S \ - a + x = 0)" a] by auto then have "rel_interior ((+) (-a) ` S) \ {}" using rel_interior_convex_nonempty_aux[of "(+) (-a) ` S"] convex_translation[of S "-a"] assms by auto then have "rel_interior S \ {}" using rel_interior_translation [of "- a"] by simp } then show ?thesis using rel_interior_empty by auto qed lemma interior_simplex_nonempty: fixes S :: "'N :: euclidean_space set" assumes "independent S" "finite S" "card S = DIM('N)" obtains a where "a \ interior (convex hull (insert 0 S))" proof - have "affine hull (insert 0 S) = UNIV" by (simp add: hull_inc affine_hull_span_0 dim_eq_full[symmetric] assms(1) assms(3) dim_eq_card_independent) moreover have "rel_interior (convex hull insert 0 S) \ {}" using rel_interior_eq_empty [of "convex hull (insert 0 S)"] by auto ultimately have "interior (convex hull insert 0 S) \ {}" by (simp add: rel_interior_interior) with that show ?thesis by auto qed lemma convex_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "convex (rel_interior S)" proof - { fix x y and u :: real assume assm: "x \ rel_interior S" "y \ rel_interior S" "0 \ u" "u \ 1" then have "x \ S" using rel_interior_subset by auto have "x - u *\<^sub>R (x-y) \ rel_interior S" proof (cases "0 = u") case False then have "0 < u" using assm by auto then show ?thesis using assm rel_interior_convex_shrink[of S y x u] assms \x \ S\ by auto next case True then show ?thesis using assm by auto qed then have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ rel_interior S" by (simp add: algebra_simps) } then show ?thesis unfolding convex_alt by auto qed lemma convex_closure_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "closure (rel_interior S) = closure S" proof - have h1: "closure (rel_interior S) \ closure S" using closure_mono[of "rel_interior S" S] rel_interior_subset[of S] by auto show ?thesis proof (cases "S = {}") case False then obtain a where a: "a \ rel_interior S" using rel_interior_eq_empty assms by auto { fix x assume x: "x \ closure S" { assume "x = a" then have "x \ closure (rel_interior S)" using a unfolding closure_def by auto } moreover { assume "x \ a" { fix e :: real assume "e > 0" define e1 where "e1 = min 1 (e/norm (x - a))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (x - a) \ e" using \x \ a\ \e > 0\ le_divide_eq[of e1 e "norm (x - a)"] by simp_all then have *: "x - e1 *\<^sub>R (x - a) \ rel_interior S" using rel_interior_closure_convex_shrink[of S a x e1] assms x a e1_def by auto have "\y. y \ rel_interior S \ y \ x \ dist y x \ e" apply (rule_tac x="x - e1 *\<^sub>R (x - a)" in exI) using * e1 dist_norm[of "x - e1 *\<^sub>R (x - a)" x] \x \ a\ apply simp done } then have "x islimpt rel_interior S" unfolding islimpt_approachable_le by auto then have "x \ closure(rel_interior S)" unfolding closure_def by auto } ultimately have "x \ closure(rel_interior S)" by auto } then show ?thesis using h1 by auto next case True then have "rel_interior S = {}" using rel_interior_empty by auto then have "closure (rel_interior S) = {}" using closure_empty by auto with True show ?thesis by auto qed qed lemma rel_interior_same_affine_hull: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "affine hull (rel_interior S) = affine hull S" by (metis assms closure_same_affine_hull convex_closure_rel_interior) lemma rel_interior_aff_dim: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "aff_dim (rel_interior S) = aff_dim S" by (metis aff_dim_affine_hull2 assms rel_interior_same_affine_hull) lemma rel_interior_rel_interior: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (rel_interior S) = rel_interior S" proof - have "openin (top_of_set (affine hull (rel_interior S))) (rel_interior S)" using openin_rel_interior[of S] rel_interior_same_affine_hull[of S] assms by auto then show ?thesis using rel_interior_def by auto qed lemma rel_interior_rel_open: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_open (rel_interior S)" unfolding rel_open_def using rel_interior_rel_interior assms by auto lemma convex_rel_interior_closure_aux: fixes x y z :: "'n::euclidean_space" assumes "0 < a" "0 < b" "(a + b) *\<^sub>R z = a *\<^sub>R x + b *\<^sub>R y" obtains e where "0 < e" "e \ 1" "z = y - e *\<^sub>R (y - x)" proof - define e where "e = a / (a + b)" have "z = (1 / (a + b)) *\<^sub>R ((a + b) *\<^sub>R z)" using assms by (simp add: eq_vector_fraction_iff) also have "\ = (1 / (a + b)) *\<^sub>R (a *\<^sub>R x + b *\<^sub>R y)" using assms scaleR_cancel_left[of "1/(a+b)" "(a + b) *\<^sub>R z" "a *\<^sub>R x + b *\<^sub>R y"] by auto also have "\ = y - e *\<^sub>R (y-x)" using e_def apply (simp add: algebra_simps) using scaleR_left_distrib[of "a/(a+b)" "b/(a+b)" y] assms add_divide_distrib[of a b "a+b"] apply auto done finally have "z = y - e *\<^sub>R (y-x)" by auto moreover have "e > 0" using e_def assms by auto moreover have "e \ 1" using e_def assms by auto ultimately show ?thesis using that[of e] by auto qed lemma convex_rel_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "rel_interior (closure S) = rel_interior S" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_eq_empty by auto next case False have "rel_interior (closure S) \ rel_interior S" using subset_rel_interior[of S "closure S"] closure_same_affine_hull closure_subset by auto moreover { fix z assume z: "z \ rel_interior (closure S)" obtain x where x: "x \ rel_interior S" using \S \ {}\ assms rel_interior_eq_empty by auto have "z \ rel_interior S" proof (cases "x = z") case True then show ?thesis using x by auto next case False obtain e where e: "e > 0" "cball z e \ affine hull closure S \ closure S" using z rel_interior_cball[of "closure S"] by auto hence *: "0 < e/norm(z-x)" using e False by auto define y where "y = z + (e/norm(z-x)) *\<^sub>R (z-x)" have yball: "y \ cball z e" using mem_cball y_def dist_norm[of z y] e by auto have "x \ affine hull closure S" using x rel_interior_subset_closure hull_inc[of x "closure S"] by blast moreover have "z \ affine hull closure S" using z rel_interior_subset hull_subset[of "closure S"] by blast ultimately have "y \ affine hull closure S" using y_def affine_affine_hull[of "closure S"] mem_affine_3_minus [of "affine hull closure S" z z x "e/norm(z-x)"] by auto then have "y \ closure S" using e yball by auto have "(1 + (e/norm(z-x))) *\<^sub>R z = (e/norm(z-x)) *\<^sub>R x + y" using y_def by (simp add: algebra_simps) then obtain e1 where "0 < e1" "e1 \ 1" "z = y - e1 *\<^sub>R (y - x)" using * convex_rel_interior_closure_aux[of "e / norm (z - x)" 1 z x y] by (auto simp add: algebra_simps) then show ?thesis using rel_interior_closure_convex_shrink assms x \y \ closure S\ by auto qed } ultimately show ?thesis by auto qed lemma convex_interior_closure: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "interior (closure S) = interior S" using closure_aff_dim[of S] interior_rel_interior_gen[of S] interior_rel_interior_gen[of "closure S"] convex_rel_interior_closure[of S] assms by auto lemma closure_eq_rel_interior_eq: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 = rel_interior S2" by (metis convex_rel_interior_closure convex_closure_rel_interior assms) lemma closure_eq_between: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" shows "closure S1 = closure S2 \ rel_interior S1 \ S2 \ S2 \ closure S1" (is "?A \ ?B") proof assume ?A then show ?B by (metis assms closure_subset convex_rel_interior_closure rel_interior_subset) next assume ?B then have "closure S1 \ closure S2" by (metis assms(1) convex_closure_rel_interior closure_mono) moreover from \?B\ have "closure S1 \ closure S2" by (metis closed_closure closure_minimal) ultimately show ?A .. qed lemma open_inter_closure_rel_interior: fixes S A :: "'n::euclidean_space set" assumes "convex S" and "open A" shows "A \ closure S = {} \ A \ rel_interior S = {}" by (metis assms convex_closure_rel_interior open_Int_closure_eq_empty) lemma rel_interior_open_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis apply (simp add: rel_interior_eq openin_open) apply (rule_tac x="ball (inverse 2 *\<^sub>R (a + b)) (norm(b - a) / 2)" in exI) apply (simp add: open_segment_as_ball) done qed lemma rel_interior_closed_segment: fixes a :: "'a :: euclidean_space" shows "rel_interior(closed_segment a b) = (if a = b then {a} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by auto next case False then show ?thesis by simp (metis closure_open_segment convex_open_segment convex_rel_interior_closure rel_interior_open_segment) qed lemmas rel_interior_segment = rel_interior_closed_segment rel_interior_open_segment lemma starlike_convex_tweak_boundary_points: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and ST: "rel_interior S \ T" and TS: "T \ closure S" shows "starlike T" proof - have "rel_interior S \ {}" by (simp add: assms rel_interior_eq_empty) then obtain a where a: "a \ rel_interior S" by blast with ST have "a \ T" by blast have *: "\x. x \ T \ open_segment a x \ rel_interior S" apply (rule rel_interior_closure_convex_segment [OF \convex S\ a]) using assms by blast show ?thesis unfolding starlike_def apply (rule bexI [OF _ \a \ T\]) apply (simp add: closed_segment_eq_open) apply (intro conjI ballI a \a \ T\ rel_interior_closure_convex_segment [OF \convex S\ a]) apply (simp add: order_trans [OF * ST]) done qed subsection\The relative frontier of a set\ definition\<^marker>\tag important\ "rel_frontier S = closure S - rel_interior S" lemma rel_frontier_empty [simp]: "rel_frontier {} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_eq_empty: fixes S :: "'n::euclidean_space set" shows "rel_frontier S = {} \ affine S" unfolding rel_frontier_def using rel_interior_subset_closure by (auto simp add: rel_interior_eq_closure [symmetric]) lemma rel_frontier_sing [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier {a} = {}" by (simp add: rel_frontier_def) lemma rel_frontier_affine_hull: fixes S :: "'a::euclidean_space set" shows "rel_frontier S \ affine hull S" using closure_affine_hull rel_frontier_def by fastforce lemma rel_frontier_cball [simp]: fixes a :: "'n::euclidean_space" shows "rel_frontier(cball a r) = (if r = 0 then {} else sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by (force simp: sphere_def) next case equal then show ?thesis by simp next case greater then show ?thesis apply simp by (metis centre_in_ball empty_iff frontier_cball frontier_def interior_cball interior_rel_interior_gen rel_frontier_def) qed lemma rel_frontier_translation: fixes a :: "'a::euclidean_space" shows "rel_frontier((\x. a + x) ` S) = (\x. a + x) ` (rel_frontier S)" by (simp add: rel_frontier_def translation_diff rel_interior_translation closure_translation) lemma closed_affine_hull [iff]: fixes S :: "'n::euclidean_space set" shows "closed (affine hull S)" by (metis affine_affine_hull affine_closed) lemma rel_frontier_nonempty_interior: fixes S :: "'n::euclidean_space set" shows "interior S \ {} \ rel_frontier S = frontier S" by (metis frontier_def interior_rel_interior_gen rel_frontier_def) lemma rel_frontier_frontier: fixes S :: "'n::euclidean_space set" shows "affine hull S = UNIV \ rel_frontier S = frontier S" by (simp add: frontier_def rel_frontier_def rel_interior_interior) lemma closest_point_in_rel_frontier: "\closed S; S \ {}; x \ affine hull S - rel_interior S\ \ closest_point S x \ rel_frontier S" by (simp add: closest_point_in_rel_interior closest_point_in_set rel_frontier_def) lemma closed_rel_frontier [iff]: fixes S :: "'n::euclidean_space set" shows "closed (rel_frontier S)" proof - have *: "closedin (top_of_set (affine hull S)) (closure S - rel_interior S)" by (simp add: closed_subset closedin_diff closure_affine_hull openin_rel_interior) show ?thesis apply (rule closedin_closed_trans[of "affine hull S" "rel_frontier S"]) unfolding rel_frontier_def using * closed_affine_hull apply auto done qed lemma closed_rel_boundary: fixes S :: "'n::euclidean_space set" shows "closed S \ closed(S - rel_interior S)" by (metis closed_rel_frontier closure_closed rel_frontier_def) lemma compact_rel_boundary: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(S - rel_interior S)" by (metis bounded_diff closed_rel_boundary closure_eq compact_closure compact_imp_closed) lemma bounded_rel_frontier: fixes S :: "'n::euclidean_space set" shows "bounded S \ bounded(rel_frontier S)" by (simp add: bounded_closure bounded_diff rel_frontier_def) lemma compact_rel_frontier_bounded: fixes S :: "'n::euclidean_space set" shows "bounded S \ compact(rel_frontier S)" using bounded_rel_frontier closed_rel_frontier compact_eq_bounded_closed by blast lemma compact_rel_frontier: fixes S :: "'n::euclidean_space set" shows "compact S \ compact(rel_frontier S)" by (meson compact_eq_bounded_closed compact_rel_frontier_bounded) lemma convex_same_rel_interior_closure: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ closure S = closure T" by (simp add: closure_eq_rel_interior_eq) lemma convex_same_rel_interior_closure_straddle: fixes S :: "'n::euclidean_space set" shows "\convex S; convex T\ \ rel_interior S = rel_interior T \ rel_interior S \ T \ T \ closure S" by (simp add: closure_eq_between convex_same_rel_interior_closure) lemma convex_rel_frontier_aff_dim: fixes S1 S2 :: "'n::euclidean_space set" assumes "convex S1" and "convex S2" and "S2 \ {}" and "S1 \ rel_frontier S2" shows "aff_dim S1 < aff_dim S2" proof - have "S1 \ closure S2" using assms unfolding rel_frontier_def by auto then have *: "affine hull S1 \ affine hull S2" using hull_mono[of "S1" "closure S2"] closure_same_affine_hull[of S2] by blast then have "aff_dim S1 \ aff_dim S2" using * aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] aff_dim_subset[of "affine hull S1" "affine hull S2"] by auto moreover { assume eq: "aff_dim S1 = aff_dim S2" then have "S1 \ {}" using aff_dim_empty[of S1] aff_dim_empty[of S2] \S2 \ {}\ by auto have **: "affine hull S1 = affine hull S2" apply (rule affine_dim_equal) using * affine_affine_hull apply auto using \S1 \ {}\ hull_subset[of S1] apply auto using eq aff_dim_affine_hull[of S1] aff_dim_affine_hull[of S2] apply auto done obtain a where a: "a \ rel_interior S1" using \S1 \ {}\ rel_interior_eq_empty assms by auto obtain T where T: "open T" "a \ T \ S1" "T \ affine hull S1 \ S1" using mem_rel_interior[of a S1] a by auto then have "a \ T \ closure S2" using a assms unfolding rel_frontier_def by auto then obtain b where b: "b \ T \ rel_interior S2" using open_inter_closure_rel_interior[of S2 T] assms T by auto then have "b \ affine hull S1" using rel_interior_subset hull_subset[of S2] ** by auto then have "b \ S1" using T b by auto then have False using b assms unfolding rel_frontier_def by auto } ultimately show ?thesis using less_le by auto qed lemma convex_rel_interior_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "z \ rel_interior S" shows "\x\affine hull S. \m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" proof - obtain e1 where e1: "e1 > 0 \ cball z e1 \ affine hull S \ S" using mem_rel_interior_cball[of z S] assms by auto { fix x assume x: "x \ affine hull S" { assume "x \ z" define m where "m = 1 + e1/norm(x-z)" hence "m > 1" using e1 \x \ z\ by auto { fix e assume e: "e > 1 \ e \ m" have "z \ affine hull S" using assms rel_interior_subset hull_subset[of S] by auto then have *: "(1 - e)*\<^sub>R x + e *\<^sub>R z \ affine hull S" using mem_affine[of "affine hull S" x z "(1-e)" e] affine_affine_hull[of S] x by auto have "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) = norm ((e - 1) *\<^sub>R (x - z))" by (simp add: algebra_simps) also have "\ = (e - 1) * norm (x-z)" using norm_scaleR e by auto also have "\ \ (m - 1) * norm (x - z)" using e mult_right_mono[of _ _ "norm(x-z)"] by auto also have "\ = (e1 / norm (x - z)) * norm (x - z)" using m_def by auto also have "\ = e1" using \x \ z\ e1 by simp finally have **: "norm (z + e *\<^sub>R x - (x + e *\<^sub>R z)) \ e1" by auto have "(1 - e)*\<^sub>R x+ e *\<^sub>R z \ cball z e1" using m_def ** unfolding cball_def dist_norm by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x+ e *\<^sub>R z \ S" using e * e1 by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" using \m> 1 \ by auto } moreover { assume "x = z" define m where "m = 1 + e1" then have "m > 1" using e1 by auto { fix e assume e: "e > 1 \ e \ m" then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e1 x \x = z\ by (auto simp add: algebra_simps) then have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using e by auto } then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using \m > 1\ by auto } ultimately have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S )" by blast } then show ?thesis by auto qed lemma convex_rel_interior_if2: fixes S :: "'n::euclidean_space set" assumes "convex S" assumes "z \ rel_interior S" shows "\x\affine hull S. \e. e > 1 \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_if[of S z] assms by auto lemma convex_rel_interior_only_if: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" assumes "\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" shows "z \ rel_interior S" proof - obtain x where x: "x \ rel_interior S" using rel_interior_eq_empty assms by auto then have "x \ S" using rel_interior_subset by auto then obtain e where e: "e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using assms by auto define y where [abs_def]: "y = (1 - e) *\<^sub>R x + e *\<^sub>R z" then have "y \ S" using e by auto define e1 where "e1 = 1/e" then have "0 < e1 \ e1 < 1" using e by auto then have "z =y - (1 - e1) *\<^sub>R (y - x)" using e1_def y_def by (auto simp add: algebra_simps) then show ?thesis using rel_interior_convex_shrink[of S x y "1-e1"] \0 < e1 \ e1 < 1\ \y \ S\ x assms by auto qed lemma convex_rel_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S "affine"] convex_rel_interior_if[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_rel_interior_iff2: fixes S :: "'n::euclidean_space set" assumes "convex S" and "S \ {}" shows "z \ rel_interior S \ (\x\affine hull S. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" using assms hull_subset[of S] convex_rel_interior_if2[of S z] convex_rel_interior_only_if[of S z] by auto lemma convex_interior_iff: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "z \ interior S \ (\x. \e. e > 0 \ z + e *\<^sub>R x \ S)" proof (cases "aff_dim S = int DIM('n)") case False { assume "z \ interior S" then have False using False interior_rel_interior_gen[of S] by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0 \ z + e1 *\<^sub>R (x - z) \ S" using r by auto obtain e2 where e2: "e2 > 0 \ z + e2 *\<^sub>R (z - x) \ S" using r by auto define x1 where [abs_def]: "x1 = z + e1 *\<^sub>R (x - z)" then have x1: "x1 \ affine hull S" using e1 hull_subset[of S] by auto define x2 where [abs_def]: "x2 = z + e2 *\<^sub>R (z - x)" then have x2: "x2 \ affine hull S" using e2 hull_subset[of S] by auto have *: "e1/(e1+e2) + e2/(e1+e2) = 1" using add_divide_distrib[of e1 e2 "e1+e2"] e1 e2 by simp then have "z = (e2/(e1+e2)) *\<^sub>R x1 + (e1/(e1+e2)) *\<^sub>R x2" using x1_def x2_def apply (auto simp add: algebra_simps) using scaleR_left_distrib[of "e1/(e1+e2)" "e2/(e1+e2)" z] apply auto done then have z: "z \ affine hull S" using mem_affine[of "affine hull S" x1 x2 "e2/(e1+e2)" "e1/(e1+e2)"] x1 x2 affine_affine_hull[of S] * by auto have "x1 - x2 = (e1 + e2) *\<^sub>R (x - z)" using x1_def x2_def by (auto simp add: algebra_simps) then have "x = z+(1/(e1+e2)) *\<^sub>R (x1-x2)" using e1 e2 by simp then have "x \ affine hull S" using mem_affine_3_minus[of "affine hull S" z x1 x2 "1/(e1+e2)"] x1 x2 z affine_affine_hull[of S] by auto } then have "affine hull S = UNIV" by auto then have "aff_dim S = int DIM('n)" using aff_dim_affine_hull[of S] by (simp add: aff_dim_UNIV) then have False using False by auto } ultimately show ?thesis by auto next case True then have "S \ {}" using aff_dim_empty[of S] by auto have *: "affine hull S = UNIV" using True affine_hull_UNIV by auto { assume "z \ interior S" then have "z \ rel_interior S" using True interior_rel_interior_gen[of S] by auto then have **: "\x. \e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ * by auto fix x obtain e1 where e1: "e1 > 1" "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z \ S" using **[rule_format, of "z-x"] by auto define e where [abs_def]: "e = e1 - 1" then have "(1 - e1) *\<^sub>R (z - x) + e1 *\<^sub>R z = z + e *\<^sub>R x" by (simp add: algebra_simps) then have "e > 0" "z + e *\<^sub>R x \ S" using e1 e_def by auto then have "\e. e > 0 \ z + e *\<^sub>R x \ S" by auto } moreover { assume r: "\x. \e. e > 0 \ z + e *\<^sub>R x \ S" { fix x obtain e1 where e1: "e1 > 0" "z + e1 *\<^sub>R (z - x) \ S" using r[rule_format, of "z-x"] by auto define e where "e = e1 + 1" then have "z + e1 *\<^sub>R (z - x) = (1 - e) *\<^sub>R x + e *\<^sub>R z" by (simp add: algebra_simps) then have "e > 1" "(1 - e)*\<^sub>R x + e *\<^sub>R z \ S" using e1 e_def by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S" by auto } then have "z \ rel_interior S" using convex_rel_interior_iff2[of S z] assms \S \ {}\ by auto then have "z \ interior S" using True interior_rel_interior_gen[of S] by auto } ultimately show ?thesis by auto qed subsubsection\<^marker>\tag unimportant\ \Relative interior and closure under common operations\ lemma rel_interior_inter_aux: "\{rel_interior S |S. S \ I} \ \I" proof - { fix y assume "y \ \{rel_interior S |S. S \ I}" then have y: "\S \ I. y \ rel_interior S" by auto { fix S assume "S \ I" then have "y \ S" using rel_interior_subset y by auto } then have "y \ \I" by auto } then show ?thesis by auto qed lemma closure_Int: "closure (\I) \ \{closure S |S. S \ I}" proof - { fix y assume "y \ \I" then have y: "\S \ I. y \ S" by auto { fix S assume "S \ I" then have "y \ closure S" using closure_subset y by auto } then have "y \ \{closure S |S. S \ I}" by auto } then have "\I \ \{closure S |S. S \ I}" by auto moreover have "closed (\{closure S |S. S \ I})" unfolding closed_Inter closed_closure by auto ultimately show ?thesis using closure_hull[of "\I"] hull_minimal[of "\I" "\{closure S |S. S \ I}" "closed"] by auto qed lemma convex_closure_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" proof - obtain x where x: "\S\I. x \ rel_interior S" using assms by auto { fix y assume "y \ \{closure S |S. S \ I}" then have y: "\S \ I. y \ closure S" by auto { assume "y = x" then have "y \ closure (\{rel_interior S |S. S \ I})" using x closure_subset[of "\{rel_interior S |S. S \ I}"] by auto } moreover { assume "y \ x" { fix e :: real assume e: "e > 0" define e1 where "e1 = min 1 (e/norm (y - x))" then have e1: "e1 > 0" "e1 \ 1" "e1 * norm (y - x) \ e" using \y \ x\ \e > 0\ le_divide_eq[of e1 e "norm (y - x)"] by simp_all define z where "z = y - e1 *\<^sub>R (y - x)" { fix S assume "S \ I" then have "z \ rel_interior S" using rel_interior_closure_convex_shrink[of S x y e1] assms x y e1 z_def by auto } then have *: "z \ \{rel_interior S |S. S \ I}" by auto have "\z. z \ \{rel_interior S |S. S \ I} \ z \ y \ dist z y \ e" apply (rule_tac x="z" in exI) using \y \ x\ z_def * e1 e dist_norm[of z y] apply simp done } then have "y islimpt \{rel_interior S |S. S \ I}" unfolding islimpt_approachable_le by blast then have "y \ closure (\{rel_interior S |S. S \ I})" unfolding closure_def by auto } ultimately have "y \ closure (\{rel_interior S |S. S \ I})" by auto } then show ?thesis by auto qed lemma convex_closure_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\I) = \{closure S |S. S \ I}" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_inter_rel_interior_same_closure: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "closure (\{rel_interior S |S. S \ I}) = closure (\I)" proof - have "\{closure S |S. S \ I} \ closure (\{rel_interior S |S. S \ I})" using convex_closure_rel_interior_inter assms by auto moreover have "closure (\{rel_interior S |S. S \ I}) \ closure (\I)" using rel_interior_inter_aux closure_mono[of "\{rel_interior S |S. S \ I}" "\I"] by auto ultimately show ?thesis using closure_Int[of I] by auto qed lemma convex_rel_interior_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" shows "rel_interior (\I) \ \{rel_interior S |S. S \ I}" proof - have "convex (\I)" using assms convex_Inter by auto moreover have "convex (\{rel_interior S |S. S \ I})" apply (rule convex_Inter) using assms convex_rel_interior apply auto done ultimately have "rel_interior (\{rel_interior S |S. S \ I}) = rel_interior (\I)" using convex_inter_rel_interior_same_closure assms closure_eq_rel_interior_eq[of "\{rel_interior S |S. S \ I}" "\I"] by blast then show ?thesis using rel_interior_subset[of "\{rel_interior S |S. S \ I}"] by auto qed lemma convex_rel_interior_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set)" and "\{rel_interior S |S. S \ I} \ {}" and "finite I" shows "rel_interior (\I) = \{rel_interior S |S. S \ I}" proof - have "\I \ {}" using assms rel_interior_inter_aux[of I] by auto have "convex (\I)" using convex_Inter assms by auto show ?thesis proof (cases "I = {}") case True then show ?thesis using Inter_empty rel_interior_UNIV by auto next case False { fix z assume z: "z \ \{rel_interior S |S. S \ I}" { fix x assume x: "x \ \I" { fix S assume S: "S \ I" then have "z \ rel_interior S" "x \ S" using z x by auto then have "\m. m > 1 \ (\e. e > 1 \ e \ m \ (1 - e)*\<^sub>R x + e *\<^sub>R z \ S)" using convex_rel_interior_if[of S z] S assms hull_subset[of S] by auto } then obtain mS where mS: "\S\I. mS S > 1 \ (\e. e > 1 \ e \ mS S \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ S)" by metis define e where "e = Min (mS ` I)" then have "e \ mS ` I" using assms \I \ {}\ by simp then have "e > 1" using mS by auto moreover have "\S\I. e \ mS S" using e_def assms by auto ultimately have "\e > 1. (1 - e) *\<^sub>R x + e *\<^sub>R z \ \I" using mS by auto } then have "z \ rel_interior (\I)" using convex_rel_interior_iff[of "\I" z] \\I \ {}\ \convex (\I)\ by auto } then show ?thesis using convex_rel_interior_inter[of I] assms by auto qed qed lemma convex_closure_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" assumes "rel_interior S \ rel_interior T \ {}" shows "closure (S \ T) = closure S \ closure T" using convex_closure_inter[of "{S,T}"] assms by auto lemma convex_rel_interior_inter_two: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "rel_interior S \ rel_interior T \ {}" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" using convex_rel_interior_finite_inter[of "{S,T}"] assms by auto lemma convex_affine_closure_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "closure (S \ T) = closure S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_closure_inter_two[of S T] assms affine_imp_convex by auto qed lemma connected_component_1_gen: fixes S :: "'a :: euclidean_space set" assumes "DIM('a) = 1" shows "connected_component S a b \ closed_segment a b \ S" unfolding connected_component_def by (metis (no_types, lifting) assms subsetD subsetI convex_contains_segment convex_segment(1) ends_in_segment connected_convex_1_gen) lemma connected_component_1: fixes S :: "real set" shows "connected_component S a b \ closed_segment a b \ S" by (simp add: connected_component_1_gen) lemma convex_affine_rel_interior_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "rel_interior S \ T \ {}" shows "rel_interior (S \ T) = rel_interior S \ T" proof - have "affine hull T = T" using assms by auto then have "rel_interior T = T" using rel_interior_affine_hull[of T] by metis moreover have "closure T = T" using assms affine_closed[of T] by auto ultimately show ?thesis using convex_rel_interior_inter_two[of S T] assms affine_imp_convex by auto qed lemma convex_affine_rel_frontier_Int: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "affine T" and "interior S \ T \ {}" shows "rel_frontier(S \ T) = frontier S \ T" using assms apply (simp add: rel_frontier_def convex_affine_closure_Int frontier_def) by (metis Diff_Int_distrib2 Int_emptyI convex_affine_closure_Int convex_affine_rel_interior_Int empty_iff interior_rel_interior_gen) lemma rel_interior_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "interior S \ T \ {}" shows "rel_interior(S \ T) = interior S \ T" proof - obtain a where aS: "a \ interior S" and aT:"a \ T" using assms by force have "rel_interior S = interior S" by (metis (no_types) aS affine_hull_nonempty_interior equals0D rel_interior_interior) then show ?thesis by (metis (no_types) affine_imp_convex assms convex_rel_interior_inter_two hull_same rel_interior_affine_hull) qed lemma closure_convex_Int_affine: fixes S :: "'a::euclidean_space set" assumes "convex S" "affine T" "rel_interior S \ T \ {}" shows "closure(S \ T) = closure S \ T" proof have "closure (S \ T) \ closure T" by (simp add: closure_mono) also have "... \ T" by (simp add: affine_closed assms) finally show "closure(S \ T) \ closure S \ T" by (simp add: closure_mono) next obtain a where "a \ rel_interior S" "a \ T" using assms by auto then have ssT: "subspace ((\x. (-a)+x) ` T)" and "a \ S" using affine_diffs_subspace rel_interior_subset assms by blast+ show "closure S \ T \ closure (S \ T)" proof fix x assume "x \ closure S \ T" show "x \ closure (S \ T)" proof (cases "x = a") case True then show ?thesis using \a \ S\ \a \ T\ closure_subset by fastforce next case False then have "x \ closure(open_segment a x)" by auto then show ?thesis using \x \ closure S \ T\ assms convex_affine_closure_Int by blast qed qed qed lemma subset_rel_interior_convex: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" and "S \ closure T" and "\ S \ rel_frontier T" shows "rel_interior S \ rel_interior T" proof - have *: "S \ closure T = S" using assms by auto have "\ rel_interior S \ rel_frontier T" using closure_mono[of "rel_interior S" "rel_frontier T"] closed_rel_frontier[of T] closure_closed[of S] convex_closure_rel_interior[of S] closure_subset[of S] assms by auto then have "rel_interior S \ rel_interior (closure T) \ {}" using assms rel_frontier_def[of T] rel_interior_subset convex_rel_interior_closure[of T] by auto then have "rel_interior S \ rel_interior T = rel_interior (S \ closure T)" using assms convex_closure convex_rel_interior_inter_two[of S "closure T"] convex_rel_interior_closure[of T] by auto also have "\ = rel_interior S" using * by auto finally show ?thesis by auto qed lemma rel_interior_convex_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" shows "f ` (rel_interior S) = rel_interior (f ` S)" proof (cases "S = {}") case True then show ?thesis using assms rel_interior_empty rel_interior_eq_empty by auto next case False interpret linear f by fact have *: "f ` (rel_interior S) \ f ` S" unfolding image_mono using rel_interior_subset by auto have "f ` S \ f ` (closure S)" unfolding image_mono using closure_subset by auto also have "\ = f ` (closure (rel_interior S))" using convex_closure_rel_interior assms by auto also have "\ \ closure (f ` (rel_interior S))" using closure_linear_image_subset assms by auto finally have "closure (f ` S) = closure (f ` rel_interior S)" using closure_mono[of "f ` S" "closure (f ` rel_interior S)"] closure_closure closure_mono[of "f ` rel_interior S" "f ` S"] * by auto then have "rel_interior (f ` S) = rel_interior (f ` rel_interior S)" using assms convex_rel_interior linear_conv_bounded_linear[of f] convex_linear_image[of _ S] convex_linear_image[of _ "rel_interior S"] closure_eq_rel_interior_eq[of "f ` S" "f ` rel_interior S"] by auto then have "rel_interior (f ` S) \ f ` rel_interior S" using rel_interior_subset by auto moreover { fix z assume "z \ f ` rel_interior S" then obtain z1 where z1: "z1 \ rel_interior S" "f z1 = z" by auto { fix x assume "x \ f ` S" then obtain x1 where x1: "x1 \ S" "f x1 = x" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1 \ S" using convex_rel_interior_iff[of S z1] \convex S\ x1 z1 by auto moreover have "f ((1 - e) *\<^sub>R x1 + e *\<^sub>R z1) = (1 - e) *\<^sub>R x + e *\<^sub>R z" using x1 z1 by (simp add: linear_add linear_scale \linear f\) ultimately have "(1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using imageI[of "(1 - e) *\<^sub>R x1 + e *\<^sub>R z1" S f] by auto then have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f ` S" using e by auto } then have "z \ rel_interior (f ` S)" using convex_rel_interior_iff[of "f ` S" z] \convex S\ \linear f\ \S \ {}\ convex_linear_image[of f S] linear_conv_bounded_linear[of f] by auto } ultimately show ?thesis by auto qed lemma rel_interior_convex_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "f -` (rel_interior S) \ {}" shows "rel_interior (f -` S) = f -` (rel_interior S)" proof - interpret linear f by fact have "S \ {}" using assms rel_interior_empty by auto have nonemp: "f -` S \ {}" by (metis assms(3) rel_interior_subset subset_empty vimage_mono) then have "S \ (range f) \ {}" by auto have conv: "convex (f -` S)" using convex_linear_vimage assms by auto then have "convex (S \ range f)" by (simp add: assms(2) convex_Int convex_linear_image linear_axioms) { fix z assume "z \ f -` (rel_interior S)" then have z: "f z \ rel_interior S" by auto { fix x assume "x \ f -` S" then have "f x \ S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R f x + e *\<^sub>R f z \ S" using convex_rel_interior_iff[of S "f z"] z assms \S \ {}\ by auto moreover have "(1 - e) *\<^sub>R f x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R x + e *\<^sub>R z)" using \linear f\ by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R z \ f -` S" using e by auto } then have "z \ rel_interior (f -` S)" using convex_rel_interior_iff[of "f -` S" z] conv nonemp by auto } moreover { fix z assume z: "z \ rel_interior (f -` S)" { fix x assume "x \ S \ range f" then obtain y where y: "f y = x" "y \ f -` S" by auto then obtain e where e: "e > 1" "(1 - e) *\<^sub>R y + e *\<^sub>R z \ f -` S" using convex_rel_interior_iff[of "f -` S" z] z conv by auto moreover have "(1 - e) *\<^sub>R x + e *\<^sub>R f z = f ((1 - e) *\<^sub>R y + e *\<^sub>R z)" using \linear f\ y by (simp add: linear_iff) ultimately have "\e. e > 1 \ (1 - e) *\<^sub>R x + e *\<^sub>R f z \ S \ range f" using e by auto } then have "f z \ rel_interior (S \ range f)" using \convex (S \ (range f))\ \S \ range f \ {}\ convex_rel_interior_iff[of "S \ (range f)" "f z"] by auto moreover have "affine (range f)" by (simp add: linear_axioms linear_subspace_image subspace_imp_affine) ultimately have "f z \ rel_interior S" using convex_affine_rel_interior_Int[of S "range f"] assms by auto then have "z \ f -` (rel_interior S)" by auto } ultimately show ?thesis by auto qed lemma rel_interior_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S \ T) = rel_interior S \ rel_interior T" proof - { assume "S = {}" then have ?thesis by auto } moreover { assume "T = {}" then have ?thesis by auto } moreover { assume "S \ {}" "T \ {}" then have ri: "rel_interior S \ {}" "rel_interior T \ {}" using rel_interior_eq_empty assms by auto then have "fst -` rel_interior S \ {}" using fst_vimage_eq_Times[of "rel_interior S"] by auto then have "rel_interior ((fst :: 'n * 'm \ 'n) -` S) = fst -` rel_interior S" using fst_linear \convex S\ rel_interior_convex_linear_preimage[of fst S] by auto then have s: "rel_interior (S \ (UNIV :: 'm set)) = rel_interior S \ UNIV" by (simp add: fst_vimage_eq_Times) from ri have "snd -` rel_interior T \ {}" using snd_vimage_eq_Times[of "rel_interior T"] by auto then have "rel_interior ((snd :: 'n * 'm \ 'm) -` T) = snd -` rel_interior T" using snd_linear \convex T\ rel_interior_convex_linear_preimage[of snd T] by auto then have t: "rel_interior ((UNIV :: 'n set) \ T) = UNIV \ rel_interior T" by (simp add: snd_vimage_eq_Times) from s t have *: "rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T) = rel_interior S \ rel_interior T" by auto have "S \ T = S \ (UNIV :: 'm set) \ (UNIV :: 'n set) \ T" by auto then have "rel_interior (S \ T) = rel_interior ((S \ (UNIV :: 'm set)) \ ((UNIV :: 'n set) \ T))" by auto also have "\ = rel_interior (S \ (UNIV :: 'm set)) \ rel_interior ((UNIV :: 'n set) \ T)" apply (subst convex_rel_interior_inter_two[of "S \ (UNIV :: 'm set)" "(UNIV :: 'n set) \ T"]) using * ri assms convex_Times apply auto done finally have ?thesis using * by auto } ultimately show ?thesis by blast qed lemma rel_interior_scaleR: fixes S :: "'n::euclidean_space set" assumes "c \ 0" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" using rel_interior_injective_linear_image[of "((*\<^sub>R) c)" S] linear_conv_bounded_linear[of "(*\<^sub>R) c"] linear_scaleR injective_scaleR[of c] assms by auto lemma rel_interior_convex_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" shows "((*\<^sub>R) c) ` (rel_interior S) = rel_interior (((*\<^sub>R) c) ` S)" by (metis assms linear_scaleR rel_interior_convex_linear_image) lemma convex_rel_open_scaleR: fixes S :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" shows "convex (((*\<^sub>R) c) ` S) \ rel_open (((*\<^sub>R) c) ` S)" by (metis assms convex_scaling rel_interior_convex_scaleR rel_open_def) lemma convex_rel_open_finite_inter: assumes "\S\I. convex (S :: 'n::euclidean_space set) \ rel_open S" and "finite I" shows "convex (\I) \ rel_open (\I)" proof (cases "\{rel_interior S |S. S \ I} = {}") case True then have "\I = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def using rel_interior_empty by auto next case False then have "rel_open (\I)" using assms unfolding rel_open_def using convex_rel_interior_finite_inter[of I] by auto then show ?thesis using convex_Inter assms by auto qed lemma convex_rel_open_linear_image: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f ` S) \ rel_open (f ` S)" by (metis assms convex_linear_image rel_interior_convex_linear_image rel_open_def) lemma convex_rel_open_linear_preimage: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" and "convex S" and "rel_open S" shows "convex (f -` S) \ rel_open (f -` S)" proof (cases "f -` (rel_interior S) = {}") case True then have "f -` S = {}" using assms unfolding rel_open_def by auto then show ?thesis unfolding rel_open_def using rel_interior_empty by auto next case False then have "rel_open (f -` S)" using assms unfolding rel_open_def using rel_interior_convex_linear_preimage[of f S] by auto then show ?thesis using convex_linear_vimage assms by auto qed lemma rel_interior_projection: fixes S :: "('m::euclidean_space \ 'n::euclidean_space) set" and f :: "'m::euclidean_space \ 'n::euclidean_space set" assumes "convex S" and "f = (\y. {z. (y, z) \ S})" shows "(y, z) \ rel_interior S \ (y \ rel_interior {y. (f y \ {})} \ z \ rel_interior (f y))" proof - { fix y assume "y \ {y. f y \ {}}" then obtain z where "(y, z) \ S" using assms by auto then have "\x. x \ S \ y = fst x" apply (rule_tac x="(y, z)" in exI) apply auto done then obtain x where "x \ S" "y = fst x" by blast then have "y \ fst ` S" unfolding image_def by auto } then have "fst ` S = {y. f y \ {}}" unfolding fst_def using assms by auto then have h1: "fst ` rel_interior S = rel_interior {y. f y \ {}}" using rel_interior_convex_linear_image[of fst S] assms fst_linear by auto { fix y assume "y \ rel_interior {y. f y \ {}}" then have "y \ fst ` rel_interior S" using h1 by auto then have *: "rel_interior S \ fst -` {y} \ {}" by auto moreover have aff: "affine (fst -` {y})" unfolding affine_alt by (simp add: algebra_simps) ultimately have **: "rel_interior (S \ fst -` {y}) = rel_interior S \ fst -` {y}" using convex_affine_rel_interior_Int[of S "fst -` {y}"] assms by auto have conv: "convex (S \ fst -` {y})" using convex_Int assms aff affine_imp_convex by auto { fix x assume "x \ f y" then have "(y, x) \ S \ (fst -` {y})" using assms by auto moreover have "x = snd (y, x)" by auto ultimately have "x \ snd ` (S \ fst -` {y})" by blast } then have "snd ` (S \ fst -` {y}) = f y" using assms by auto then have ***: "rel_interior (f y) = snd ` rel_interior (S \ fst -` {y})" using rel_interior_convex_linear_image[of snd "S \ fst -` {y}"] snd_linear conv by auto { fix z assume "z \ rel_interior (f y)" then have "z \ snd ` rel_interior (S \ fst -` {y})" using *** by auto moreover have "{y} = fst ` rel_interior (S \ fst -` {y})" using * ** rel_interior_subset by auto ultimately have "(y, z) \ rel_interior (S \ fst -` {y})" by force then have "(y,z) \ rel_interior S" using ** by auto } moreover { fix z assume "(y, z) \ rel_interior S" then have "(y, z) \ rel_interior (S \ fst -` {y})" using ** by auto then have "z \ snd ` rel_interior (S \ fst -` {y})" by (metis Range_iff snd_eq_Range) then have "z \ rel_interior (f y)" using *** by auto } ultimately have "\z. (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto } then have h2: "\y z. y \ rel_interior {t. f t \ {}} \ (y, z) \ rel_interior S \ z \ rel_interior (f y)" by auto { fix y z assume asm: "(y, z) \ rel_interior S" then have "y \ fst ` rel_interior S" by (metis Domain_iff fst_eq_Domain) then have "y \ rel_interior {t. f t \ {}}" using h1 by auto then have "y \ rel_interior {t. f t \ {}}" and "(z \ rel_interior (f y))" using h2 asm by auto } then show ?thesis using h2 by blast qed lemma rel_frontier_Times: fixes S :: "'n::euclidean_space set" and T :: "'m::euclidean_space set" assumes "convex S" and "convex T" shows "rel_frontier S \ rel_frontier T \ rel_frontier (S \ T)" by (force simp: rel_frontier_def rel_interior_Times assms closure_Times) subsubsection\<^marker>\tag unimportant\ \Relative interior of convex cone\ lemma cone_rel_interior: fixes S :: "'m::euclidean_space set" assumes "cone S" shows "cone ({0} \ rel_interior S)" proof (cases "S = {}") case True then show ?thesis by (simp add: rel_interior_empty cone_0) next case False then have *: "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have *: "0 \ ({0} \ rel_interior S)" and "\c. c > 0 \ (*\<^sub>R) c ` ({0} \ rel_interior S) = ({0} \ rel_interior S)" by (auto simp add: rel_interior_scaleR) then show ?thesis using cone_iff[of "{0} \ rel_interior S"] by auto qed lemma rel_interior_convex_cone_aux: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) \ c > 0 \ x \ (((*\<^sub>R) c) ` (rel_interior S))" proof (cases "S = {}") case True then show ?thesis by (simp add: rel_interior_empty cone_hull_empty) next case False then obtain s where "s \ S" by auto have conv: "convex ({(1 :: real)} \ S)" using convex_Times[of "{(1 :: real)}" S] assms convex_singleton[of "1 :: real"] by auto define f where "f y = {z. (y, z) \ cone hull ({1 :: real} \ S)}" for y then have *: "(c, x) \ rel_interior (cone hull ({(1 :: real)} \ S)) = (c \ rel_interior {y. f y \ {}} \ x \ rel_interior (f c))" apply (subst rel_interior_projection[of "cone hull ({(1 :: real)} \ S)" f c x]) using convex_cone_hull[of "{(1 :: real)} \ S"] conv apply auto done { fix y :: real assume "y \ 0" then have "y *\<^sub>R (1,s) \ cone hull ({1 :: real} \ S)" using cone_hull_expl[of "{(1 :: real)} \ S"] \s \ S\ by auto then have "f y \ {}" using f_def by auto } then have "{y. f y \ {}} = {0..}" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have **: "rel_interior {y. f y \ {}} = {0<..}" using rel_interior_real_semiline by auto { fix c :: real assume "c > 0" then have "f c = ((*\<^sub>R) c ` S)" using f_def cone_hull_expl[of "{1 :: real} \ S"] by auto then have "rel_interior (f c) = (*\<^sub>R) c ` rel_interior S" using rel_interior_convex_scaleR[of S c] assms by auto } then show ?thesis using * ** by auto qed lemma rel_interior_convex_cone: fixes S :: "'m::euclidean_space set" assumes "convex S" shows "rel_interior (cone hull ({1 :: real} \ S)) = {(c, c *\<^sub>R x) | c x. c > 0 \ x \ rel_interior S}" (is "?lhs = ?rhs") proof - { fix z assume "z \ ?lhs" have *: "z = (fst z, snd z)" by auto then have "z \ ?rhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms \z \ ?lhs\ by fastforce } moreover { fix z assume "z \ ?rhs" then have "z \ ?lhs" using rel_interior_convex_cone_aux[of S "fst z" "snd z"] assms by auto } ultimately show ?thesis by blast qed lemma convex_hull_finite_union: assumes "finite I" assumes "\i\I. convex (S i) \ (S i) \ {}" shows "convex hull (\(S ` I)) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)}" (is "?lhs = ?rhs") proof - have "?lhs \ ?rhs" proof fix x assume "x \ ?rhs" then obtain c s where *: "sum (\i. c i *\<^sub>R s i) I = x" "sum c I = 1" "(\i\I. c i \ 0) \ (\i\I. s i \ S i)" by auto then have "\i\I. s i \ convex hull (\(S ` I))" using hull_subset[of "\(S ` I)" convex] by auto then show "x \ ?lhs" unfolding *(1)[symmetric] apply (subst convex_sum[of I "convex hull \(S ` I)" c s]) using * assms convex_convex_hull apply auto done qed { fix i assume "i \ I" with assms have "\p. p \ S i" by auto } then obtain p where p: "\i\I. p i \ S i" by metis { fix i assume "i \ I" { fix x assume "x \ S i" define c where "c j = (if j = i then 1::real else 0)" for j then have *: "sum c I = 1" using \finite I\ \i \ I\ sum.delta[of I i "\j::'a. 1::real"] by auto define s where "s j = (if j = i then x else p j)" for j then have "\j. c j *\<^sub>R s j = (if j = i then x else 0)" using c_def by (auto simp add: algebra_simps) then have "x = sum (\i. c i *\<^sub>R s i) I" using s_def c_def \finite I\ \i \ I\ sum.delta[of I i "\j::'a. x"] by auto then have "x \ ?rhs" apply auto apply (rule_tac x = c in exI) apply (rule_tac x = s in exI) using * c_def s_def p \x \ S i\ apply auto done } then have "?rhs \ S i" by auto } then have *: "?rhs \ \(S ` I)" by auto { fix u v :: real assume uv: "u \ 0 \ v \ 0 \ u + v = 1" fix x y assume xy: "x \ ?rhs \ y \ ?rhs" from xy obtain c s where xc: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. s i \ S i)" by auto from xy obtain d t where yc: "y = sum (\i. d i *\<^sub>R t i) I \ (\i\I. d i \ 0) \ sum d I = 1 \ (\i\I. t i \ S i)" by auto define e where "e i = u * c i + v * d i" for i have ge0: "\i\I. e i \ 0" using e_def xc yc uv by simp have "sum (\i. u * c i) I = u * sum c I" by (simp add: sum_distrib_left) moreover have "sum (\i. v * d i) I = v * sum d I" by (simp add: sum_distrib_left) ultimately have sum1: "sum e I = 1" using e_def xc yc uv by (simp add: sum.distrib) define q where "q i = (if e i = 0 then p i else (u * c i / e i) *\<^sub>R s i + (v * d i / e i) *\<^sub>R t i)" for i { fix i assume i: "i \ I" have "q i \ S i" proof (cases "e i = 0") case True then show ?thesis using i p q_def by auto next case False then show ?thesis using mem_convex_alt[of "S i" "s i" "t i" "u * (c i)" "v * (d i)"] mult_nonneg_nonneg[of u "c i"] mult_nonneg_nonneg[of v "d i"] assms q_def e_def i False xc yc uv by (auto simp del: mult_nonneg_nonneg) qed } then have qs: "\i\I. q i \ S i" by auto { fix i assume i: "i \ I" have "(u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" proof (cases "e i = 0") case True have ge: "u * (c i) \ 0 \ v * d i \ 0" using xc yc uv i by simp moreover from ge have "u * c i \ 0 \ v * d i \ 0" using True e_def i by simp ultimately have "u * c i = 0 \ v * d i = 0" by auto with True show ?thesis by auto next case False then have "(u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i) = q i" using q_def by auto then have "e i *\<^sub>R ((u * (c i)/(e i))*\<^sub>R (s i)+(v * (d i)/(e i))*\<^sub>R (t i)) = (e i) *\<^sub>R (q i)" by auto with False show ?thesis by (simp add: algebra_simps) qed } then have *: "\i\I. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i = e i *\<^sub>R q i" by auto have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (u * c i) *\<^sub>R s i + (v * d i) *\<^sub>R t i) I" using xc yc by (simp add: algebra_simps scaleR_right.sum sum.distrib) also have "\ = sum (\i. e i *\<^sub>R q i) I" using * by auto finally have "u *\<^sub>R x + v *\<^sub>R y = sum (\i. (e i) *\<^sub>R (q i)) I" by auto then have "u *\<^sub>R x + v *\<^sub>R y \ ?rhs" using ge0 sum1 qs by auto } then have "convex ?rhs" unfolding convex_def by auto then show ?thesis using \?lhs \ ?rhs\ * hull_minimal[of "\(S ` I)" ?rhs convex] by blast qed lemma convex_hull_union_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "S \ {}" and "convex T" and "T \ {}" shows "convex hull (S \ T) = {u *\<^sub>R s + v *\<^sub>R t | u v s t. u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T}" (is "?lhs = ?rhs") proof define I :: "nat set" where "I = {1, 2}" define s where "s i = (if i = (1::nat) then S else T)" for i have "\(s ` I) = S \ T" using s_def I_def by auto then have "convex hull (\(s ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(s ` I) = {\ i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)}" apply (subst convex_hull_finite_union[of I s]) using assms s_def I_def apply auto done moreover have "{\i\I. c i *\<^sub>R sa i | c sa. (\i\I. 0 \ c i) \ sum c I = 1 \ (\i\I. sa i \ s i)} \ ?rhs" using s_def I_def by auto ultimately show "?lhs \ ?rhs" by auto { fix x assume "x \ ?rhs" then obtain u v s t where *: "x = u *\<^sub>R s + v *\<^sub>R t \ u \ 0 \ v \ 0 \ u + v = 1 \ s \ S \ t \ T" by auto then have "x \ convex hull {s, t}" using convex_hull_2[of s t] by auto then have "x \ convex hull (S \ T)" using * hull_mono[of "{s, t}" "S \ T"] by auto } then show "?lhs \ ?rhs" by blast qed proposition ray_to_rel_frontier: fixes a :: "'a::real_inner" assumes "bounded S" and a: "a \ rel_interior S" and aff: "(a + l) \ affine hull S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ rel_frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ rel_interior S" proof - have aaff: "a \ affine hull S" by (meson a hull_subset rel_interior_subset rev_subsetD) let ?D = "{d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" obtain B where "B > 0" and B: "S \ ball a B" using bounded_subset_ballD [OF \bounded S\] by blast have "a + (B / norm l) *\<^sub>R l \ ball a B" by (simp add: dist_norm \l \ 0\) with B have "a + (B / norm l) *\<^sub>R l \ rel_interior S" using rel_interior_subset subsetCE by blast with \B > 0\ \l \ 0\ have nonMT: "?D \ {}" using divide_pos_pos zero_less_norm_iff by fastforce have bdd: "bdd_below ?D" by (metis (no_types, lifting) bdd_belowI le_less mem_Collect_eq) have relin_Ex: "\x. x \ rel_interior S \ \e>0. \x'\affine hull S. dist x' x < e \ x' \ rel_interior S" using openin_rel_interior [of S] by (simp add: openin_euclidean_subtopology_iff) define d where "d = Inf ?D" obtain \ where "0 < \" and \: "\\. \0 \ \; \ < \\ \ (a + \ *\<^sub>R l) \ rel_interior S" proof - obtain e where "e>0" and e: "\x'. x' \ affine hull S \ dist x' a < e \ x' \ rel_interior S" using relin_Ex a by blast show thesis proof (rule_tac \ = "e / norm l" in that) show "0 < e / norm l" by (simp add: \0 < e\ \l \ 0\) next show "a + \ *\<^sub>R l \ rel_interior S" if "0 \ \" "\ < e / norm l" for \ proof (rule e) show "a + \ *\<^sub>R l \ affine hull S" by (metis (no_types) add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show "dist (a + \ *\<^sub>R l) a < e" using that by (simp add: \l \ 0\ dist_norm pos_less_divide_eq) qed qed qed have inint: "\e. \0 \ e; e < d\ \ a + e *\<^sub>R l \ rel_interior S" unfolding d_def using cInf_lower [OF _ bdd] by (metis (no_types, lifting) a add.right_neutral le_less mem_Collect_eq not_less real_vector.scale_zero_left) have "\ \ d" unfolding d_def apply (rule cInf_greatest [OF nonMT]) using \ dual_order.strict_implies_order le_less_linear by blast with \0 < \\ have "0 < d" by simp have "a + d *\<^sub>R l \ rel_interior S" proof assume adl: "a + d *\<^sub>R l \ rel_interior S" obtain e where "e > 0" and e: "\x'. x' \ affine hull S \ dist x' (a + d *\<^sub>R l) < e \ x' \ rel_interior S" using relin_Ex adl by blast have "d + e / norm l \ Inf {d. 0 < d \ a + d *\<^sub>R l \ rel_interior S}" proof (rule cInf_greatest [OF nonMT], clarsimp) fix x::real assume "0 < x" and nonrel: "a + x *\<^sub>R l \ rel_interior S" show "d + e / norm l \ x" proof (cases "x < d") case True with inint nonrel \0 < x\ show ?thesis by auto next case False then have dle: "x < d + e / norm l \ dist (a + x *\<^sub>R l) (a + d *\<^sub>R l) < e" by (simp add: field_simps \l \ 0\) have ain: "a + x *\<^sub>R l \ affine hull S" by (metis add_diff_cancel_left' aff affine_affine_hull mem_affine_3_minus aaff) show ?thesis using e [OF ain] nonrel dle by force qed qed then show False using \0 < e\ \l \ 0\ by (simp add: d_def [symmetric] field_simps) qed moreover have "a + d *\<^sub>R l \ closure S" proof (clarsimp simp: closure_approachable) fix \::real assume "0 < \" have 1: "a + (d - min d (\ / 2 / norm l)) *\<^sub>R l \ S" apply (rule subsetD [OF rel_interior_subset inint]) using \l \ 0\ \0 < d\ \0 < \\ by auto have "norm l * min d (\ / (norm l * 2)) \ norm l * (\ / (norm l * 2))" by (metis min_def mult_left_mono norm_ge_zero order_refl) also have "... < \" using \l \ 0\ \0 < \\ by (simp add: field_simps) finally have 2: "norm l * min d (\ / (norm l * 2)) < \" . show "\y\S. dist y (a + d *\<^sub>R l) < \" apply (rule_tac x="a + (d - min d (\ / 2 / norm l)) *\<^sub>R l" in bexI) using 1 2 \0 < d\ \0 < \\ apply (auto simp: algebra_simps) done qed ultimately have infront: "a + d *\<^sub>R l \ rel_frontier S" by (simp add: rel_frontier_def) show ?thesis by (rule that [OF \0 < d\ infront inint]) qed corollary ray_to_frontier: fixes a :: "'a::euclidean_space" assumes "bounded S" and a: "a \ interior S" and "l \ 0" obtains d where "0 < d" "(a + d *\<^sub>R l) \ frontier S" "\e. \0 \ e; e < d\ \ (a + e *\<^sub>R l) \ interior S" proof - have "interior S = rel_interior S" using a rel_interior_nonempty_interior by auto then have "a \ rel_interior S" using a by simp then show ?thesis apply (rule ray_to_rel_frontier [OF \bounded S\ _ _ \l \ 0\]) using a affine_hull_nonempty_interior apply blast by (simp add: \interior S = rel_interior S\ frontier_def rel_frontier_def that) qed lemma segment_to_rel_frontier_aux: fixes x :: "'a::euclidean_space" assumes "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "x \ y" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof - have "x + (y - x) \ affine hull S" using hull_inc [OF y] by auto then obtain d where "0 < d" and df: "(x + d *\<^sub>R (y-x)) \ rel_frontier S" and di: "\e. \0 \ e; e < d\ \ (x + e *\<^sub>R (y-x)) \ rel_interior S" by (rule ray_to_rel_frontier [OF \bounded S\ x]) (use xy in auto) show ?thesis proof show "x + d *\<^sub>R (y - x) \ rel_frontier S" by (simp add: df) next have "open_segment x y \ rel_interior S" using rel_interior_closure_convex_segment [OF \convex S\ x] closure_subset y by blast moreover have "x + d *\<^sub>R (y - x) \ open_segment x y" if "d < 1" using xy apply (auto simp: in_segment) apply (rule_tac x="d" in exI) using \0 < d\ that apply (auto simp: algebra_simps) done ultimately have "1 \ d" using df rel_frontier_def by fastforce moreover have "x = (1 / d) *\<^sub>R x + ((d - 1) / d) *\<^sub>R x" by (metis \0 < d\ add.commute add_divide_distrib diff_add_cancel divide_self_if less_irrefl scaleR_add_left scaleR_one) ultimately show "y \ closed_segment x (x + d *\<^sub>R (y - x))" apply (auto simp: in_segment) apply (rule_tac x="1/d" in exI) apply (auto simp: algebra_simps) done next show "open_segment x (x + d *\<^sub>R (y - x)) \ rel_interior S" apply (rule rel_interior_closure_convex_segment [OF \convex S\ x]) using df rel_frontier_def by auto qed qed lemma segment_to_rel_frontier: fixes x :: "'a::euclidean_space" assumes S: "convex S" "bounded S" and x: "x \ rel_interior S" and y: "y \ S" and xy: "\(x = y \ S = {x})" obtains z where "z \ rel_frontier S" "y \ closed_segment x z" "open_segment x z \ rel_interior S" proof (cases "x=y") case True with xy have "S \ {x}" by blast with True show ?thesis by (metis Set.set_insert all_not_in_conv ends_in_segment(1) insert_iff segment_to_rel_frontier_aux[OF S x] that y) next case False then show ?thesis using segment_to_rel_frontier_aux [OF S x y] that by blast qed proposition rel_frontier_not_sing: fixes a :: "'a::euclidean_space" assumes "bounded S" shows "rel_frontier S \ {a}" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain z where "z \ S" by blast then show ?thesis proof (cases "S = {z}") case True then show ?thesis by simp next case False then obtain w where "w \ S" "w \ z" using \z \ S\ by blast show ?thesis proof assume "rel_frontier S = {a}" then consider "w \ rel_frontier S" | "z \ rel_frontier S" using \w \ z\ by auto then show False proof cases case 1 then have w: "w \ rel_interior S" using \w \ S\ closure_subset rel_frontier_def by fastforce have "w + (w - z) \ affine hull S" by (metis \w \ S\ \z \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(w + e *\<^sub>R (w - z)) \ rel_frontier S" using \w \ z\ \z \ S\ by (metis assms ray_to_rel_frontier right_minus_eq w) moreover obtain d where "0 < d" "(w + d *\<^sub>R (z - w)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ w, of "1 *\<^sub>R (z - w)"] \w \ z\ \z \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (z - w) = e *\<^sub>R (w - z)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) next case 2 then have z: "z \ rel_interior S" using \z \ S\ closure_subset rel_frontier_def by fastforce have "z + (z - w) \ affine hull S" by (metis \z \ S\ \w \ S\ affine_affine_hull hull_inc mem_affine_3_minus scaleR_one) then obtain e where "0 < e" "(z + e *\<^sub>R (z - w)) \ rel_frontier S" using \w \ z\ \w \ S\ by (metis assms ray_to_rel_frontier right_minus_eq z) moreover obtain d where "0 < d" "(z + d *\<^sub>R (w - z)) \ rel_frontier S" using ray_to_rel_frontier [OF \bounded S\ z, of "1 *\<^sub>R (w - z)"] \w \ z\ \w \ S\ by (metis add.commute add.right_neutral diff_add_cancel hull_inc scaleR_one) ultimately have "d *\<^sub>R (w - z) = e *\<^sub>R (z - w)" using \rel_frontier S = {a}\ by force moreover have "e \ -d " using \0 < e\ \0 < d\ by force ultimately show False by (metis (no_types, lifting) \w \ z\ eq_iff_diff_eq_0 minus_diff_eq real_vector.scale_cancel_right real_vector.scale_minus_right scaleR_left.minus) qed qed qed qed subsection\<^marker>\tag unimportant\ \Convexity on direct sums\ lemma closure_sum: fixes S T :: "'a::real_normed_vector set" shows "closure S + closure T \ closure (S + T)" unfolding set_plus_image closure_Times [symmetric] split_def by (intro closure_bounded_linear_image_subset bounded_linear_add bounded_linear_fst bounded_linear_snd) lemma rel_interior_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "convex T" shows "rel_interior (S + T) = rel_interior S + rel_interior T" proof - have "rel_interior S + rel_interior T = (\(x,y). x + y) ` (rel_interior S \ rel_interior T)" by (simp add: set_plus_image) also have "\ = (\(x,y). x + y) ` rel_interior (S \ T)" using rel_interior_Times assms by auto also have "\ = rel_interior (S + T)" using fst_snd_linear convex_Times assms rel_interior_convex_linear_image[of "(\(x,y). x + y)" "S \ T"] by (auto simp add: set_plus_image) finally show ?thesis .. qed lemma rel_interior_sum_gen: fixes S :: "'a \ 'n::euclidean_space set" assumes "\i\I. convex (S i)" shows "rel_interior (sum S I) = sum (\i. rel_interior (S i)) I" apply (subst sum_set_cond_linear[of convex]) using rel_interior_sum rel_interior_sing[of "0"] assms apply (auto simp add: convex_set_plus) done lemma convex_rel_open_direct_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S \ T) \ rel_open (S \ T)" by (metis assms convex_Times rel_interior_Times rel_open_def) lemma convex_rel_open_sum: fixes S T :: "'n::euclidean_space set" assumes "convex S" and "rel_open S" and "convex T" and "rel_open T" shows "convex (S + T) \ rel_open (S + T)" by (metis assms convex_set_plus rel_interior_sum rel_open_def) lemma convex_hull_finite_union_cones: assumes "finite I" and "I \ {}" assumes "\i\I. convex (S i) \ cone (S i) \ S i \ {}" shows "convex hull (\(S ` I)) = sum S I" (is "?lhs = ?rhs") proof - { fix x assume "x \ ?lhs" then obtain c xs where x: "x = sum (\i. c i *\<^sub>R xs i) I \ (\i\I. c i \ 0) \ sum c I = 1 \ (\i\I. xs i \ S i)" using convex_hull_finite_union[of I S] assms by auto define s where "s i = c i *\<^sub>R xs i" for i { fix i assume "i \ I" then have "s i \ S i" using s_def x assms mem_cone[of "S i" "xs i" "c i"] by auto } then have "\i\I. s i \ S i" by auto moreover have "x = sum s I" using x s_def by auto ultimately have "x \ ?rhs" using set_sum_alt[of I S] assms by auto } moreover { fix x assume "x \ ?rhs" then obtain s where x: "x = sum s I \ (\i\I. s i \ S i)" using set_sum_alt[of I S] assms by auto define xs where "xs i = of_nat(card I) *\<^sub>R s i" for i then have "x = sum (\i. ((1 :: real) / of_nat(card I)) *\<^sub>R xs i) I" using x assms by auto moreover have "\i\I. xs i \ S i" using x xs_def assms by (simp add: cone_def) moreover have "\i\I. (1 :: real) / of_nat (card I) \ 0" by auto moreover have "sum (\i. (1 :: real) / of_nat (card I)) I = 1" using assms by auto ultimately have "x \ ?lhs" apply (subst convex_hull_finite_union[of I S]) using assms apply blast using assms apply blast apply rule apply (rule_tac x = "(\i. (1 :: real) / of_nat (card I))" in exI) apply auto done } ultimately show ?thesis by auto qed lemma convex_hull_union_cones_two: fixes S T :: "'m::euclidean_space set" assumes "convex S" and "cone S" and "S \ {}" assumes "convex T" and "cone T" and "T \ {}" shows "convex hull (S \ T) = S + T" proof - define I :: "nat set" where "I = {1, 2}" define A where "A i = (if i = (1::nat) then S else T)" for i have "\(A ` I) = S \ T" using A_def I_def by auto then have "convex hull (\(A ` I)) = convex hull (S \ T)" by auto moreover have "convex hull \(A ` I) = sum A I" apply (subst convex_hull_finite_union_cones[of I A]) using assms A_def I_def apply auto done moreover have "sum A I = S + T" using A_def I_def unfolding set_plus_def apply auto unfolding set_plus_def apply auto done ultimately show ?thesis by auto qed lemma rel_interior_convex_hull_union: fixes S :: "'a \ 'n::euclidean_space set" assumes "finite I" and "\i\I. convex (S i) \ S i \ {}" shows "rel_interior (convex hull (\(S ` I))) = {sum (\i. c i *\<^sub>R s i) I | c s. (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior(S i))}" (is "?lhs = ?rhs") proof (cases "I = {}") case True then show ?thesis using convex_hull_empty rel_interior_empty by auto next case False define C0 where "C0 = convex hull (\(S ` I))" have "\i\I. C0 \ S i" unfolding C0_def using hull_subset[of "\(S ` I)"] by auto define K0 where "K0 = cone hull ({1 :: real} \ C0)" define K where "K i = cone hull ({1 :: real} \ S i)" for i have "\i\I. K i \ {}" unfolding K_def using assms by (simp add: cone_hull_empty_iff[symmetric]) { fix i assume "i \ I" then have "convex (K i)" unfolding K_def apply (subst convex_cone_hull) apply (subst convex_Times) using assms apply auto done } then have convK: "\i\I. convex (K i)" by auto { fix i assume "i \ I" then have "K0 \ K i" unfolding K0_def K_def apply (subst hull_mono) using \\i\I. C0 \ S i\ apply auto done } then have "K0 \ \(K ` I)" by auto moreover have "convex K0" unfolding K0_def apply (subst convex_cone_hull) apply (subst convex_Times) unfolding C0_def using convex_convex_hull apply auto done ultimately have geq: "K0 \ convex hull (\(K ` I))" using hull_minimal[of _ "K0" "convex"] by blast have "\i\I. K i \ {1 :: real} \ S i" using K_def by (simp add: hull_subset) then have "\(K ` I) \ {1 :: real} \ \(S ` I)" by auto then have "convex hull \(K ` I) \ convex hull ({1 :: real} \ \(S ` I))" by (simp add: hull_mono) then have "convex hull \(K ` I) \ {1 :: real} \ C0" unfolding C0_def using convex_hull_Times[of "{(1 :: real)}" "\(S ` I)"] convex_hull_singleton by auto moreover have "cone (convex hull (\(K ` I)))" apply (subst cone_convex_hull) using cone_Union[of "K ` I"] apply auto unfolding K_def using cone_cone_hull apply auto done ultimately have "convex hull (\(K ` I)) \ K0" unfolding K0_def using hull_minimal[of _ "convex hull (\(K ` I))" "cone"] by blast then have "K0 = convex hull (\(K ` I))" using geq by auto also have "\ = sum K I" apply (subst convex_hull_finite_union_cones[of I K]) using assms apply blast using False apply blast unfolding K_def apply rule apply (subst convex_cone_hull) apply (subst convex_Times) using assms cone_cone_hull \\i\I. K i \ {}\ K_def apply auto done finally have "K0 = sum K I" by auto then have *: "rel_interior K0 = sum (\i. (rel_interior (K i))) I" using rel_interior_sum_gen[of I K] convK by auto { fix x assume "x \ ?lhs" then have "(1::real, x) \ rel_interior K0" using K0_def C0_def rel_interior_convex_cone_aux[of C0 "1::real" x] convex_convex_hull by auto then obtain k where k: "(1::real, x) = sum k I \ (\i\I. k i \ rel_interior (K i))" using \finite I\ * set_sum_alt[of I "\i. rel_interior (K i)"] by auto { fix i assume "i \ I" then have "convex (S i) \ k i \ rel_interior (cone hull {1} \ S i)" using k K_def assms by auto then have "\ci si. k i = (ci, ci *\<^sub>R si) \ 0 < ci \ si \ rel_interior (S i)" using rel_interior_convex_cone[of "S i"] by auto } then obtain c s where cs: "\i\I. k i = (c i, c i *\<^sub>R s i) \ 0 < c i \ s i \ rel_interior (S i)" by metis then have "x = (\i\I. c i *\<^sub>R s i) \ sum c I = 1" using k by (simp add: sum_prod) then have "x \ ?rhs" using k cs by auto } moreover { fix x assume "x \ ?rhs" then obtain c s where cs: "x = sum (\i. c i *\<^sub>R s i) I \ (\i\I. c i > 0) \ sum c I = 1 \ (\i\I. s i \ rel_interior (S i))" by auto define k where "k i = (c i, c i *\<^sub>R s i)" for i { fix i assume "i \ I" then have "k i \ rel_interior (K i)" using k_def K_def assms cs rel_interior_convex_cone[of "S i"] by auto } then have "(1::real, x) \ rel_interior K0" using K0_def * set_sum_alt[of I "(\i. rel_interior (K i))"] assms k_def cs apply auto apply (rule_tac x = k in exI) apply (simp add: sum_prod) done then have "x \ ?lhs" using K0_def C0_def rel_interior_convex_cone_aux[of C0 1 x] by auto } ultimately show ?thesis by blast qed lemma convex_le_Inf_differential: fixes f :: "real \ real" assumes "convex_on I f" and "x \ interior I" and "y \ I" shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" (is "_ \ _ + Inf (?F x) * (y - x)") proof (cases rule: linorder_cases) assume "x < y" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = min (x + e / 2) ((x + y) / 2)" ultimately have "x < t" "t < y" "t \ ball x e" by (auto simp: dist_real_def field_simps split: split_min) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where "0 < e" "ball x e \ interior I" . moreover define K where "K = x - e / 2" with \0 < e\ have "K \ ball x e" "K < x" by (auto simp: dist_real_def) ultimately have "K \ I" "K < x" "x \ I" using interior_subset[of I] \x \ interior I\ by auto have "Inf (?F x) \ (f x - f y) / (x - y)" proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using \t \ I\ \x < t\ by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" using \convex_on I f\ \x \ I\ \y \ I\ \x < t\ \t < y\ by (rule convex_on_diff) next fix y assume "y \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \K \ I\ _ \K < x\ _]] show "(f K - f x) / (K - x) \ y" by auto qed then show ?thesis using \x < y\ by (simp add: field_simps) next assume "y < x" moreover have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . moreover define t where "t = x + e / 2" ultimately have "x < t" "t \ ball x e" by (auto simp: dist_real_def field_simps) with \x \ interior I\ e interior_subset[of I] have "t \ I" "x \ I" by auto have "(f x - f y) / (x - y) \ Inf (?F x)" proof (rule cInf_greatest) have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" using \y < x\ by (auto simp: field_simps) also fix z assume "z \ ?F x" with order_trans[OF convex_on_diff[OF \convex_on I f\ \y \ I\ _ \y < x\]] have "(f y - f x) / (y - x) \ z" by auto finally show "(f x - f y) / (x - y) \ z" . next have "open (interior I)" by auto from openE[OF this \x \ interior I\] obtain e where e: "0 < e" "ball x e \ interior I" . then have "x + e / 2 \ ball x e" by (auto simp: dist_real_def) with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto then show "?F x \ {}" by blast qed then show ?thesis using \y < x\ by (simp add: field_simps) qed simp subsection\<^marker>\tag unimportant\\Explicit formulas for interior and relative interior of convex hull\ lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma affine_independent_convex_affine_hull: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "t \ s" shows "convex hull t = affine hull t \ convex hull s" proof - have fin: "finite s" "finite t" using assms aff_independent_finite finite_subset by auto { fix u v x assume uv: "sum u t = 1" "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\v\t. u v *\<^sub>R v)" "x \ t" then have s: "s = (s - t) \ t" \ \split into separate cases\ using assms by auto have [simp]: "(\x\t. v x *\<^sub>R x) + (\x\s - t. v x *\<^sub>R x) = (\x\t. u x *\<^sub>R x)" "sum v t + sum v (s - t) = 1" using uv fin s by (auto simp: sum.union_disjoint [symmetric] Un_commute) have "(\x\s. if x \ t then v x - u x else v x) = 0" "(\x\s. (if x \ t then v x - u x else v x) *\<^sub>R x) = 0" using uv fin by (subst s, subst sum.union_disjoint, auto simp: algebra_simps sum_subtractf)+ } note [simp] = this have "convex hull t \ affine hull t" using convex_hull_subset_affine_hull by blast moreover have "convex hull t \ convex hull s" using assms hull_mono by blast moreover have "affine hull t \ convex hull s \ convex hull t" using assms apply (simp add: convex_hull_finite affine_hull_finite fin affine_dependent_explicit) apply (drule_tac x=s in spec) apply (auto simp: fin) apply (rule_tac x=u in exI) apply (rename_tac v) apply (drule_tac x="\x. if x \ t then v x - u x else v x" in spec) apply (force)+ done ultimately show ?thesis by blast qed lemma affine_independent_span_eq: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" "card s = Suc (DIM ('a))" shows "affine hull s = UNIV" proof (cases "s = {}") case True then show ?thesis using assms by simp next case False then obtain a t where t: "a \ t" "s = insert a t" by blast then have fin: "finite t" using assms by (metis finite_insert aff_independent_finite) show ?thesis using assms t fin apply (simp add: affine_dependent_iff_dependent affine_hull_insert_span_gen) apply (rule subset_antisym) apply force apply (rule Fun.vimage_subsetD) apply (metis add.commute diff_add_cancel surj_def) apply (rule card_ge_dim_independent) apply (auto simp: card_image inj_on_def dim_subset_UNIV) done qed lemma affine_independent_span_gt: fixes s :: "'a::euclidean_space set" assumes ind: "\ affine_dependent s" and dim: "DIM ('a) < card s" shows "affine hull s = UNIV" apply (rule affine_independent_span_eq [OF ind]) apply (rule antisym) using assms apply auto apply (metis add_2_eq_Suc' not_less_eq_eq affine_dependent_biggerset aff_independent_finite) done lemma empty_interior_affine_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(affine hull s) = {}" using assms apply (induct s rule: finite_induct) apply (simp_all add: affine_dependent_iff_dependent affine_hull_insert_span_gen interior_translation) apply (rule empty_interior_lowdim) by (auto simp: Suc_le_lessD card_image_le dual_order.trans intro!: dim_le_card'[THEN le_less_trans]) lemma empty_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" and dim: "card s \ DIM ('a)" shows "interior(convex hull s) = {}" by (metis Diff_empty Diff_eq_empty_iff convex_hull_subset_affine_hull interior_mono empty_interior_affine_hull [OF assms]) lemma explicit_subset_rel_interior_convex_hull: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma explicit_subset_rel_interior_convex_hull_minimal: fixes s :: "'a::euclidean_space set" shows "finite s \ {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y} \ rel_interior (convex hull s)" by (force simp add: rel_interior_convex_hull_union [where S="\x. {x}" and I=s, simplified]) lemma rel_interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_interior(convex hull s) = {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" (is "?lhs = ?rhs") proof show "?rhs \ ?lhs" by (simp add: aff_independent_finite explicit_subset_rel_interior_convex_hull_minimal assms) next show "?lhs \ ?rhs" proof (cases "\a. s = {a}") case True then show "?lhs \ ?rhs" by force next case False have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix a b and d::real assume ab: "a \ s" "b \ s" "a \ b" then have s: "s = (s - {a,b}) \ {a,b}" \ \split into separate cases\ by auto have "(\x\s. if x = a then - d else if x = b then d else 0) = 0" "(\x\s. (if x = a then - d else if x = b then d else 0) *\<^sub>R x) = d *\<^sub>R b - d *\<^sub>R a" using ab fs by (subst s, subst sum.union_disjoint, auto)+ } note [simp] = this { fix y assume y: "y \ convex hull s" "y \ ?rhs" { fix u T a assume ua: "\x\s. 0 \ u x" "sum u s = 1" "\ 0 < u a" "a \ s" and yT: "y = (\x\s. u x *\<^sub>R x)" "y \ T" "open T" and sb: "T \ affine hull s \ {w. \u. (\x\s. 0 \ u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = w}" have ua0: "u a = 0" using ua by auto obtain b where b: "b\s" "a \ b" using ua False by auto obtain e where e: "0 < e" "ball (\x\s. u x *\<^sub>R x) e \ T" using yT by (auto elim: openE) with b obtain d where d: "0 < d" "norm(d *\<^sub>R (a-b)) < e" by (auto intro: that [of "e / 2 / norm(a-b)"]) have "(\x\s. u x *\<^sub>R x) \ affine hull s" using yT y by (metis affine_hull_convex_hull hull_redundant_eq) then have "(\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b) \ affine hull s" using ua b by (auto simp: hull_inc intro: mem_affine_3_minus2) then have "y - d *\<^sub>R (a - b) \ T \ affine hull s" using d e yT by auto then obtain v where "\x\s. 0 \ v x" "sum v s = 1" "(\x\s. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x) - d *\<^sub>R (a - b)" using subsetD [OF sb] yT by auto then have False using assms apply (simp add: affine_dependent_explicit_finite fs) apply (drule_tac x="\x. (v x - u x) - (if x = a then -d else if x = b then d else 0)" in spec) using ua b d apply (auto simp: algebra_simps sum_subtractf sum.distrib) done } note * = this have "y \ rel_interior (convex hull s)" using y apply (simp add: mem_rel_interior affine_hull_convex_hull) apply (auto simp: convex_hull_finite [OF fs]) apply (drule_tac x=u in spec) apply (auto intro: *) done } with rel_interior_subset show "?lhs \ ?rhs" by blast qed qed lemma interior_convex_hull_explicit_minimal: fixes s :: "'a::euclidean_space set" shows "\ affine_dependent s ==> interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" apply (simp add: aff_independent_finite empty_interior_convex_hull, clarify) apply (rule trans [of _ "rel_interior(convex hull s)"]) apply (simp add: affine_independent_span_gt rel_interior_interior) by (simp add: rel_interior_convex_hull_explicit) lemma interior_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "interior(convex hull s) = (if card(s) \ DIM('a) then {} else {y. \u. (\x \ s. 0 < u x \ u x < 1) \ sum u s = 1 \ (\x\s. u x *\<^sub>R x) = y})" proof - { fix u :: "'a \ real" and a assume "card Basis < card s" and u: "\x. x\s \ 0 < u x" "sum u s = 1" and a: "a \ s" then have cs: "Suc 0 < card s" by (metis DIM_positive less_trans_Suc) obtain b where b: "b \ s" "a \ b" proof (cases "s \ {a}") case True then show thesis using cs subset_singletonD by fastforce next case False then show thesis by (blast intro: that) qed have "u a + u b \ sum u {a,b}" using a b by simp also have "... \ sum u s" apply (rule Groups_Big.sum_mono2) using a b u apply (auto simp: less_imp_le aff_independent_finite assms) done finally have "u a < 1" using \b \ s\ u by fastforce } note [simp] = this show ?thesis using assms apply (auto simp: interior_convex_hull_explicit_minimal) apply (rule_tac x=u in exI) apply (auto simp: not_le) done qed lemma interior_closed_segment_ge2: fixes a :: "'a::euclidean_space" assumes "2 \ DIM('a)" shows "interior(closed_segment a b) = {}" using assms unfolding segment_convex_hull proof - have "card {a, b} \ DIM('a)" using assms by (simp add: card_insert_if linear not_less_eq_eq numeral_2_eq_2) then show "interior (convex hull {a, b}) = {}" by (metis empty_interior_convex_hull finite.insertI finite.emptyI) qed lemma interior_open_segment: fixes a :: "'a::euclidean_space" shows "interior(open_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (simp add: not_le, intro conjI impI) assume "2 \ DIM('a)" then show "interior (open_segment a b) = {}" apply (simp add: segment_convex_hull open_segment_def) apply (metis Diff_subset interior_mono segment_convex_hull subset_empty interior_closed_segment_ge2) done next assume le2: "DIM('a) < 2" show "interior (open_segment a b) = open_segment a b" proof (cases "a = b") case True then show ?thesis by auto next case False with le2 have "affine hull (open_segment a b) = UNIV" apply simp apply (rule affine_independent_span_gt) apply (simp_all add: affine_dependent_def insert_Diff_if) done then show "interior (open_segment a b) = open_segment a b" using rel_interior_interior rel_interior_open_segment by blast qed qed lemma interior_closed_segment: fixes a :: "'a::euclidean_space" shows "interior(closed_segment a b) = (if 2 \ DIM('a) then {} else open_segment a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then have "closure (open_segment a b) = closed_segment a b" by simp then show ?thesis by (metis (no_types) convex_interior_closure convex_open_segment interior_open_segment) qed lemmas interior_segment = interior_closed_segment interior_open_segment lemma closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b = closed_segment c d \ {a,b} = {c,d}" proof assume abcd: "closed_segment a b = closed_segment c d" show "{a,b} = {c,d}" proof (cases "a=b \ c=d") case True with abcd show ?thesis by force next case False then have neq: "a \ b \ c \ d" by force have *: "closed_segment c d - {a, b} = rel_interior (closed_segment c d)" using neq abcd by (metis (no_types) open_segment_def rel_interior_closed_segment) have "b \ {c, d}" proof - have "insert b (closed_segment c d) = closed_segment c d" using abcd by blast then show ?thesis by (metis DiffD2 Diff_insert2 False * insertI1 insert_Diff_if open_segment_def rel_interior_closed_segment) qed moreover have "a \ {c, d}" by (metis Diff_iff False * abcd ends_in_segment(1) insertI1 open_segment_def rel_interior_closed_segment) ultimately show "{a, b} = {c, d}" using neq by fastforce qed next assume "{a,b} = {c,d}" then show "closed_segment a b = closed_segment c d" by (simp add: segment_convex_hull) qed lemma closed_open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "closed_segment a b \ open_segment c d" by (metis DiffE closed_segment_neq_empty closure_closed_segment closure_open_segment ends_in_segment(1) insertI1 open_segment_def) lemma open_closed_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b \ closed_segment c d" using closed_open_segment_eq by blast lemma open_segment_eq [simp]: fixes a :: "'a::euclidean_space" shows "open_segment a b = open_segment c d \ a = b \ c = d \ {a,b} = {c,d}" (is "?lhs = ?rhs") proof assume abcd: ?lhs show ?rhs proof (cases "a=b \ c=d") case True with abcd show ?thesis using finite_open_segment by fastforce next case False then have a2: "a \ b \ c \ d" by force with abcd show ?rhs unfolding open_segment_def by (metis (no_types) abcd closed_segment_eq closure_open_segment) qed next assume ?rhs then show ?lhs by (metis Diff_cancel convex_hull_singleton insert_absorb2 open_segment_def segment_convex_hull) qed subsection\<^marker>\tag unimportant\\Similar results for closure and (relative or absolute) frontier\ lemma closure_convex_hull [simp]: fixes s :: "'a::euclidean_space set" shows "compact s ==> closure(convex hull s) = convex hull s" by (simp add: compact_imp_closed compact_convex_hull) lemma rel_frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (\x \ s. u x = 0) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis apply (simp add: rel_frontier_def finite_imp_compact rel_interior_convex_hull_explicit assms fs) apply (auto simp: convex_hull_finite fs) apply (drule_tac x=u in spec) apply (rule_tac x=u in exI) apply force apply (rename_tac v) apply (rule notE [OF assms]) apply (simp add: affine_dependent_explicit) apply (rule_tac x=s in exI) apply (auto simp: fs) apply (rule_tac x = "\x. u x - v x" in exI) apply (force simp: sum_subtractf scaleR_diff_left) done qed lemma frontier_convex_hull_explicit: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = {y. \u. (\x \ s. 0 \ u x) \ (DIM ('a) < card s \ (\x \ s. u x = 0)) \ sum u s = 1 \ sum (\x. u x *\<^sub>R x) s = y}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) show ?thesis proof (cases "DIM ('a) < card s") case True with assms fs show ?thesis by (simp add: rel_frontier_def frontier_def rel_frontier_convex_hull_explicit [symmetric] interior_convex_hull_explicit_minimal rel_interior_convex_hull_explicit) next case False then have "card s \ DIM ('a)" by linarith then show ?thesis using assms fs apply (simp add: frontier_def interior_convex_hull_explicit finite_imp_compact) apply (simp add: convex_hull_finite) done qed qed lemma rel_frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "rel_frontier(convex hull s) = \{convex hull (s - {x}) |x. x \ s}" proof - have fs: "finite s" using assms by (simp add: aff_independent_finite) { fix u a have "\x\s. 0 \ u x \ a \ s \ u a = 0 \ sum u s = 1 \ \x v. x \ s \ (\x\s - {x}. 0 \ v x) \ sum v (s - {x}) = 1 \ (\x\s - {x}. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" apply (rule_tac x=a in exI) apply (rule_tac x=u in exI) apply (simp add: Groups_Big.sum_diff1 fs) done } moreover { fix a u have "a \ s \ \x\s - {a}. 0 \ u x \ sum u (s - {a}) = 1 \ \v. (\x\s. 0 \ v x) \ (\x\s. v x = 0) \ sum v s = 1 \ (\x\s. v x *\<^sub>R x) = (\x\s - {a}. u x *\<^sub>R x)" apply (rule_tac x="\x. if x = a then 0 else u x" in exI) apply (auto simp: sum.If_cases Diff_eq if_smult fs) done } ultimately show ?thesis using assms apply (simp add: rel_frontier_convex_hull_explicit) apply (simp add: convex_hull_finite fs Union_SetCompr_eq, auto) done qed lemma frontier_convex_hull_eq_rel_frontier: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else rel_frontier(convex hull s))" using assms unfolding rel_frontier_def frontier_def by (simp add: affine_independent_span_gt rel_interior_interior finite_imp_compact empty_interior_convex_hull aff_independent_finite) lemma frontier_convex_hull_cases: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent s" shows "frontier(convex hull s) = (if card s \ DIM ('a) then convex hull s else \{convex hull (s - {x}) |x. x \ s})" by (simp add: assms frontier_convex_hull_eq_rel_frontier rel_frontier_convex_hull_cases) lemma in_frontier_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ frontier(convex hull s)" proof (cases "affine_dependent s") case True with assms show ?thesis apply (auto simp: affine_dependent_def frontier_def finite_imp_compact hull_inc) by (metis card.insert_remove convex_hull_subset_affine_hull empty_interior_affine_hull finite_Diff hull_redundant insert_Diff insert_Diff_single insert_not_empty interior_mono not_less_eq_eq subset_empty) next case False { assume "card s = Suc (card Basis)" then have cs: "Suc 0 < card s" - by (simp add: DIM_positive) + by (simp) with subset_singletonD have "\y \ s. y \ x" by (cases "s \ {x}") fastforce+ } note [dest!] = this show ?thesis using assms unfolding frontier_convex_hull_cases [OF False] Union_SetCompr_eq by (auto simp: le_Suc_eq hull_inc) qed lemma not_in_interior_convex_hull: fixes s :: "'a::euclidean_space set" assumes "finite s" "card s \ Suc (DIM ('a))" "x \ s" shows "x \ interior(convex hull s)" using in_frontier_convex_hull [OF assms] by (metis Diff_iff frontier_def) lemma interior_convex_hull_eq_empty: fixes s :: "'a::euclidean_space set" assumes "card s = Suc (DIM ('a))" shows "interior(convex hull s) = {} \ affine_dependent s" proof - { fix a b assume ab: "a \ interior (convex hull s)" "b \ s" "b \ affine hull (s - {b})" then have "interior(affine hull s) = {}" using assms by (metis DIM_positive One_nat_def Suc_mono card.remove card_infinite empty_interior_affine_hull eq_iff hull_redundant insert_Diff not_less zero_le_one) then have False using ab by (metis convex_hull_subset_affine_hull equals0D interior_mono subset_eq) } then show ?thesis using assms apply auto apply (metis UNIV_I affine_hull_convex_hull affine_hull_empty affine_independent_span_eq convex_convex_hull empty_iff rel_interior_interior rel_interior_same_affine_hull) apply (auto simp: affine_dependent_def) done qed subsection \Coplanarity, and collinearity in terms of affine hull\ definition\<^marker>\tag important\ coplanar where "coplanar s \ \u v w. s \ affine hull {u,v,w}" lemma collinear_affine_hull: "collinear s \ (\u v. s \ affine hull {u,v})" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain x where x: "x \ s" by auto { fix u assume *: "\x y. \x\s; y\s\ \ \c. x - y = c *\<^sub>R u" have "\u v. s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" apply (rule_tac x=x in exI) apply (rule_tac x="x+u" in exI, clarify) apply (erule exE [OF * [OF x]]) apply (rename_tac c) apply (rule_tac x="1+c" in exI) apply (rule_tac x="-c" in exI) apply (simp add: algebra_simps) done } moreover { fix u v x y assume *: "s \ {a *\<^sub>R u + b *\<^sub>R v |a b. a + b = 1}" have "x\s \ y\s \ \c. x - y = c *\<^sub>R (v-u)" apply (drule subsetD [OF *])+ apply simp apply clarify apply (rename_tac r1 r2) apply (rule_tac x="r1-r2" in exI) apply (simp add: algebra_simps) apply (metis scaleR_left.add) done } ultimately show ?thesis unfolding collinear_def affine_hull_2 by blast qed lemma collinear_closed_segment [simp]: "collinear (closed_segment a b)" by (metis affine_hull_convex_hull collinear_affine_hull hull_subset segment_convex_hull) lemma collinear_open_segment [simp]: "collinear (open_segment a b)" unfolding open_segment_def by (metis convex_hull_subset_affine_hull segment_convex_hull dual_order.trans convex_hull_subset_affine_hull Diff_subset collinear_affine_hull) lemma collinear_between_cases: fixes c :: "'a::euclidean_space" shows "collinear {a,b,c} \ between (b,c) a \ between (c,a) b \ between (a,b) c" (is "?lhs = ?rhs") proof assume ?lhs then obtain u v where uv: "\x. x \ {a, b, c} \ \c. x = u + c *\<^sub>R v" by (auto simp: collinear_alt) show ?rhs using uv [of a] uv [of b] uv [of c] by (auto simp: between_1) next assume ?rhs then show ?lhs unfolding between_mem_convex_hull by (metis (no_types, hide_lams) collinear_closed_segment collinear_subset hull_redundant hull_subset insert_commute segment_convex_hull) qed lemma subset_continuous_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes "continuous_on (closed_segment a b) f" shows "closed_segment (f a) (f b) \ image f (closed_segment a b)" by (metis connected_segment convex_contains_segment ends_in_segment imageI is_interval_connected_1 is_interval_convex connected_continuous_image [OF assms]) lemma continuous_injective_image_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (closed_segment a b) = closed_segment (f a) (f b)" proof show "closed_segment (f a) (f b) \ f ` closed_segment a b" by (metis subset_continuous_image_segment_1 contf) show "f ` closed_segment a b \ closed_segment (f a) (f b)" proof (cases "a = b") case True then show ?thesis by auto next case False then have fnot: "f a \ f b" using inj_onD injf by fastforce moreover have "f a \ open_segment (f c) (f b)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def) assume fa: "f a \ closed_segment (f c) (f b)" moreover have "closed_segment (f c) (f b) \ f ` closed_segment c b" by (meson closed_segment_subset contf continuous_on_subset convex_closed_segment ends_in_segment(2) subset_continuous_image_segment_1 that) ultimately have "f a \ f ` closed_segment c b" by blast then have a: "a \ closed_segment c b" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have cb: "closed_segment c b \ closed_segment a b" by (simp add: closed_segment_subset that) show "f a = f c" proof (rule between_antisym) show "between (f c, f b) (f a)" by (simp add: between_mem_segment fa) show "between (f a, f b) (f c)" by (metis a cb between_antisym between_mem_segment between_triv1 subset_iff) qed qed moreover have "f b \ open_segment (f a) (f c)" if c: "c \ closed_segment a b" for c proof (clarsimp simp add: open_segment_def fnot eq_commute) assume fb: "f b \ closed_segment (f a) (f c)" moreover have "closed_segment (f a) (f c) \ f ` closed_segment a c" by (meson contf continuous_on_subset ends_in_segment(1) subset_closed_segment subset_continuous_image_segment_1 that) ultimately have "f b \ f ` closed_segment a c" by blast then have b: "b \ closed_segment a c" by (meson ends_in_segment inj_on_image_mem_iff_alt injf subset_closed_segment that) have ca: "closed_segment a c \ closed_segment a b" by (simp add: closed_segment_subset that) show "f b = f c" proof (rule between_antisym) show "between (f c, f a) (f b)" by (simp add: between_commute between_mem_segment fb) show "between (f b, f a) (f c)" by (metis b between_antisym between_commute between_mem_segment between_triv2 that) qed qed ultimately show ?thesis by (force simp: closed_segment_eq_real_ivl open_segment_eq_real_ivl split: if_split_asm) qed qed lemma continuous_injective_image_open_segment_1: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and injf: "inj_on f (closed_segment a b)" shows "f ` (open_segment a b) = open_segment (f a) (f b)" proof - have "f ` (open_segment a b) = f ` (closed_segment a b) - {f a, f b}" by (metis (no_types, hide_lams) empty_subsetI ends_in_segment image_insert image_is_empty inj_on_image_set_diff injf insert_subset open_segment_def segment_open_subset_closed) also have "... = open_segment (f a) (f b)" using continuous_injective_image_segment_1 [OF assms] by (simp add: open_segment_def inj_on_image_set_diff [OF injf]) finally show ?thesis . qed lemma collinear_imp_coplanar: "collinear s ==> coplanar s" by (metis collinear_affine_hull coplanar_def insert_absorb2) lemma collinear_small: assumes "finite s" "card s \ 2" shows "collinear s" proof - have "card s = 0 \ card s = 1 \ card s = 2" using assms by linarith then show ?thesis using assms using card_eq_SucD by auto (metis collinear_2 numeral_2_eq_2) qed lemma coplanar_small: assumes "finite s" "card s \ 3" shows "coplanar s" proof - have "card s \ 2 \ card s = Suc (Suc (Suc 0))" using assms by linarith then show ?thesis using assms apply safe apply (simp add: collinear_small collinear_imp_coplanar) apply (safe dest!: card_eq_SucD) apply (auto simp: coplanar_def) apply (metis hull_subset insert_subset) done qed lemma coplanar_empty: "coplanar {}" by (simp add: coplanar_small) lemma coplanar_sing: "coplanar {a}" by (simp add: coplanar_small) lemma coplanar_2: "coplanar {a,b}" by (auto simp: card_insert_if coplanar_small) lemma coplanar_3: "coplanar {a,b,c}" by (auto simp: card_insert_if coplanar_small) lemma collinear_affine_hull_collinear: "collinear(affine hull s) \ collinear s" unfolding collinear_affine_hull by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_affine_hull_coplanar: "coplanar(affine hull s) \ coplanar s" unfolding coplanar_def by (metis affine_affine_hull subset_hull hull_hull hull_mono) lemma coplanar_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "coplanar s" "linear f" shows "coplanar(f ` s)" proof - { fix u v w assume "s \ affine hull {u, v, w}" then have "f ` s \ f ` (affine hull {u, v, w})" by (simp add: image_mono) then have "f ` s \ affine hull (f ` {u, v, w})" by (metis assms(2) linear_conv_bounded_linear affine_hull_linear_image) } then show ?thesis by auto (meson assms(1) coplanar_def) qed lemma coplanar_translation_imp: "coplanar s \ coplanar ((\x. a + x) ` s)" unfolding coplanar_def apply clarify apply (rule_tac x="u+a" in exI) apply (rule_tac x="v+a" in exI) apply (rule_tac x="w+a" in exI) using affine_hull_translation [of a "{u,v,w}" for u v w] apply (force simp: add.commute) done lemma coplanar_translation_eq: "coplanar((\x. a + x) ` s) \ coplanar s" by (metis (no_types) coplanar_translation_imp translation_galois) lemma coplanar_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "coplanar(f ` s) = coplanar s" proof assume "coplanar s" then show "coplanar (f ` s)" unfolding coplanar_def using affine_hull_linear_image [of f "{u,v,w}" for u v w] assms by (meson coplanar_def coplanar_linear_image) next obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF assms] by blast assume "coplanar (f ` s)" then obtain u v w where "f ` s \ affine hull {u, v, w}" by (auto simp: coplanar_def) then have "g ` f ` s \ g ` (affine hull {u, v, w})" by blast then have "s \ g ` (affine hull {u, v, w})" using g by (simp add: Fun.image_comp) then show "coplanar s" unfolding coplanar_def using affine_hull_linear_image [of g "{u,v,w}" for u v w] \linear g\ linear_conv_bounded_linear by fastforce qed (*The HOL Light proof is simply MATCH_ACCEPT_TAC(LINEAR_INVARIANT_RULE COPLANAR_LINEAR_IMAGE));; *) lemma coplanar_subset: "\coplanar t; s \ t\ \ coplanar s" by (meson coplanar_def order_trans) lemma affine_hull_3_imp_collinear: "c \ affine hull {a,b} \ collinear {a,b,c}" by (metis collinear_2 collinear_affine_hull_collinear hull_redundant insert_commute) lemma collinear_3_imp_in_affine_hull: "\collinear {a,b,c}; a \ b\ \ c \ affine hull {a,b}" unfolding collinear_def apply clarify apply (frule_tac x=b in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (drule_tac x=c in bspec, blast, drule_tac x=a in bspec, blast, erule exE) apply (rename_tac y x) apply (simp add: affine_hull_2) apply (rule_tac x="1 - x/y" in exI) apply (simp add: algebra_simps) done lemma collinear_3_affine_hull: assumes "a \ b" shows "collinear {a,b,c} \ c \ affine hull {a,b}" using affine_hull_3_imp_collinear assms collinear_3_imp_in_affine_hull by blast lemma collinear_3_eq_affine_dependent: "collinear{a,b,c} \ a = b \ a = c \ b = c \ affine_dependent {a,b,c}" apply (case_tac "a=b", simp) apply (case_tac "a=c") apply (simp add: insert_commute) apply (case_tac "b=c") apply (simp add: insert_commute) apply (auto simp: affine_dependent_def collinear_3_affine_hull insert_Diff_if) apply (metis collinear_3_affine_hull insert_commute)+ done lemma affine_dependent_imp_collinear_3: "affine_dependent {a,b,c} \ collinear{a,b,c}" by (simp add: collinear_3_eq_affine_dependent) lemma collinear_3: "NO_MATCH 0 x \ collinear {x,y,z} \ collinear {0, x-y, z-y}" by (auto simp add: collinear_def) lemma collinear_3_expand: "collinear{a,b,c} \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" proof - have "collinear{a,b,c} = collinear{a,c,b}" by (simp add: insert_commute) also have "... = collinear {0, a - c, b - c}" by (simp add: collinear_3) also have "... \ (a = c \ b = c \ (\ca. b - c = ca *\<^sub>R (a - c)))" by (simp add: collinear_lemma) also have "... \ a = c \ (\u. b = u *\<^sub>R a + (1 - u) *\<^sub>R c)" by (cases "a = c \ b = c") (auto simp: algebra_simps) finally show ?thesis . qed lemma collinear_aff_dim: "collinear S \ aff_dim S \ 1" proof assume "collinear S" then obtain u and v :: "'a" where "aff_dim S \ aff_dim {u,v}" by (metis \collinear S\ aff_dim_affine_hull aff_dim_subset collinear_affine_hull) then show "aff_dim S \ 1" using order_trans by fastforce next assume "aff_dim S \ 1" then have le1: "aff_dim (affine hull S) \ 1" by simp obtain B where "B \ S" and B: "\ affine_dependent B" "affine hull S = affine hull B" using affine_basis_exists [of S] by auto then have "finite B" "card B \ 2" using B le1 by (auto simp: affine_independent_iff_card) then have "collinear B" by (rule collinear_small) then show "collinear S" by (metis \affine hull S = affine hull B\ collinear_affine_hull_collinear) qed lemma collinear_midpoint: "collinear{a,midpoint a b,b}" apply (auto simp: collinear_3 collinear_lemma) apply (drule_tac x="-1" in spec) apply (simp add: algebra_simps) done lemma midpoint_collinear: fixes a b c :: "'a::real_normed_vector" assumes "a \ c" shows "b = midpoint a c \ collinear{a,b,c} \ dist a b = dist b c" proof - have *: "a - (u *\<^sub>R a + (1 - u) *\<^sub>R c) = (1 - u) *\<^sub>R (a - c)" "u *\<^sub>R a + (1 - u) *\<^sub>R c - c = u *\<^sub>R (a - c)" "\1 - u\ = \u\ \ u = 1/2" for u::real by (auto simp: algebra_simps) have "b = midpoint a c \ collinear{a,b,c} " using collinear_midpoint by blast moreover have "collinear{a,b,c} \ b = midpoint a c \ dist a b = dist b c" apply (auto simp: collinear_3_expand assms dist_midpoint) apply (simp add: dist_norm * assms midpoint_def del: divide_const_simps) apply (simp add: algebra_simps) done ultimately show ?thesis by blast qed lemma between_imp_collinear: fixes x :: "'a :: euclidean_space" assumes "between (a,b) x" shows "collinear {a,x,b}" proof (cases "x = a \ x = b \ a = b") case True with assms show ?thesis by (auto simp: dist_commute) next case False with assms show ?thesis apply (auto simp: collinear_3 collinear_lemma between_norm) apply (drule_tac x="-(norm(b - x) / norm(x - a))" in spec) apply (simp add: vector_add_divide_simps eq_vector_fraction_iff real_vector.scale_minus_right [symmetric]) done qed lemma midpoint_between: fixes a b :: "'a::euclidean_space" shows "b = midpoint a c \ between (a,c) b \ dist a b = dist b c" proof (cases "a = c") case True then show ?thesis by (auto simp: dist_commute) next case False show ?thesis apply (rule iffI) apply (simp add: between_midpoint(1) dist_midpoint) using False between_imp_collinear midpoint_collinear by blast qed lemma collinear_triples: assumes "a \ b" shows "collinear(insert a (insert b S)) \ (\x \ S. collinear{a,b,x})" (is "?lhs = ?rhs") proof safe fix x assume ?lhs and "x \ S" then show "collinear {a, b, x}" using collinear_subset by force next assume ?rhs then have "\x \ S. collinear{a,x,b}" by (simp add: insert_commute) then have *: "\u. x = u *\<^sub>R a + (1 - u) *\<^sub>R b" if "x \ (insert a (insert b S))" for x using that assms collinear_3_expand by fastforce+ show ?lhs unfolding collinear_def apply (rule_tac x="b-a" in exI) apply (clarify dest!: *) by (metis (no_types, hide_lams) add.commute diff_add_cancel diff_diff_eq2 real_vector.scale_right_diff_distrib scaleR_left.diff) qed lemma collinear_4_3: assumes "a \ b" shows "collinear {a,b,c,d} \ collinear{a,b,c} \ collinear{a,b,d}" using collinear_triples [OF assms, of "{c,d}"] by (force simp:) lemma collinear_3_trans: assumes "collinear{a,b,c}" "collinear{b,c,d}" "b \ c" shows "collinear{a,b,d}" proof - have "collinear{b,c,a,d}" by (metis (full_types) assms collinear_4_3 insert_commute) then show ?thesis by (simp add: collinear_subset) qed lemma affine_hull_eq_empty [simp]: "affine hull S = {} \ S = {}" using affine_hull_nonempty by blast lemma affine_hull_2_alt: fixes a b :: "'a::real_vector" shows "affine hull {a,b} = range (\u. a + u *\<^sub>R (b - a))" apply (simp add: affine_hull_2, safe) apply (rule_tac x=v in image_eqI) apply (simp add: algebra_simps) apply (metis scaleR_add_left scaleR_one, simp) apply (rule_tac x="1-u" in exI) apply (simp add: algebra_simps) done lemma interior_convex_hull_3_minimal: fixes a :: "'a::euclidean_space" shows "\\ collinear{a,b,c}; DIM('a) = 2\ \ interior(convex hull {a,b,c}) = {v. \x y z. 0 < x \ 0 < y \ 0 < z \ x + y + z = 1 \ x *\<^sub>R a + y *\<^sub>R b + z *\<^sub>R c = v}" apply (simp add: collinear_3_eq_affine_dependent interior_convex_hull_explicit_minimal, safe) apply (rule_tac x="u a" in exI, simp) apply (rule_tac x="u b" in exI, simp) apply (rule_tac x="u c" in exI, simp) apply (rename_tac uu x y z) apply (rule_tac x="\r. (if r=a then x else if r=b then y else if r=c then z else 0)" in exI) apply simp done subsection\<^marker>\tag unimportant\\Basic lemmas about hyperplanes and halfspaces\ lemma halfspace_Int_eq: "{x. a \ x \ b} \ {x. b \ a \ x} = {x. a \ x = b}" "{x. b \ a \ x} \ {x. a \ x \ b} = {x. a \ x = b}" by auto lemma hyperplane_eq_Ex: assumes "a \ 0" obtains x where "a \ x = b" by (rule_tac x = "(b / (a \ a)) *\<^sub>R a" in that) (simp add: assms) lemma hyperplane_eq_empty: "{x. a \ x = b} = {} \ a = 0 \ b \ 0" using hyperplane_eq_Ex apply auto[1] using inner_zero_right by blast lemma hyperplane_eq_UNIV: "{x. a \ x = b} = UNIV \ a = 0 \ b = 0" proof - have "UNIV \ {x. a \ x = b} \ a = 0 \ b = 0" apply (drule_tac c = "((b+1) / (a \ a)) *\<^sub>R a" in subsetD) apply simp_all by (metis add_cancel_right_right zero_neq_one) then show ?thesis by force qed lemma halfspace_eq_empty_lt: "{x. a \ x < b} = {} \ a = 0 \ b \ 0" proof - have "{x. a \ x < b} \ {} \ a = 0 \ b \ 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_gt: "{x. a \ x > b} = {} \ a = 0 \ b \ 0" using halfspace_eq_empty_lt [of "-a" "-b"] by simp lemma halfspace_eq_empty_le: "{x. a \ x \ b} = {} \ a = 0 \ b < 0" proof - have "{x. a \ x \ b} \ {} \ a = 0 \ b < 0" apply (rule ccontr) apply (drule_tac c = "((b-1) / (a \ a)) *\<^sub>R a" in subsetD) apply force+ done then show ?thesis by force qed lemma halfspace_eq_empty_ge: "{x. a \ x \ b} = {} \ a = 0 \ b > 0" using halfspace_eq_empty_le [of "-a" "-b"] by simp subsection\<^marker>\tag unimportant\\Use set distance for an easy proof of separation properties\ proposition\<^marker>\tag unimportant\ separation_closures: fixes S :: "'a::euclidean_space set" assumes "S \ closure T = {}" "T \ closure S = {}" obtains U V where "U \ V = {}" "open U" "open V" "S \ U" "T \ V" proof (cases "S = {} \ T = {}") case True with that show ?thesis by auto next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on UNIV f" unfolding f_def by (intro continuous_intros continuous_on_setdist) show ?thesis proof (rule_tac U = "{x. f x > 0}" and V = "{x. f x < 0}" in that) show "{x. 0 < f x} \ {x. f x < 0} = {}" by auto show "open {x. 0 < f x}" - by (simp add: open_Collect_less contf continuous_on_const) + by (simp add: open_Collect_less contf) show "open {x. f x < 0}" - by (simp add: open_Collect_less contf continuous_on_const) + by (simp add: open_Collect_less contf) show "S \ {x. 0 < f x}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) show "T \ {x. f x < 0}" apply (clarsimp simp add: f_def setdist_sing_in_set) using assms by (metis False IntI empty_iff le_less setdist_eq_0_sing_2 setdist_pos_le setdist_sym) qed qed lemma separation_normal: fixes S :: "'a::euclidean_space set" assumes "closed S" "closed T" "S \ T = {}" obtains U V where "open U" "open V" "S \ U" "T \ V" "U \ V = {}" using separation_closures [of S T] by (metis assms closure_closed disjnt_def inf_commute) lemma separation_normal_local: fixes S :: "'a::euclidean_space set" assumes US: "closedin (top_of_set U) S" and UT: "closedin (top_of_set U) T" and "S \ T = {}" obtains S' T' where "openin (top_of_set U) S'" "openin (top_of_set U) T'" "S \ S'" "T \ T'" "S' \ T' = {}" proof (cases "S = {} \ T = {}") case True with that show ?thesis using UT US by (blast dest: closedin_subset) next case False define f where "f \ \x. setdist {x} T - setdist {x} S" have contf: "continuous_on U f" unfolding f_def by (intro continuous_intros) show ?thesis proof (rule_tac S' = "(U \ f -` {0<..})" and T' = "(U \ f -` {..<0})" in that) show "(U \ f -` {0<..}) \ (U \ f -` {..<0}) = {}" by auto show "openin (top_of_set U) (U \ f -` {0<..})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next show "openin (top_of_set U) (U \ f -` {..<0})" by (rule continuous_openin_preimage [where T=UNIV]) (simp_all add: contf) next have "S \ U" "T \ U" using closedin_imp_subset assms by blast+ then show "S \ U \ f -` {0<..}" "T \ U \ f -` {..<0}" using assms False by (force simp add: f_def setdist_sing_in_set intro!: setdist_gt_0_closedin)+ qed qed lemma separation_normal_compact: fixes S :: "'a::euclidean_space set" assumes "compact S" "closed T" "S \ T = {}" obtains U V where "open U" "compact(closure U)" "open V" "S \ U" "T \ V" "U \ V = {}" proof - have "closed S" "bounded S" using assms by (auto simp: compact_eq_bounded_closed) then obtain r where "r>0" and r: "S \ ball 0 r" by (auto dest!: bounded_subset_ballD) have **: "closed (T \ - ball 0 r)" "S \ (T \ - ball 0 r) = {}" using assms r by blast+ then show ?thesis apply (rule separation_normal [OF \closed S\]) apply (rule_tac U=U and V=V in that) by auto (meson bounded_ball bounded_subset compl_le_swap2 disjoint_eq_subset_Compl) qed subsection\Connectedness of the intersection of a chain\ proposition connected_chain: fixes \ :: "'a :: euclidean_space set set" assumes cc: "\S. S \ \ \ compact S \ connected S" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof (cases "\ = {}") case True then show ?thesis by auto next case False then have cf: "compact(\\)" by (simp add: cc compact_Inter) have False if AB: "closed A" "closed B" "A \ B = {}" and ABeq: "A \ B = \\" and "A \ {}" "B \ {}" for A B proof - obtain U V where "open U" "open V" "A \ U" "B \ V" "U \ V = {}" using separation_normal [OF AB] by metis obtain K where "K \ \" "compact K" using cc False by blast then obtain N where "open N" and "K \ N" by blast let ?\ = "insert (U \ V) ((\S. N - S) ` \)" obtain \ where "\ \ ?\" "finite \" "K \ \\" proof (rule compactE [OF \compact K\]) show "K \ \(insert (U \ V) ((-) N ` \))" using \K \ N\ ABeq \A \ U\ \B \ V\ by auto show "\B. B \ insert (U \ V) ((-) N ` \) \ open B" by (auto simp: \open U\ \open V\ open_Un \open N\ cc compact_imp_closed open_Diff) qed then have "finite(\ - {U \ V})" by blast moreover have "\ - {U \ V} \ (\S. N - S) ` \" using \\ \ ?\\ by blast ultimately obtain \ where "\ \ \" "finite \" and Deq: "\ - {U \ V} = (\S. N-S) ` \" using finite_subset_image by metis obtain J where "J \ \" and J: "(\S\\. N - S) \ N - J" proof (cases "\ = {}") case True with \\ \ {}\ that show ?thesis by auto next case False have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (meson \\ \ \\ in_mono local.linear) with \finite \\ \\ \ {}\ have "\J \ \. (\S\\. N - S) \ N - J" proof induction case (insert X \) show ?case proof (cases "\ = {}") case True then show ?thesis by auto next case False then have "\S T. \S \ \; T \ \\ \ S \ T \ T \ S" by (simp add: insert.prems) with insert.IH False obtain J where "J \ \" and J: "(\Y\\. N - Y) \ N - J" by metis have "N - J \ N - X \ N - X \ N - J" by (meson Diff_mono \J \ \\ insert.prems(2) insert_iff order_refl) then show ?thesis proof assume "N - J \ N - X" with J show ?thesis by auto next assume "N - X \ N - J" with J have "N - X \ \ ((-) N ` \) \ N - J" by auto with \J \ \\ show ?thesis by blast qed qed qed simp with \\ \ \\ show ?thesis by (blast intro: that) qed have "K \ \(insert (U \ V) (\ - {U \ V}))" using \K \ \\\ by auto also have "... \ (U \ V) \ (N - J)" by (metis (no_types, hide_lams) Deq Un_subset_iff Un_upper2 J Union_insert order_trans sup_ge1) finally have "J \ K \ U \ V" by blast moreover have "connected(J \ K)" by (metis Int_absorb1 \J \ \\ \K \ \\ cc inf.orderE local.linear) moreover have "U \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \A \ {}\ \A \ U\ by blast moreover have "V \ (J \ K) \ {}" using ABeq \J \ \\ \K \ \\ \B \ {}\ \B \ V\ by blast ultimately show False using connectedD [of "J \ K" U V] \open U\ \open V\ \U \ V = {}\ by auto qed with cf show ?thesis by (auto simp: connected_closed_set compact_imp_closed) qed lemma connected_chain_gen: fixes \ :: "'a :: euclidean_space set set" assumes X: "X \ \" "compact X" and cc: "\T. T \ \ \ closed T \ connected T" and linear: "\S T. S \ \ \ T \ \ \ S \ T \ T \ S" shows "connected(\\)" proof - have "\\ = (\T\\. X \ T)" using X by blast moreover have "connected (\T\\. X \ T)" proof (rule connected_chain) show "\T. T \ (\) X ` \ \ compact T \ connected T" using cc X by auto (metis inf.absorb2 inf.orderE local.linear) show "\S T. S \ (\) X ` \ \ T \ (\) X ` \ \ S \ T \ T \ S" using local.linear by blast qed ultimately show ?thesis by metis qed lemma connected_nest: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. compact(S n)" "\n. connected(S n)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain) using S apply blast by (metis image_iff le_cases nest) lemma connected_nest_gen: fixes S :: "'a::linorder \ 'b::euclidean_space set" assumes S: "\n. closed(S n)" "\n. connected(S n)" "compact(S k)" and nest: "\m n. m \ n \ S n \ S m" shows "connected(\ (range S))" apply (rule connected_chain_gen [of "S k"]) using S apply auto by (meson le_cases nest subsetCE) subsection\Proper maps, including projections out of compact sets\ lemma finite_indexed_bound: assumes A: "finite A" "\x. x \ A \ \n::'a::linorder. P x n" shows "\m. \x \ A. \k\m. P x k" using A proof (induction A) case empty then show ?case by force next case (insert a A) then obtain m n where "\x \ A. \k\m. P x k" "P a n" by force then show ?case apply (rule_tac x="max m n" in exI, safe) using max.cobounded2 apply blast by (meson le_max_iff_disj) qed proposition proper_map: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes "closedin (top_of_set S) K" and com: "\U. \U \ T; compact U\ \ compact (S \ f -` U)" and "f ` S \ T" shows "closedin (top_of_set T) (f ` K)" proof - have "K \ S" using assms closedin_imp_subset by metis obtain C where "closed C" and Keq: "K = S \ C" using assms by (auto simp: closedin_closed) have *: "y \ f ` K" if "y \ T" and y: "y islimpt f ` K" for y proof - obtain h where "\n. (\x\K. h n = f x) \ h n \ y" "inj h" and hlim: "(h \ y) sequentially" using \y \ T\ y by (force simp: limpt_sequential_inj) then obtain X where X: "\n. X n \ K \ h n = f (X n) \ h n \ y" by metis then have fX: "\n. f (X n) = h n" by metis have "compact (C \ (S \ f -` insert y (range (\i. f(X(n + i))))))" for n apply (rule closed_Int_compact [OF \closed C\]) apply (rule com) using X \K \ S\ \f ` S \ T\ \y \ T\ apply blast apply (rule compact_sequence_with_limit) apply (simp add: fX add.commute [of n] LIMSEQ_ignore_initial_segment [OF hlim]) done then have comf: "compact {a \ K. f a \ insert y (range (\i. f(X(n + i))))}" for n by (simp add: Keq Int_def conj_commute) have ne: "\\ \ {}" if "finite \" and \: "\t. t \ \ \ (\n. t = {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" for \ proof - obtain m where m: "\t. t \ \ \ \k\m. t = {a \ K. f a \ insert y (range (\i. f (X (k + i))))}" apply (rule exE) apply (rule finite_indexed_bound [OF \finite \\ \], assumption, force) done have "X m \ \\" using X le_Suc_ex by (fastforce dest: m) then show ?thesis by blast qed have "\{{a. a \ K \ f a \ insert y (range (\i. f(X(n + i))))} |n. n \ UNIV} \ {}" apply (rule compact_fip_Heine_Borel) using comf apply force using ne apply (simp add: subset_iff del: insert_iff) done then have "\x. x \ (\n. {a \ K. f a \ insert y (range (\i. f (X (n + i))))})" by blast then show ?thesis apply (simp add: image_iff fX) by (metis \inj h\ le_add1 not_less_eq_eq rangeI range_ex1_eq) qed with assms closedin_subset show ?thesis by (force simp: closedin_limpt) qed lemma compact_continuous_image_eq: fixes f :: "'a::heine_borel \ 'b::heine_borel" assumes f: "inj_on f S" shows "continuous_on S f \ (\T. compact T \ T \ S \ compact(f ` T))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (metis continuous_on_subset compact_continuous_image) next assume RHS: ?rhs obtain g where gf: "\x. x \ S \ g (f x) = x" by (metis inv_into_f_f f) then have *: "(S \ f -` U) = g ` U" if "U \ f ` S" for U using that by fastforce have gfim: "g ` f ` S \ S" using gf by auto have **: "compact (f ` S \ g -` C)" if C: "C \ S" "compact C" for C proof - obtain h where "h C \ C \ h C \ S \ compact (f ` C)" by (force simp: C RHS) moreover have "f ` C = (f ` S \ g -` C)" using C gf by auto ultimately show ?thesis using C by auto qed show ?lhs using proper_map [OF _ _ gfim] ** by (simp add: continuous_on_closed * closedin_imp_subset) qed subsection\<^marker>\tag unimportant\\Trivial fact: convexity equals connectedness for collinear sets\ lemma convex_connected_collinear: fixes S :: "'a::euclidean_space set" assumes "collinear S" shows "convex S \ connected S" proof assume "convex S" then show "connected S" using convex_connected by blast next assume S: "connected S" show "convex S" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain a where "a \ S" by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - a = c *\<^sub>R z" by (meson \a \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - a = f x *\<^sub>R z" by metis then have inj_f: "inj_on f (affine hull S)" by (metis diff_add_cancel inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - a" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - a" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have conn_fS: "connected (f ` S)" by (meson S connected_continuous_image continuous_on_subset hull_subset) show ?thesis proof (clarsimp simp: convex_contains_segment) fix x y z assume "x \ S" "y \ S" "z \ closed_segment x y" have False if "z \ S" proof - have "f ` (closed_segment x y) = closed_segment (f x) (f y)" apply (rule continuous_injective_image_segment_1) apply (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \x \ S\ \y \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have fz: "f z \ closed_segment (f x) (f y)" using \z \ closed_segment x y\ by blast have "z \ affine hull S" by (meson \x \ S\ \y \ S\ \z \ closed_segment x y\ convex_affine_hull convex_contains_segment hull_inc subset_eq) then have fz_notin: "f z \ f ` S" using hull_subset inj_f inj_onD that by fastforce moreover have "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" proof - have "{.. f ` {x,y} \ {}" "{f z<..} \ f ` {x,y} \ {}" using fz fz_notin \x \ S\ \y \ S\ apply (auto simp: closed_segment_eq_real_ivl split: if_split_asm) apply (metis image_eqI less_eq_real_def)+ done then show "{.. f ` S \ {}" "{f z<..} \ f ` S \ {}" using \x \ S\ \y \ S\ by blast+ qed ultimately show False using connectedD [OF conn_fS, of "{.. S" by meson qed qed qed lemma compact_convex_collinear_segment_alt: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "connected S" "collinear S" obtains a b where "S = closed_segment a b" proof - obtain \ where "\ \ S" using \S \ {}\ by auto have "collinear (affine hull S)" by (simp add: assms collinear_affine_hull_collinear) then obtain z where "z \ 0" "\x. x \ affine hull S \ \c. x - \ = c *\<^sub>R z" by (meson \\ \ S\ collinear hull_inc) then obtain f where f: "\x. x \ affine hull S \ x - \ = f x *\<^sub>R z" by metis let ?g = "\r. r *\<^sub>R z + \" have gf: "?g (f x) = x" if "x \ affine hull S" for x by (metis diff_add_cancel f that) then have inj_f: "inj_on f (affine hull S)" by (metis inj_onI) have diff: "x - y = (f x - f y) *\<^sub>R z" if x: "x \ affine hull S" and y: "y \ affine hull S" for x y proof - have "f x *\<^sub>R z = x - \" by (simp add: f hull_inc x) moreover have "f y *\<^sub>R z = y - \" by (simp add: f hull_inc y) ultimately show ?thesis by (simp add: scaleR_left.diff) qed have cont_f: "continuous_on (affine hull S) f" apply (clarsimp simp: dist_norm continuous_on_iff diff) by (metis \z \ 0\ mult.commute mult_less_cancel_left_pos norm_minus_commute real_norm_def zero_less_mult_iff zero_less_norm_iff) then have "connected (f ` S)" by (meson \connected S\ connected_continuous_image continuous_on_subset hull_subset) moreover have "compact (f ` S)" by (meson \compact S\ compact_continuous_image_eq cont_f hull_subset inj_f) ultimately obtain x y where "f ` S = {x..y}" by (meson connected_compact_interval_1) then have fS_eq: "f ` S = closed_segment x y" using \S \ {}\ closed_segment_eq_real_ivl by auto obtain a b where "a \ S" "f a = x" "b \ S" "f b = y" by (metis (full_types) ends_in_segment fS_eq imageE) have "f ` (closed_segment a b) = closed_segment (f a) (f b)" apply (rule continuous_injective_image_segment_1) apply (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc continuous_on_subset [OF cont_f]) by (meson \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment hull_inc inj_on_subset [OF inj_f]) then have "f ` (closed_segment a b) = f ` S" by (simp add: \f a = x\ \f b = y\ fS_eq) then have "?g ` f ` (closed_segment a b) = ?g ` f ` S" by simp moreover have "(\x. f x *\<^sub>R z + \) ` closed_segment a b = closed_segment a b" apply safe apply (metis (mono_tags, hide_lams) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_inc subsetCE) by (metis (mono_tags, lifting) \a \ S\ \b \ S\ convex_affine_hull convex_contains_segment gf hull_subset image_iff subsetCE) ultimately have "closed_segment a b = S" using gf by (simp add: image_comp o_def hull_inc cong: image_cong) then show ?thesis using that by blast qed lemma compact_convex_collinear_segment: fixes S :: "'a::euclidean_space set" assumes "S \ {}" "compact S" "convex S" "collinear S" obtains a b where "S = closed_segment a b" using assms convex_connected_collinear compact_convex_collinear_segment_alt by blast lemma proper_map_from_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and imf: "f ` S \ T" and "compact S" "closedin (top_of_set T) K" shows "compact (S \ f -` K)" by (rule closedin_compact [OF \compact S\] continuous_closedin_preimage_gen assms)+ lemma proper_map_fst: assumes "compact T" "K \ S" "compact K" shows "compact (S \ T \ fst -` K)" proof - have "(S \ T \ fst -` K) = K \ T" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_fst: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact T" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set S) (fst ` c)" proof - have *: "fst ` (S \ T) \ S" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_fst assms) qed lemma proper_map_snd: assumes "compact S" "K \ T" "compact K" shows "compact (S \ T \ snd -` K)" proof - have "(S \ T \ snd -` K) = S \ K" using assms by auto then show ?thesis by (simp add: assms compact_Times) qed lemma closed_map_snd: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" "closedin (top_of_set (S \ T)) c" shows "closedin (top_of_set T) (snd ` c)" proof - have *: "snd ` (S \ T) \ T" by auto show ?thesis using proper_map [OF _ _ *] by (simp add: proper_map_snd assms) qed lemma closedin_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "compact S" and clo: "closedin (top_of_set (S \ T)) U" shows "closedin (top_of_set T) {y. \x. x \ S \ (x, y) \ U}" proof - have "U \ S \ T" by (metis clo closedin_imp_subset) then have "{y. \x. x \ S \ (x, y) \ U} = snd ` U" by force moreover have "closedin (top_of_set T) (snd ` U)" by (rule closed_map_snd [OF assms]) ultimately show ?thesis by simp qed lemma closed_compact_projection: fixes S :: "'a::euclidean_space set" and T :: "('a * 'b::euclidean_space) set" assumes "compact S" and clo: "closed T" shows "closed {y. \x. x \ S \ (x, y) \ T}" proof - have *: "{y. \x. x \ S \ Pair x y \ T} = {y. \x. x \ S \ Pair x y \ ((S \ UNIV) \ T)}" by auto show ?thesis apply (subst *) apply (rule closedin_closed_trans [OF _ closed_UNIV]) apply (rule closedin_compact_projection [OF \compact S\]) by (simp add: clo closedin_closed_Int) qed subsubsection\<^marker>\tag unimportant\\Representing affine hull as a finite intersection of hyperplanes\ proposition\<^marker>\tag unimportant\ affine_hull_convex_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "convex S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" proof show "affine hull (S \ T) \ affine hull S" by (simp add: hull_mono) next obtain a where "a \ S" "a \ T" and at: "a \ interior T" using assms interior_subset by blast then obtain e where "e > 0" and e: "cball a e \ T" using mem_interior_cball by blast have *: "x \ (+) a ` span ((\x. x - a) ` (S \ T))" if "x \ S" for x proof (cases "x = a") case True with that span_0 eq_add_iff image_def mem_Collect_eq show ?thesis by blast next case False define k where "k = min (1/2) (e / norm (x-a))" have k: "0 < k" "k < 1" using \e > 0\ False by (auto simp: k_def) then have xa: "(x-a) = inverse k *\<^sub>R k *\<^sub>R (x-a)" by simp have "e / norm (x - a) \ k" using k_def by linarith then have "a + k *\<^sub>R (x - a) \ cball a e" using \0 < k\ False by (simp add: dist_norm) (simp add: field_simps) then have T: "a + k *\<^sub>R (x - a) \ T" using e by blast have S: "a + k *\<^sub>R (x - a) \ S" using k \a \ S\ convexD [OF \convex S\ \a \ S\ \x \ S\, of "1-k" k] by (simp add: algebra_simps) have "inverse k *\<^sub>R k *\<^sub>R (x-a) \ span ((\x. x - a) ` (S \ T))" apply (rule span_mul) apply (rule span_base) apply (rule image_eqI [where x = "a + k *\<^sub>R (x - a)"]) apply (auto simp: S T) done with xa image_iff show ?thesis by fastforce qed show "affine hull S \ affine hull (S \ T)" apply (simp add: subset_hull) apply (simp add: \a \ S\ \a \ T\ hull_inc affine_hull_span_gen [of a]) apply (force simp: *) done qed corollary affine_hull_convex_Int_open: fixes S :: "'a::real_normed_vector set" assumes "convex S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using affine_hull_convex_Int_nonempty_interior assms interior_eq by blast corollary affine_hull_affine_Int_nonempty_interior: fixes S :: "'a::real_normed_vector set" assumes "affine S" "S \ interior T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_nonempty_interior affine_imp_convex assms) corollary affine_hull_affine_Int_open: fixes S :: "'a::real_normed_vector set" assumes "affine S" "open T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" by (simp add: affine_hull_convex_Int_open affine_imp_convex assms) corollary affine_hull_convex_Int_openin: fixes S :: "'a::real_normed_vector set" assumes "convex S" "openin (top_of_set (affine hull S)) T" "S \ T \ {}" shows "affine hull (S \ T) = affine hull S" using assms unfolding openin_open by (metis affine_hull_convex_Int_open hull_subset inf.orderE inf_assoc) corollary affine_hull_openin: fixes S :: "'a::real_normed_vector set" assumes "openin (top_of_set (affine hull T)) S" "S \ {}" shows "affine hull S = affine hull T" using assms unfolding openin_open by (metis affine_affine_hull affine_hull_affine_Int_open hull_hull) corollary affine_hull_open: fixes S :: "'a::real_normed_vector set" assumes "open S" "S \ {}" shows "affine hull S = UNIV" by (metis affine_hull_convex_Int_nonempty_interior assms convex_UNIV hull_UNIV inf_top.left_neutral interior_open) lemma aff_dim_convex_Int_nonempty_interior: fixes S :: "'a::euclidean_space set" shows "\convex S; S \ interior T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_affine_hull2 affine_hull_convex_Int_nonempty_interior by blast lemma aff_dim_convex_Int_open: fixes S :: "'a::euclidean_space set" shows "\convex S; open T; S \ T \ {}\ \ aff_dim(S \ T) = aff_dim S" using aff_dim_convex_Int_nonempty_interior interior_eq by blast lemma affine_hull_Diff: fixes S:: "'a::real_normed_vector set" assumes ope: "openin (top_of_set (affine hull S)) S" and "finite F" "F \ S" shows "affine hull (S - F) = affine hull S" proof - have clo: "closedin (top_of_set S) F" using assms finite_imp_closedin by auto moreover have "S - F \ {}" using assms by auto ultimately show ?thesis by (metis ope closedin_def topspace_euclidean_subtopology affine_hull_openin openin_trans) qed lemma affine_hull_halfspace_lt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x < r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_lt [of a r] by (simp add: open_halfspace_lt affine_hull_open) lemma affine_hull_halfspace_le: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r < 0 then {} else UNIV)" proof (cases "a = 0") case True then show ?thesis by simp next case False then have "affine hull closure {x. a \ x < r} = UNIV" using affine_hull_halfspace_lt closure_same_affine_hull by fastforce moreover have "{x. a \ x < r} \ {x. a \ x \ r}" by (simp add: Collect_mono) ultimately show ?thesis using False antisym_conv hull_mono top_greatest by (metis affine_hull_halfspace_lt) qed lemma affine_hull_halfspace_gt: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x > r} = (if a = 0 \ r \ 0 then {} else UNIV)" using halfspace_eq_empty_gt [of r a] by (simp add: open_halfspace_gt affine_hull_open) lemma affine_hull_halfspace_ge: fixes a :: "'a::euclidean_space" shows "affine hull {x. a \ x \ r} = (if a = 0 \ r > 0 then {} else UNIV)" using affine_hull_halfspace_le [of "-a" "-r"] by simp lemma aff_dim_halfspace_lt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x < r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_lt open_halfspace_lt) lemma aff_dim_halfspace_le: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r < 0 then -1 else DIM('a))" proof - have "int (DIM('a)) = aff_dim (UNIV::'a set)" by (simp add: aff_dim_UNIV) then have "aff_dim (affine hull {x. a \ x \ r}) = DIM('a)" if "(a = 0 \ r \ 0)" using that by (simp add: affine_hull_halfspace_le not_less) then show ?thesis by (force simp: aff_dim_affine_hull) qed lemma aff_dim_halfspace_gt: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x > r} = (if a = 0 \ r \ 0 then -1 else DIM('a))" by simp (metis aff_dim_open halfspace_eq_empty_gt open_halfspace_gt) lemma aff_dim_halfspace_ge: fixes a :: "'a::euclidean_space" shows "aff_dim {x. a \ x \ r} = (if a = 0 \ r > 0 then -1 else DIM('a))" using aff_dim_halfspace_le [of "-a" "-r"] by simp proposition aff_dim_eq_hyperplane: fixes S :: "'a::euclidean_space set" shows "aff_dim S = DIM('a) - 1 \ (\a b. a \ 0 \ affine hull S = {x. a \ x = b})" proof (cases "S = {}") case True then show ?thesis by (auto simp: dest: hyperplane_eq_Ex) next case False then obtain c where "c \ S" by blast show ?thesis proof (cases "c = 0") case True show ?thesis using span_zero [of S] apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def) apply (auto simp add: \c = 0\) done next case False have xc_im: "x \ (+) c ` {y. a \ y = 0}" if "a \ x = a \ c" for a x proof - have "\y. a \ y = 0 \ c + y = x" by (metis that add.commute diff_add_cancel inner_commute inner_diff_left right_minus_eq) then show "x \ (+) c ` {y. a \ y = 0}" by blast qed have 2: "span ((\x. x - c) ` S) = {x. a \ x = 0}" if "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = b}" for a b proof - have "b = a \ c" using span_0 that by fastforce with that have "(+) c ` span ((\x. x - c) ` S) = {x. a \ x = a \ c}" by simp then have "span ((\x. x - c) ` S) = (\x. x - c) ` {x. a \ x = a \ c}" by (metis (no_types) image_cong translation_galois uminus_add_conv_diff) also have "... = {x. a \ x = 0}" by (force simp: inner_distrib inner_diff_right intro: image_eqI [where x="x+c" for x]) finally show ?thesis . qed show ?thesis apply (simp add: aff_dim_eq_dim [of c] affine_hull_span_gen [of c] \c \ S\ hull_inc dim_eq_hyperplane del: One_nat_def cong: image_cong_simp, safe) apply (fastforce simp add: inner_distrib intro: xc_im) apply (force simp: intro!: 2) done qed qed corollary aff_dim_hyperplane [simp]: fixes a :: "'a::euclidean_space" shows "a \ 0 \ aff_dim {x. a \ x = r} = DIM('a) - 1" by (metis aff_dim_eq_hyperplane affine_hull_eq affine_hyperplane) subsection\<^marker>\tag unimportant\\Some stepping theorems\ lemma aff_dim_insert: fixes a :: "'a::euclidean_space" shows "aff_dim (insert a S) = (if a \ affine hull S then aff_dim S else aff_dim S + 1)" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain x s' where S: "S = insert x s'" "x \ s'" by (meson Set.set_insert all_not_in_conv) show ?thesis using S apply (simp add: hull_redundant cong: aff_dim_affine_hull2) apply (simp add: affine_hull_insert_span_gen hull_inc) by (force simp add: span_zero insert_commute [of a] hull_inc aff_dim_eq_dim [of x] dim_insert cong: image_cong_simp) qed lemma affine_dependent_choose: fixes a :: "'a :: euclidean_space" assumes "\(affine_dependent S)" shows "affine_dependent(insert a S) \ a \ S \ a \ affine hull S" (is "?lhs = ?rhs") proof safe assume "affine_dependent (insert a S)" and "a \ S" then show "False" using \a \ S\ assms insert_absorb by fastforce next assume lhs: "affine_dependent (insert a S)" then have "a \ S" by (metis (no_types) assms insert_absorb) moreover have "finite S" using affine_independent_iff_card assms by blast moreover have "aff_dim (insert a S) \ int (card S)" using \finite S\ affine_independent_iff_card \a \ S\ lhs by fastforce ultimately show "a \ affine hull S" by (metis aff_dim_affine_independent aff_dim_insert assms) next assume "a \ S" and "a \ affine hull S" show "affine_dependent (insert a S)" by (simp add: \a \ affine hull S\ \a \ S\ affine_dependent_def) qed lemma affine_independent_insert: fixes a :: "'a :: euclidean_space" shows "\\ affine_dependent S; a \ affine hull S\ \ \ affine_dependent(insert a S)" by (simp add: affine_dependent_choose) lemma subspace_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "subspace S" shows "bounded S \ S = {0}" proof - have "False" if "bounded S" "x \ S" "x \ 0" for x proof - obtain B where B: "\y. y \ S \ norm y < B" "B > 0" using \bounded S\ by (force simp: bounded_pos_less) have "(B / norm x) *\<^sub>R x \ S" using assms subspace_mul \x \ S\ by auto moreover have "norm ((B / norm x) *\<^sub>R x) = B" using that B by (simp add: algebra_simps) ultimately show False using B by force qed then have "bounded S \ S = {0}" using assms subspace_0 by fastforce then show ?thesis by blast qed lemma affine_bounded_eq_trivial: fixes S :: "'a::real_normed_vector set" assumes "affine S" shows "bounded S \ S = {} \ (\a. S = {a})" proof (cases "S = {}") case True then show ?thesis by simp next case False then obtain b where "b \ S" by blast with False assms show ?thesis apply safe using affine_diffs_subspace [OF assms \b \ S\] apply (metis (no_types, lifting) subspace_bounded_eq_trivial ab_left_minus bounded_translation image_empty image_insert translation_invert) apply force done qed lemma affine_bounded_eq_lowdim: fixes S :: "'a::euclidean_space set" assumes "affine S" shows "bounded S \ aff_dim S \ 0" apply safe using affine_bounded_eq_trivial assms apply fastforce by (metis aff_dim_sing aff_dim_subset affine_dim_equal affine_sing all_not_in_conv assms bounded_empty bounded_insert dual_order.antisym empty_subsetI insert_subset) lemma bounded_hyperplane_eq_trivial_0: fixes a :: "'a::euclidean_space" assumes "a \ 0" shows "bounded {x. a \ x = 0} \ DIM('a) = 1" proof assume "bounded {x. a \ x = 0}" then have "aff_dim {x. a \ x = 0} \ 0" by (simp add: affine_bounded_eq_lowdim affine_hyperplane) with assms show "DIM('a) = 1" by (simp add: le_Suc_eq aff_dim_hyperplane) next assume "DIM('a) = 1" then show "bounded {x. a \ x = 0}" by (simp add: aff_dim_hyperplane affine_bounded_eq_lowdim affine_hyperplane assms) qed lemma bounded_hyperplane_eq_trivial: fixes a :: "'a::euclidean_space" shows "bounded {x. a \ x = r} \ (if a = 0 then r \ 0 else DIM('a) = 1)" proof (simp add: bounded_hyperplane_eq_trivial_0, clarify) assume "r \ 0" "a \ 0" have "aff_dim {x. y \ x = 0} = aff_dim {x. a \ x = r}" if "y \ 0" for y::'a by (metis that \a \ 0\ aff_dim_hyperplane) then show "bounded {x. a \ x = r} = (DIM('a) = Suc 0)" by (metis One_nat_def \a \ 0\ affine_bounded_eq_lowdim affine_hyperplane bounded_hyperplane_eq_trivial_0) qed subsection\<^marker>\tag unimportant\\General case without assuming closure and getting non-strict separation\ proposition\<^marker>\tag unimportant\ separating_hyperplane_closed_point_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "z \ S" obtains a b where "a \ S" "(a - z) \ z < b" "\x. x \ S \ b < (a - z) \ x" proof - obtain y where "y \ S" and y: "\u. u \ S \ dist z y \ dist z u" using distance_attains_inf [of S z] assms by auto then have *: "(y - z) \ z < (y - z) \ z + (norm (y - z))\<^sup>2 / 2" using \y \ S\ \z \ S\ by auto show ?thesis proof (rule that [OF \y \ S\ *]) fix x assume "x \ S" have yz: "0 < (y - z) \ (y - z)" using \y \ S\ \z \ S\ by auto { assume 0: "0 < ((z - y) \ (x - y))" with any_closest_point_dot [OF \convex S\ \closed S\] have False using y \x \ S\ \y \ S\ not_less by blast } then have "0 \ ((y - z) \ (x - y))" by (force simp: not_less inner_diff_left) with yz have "0 < 2 * ((y - z) \ (x - y)) + (y - z) \ (y - z)" by (simp add: algebra_simps) then show "(y - z) \ z + (norm (y - z))\<^sup>2 / 2 < (y - z) \ x" by (simp add: field_simps inner_diff_left inner_diff_right dot_square_norm [symmetric]) qed qed lemma separating_hyperplane_closed_0_inset: fixes S :: "'a::euclidean_space set" assumes "convex S" "closed S" "S \ {}" "0 \ S" obtains a b where "a \ S" "a \ 0" "0 < b" "\x. x \ S \ a \ x > b" using separating_hyperplane_closed_point_inset [OF assms] by simp (metis \0 \ S\) proposition\<^marker>\tag unimportant\ separating_hyperplane_set_0_inspan: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" "0 \ S" obtains a where "a \ span S" "a \ 0" "\x. x \ S \ 0 \ a \ x" proof - define k where [abs_def]: "k c = {x. 0 \ c \ x}" for c :: 'a have *: "span S \ frontier (cball 0 1) \ \f' \ {}" if f': "finite f'" "f' \ k ` S" for f' proof - obtain C where "C \ S" "finite C" and C: "f' = k ` C" using finite_subset_image [OF f'] by blast obtain a where "a \ S" "a \ 0" using \S \ {}\ \0 \ S\ ex_in_conv by blast then have "norm (a /\<^sub>R (norm a)) = 1" by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp show ?thesis proof (cases "C = {}") case True with C ass show ?thesis by auto next case False have "closed (convex hull C)" using \finite C\ compact_eq_bounded_closed finite_imp_compact_convex_hull by auto moreover have "convex hull C \ {}" by (simp add: False) moreover have "0 \ convex hull C" by (metis \C \ S\ \convex S\ \0 \ S\ convex_hull_subset hull_same insert_absorb insert_subset) ultimately obtain a b where "a \ convex hull C" "a \ 0" "0 < b" and ab: "\x. x \ convex hull C \ a \ x > b" using separating_hyperplane_closed_0_inset by blast then have "a \ S" by (metis \C \ S\ assms(1) subsetCE subset_hull) moreover have "norm (a /\<^sub>R (norm a)) = 1" using \a \ 0\ by simp moreover have "a /\<^sub>R (norm a) \ span S" by (simp add: \a \ S\ span_scale span_base) ultimately have ass: "a /\<^sub>R (norm a) \ span S \ sphere 0 1" by simp have aa: "a /\<^sub>R (norm a) \ (\c\C. {x. 0 \ c \ x})" apply (clarsimp simp add: field_split_simps) using ab \0 < b\ by (metis hull_inc inner_commute less_eq_real_def less_trans) show ?thesis apply (simp add: C k_def) using ass aa Int_iff empty_iff by blast qed qed have "(span S \ frontier(cball 0 1)) \ (\ (k ` S)) \ {}" apply (rule compact_imp_fip) apply (blast intro: compact_cball) using closed_halfspace_ge k_def apply blast apply (metis *) done then show ?thesis unfolding set_eq_iff k_def by simp (metis inner_commute norm_eq_zero that zero_neq_one) qed lemma separating_hyperplane_set_point_inaff: fixes S :: "'a::euclidean_space set" assumes "convex S" "S \ {}" and zno: "z \ S" obtains a b where "(z + a) \ affine hull (insert z S)" and "a \ 0" and "a \ z \ b" and "\x. x \ S \ a \ x \ b" proof - from separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] have "convex ((+) (- z) ` S)" using \convex S\ by simp moreover have "(+) (- z) ` S \ {}" by (simp add: \S \ {}\) moreover have "0 \ (+) (- z) ` S" using zno by auto ultimately obtain a where "a \ span ((+) (- z) ` S)" "a \ 0" and a: "\x. x \ ((+) (- z) ` S) \ 0 \ a \ x" using separating_hyperplane_set_0_inspan [of "image (\x. -z + x) S"] by blast then have szx: "\x. x \ S \ a \ z \ a \ x" by (metis (no_types, lifting) imageI inner_minus_right inner_right_distrib minus_add neg_le_0_iff_le neg_le_iff_le real_add_le_0_iff) show ?thesis apply (rule_tac a=a and b = "a \ z" in that, simp_all) using \a \ span ((+) (- z) ` S)\ affine_hull_insert_span_gen apply blast apply (simp_all add: \a \ 0\ szx) done qed proposition\<^marker>\tag unimportant\ supporting_hyperplane_rel_boundary: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ S" and xno: "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" proof - obtain a b where aff: "(x + a) \ affine hull (insert x (rel_interior S))" and "a \ 0" and "a \ x \ b" and ageb: "\u. u \ (rel_interior S) \ a \ u \ b" using separating_hyperplane_set_point_inaff [of "rel_interior S" x] assms by (auto simp: rel_interior_eq_empty convex_rel_interior) have le_ay: "a \ x \ a \ y" if "y \ S" for y proof - have con: "continuous_on (closure (rel_interior S)) ((\) a)" by (rule continuous_intros continuous_on_subset | blast)+ have y: "y \ closure (rel_interior S)" using \convex S\ closure_def convex_closure_rel_interior \y \ S\ by fastforce show ?thesis using continuous_ge_on_closure [OF con y] ageb \a \ x \ b\ by fastforce qed have 3: "a \ x < a \ y" if "y \ rel_interior S" for y proof - obtain e where "0 < e" "y \ S" and e: "cball y e \ affine hull S \ S" using \y \ rel_interior S\ by (force simp: rel_interior_cball) define y' where "y' = y - (e / norm a) *\<^sub>R ((x + a) - x)" have "y' \ cball y e" unfolding y'_def using \0 < e\ by force moreover have "y' \ affine hull S" unfolding y'_def by (metis \x \ S\ \y \ S\ \convex S\ aff affine_affine_hull hull_redundant rel_interior_same_affine_hull hull_inc mem_affine_3_minus2) ultimately have "y' \ S" using e by auto have "a \ x \ a \ y" using le_ay \a \ 0\ \y \ S\ by blast moreover have "a \ x \ a \ y" using le_ay [OF \y' \ S\] \a \ 0\ apply (simp add: y'_def inner_diff dot_square_norm power2_eq_square) by (metis \0 < e\ add_le_same_cancel1 inner_commute inner_real_def inner_zero_left le_diff_eq norm_le_zero_iff real_mult_le_cancel_iff2) ultimately show ?thesis by force qed show ?thesis by (rule that [OF \a \ 0\ le_ay 3]) qed lemma supporting_hyperplane_relative_frontier: fixes S :: "'a::euclidean_space set" assumes "convex S" "x \ closure S" "x \ rel_interior S" obtains a where "a \ 0" and "\y. y \ closure S \ a \ x \ a \ y" and "\y. y \ rel_interior S \ a \ x < a \ y" using supporting_hyperplane_rel_boundary [of "closure S" x] by (metis assms convex_closure convex_rel_interior_closure) subsection\<^marker>\tag unimportant\\ Some results on decomposing convex hulls: intersections, simplicial subdivision\ lemma fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows convex_hull_Int_subset: "convex hull s \ convex hull t \ convex hull (s \ t)" (is ?C) and affine_hull_Int_subset: "affine hull s \ affine hull t \ affine hull (s \ t)" (is ?A) proof - have [simp]: "finite s" "finite t" using aff_independent_finite assms by blast+ have "sum u (s \ t) = 1 \ (\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" if [simp]: "sum u s = 1" "sum v t = 1" and eq: "(\x\t. v x *\<^sub>R x) = (\x\s. u x *\<^sub>R x)" for u v proof - define f where "f x = (if x \ s then u x else 0) - (if x \ t then v x else 0)" for x have "sum f (s \ t) = 0" apply (simp add: f_def sum_Un sum_subtractf) apply (simp add: sum.inter_restrict [symmetric] Int_commute) done moreover have "(\x\(s \ t). f x *\<^sub>R x) = 0" apply (simp add: f_def sum_Un scaleR_left_diff_distrib sum_subtractf) apply (simp add: if_smult sum.inter_restrict [symmetric] Int_commute eq cong del: if_weak_cong) done ultimately have "\v. v \ s \ t \ f v = 0" using aff_independent_finite assms unfolding affine_dependent_explicit by blast then have u [simp]: "\x. x \ s \ u x = (if x \ t then v x else 0)" by (simp add: f_def) presburger have "sum u (s \ t) = sum u s" by (simp add: sum.inter_restrict) then have "sum u (s \ t) = 1" using that by linarith moreover have "(\v\s \ t. u v *\<^sub>R v) = (\v\s. u v *\<^sub>R v)" by (auto simp: if_smult sum.inter_restrict intro: sum.cong) ultimately show ?thesis by force qed then show ?A ?C by (auto simp: convex_hull_finite affine_hull_finite) qed proposition\<^marker>\tag unimportant\ affine_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "affine hull (s \ t) = affine hull s \ affine hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: affine_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ convex_hull_Int: fixes s :: "'a::euclidean_space set" assumes "\ affine_dependent(s \ t)" shows "convex hull (s \ t) = convex hull s \ convex hull t" apply (rule subset_antisym) apply (simp add: hull_mono) by (simp add: convex_hull_Int_subset assms) proposition\<^marker>\tag unimportant\ fixes s :: "'a::euclidean_space set set" assumes "\ affine_dependent (\s)" shows affine_hull_Inter: "affine hull (\s) = (\t\s. affine hull t)" (is "?A") and convex_hull_Inter: "convex hull (\s) = (\t\s. convex hull t)" (is "?C") proof - have "finite s" using aff_independent_finite assms finite_UnionD by blast then have "?A \ ?C" using assms proof (induction s rule: finite_induct) case empty then show ?case by auto next case (insert t F) then show ?case proof (cases "F={}") case True then show ?thesis by simp next case False with "insert.prems" have [simp]: "\ affine_dependent (t \ \F)" by (auto intro: affine_dependent_subset) have [simp]: "\ affine_dependent (\F)" using affine_independent_subset insert.prems by fastforce show ?thesis by (simp add: affine_hull_Int convex_hull_Int insert.IH) qed qed then show "?A" "?C" by auto qed proposition\<^marker>\tag unimportant\ in_convex_hull_exchange_unique: fixes S :: "'a::euclidean_space set" assumes naff: "\ affine_dependent S" and a: "a \ convex hull S" and S: "T \ S" "T' \ S" and x: "x \ convex hull (insert a T)" and x': "x \ convex hull (insert a T')" shows "x \ convex hull (insert a (T \ T'))" proof (cases "a \ S") case True then have "\ affine_dependent (insert a T \ insert a T')" using affine_dependent_subset assms by auto then have "x \ convex hull (insert a T \ insert a T')" by (metis IntI convex_hull_Int x x') then show ?thesis by simp next case False then have anot: "a \ T" "a \ T'" using assms by auto have [simp]: "finite S" by (simp add: aff_independent_finite assms) then obtain b where b0: "\s. s \ S \ 0 \ b s" and b1: "sum b S = 1" and aeq: "a = (\s\S. b s *\<^sub>R s)" using a by (auto simp: convex_hull_finite) have fin [simp]: "finite T" "finite T'" using assms infinite_super \finite S\ by blast+ then obtain c c' where c0: "\t. t \ insert a T \ 0 \ c t" and c1: "sum c (insert a T) = 1" and xeq: "x = (\t \ insert a T. c t *\<^sub>R t)" and c'0: "\t. t \ insert a T' \ 0 \ c' t" and c'1: "sum c' (insert a T') = 1" and x'eq: "x = (\t \ insert a T'. c' t *\<^sub>R t)" using x x' by (auto simp: convex_hull_finite) with fin anot have sumTT': "sum c T = 1 - c a" "sum c' T' = 1 - c' a" and wsumT: "(\t \ T. c t *\<^sub>R t) = x - c a *\<^sub>R a" by simp_all have wsumT': "(\t \ T'. c' t *\<^sub>R t) = x - c' a *\<^sub>R a" using x'eq fin anot by simp define cc where "cc \ \x. if x \ T then c x else 0" define cc' where "cc' \ \x. if x \ T' then c' x else 0" define dd where "dd \ \x. cc x - cc' x + (c a - c' a) * b x" have sumSS': "sum cc S = 1 - c a" "sum cc' S = 1 - c' a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 sum_subtractf sum.inter_restrict [symmetric] sumTT') have wsumSS: "(\t \ S. cc t *\<^sub>R t) = x - c a *\<^sub>R a" "(\t \ S. cc' t *\<^sub>R t) = x - c' a *\<^sub>R a" unfolding cc_def cc'_def using S by (simp_all add: Int_absorb1 Int_absorb2 if_smult sum.inter_restrict [symmetric] wsumT wsumT' cong: if_cong) have sum_dd0: "sum dd S = 0" unfolding dd_def using S by (simp add: sumSS' comm_monoid_add_class.sum.distrib sum_subtractf algebra_simps sum_distrib_right [symmetric] b1) have "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R (\v\S. b v *\<^sub>R v)" for x by (simp add: pth_5 real_vector.scale_sum_right mult.commute) then have *: "(\v\S. (b v * x) *\<^sub>R v) = x *\<^sub>R a" for x using aeq by blast have "(\v \ S. dd v *\<^sub>R v) = 0" unfolding dd_def using S by (simp add: * wsumSS sum.distrib sum_subtractf algebra_simps) then have dd0: "dd v = 0" if "v \ S" for v using naff that \finite S\ sum_dd0 unfolding affine_dependent_explicit apply (simp only: not_ex) apply (drule_tac x=S in spec) apply (drule_tac x=dd in spec, simp) done consider "c' a \ c a" | "c a \ c' a" by linarith then show ?thesis proof cases case 1 then have "sum cc S \ sum cc' S" by (simp add: sumSS') then have le: "cc x \ cc' x" if "x \ S" for x using dd0 [OF that] 1 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc(a := c a)) x" by (simp add: c0 cc_def) show "0 \ (cc(a := c a)) a" by (simp add: c0) have "sum (cc(a := c a)) (insert a (T \ T')) = c a + sum (cc(a := c a)) (T \ T')" by (simp add: anot) also have "... = c a + sum (cc(a := c a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a + (1 - c a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc(a := c a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = c a *\<^sub>R a + (\x \ T \ T'. (cc(a := c a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c a *\<^sub>R a + (\x \ S. (cc(a := c a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc(a := c a)) x *\<^sub>R x) = x" by simp qed next case 2 then have "sum cc' S \ sum cc S" by (simp add: sumSS') then have le: "cc' x \ cc x" if "x \ S" for x using dd0 [OF that] 2 b0 mult_left_mono that by (fastforce simp add: dd_def algebra_simps) have cc0: "cc' x = 0" if "x \ S" "x \ T \ T'" for x using le [OF \x \ S\] that c'0 by (force simp: cc_def cc'_def split: if_split_asm) show ?thesis proof (simp add: convex_hull_finite, intro exI conjI) show "\x\T \ T'. 0 \ (cc'(a := c' a)) x" by (simp add: c'0 cc'_def) show "0 \ (cc'(a := c' a)) a" by (simp add: c'0) have "sum (cc'(a := c' a)) (insert a (T \ T')) = c' a + sum (cc'(a := c' a)) (T \ T')" by (simp add: anot) also have "... = c' a + sum (cc'(a := c' a)) S" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c' a + (1 - c' a)" by (metis \a \ S\ fun_upd_other sum.cong sumSS') finally show "sum (cc'(a := c' a)) (insert a (T \ T')) = 1" by simp have "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = c' a *\<^sub>R a + (\x \ T \ T'. (cc'(a := c' a)) x *\<^sub>R x)" by (simp add: anot) also have "... = c' a *\<^sub>R a + (\x \ S. (cc'(a := c' a)) x *\<^sub>R x)" apply simp apply (rule sum.mono_neutral_left) using \T \ S\ apply (auto simp: \a \ S\ cc0) done also have "... = c a *\<^sub>R a + x - c a *\<^sub>R a" by (simp add: wsumSS \a \ S\ if_smult sum_delta_notmem) finally show "(\x\insert a (T \ T'). (cc'(a := c' a)) x *\<^sub>R x) = x" by simp qed qed qed corollary\<^marker>\tag unimportant\ convex_hull_exchange_Int: fixes a :: "'a::euclidean_space" assumes "\ affine_dependent S" "a \ convex hull S" "T \ S" "T' \ S" shows "(convex hull (insert a T)) \ (convex hull (insert a T')) = convex hull (insert a (T \ T'))" apply (rule subset_antisym) using in_convex_hull_exchange_unique assms apply blast by (metis hull_mono inf_le1 inf_le2 insert_inter_insert le_inf_iff) lemma Int_closed_segment: fixes b :: "'a::euclidean_space" assumes "b \ closed_segment a c \ \ collinear{a,b,c}" shows "closed_segment a b \ closed_segment b c = {b}" proof (cases "c = a") case True then show ?thesis using assms collinear_3_eq_affine_dependent by fastforce next case False from assms show ?thesis proof assume "b \ closed_segment a c" moreover have "\ affine_dependent {a, c}" by (simp add: affine_independent_2) ultimately show ?thesis using False convex_hull_exchange_Int [of "{a,c}" b "{a}" "{c}"] by (simp add: segment_convex_hull insert_commute) next assume ncoll: "\ collinear {a, b, c}" have False if "closed_segment a b \ closed_segment b c \ {b}" proof - have "b \ closed_segment a b" and "b \ closed_segment b c" by auto with that obtain d where "b \ d" "d \ closed_segment a b" "d \ closed_segment b c" by force then have d: "collinear {a, d, b}" "collinear {b, d, c}" by (auto simp: between_mem_segment between_imp_collinear) have "collinear {a, b, c}" apply (rule collinear_3_trans [OF _ _ \b \ d\]) using d by (auto simp: insert_commute) with ncoll show False .. qed then show ?thesis by blast qed qed lemma affine_hull_finite_intersection_hyperplanes: fixes s :: "'a::euclidean_space set" obtains f where "finite f" "of_nat (card f) + aff_dim s = DIM('a)" "affine hull s = \f" "\h. h \ f \ \a b. a \ 0 \ h = {x. a \ x = b}" proof - obtain b where "b \ s" and indb: "\ affine_dependent b" and eq: "affine hull s = affine hull b" using affine_basis_exists by blast obtain c where indc: "\ affine_dependent c" and "b \ c" and affc: "affine hull c = UNIV" by (metis extend_to_affine_basis affine_UNIV hull_same indb subset_UNIV) then have "finite c" by (simp add: aff_independent_finite) then have fbc: "finite b" "card b \ card c" using \b \ c\ infinite_super by (auto simp: card_mono) have imeq: "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b)) = ((\a. affine hull (c - {a})) ` (c - b))" by blast have card1: "card ((\a. affine hull (c - {a})) ` (c - b)) = card (c - b)" apply (rule card_image [OF inj_onI]) by (metis Diff_eq_empty_iff Diff_iff indc affine_dependent_def hull_subset insert_iff) have card2: "(card (c - b)) + aff_dim s = DIM('a)" proof - have aff: "aff_dim (UNIV::'a set) = aff_dim c" by (metis aff_dim_affine_hull affc) have "aff_dim b = aff_dim s" by (metis (no_types) aff_dim_affine_hull eq) then have "int (card b) = 1 + aff_dim s" by (simp add: aff_dim_affine_independent indb) then show ?thesis using fbc aff by (simp add: \\ affine_dependent c\ \b \ c\ aff_dim_affine_independent aff_dim_UNIV card_Diff_subset of_nat_diff) qed show ?thesis proof (cases "c = b") case True show ?thesis apply (rule_tac f="{}" in that) using True affc apply (simp_all add: eq [symmetric]) by (metis aff_dim_UNIV aff_dim_affine_hull) next case False have ind: "\ affine_dependent (\a\c - b. c - {a})" by (rule affine_independent_subset [OF indc]) auto have affeq: "affine hull s = (\x\(\a. c - {a}) ` (c - b). affine hull x)" using \b \ c\ False apply (subst affine_hull_Inter [OF ind, symmetric]) apply (simp add: eq double_diff) done have *: "1 + aff_dim (c - {t}) = int (DIM('a))" if t: "t \ c" for t proof - have "insert t c = c" using t by blast then show ?thesis by (metis (full_types) add.commute aff_dim_affine_hull aff_dim_insert aff_dim_UNIV affc affine_dependent_def indc insert_Diff_single t) qed show ?thesis apply (rule_tac f = "(\x. affine hull x) ` ((\a. c - {a}) ` (c - b))" in that) using \finite c\ apply blast apply (simp add: imeq card1 card2) apply (simp add: affeq, clarify) apply (metis DIM_positive One_nat_def Suc_leI add_diff_cancel_left' of_nat_1 aff_dim_eq_hyperplane of_nat_diff *) done qed qed lemma affine_hyperplane_sums_eq_UNIV_0: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "0 \ S" and "w \ S" and "a \ w \ 0" shows "{x + y| x y. x \ S \ a \ y = 0} = UNIV" proof - have "subspace S" by (simp add: assms subspace_affine) have span1: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" apply (rule span_mono) using \0 \ S\ add.left_neutral by force have "w \ span {y. a \ y = 0}" using \a \ w \ 0\ span_induct subspace_hyperplane by auto moreover have "w \ span {x + y |x y. x \ S \ a \ y = 0}" using \w \ S\ by (metis (mono_tags, lifting) inner_zero_right mem_Collect_eq pth_d span_base) ultimately have span2: "span {y. a \ y = 0} \ span {x + y |x y. x \ S \ a \ y = 0}" by blast have "a \ 0" using assms inner_zero_left by blast then have "DIM('a) - 1 = dim {y. a \ y = 0}" by (simp add: dim_hyperplane) also have "... < dim {x + y |x y. x \ S \ a \ y = 0}" using span1 span2 by (blast intro: dim_psubset) finally have DIM_lt: "DIM('a) - 1 < dim {x + y |x y. x \ S \ a \ y = 0}" . have subs: "subspace {x + y| x y. x \ S \ a \ y = 0}" using subspace_sums [OF \subspace S\ subspace_hyperplane] by simp moreover have "span {x + y| x y. x \ S \ a \ y = 0} = UNIV" apply (rule dim_eq_full [THEN iffD1]) apply (rule antisym [OF dim_subset_UNIV]) using DIM_lt apply simp done ultimately show ?thesis by (simp add: subs) (metis (lifting) span_eq_iff subs) qed proposition\<^marker>\tag unimportant\ affine_hyperplane_sums_eq_UNIV: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {v. a \ v = b} \ {}" and "S - {v. a \ v = b} \ {}" shows "{x + y| x y. x \ S \ a \ y = b} = UNIV" proof (cases "a = 0") case True with assms show ?thesis by (auto simp: if_splits) next case False obtain c where "c \ S" and c: "a \ c = b" using assms by force with affine_diffs_subspace [OF \affine S\] have "subspace ((+) (- c) ` S)" by blast then have aff: "affine ((+) (- c) ` S)" by (simp add: subspace_imp_affine) have 0: "0 \ (+) (- c) ` S" by (simp add: \c \ S\) obtain d where "d \ S" and "a \ d \ b" and dc: "d-c \ (+) (- c) ` S" using assms by auto then have adc: "a \ (d - c) \ 0" by (simp add: c inner_diff_right) let ?U = "(+) (c+c) ` {x + y |x y. x \ (+) (- c) ` S \ a \ y = 0}" have "u + v \ (+) (c + c) ` {x + v |x v. x \ (+) (- c) ` S \ a \ v = 0}" if "u \ S" "b = a \ v" for u v apply (rule_tac x="u+v-c-c" in image_eqI) apply (simp_all add: algebra_simps) apply (rule_tac x="u-c" in exI) apply (rule_tac x="v-c" in exI) apply (simp add: algebra_simps that c) done moreover have "\a \ v = 0; u \ S\ \ \x ya. v + (u + c) = x + ya \ x \ S \ a \ ya = b" for v u by (metis add.left_commute c inner_right_distrib pth_d) ultimately have "{x + y |x y. x \ S \ a \ y = b} = ?U" by (fastforce simp: algebra_simps) also have "... = range ((+) (c + c))" by (simp only: affine_hyperplane_sums_eq_UNIV_0 [OF aff 0 dc adc]) also have "... = UNIV" by simp finally show ?thesis . qed lemma aff_dim_sums_Int_0: assumes "affine S" and "affine T" and "0 \ S" "0 \ T" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - have "0 \ {x + y |x y. x \ S \ y \ T}" using assms by force then have 0: "0 \ affine hull {x + y |x y. x \ S \ y \ T}" by (metis (lifting) hull_inc) have sub: "subspace S" "subspace T" using assms by (auto simp: subspace_affine) show ?thesis using dim_sums_Int [OF sub] by (simp add: aff_dim_zero assms 0 hull_inc) qed proposition aff_dim_sums_Int: assumes "affine S" and "affine T" and "S \ T \ {}" shows "aff_dim {x + y| x y. x \ S \ y \ T} = (aff_dim S + aff_dim T) - aff_dim(S \ T)" proof - obtain a where a: "a \ S" "a \ T" using assms by force have aff: "affine ((+) (-a) ` S)" "affine ((+) (-a) ` T)" using affine_translation [symmetric, of "- a"] assms by (simp_all cong: image_cong_simp) have zero: "0 \ ((+) (-a) ` S)" "0 \ ((+) (-a) ` T)" using a assms by auto have "{x + y |x y. x \ (+) (- a) ` S \ y \ (+) (- a) ` T} = (+) (- 2 *\<^sub>R a) ` {x + y| x y. x \ S \ y \ T}" by (force simp: algebra_simps scaleR_2) moreover have "(+) (- a) ` S \ (+) (- a) ` T = (+) (- a) ` (S \ T)" by auto ultimately show ?thesis using aff_dim_sums_Int_0 [OF aff zero] aff_dim_translation_eq by (metis (lifting)) qed lemma aff_dim_affine_Int_hyperplane: fixes a :: "'a::euclidean_space" assumes "affine S" shows "aff_dim(S \ {x. a \ x = b}) = (if S \ {v. a \ v = b} = {} then - 1 else if S \ {v. a \ v = b} then aff_dim S else aff_dim S - 1)" proof (cases "a = 0") case True with assms show ?thesis by auto next case False then have "aff_dim (S \ {x. a \ x = b}) = aff_dim S - 1" if "x \ S" "a \ x \ b" and non: "S \ {v. a \ v = b} \ {}" for x proof - have [simp]: "{x + y| x y. x \ S \ a \ y = b} = UNIV" using affine_hyperplane_sums_eq_UNIV [OF assms non] that by blast show ?thesis using aff_dim_sums_Int [OF assms affine_hyperplane non] by (simp add: of_nat_diff False) qed then show ?thesis by (metis (mono_tags, lifting) inf.orderE aff_dim_empty_eq mem_Collect_eq subsetI) qed lemma aff_dim_lt_full: fixes S :: "'a::euclidean_space set" shows "aff_dim S < DIM('a) \ (affine hull S \ UNIV)" by (metis (no_types) aff_dim_affine_hull aff_dim_le_DIM aff_dim_UNIV affine_hull_UNIV less_le) lemma aff_dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "affine T" "S \ {}" shows "aff_dim S = aff_dim T" proof - show ?thesis proof (rule order_antisym) show "aff_dim S \ aff_dim T" by (blast intro: aff_dim_subset [OF openin_imp_subset] ope) next obtain a where "a \ S" using \S \ {}\ by blast have "S \ T" using ope openin_imp_subset by auto then have "a \ T" using \a \ S\ by auto then have subT': "subspace ((\x. - a + x) ` T)" using affine_diffs_subspace \affine T\ by auto then obtain B where Bsub: "B \ ((\x. - a + x) ` T)" and po: "pairwise orthogonal B" and eq1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim ((\x. - a + x) ` T)" and spanB: "span B = ((\x. - a + x) ` T)" by (rule orthonormal_basis_subspace) auto obtain e where "0 < e" and e: "cball a e \ T \ S" by (meson \a \ S\ openin_contains_cball ope) have "aff_dim T = aff_dim ((\x. - a + x) ` T)" by (metis aff_dim_translation_eq) also have "... = dim ((\x. - a + x) ` T)" using aff_dim_subspace subT' by blast also have "... = card B" by (simp add: cardB) also have "... = card ((\x. e *\<^sub>R x) ` B)" using \0 < e\ by (force simp: inj_on_def card_image) also have "... \ dim ((\x. - a + x) ` S)" proof (simp, rule independent_card_le_dim) have e': "cball 0 e \ (\x. x - a) ` T \ (\x. x - a) ` S" using e by (auto simp: dist_norm norm_minus_commute subset_eq) have "(\x. e *\<^sub>R x) ` B \ cball 0 e \ (\x. x - a) ` T" using Bsub \0 < e\ eq1 subT' \a \ T\ by (auto simp: subspace_def) then show "(\x. e *\<^sub>R x) ` B \ (\x. x - a) ` S" using e' by blast show "independent ((\x. e *\<^sub>R x) ` B)" using linear_scale_self \independent B\ apply (rule linear_independent_injective_image) using \0 < e\ inj_on_def by fastforce qed also have "... = aff_dim S" using \a \ S\ aff_dim_eq_dim hull_inc by (force cong: image_cong_simp) finally show "aff_dim T \ aff_dim S" . qed qed lemma dim_openin: fixes S :: "'a::euclidean_space set" assumes ope: "openin (top_of_set T) S" and "subspace T" "S \ {}" shows "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis ope dim_subset openin_subset topspace_euclidean_subtopology) next have "dim T = aff_dim S" using aff_dim_openin by (metis aff_dim_subspace \subspace T\ \S \ {}\ ope subspace_affine) also have "... \ dim S" by (metis aff_dim_subset aff_dim_subspace dim_span span_superset subspace_span) finally show "dim T \ dim S" by simp qed subsection\Lower-dimensional affine subsets are nowhere dense\ proposition dense_complement_subspace: fixes S :: "'a :: euclidean_space set" assumes dim_less: "dim T < dim S" and "subspace S" shows "closure(S - T) = S" proof - have "closure(S - U) = S" if "dim U < dim S" "U \ S" for U proof - have "span U \ span S" by (metis neq_iff psubsetI span_eq_dim span_mono that) then obtain a where "a \ 0" "a \ span S" and a: "\y. y \ span U \ orthogonal a y" using orthogonal_to_subspace_exists_gen by metis show ?thesis proof have "closed S" by (simp add: \subspace S\ closed_subspace) then show "closure (S - U) \ S" by (simp add: closure_minimal) show "S \ closure (S - U)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" show "\y\S - U. dist y x < e" proof (cases "x \ U") case True let ?y = "x + (e/2 / norm a) *\<^sub>R a" show ?thesis proof show "dist ?y x < e" using \0 < e\ by (simp add: dist_norm) next have "?y \ S" by (metis \a \ span S\ \x \ S\ assms(2) span_eq_iff subspace_add subspace_scale) moreover have "?y \ U" proof - have "e/2 / norm a \ 0" using \0 < e\ \a \ 0\ by auto then show ?thesis by (metis True \a \ 0\ a orthogonal_scaleR orthogonal_self real_vector.scale_eq_0_iff span_add_eq span_base) qed ultimately show "?y \ S - U" by blast qed next case False with \0 < e\ \x \ S\ show ?thesis by force qed qed qed qed moreover have "S - S \ T = S-T" by blast moreover have "dim (S \ T) < dim S" by (metis dim_less dim_subset inf.cobounded2 inf.orderE inf.strict_boundedE not_le) ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ dense_complement_affine: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and "affine S" shows "closure(S - T) = S" proof (cases "S \ T = {}") case True then show ?thesis by (metis Diff_triv affine_hull_eq \affine S\ closure_same_affine_hull closure_subset hull_subset subset_antisym) next case False then obtain z where z: "z \ S \ T" by blast then have "subspace ((+) (- z) ` S)" by (meson IntD1 affine_diffs_subspace \affine S\) moreover have "int (dim ((+) (- z) ` T)) < int (dim ((+) (- z) ` S))" thm aff_dim_eq_dim using z less by (simp add: aff_dim_eq_dim_subtract [of z] hull_inc cong: image_cong_simp) ultimately have "closure(((+) (- z) ` S) - ((+) (- z) ` T)) = ((+) (- z) ` S)" by (simp add: dense_complement_subspace) then show ?thesis by (metis closure_translation translation_diff translation_invert) qed corollary\<^marker>\tag unimportant\ dense_complement_openin_affine_hull: fixes S :: "'a :: euclidean_space set" assumes less: "aff_dim T < aff_dim S" and ope: "openin (top_of_set (affine hull S)) S" shows "closure(S - T) = closure S" proof - have "affine hull S - T \ affine hull S" by blast then have "closure (S \ closure (affine hull S - T)) = closure (S \ (affine hull S - T))" by (rule closure_openin_Int_closure [OF ope]) then show ?thesis by (metis Int_Diff aff_dim_affine_hull affine_affine_hull dense_complement_affine hull_subset inf.orderE less) qed corollary\<^marker>\tag unimportant\ dense_complement_convex: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" shows "closure(S - T) = closure S" proof show "closure (S - T) \ closure S" by (simp add: closure_mono) have "closure (rel_interior S - T) = closure (rel_interior S)" apply (rule dense_complement_openin_affine_hull) apply (simp add: assms rel_interior_aff_dim) using \convex S\ rel_interior_rel_open rel_open by blast then show "closure S \ closure (S - T)" by (metis Diff_mono \convex S\ closure_mono convex_closure_rel_interior order_refl rel_interior_subset) qed corollary\<^marker>\tag unimportant\ dense_complement_convex_closed: fixes S :: "'a :: euclidean_space set" assumes "aff_dim T < aff_dim S" "convex S" "closed S" shows "closure(S - T) = S" by (simp add: assms dense_complement_convex) subsection\<^marker>\tag unimportant\\Parallel slices, etc\ text\ If we take a slice out of a set, we can do it perpendicularly, with the normal vector to the slice parallel to the affine hull.\ proposition\<^marker>\tag unimportant\ affine_parallel_slice: fixes S :: "'a :: euclidean_space set" assumes "affine S" and "S \ {x. a \ x \ b} \ {}" and "\ (S \ {x. a \ x \ b})" obtains a' b' where "a' \ 0" "S \ {x. a' \ x \ b'} = S \ {x. a \ x \ b}" "S \ {x. a' \ x = b'} = S \ {x. a \ x = b}" "\w. w \ S \ (w + a') \ S" proof (cases "S \ {x. a \ x = b} = {}") case True then obtain u v where "u \ S" "v \ S" "a \ u \ b" "a \ v > b" using assms by (auto simp: not_le) define \ where "\ = u + ((b - a \ u) / (a \ v - a \ u)) *\<^sub>R (v - u)" have "\ \ S" by (simp add: \_def \u \ S\ \v \ S\ \affine S\ mem_affine_3_minus) moreover have "a \ \ = b" using \a \ u \ b\ \b < a \ v\ by (simp add: \_def algebra_simps) (simp add: field_simps) ultimately have False using True by force then show ?thesis .. next case False then obtain z where "z \ S" and z: "a \ z = b" using assms by auto with affine_diffs_subspace [OF \affine S\] have sub: "subspace ((+) (- z) ` S)" by blast then have aff: "affine ((+) (- z) ` S)" and span: "span ((+) (- z) ` S) = ((+) (- z) ` S)" by (auto simp: subspace_imp_affine) obtain a' a'' where a': "a' \ span ((+) (- z) ` S)" and a: "a = a' + a''" and "\w. w \ span ((+) (- z) ` S) \ orthogonal a'' w" using orthogonal_subspace_decomp_exists [of "(+) (- z) ` S" "a"] by metis then have "\w. w \ S \ a'' \ (w-z) = 0" by (simp add: span_base orthogonal_def) then have a'': "\w. w \ S \ a'' \ w = (a - a') \ z" by (simp add: a inner_diff_right) then have ba'': "\w. w \ S \ a'' \ w = b - a' \ z" by (simp add: inner_diff_left z) have "\w. w \ (+) (- z) ` S \ (w + a') \ (+) (- z) ` S" by (metis subspace_add a' span_eq_iff sub) then have Sclo: "\w. w \ S \ (w + a') \ S" by fastforce show ?thesis proof (cases "a' = 0") case True with a assms True a'' diff_zero less_irrefl show ?thesis by auto next case False show ?thesis apply (rule_tac a' = "a'" and b' = "a' \ z" in that) apply (auto simp: a ba'' inner_left_distrib False Sclo) done qed qed lemma diffs_affine_hull_span: assumes "a \ S" shows "{x - a |x. x \ affine hull S} = span {x - a |x. x \ S}" proof - have *: "((\x. x - a) ` (S - {a})) = {x. x + a \ S} - {0}" by (auto simp: algebra_simps) show ?thesis apply (simp add: affine_hull_span2 [OF assms] *) apply (auto simp: algebra_simps) done qed lemma aff_dim_dim_affine_diffs: fixes S :: "'a :: euclidean_space set" assumes "affine S" "a \ S" shows "aff_dim S = dim {x - a |x. x \ S}" proof - obtain B where aff: "affine hull B = affine hull S" and ind: "\ affine_dependent B" and card: "of_nat (card B) = aff_dim S + 1" using aff_dim_basis_exists by blast then have "B \ {}" using assms by (metis affine_hull_eq_empty ex_in_conv) then obtain c where "c \ B" by auto then have "c \ S" by (metis aff affine_hull_eq \affine S\ hull_inc) have xy: "x - c = y - a \ y = x + 1 *\<^sub>R (a - c)" for x y c and a::'a by (auto simp: algebra_simps) have *: "{x - c |x. x \ S} = {x - a |x. x \ S}" apply safe apply (simp_all only: xy) using mem_affine_3_minus [OF \affine S\] \a \ S\ \c \ S\ apply blast+ done have affS: "affine hull S = S" by (simp add: \affine S\) have "aff_dim S = of_nat (card B) - 1" using card by simp also have "... = dim {x - c |x. x \ B}" by (simp add: affine_independent_card_dim_diffs [OF ind \c \ B\]) also have "... = dim {x - c | x. x \ affine hull B}" by (simp add: diffs_affine_hull_span \c \ B\ dim_span) also have "... = dim {x - a |x. x \ S}" by (simp add: affS aff *) finally show ?thesis . qed lemma aff_dim_linear_image_le: assumes "linear f" shows "aff_dim(f ` S) \ aff_dim S" proof - have "aff_dim (f ` T) \ aff_dim T" if "affine T" for T proof (cases "T = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False then obtain a where "a \ T" by auto have 1: "((\x. x - f a) ` f ` T) = {x - f a |x. x \ f ` T}" by auto have 2: "{x - f a| x. x \ f ` T} = f ` {x - a| x. x \ T}" by (force simp: linear_diff [OF assms]) have "aff_dim (f ` T) = int (dim {x - f a |x. x \ f ` T})" by (simp add: \a \ T\ hull_inc aff_dim_eq_dim [of "f a"] 1 cong: image_cong_simp) also have "... = int (dim (f ` {x - a| x. x \ T}))" by (force simp: linear_diff [OF assms] 2) also have "... \ int (dim {x - a| x. x \ T})" by (simp add: dim_image_le [OF assms]) also have "... \ aff_dim T" by (simp add: aff_dim_dim_affine_diffs [symmetric] \a \ T\ \affine T\) finally show ?thesis . qed then have "aff_dim (f ` (affine hull S)) \ aff_dim (affine hull S)" using affine_affine_hull [of S] by blast then show ?thesis using affine_hull_linear_image assms linear_conv_bounded_linear by fastforce qed lemma aff_dim_injective_linear_image [simp]: assumes "linear f" "inj f" shows "aff_dim (f ` S) = aff_dim S" proof (rule antisym) show "aff_dim (f ` S) \ aff_dim S" by (simp add: aff_dim_linear_image_le assms(1)) next obtain g where "linear g" "g \ f = id" using assms(1) assms(2) linear_injective_left_inverse by blast then have "aff_dim S \ aff_dim(g ` f ` S)" by (simp add: image_comp) also have "... \ aff_dim (f ` S)" by (simp add: \linear g\ aff_dim_linear_image_le) finally show "aff_dim S \ aff_dim (f ` S)" . qed lemma choose_affine_subset: assumes "affine S" "-1 \ d" and dle: "d \ aff_dim S" obtains T where "affine T" "T \ S" "aff_dim T = d" proof (cases "d = -1 \ S={}") case True with assms show ?thesis by (metis aff_dim_empty affine_empty bot.extremum that eq_iff) next case False with assms obtain a where "a \ S" "0 \ d" by auto with assms have ss: "subspace ((+) (- a) ` S)" by (simp add: affine_diffs_subspace_subtract cong: image_cong_simp) have "nat d \ dim ((+) (- a) ` S)" by (metis aff_dim_subspace aff_dim_translation_eq dle nat_int nat_mono ss) then obtain T where "subspace T" and Tsb: "T \ span ((+) (- a) ` S)" and Tdim: "dim T = nat d" using choose_subspace_of_subspace [of "nat d" "(+) (- a) ` S"] by blast then have "affine T" using subspace_affine by blast then have "affine ((+) a ` T)" by (metis affine_hull_eq affine_hull_translation) moreover have "(+) a ` T \ S" proof - have "T \ (+) (- a) ` S" by (metis (no_types) span_eq_iff Tsb ss) then show "(+) a ` T \ S" using add_ac by auto qed moreover have "aff_dim ((+) a ` T) = d" by (simp add: aff_dim_subspace Tdim \0 \ d\ \subspace T\ aff_dim_translation_eq) ultimately show ?thesis by (rule that) qed subsection\Paracompactness\ proposition paracompact: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "S \ \\" and opC: "\T. T \ \ \ open T" obtains \' where "S \ \ \'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. x \ S \ \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" proof (cases "S = {}") case True with that show ?thesis by blast next case False have "\T U. x \ U \ open U \ closure U \ T \ T \ \" if "x \ S" for x proof - obtain T where "x \ T" "T \ \" "open T" using assms \x \ S\ by blast then obtain e where "e > 0" "cball x e \ T" by (force simp: open_contains_cball) then show ?thesis apply (rule_tac x = T in exI) apply (rule_tac x = "ball x e" in exI) using \T \ \\ apply (simp add: closure_minimal) using closed_cball closure_minimal by blast qed then obtain F G where Gin: "x \ G x" and oG: "open (G x)" and clos: "closure (G x) \ F x" and Fin: "F x \ \" if "x \ S" for x by metis then obtain \ where "\ \ G ` S" "countable \" "\\ = \(G ` S)" using Lindelof [of "G ` S"] by (metis image_iff) then obtain K where K: "K \ S" "countable K" and eq: "\(G ` K) = \(G ` S)" by (metis countable_subset_image) with False Gin have "K \ {}" by force then obtain a :: "nat \ 'a" where "range a = K" by (metis range_from_nat_into \countable K\) then have odif: "\n. open (F (a n) - \{closure (G (a m)) |m. m < n})" using \K \ S\ Fin opC by (fastforce simp add:) let ?C = "range (\n. F(a n) - \{closure(G(a m)) |m. m < n})" have enum_S: "\n. x \ F(a n) \ x \ G(a n)" if "x \ S" for x proof - have "\y \ K. x \ G y" using eq that Gin by fastforce then show ?thesis using clos K \range a = K\ closure_subset by blast qed have 1: "S \ Union ?C" proof fix x assume "x \ S" define n where "n \ LEAST n. x \ F(a n)" have n: "x \ F(a n)" using enum_S [OF \x \ S\] by (force simp: n_def intro: LeastI) have notn: "x \ F(a m)" if "m < n" for m using that not_less_Least by (force simp: n_def) then have "x \ \{closure (G (a m)) |m. m < n}" using n \K \ S\ \range a = K\ clos notn by fastforce with n show "x \ Union ?C" by blast qed have 3: "\V. open V \ x \ V \ finite {U. U \ ?C \ (U \ V \ {})}" if "x \ S" for x proof - obtain n where n: "x \ F(a n)" "x \ G(a n)" using \x \ S\ enum_S by auto have "{U \ ?C. U \ G (a n) \ {}} \ (\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n" proof clarsimp fix k assume "(F (a k) - \{closure (G (a m)) |m. m < k}) \ G (a n) \ {}" then have "k \ n" by auto (metis closure_subset not_le subsetCE) then show "F (a k) - \{closure (G (a m)) |m. m < k} \ (\n. F (a n) - \{closure (G (a m)) |m. m < n}) ` {..n}" by force qed moreover have "finite ((\n. F(a n) - \{closure(G(a m)) |m. m < n}) ` atMost n)" by force ultimately have *: "finite {U \ ?C. U \ G (a n) \ {}}" using finite_subset by blast show ?thesis apply (rule_tac x="G (a n)" in exI) apply (intro conjI oG n *) using \K \ S\ \range a = K\ apply blast done qed show ?thesis apply (rule that [OF 1 _ 3]) using Fin \K \ S\ \range a = K\ apply (auto simp: odif) done qed corollary paracompact_closedin: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes cin: "closedin (top_of_set U) S" and oin: "\T. T \ \ \ openin (top_of_set U) T" and "S \ \\" obtains \' where "S \ \ \'" and "\V. V \ \' \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" and "\x. x \ U \ \V. openin (top_of_set U) V \ x \ V \ finite {X. X \ \' \ (X \ V \ {})}" proof - have "\Z. open Z \ (T = U \ Z)" if "T \ \" for T using oin [OF that] by (auto simp: openin_open) then obtain F where opF: "open (F T)" and intF: "U \ F T = T" if "T \ \" for T by metis obtain K where K: "closed K" "U \ K = S" using cin by (auto simp: closedin_closed) have 1: "U \ \(insert (- K) (F ` \))" by clarsimp (metis Int_iff Union_iff \U \ K = S\ \S \ \\\ subsetD intF) have 2: "\T. T \ insert (- K) (F ` \) \ open T" using \closed K\ by (auto simp: opF) obtain \ where "U \ \\" and D1: "\U. U \ \ \ open U \ (\T. T \ insert (- K) (F ` \) \ U \ T)" and D2: "\x. x \ U \ \V. open V \ x \ V \ finite {U \ \. U \ V \ {}}" by (blast intro: paracompact [OF 1 2]) let ?C = "{U \ V |V. V \ \ \ (V \ K \ {})}" show ?thesis proof (rule_tac \' = "{U \ V |V. V \ \ \ (V \ K \ {})}" in that) show "S \ \?C" using \U \ K = S\ \U \ \\\ K by (blast dest!: subsetD) show "\V. V \ ?C \ openin (top_of_set U) V \ (\T. T \ \ \ V \ T)" using D1 intF by fastforce have *: "{X. (\V. X = U \ V \ V \ \ \ V \ K \ {}) \ X \ (U \ V) \ {}} \ (\x. U \ x) ` {U \ \. U \ V \ {}}" for V by blast show "\V. openin (top_of_set U) V \ x \ V \ finite {X \ ?C. X \ V \ {}}" if "x \ U" for x using D2 [OF that] apply clarify apply (rule_tac x="U \ V" in exI) apply (auto intro: that finite_subset [OF *]) done qed qed corollary\<^marker>\tag unimportant\ paracompact_closed: fixes S :: "'a :: {metric_space,second_countable_topology} set" assumes "closed S" and opC: "\T. T \ \ \ open T" and "S \ \\" obtains \' where "S \ \\'" and "\U. U \ \' \ open U \ (\T. T \ \ \ U \ T)" and "\x. \V. open V \ x \ V \ finite {U. U \ \' \ (U \ V \ {})}" by (rule paracompact_closedin [of UNIV S \]) (auto simp: assms) subsection\<^marker>\tag unimportant\\Closed-graph characterization of continuity\ lemma continuous_closed_graph_gen: fixes T :: "'b::real_normed_vector set" assumes contf: "continuous_on S f" and fim: "f ` S \ T" shows "closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" proof - have eq: "((\x. Pair x (f x)) ` S) =(S \ T \ (\z. (f \ fst)z - snd z) -` {0})" using fim by auto show ?thesis apply (subst eq) apply (intro continuous_intros continuous_closedin_preimage continuous_on_subset [OF contf]) by auto qed lemma continuous_closed_graph_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" shows "continuous_on S f \ closedin (top_of_set (S \ T)) ((\x. Pair x (f x)) ` S)" (is "?lhs = ?rhs") proof - have "?lhs" if ?rhs proof (clarsimp simp add: continuous_on_closed_gen [OF fim]) fix U assume U: "closedin (top_of_set T) U" have eq: "(S \ f -` U) = fst ` (((\x. Pair x (f x)) ` S) \ (S \ U))" by (force simp: image_iff) show "closedin (top_of_set S) (S \ f -` U)" by (simp add: U closedin_Int closedin_Times closed_map_fst [OF \compact T\] that eq) qed with continuous_closed_graph_gen assms show ?thesis by blast qed lemma continuous_closed_graph: fixes f :: "'a::topological_space \ 'b::real_normed_vector" assumes "closed S" and contf: "continuous_on S f" shows "closed ((\x. Pair x (f x)) ` S)" apply (rule closedin_closed_trans) apply (rule continuous_closed_graph_gen [OF contf subset_UNIV]) by (simp add: \closed S\ closed_Times) lemma continuous_from_closed_graph: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact T" and fim: "f ` S \ T" and clo: "closed ((\x. Pair x (f x)) ` S)" shows "continuous_on S f" using fim clo by (auto intro: closed_subset simp: continuous_closed_graph_eq [OF \compact T\ fim]) lemma continuous_on_Un_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T f" shows "continuous_on (S \ T) f" using pasting_lemma [of "{S,T}" "top_of_set (S \ T)" id euclidean "\i. f" f] contf contg opS opT by (simp add: subtopology_subtopology) (metis inf.absorb2 openin_imp_subset) lemma continuous_on_cases_local_open: assumes opS: "openin (top_of_set (S \ T)) S" and opT: "openin (top_of_set (S \ T)) T" and contf: "continuous_on S f" and contg: "continuous_on T g" and fg: "\x. x \ S \ \P x \ x \ T \ P x \ f x = g x" shows "continuous_on (S \ T) (\x. if P x then f x else g x)" proof - have "\x. x \ S \ (if P x then f x else g x) = f x" "\x. x \ T \ (if P x then f x else g x) = g x" by (simp_all add: fg) then have "continuous_on S (\x. if P x then f x else g x)" "continuous_on T (\x. if P x then f x else g x)" by (simp_all add: contf contg cong: continuous_on_cong) then show ?thesis by (rule continuous_on_Un_local_open [OF opS opT]) qed lemma continuous_map_cases_le: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed lemma continuous_map_cases_lt: assumes contp: "continuous_map X euclideanreal p" and contq: "continuous_map X euclideanreal q" and contf: "continuous_map (subtopology X {x. x \ topspace X \ p x \ q x}) Y f" and contg: "continuous_map (subtopology X {x. x \ topspace X \ q x \ p x}) Y g" and fg: "\x. \x \ topspace X; p x = q x\ \ f x = g x" shows "continuous_map X Y (\x. if p x < q x then f x else g x)" proof - have "continuous_map X Y (\x. if q x - p x \ {0<..} then f x else g x)" proof (rule continuous_map_cases_function) show "continuous_map X euclideanreal (\x. q x - p x)" by (intro contp contq continuous_intros) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of {0<..}}) Y f" by (simp add: contf) show "continuous_map (subtopology X {x \ topspace X. q x - p x \ euclideanreal closure_of (topspace euclideanreal - {0<..})}) Y g" by (simp add: contg flip: Compl_eq_Diff_UNIV) qed (auto simp: fg) then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\The union of two collinear segments is another segment\ proposition\<^marker>\tag unimportant\ in_convex_hull_exchange: fixes a :: "'a::euclidean_space" assumes a: "a \ convex hull S" and xS: "x \ convex hull S" obtains b where "b \ S" "x \ convex hull (insert a (S - {b}))" proof (cases "a \ S") case True with xS insert_Diff that show ?thesis by fastforce next case False show ?thesis proof (cases "finite S \ card S \ Suc (DIM('a))") case True then obtain u where u0: "\i. i \ S \ 0 \ u i" and u1: "sum u S = 1" and ua: "(\i\S. u i *\<^sub>R i) = a" using a by (auto simp: convex_hull_finite) obtain v where v0: "\i. i \ S \ 0 \ v i" and v1: "sum v S = 1" and vx: "(\i\S. v i *\<^sub>R i) = x" using True xS by (auto simp: convex_hull_finite) show ?thesis proof (cases "\b. b \ S \ v b = 0") case True then obtain b where b: "b \ S" "v b = 0" by blast show ?thesis proof have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x = a then 0 else v x) = (\x \ S - {b}. if x = a then 0 else v x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x)" by (metis \v b = 0\ diff_zero sum.infinite sum_diff1 u1 zero_neq_one) finally show "(\x\insert a (S - {b}). if x = a then 0 else v x) = 1" by (simp add: v1) show "\x. x \ insert a (S - {b}) \ 0 \ (if x = a then 0 else v x)" by (auto simp: v0) have "(\x \ insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = (\x \ S - {b}. (if x = a then 0 else v x) *\<^sub>R x)" apply (rule sum.mono_neutral_right) using fin by auto also have "... = (\x \ S - {b}. v x *\<^sub>R x)" using b False by (auto intro!: sum.cong split: if_split_asm) also have "... = (\x\S. v x *\<^sub>R x)" by (metis (no_types, lifting) b(2) diff_zero fin finite.emptyI finite_Diff2 finite_insert scale_eq_0_iff sum_diff1) finally show "(\x\insert a (S - {b}). (if x = a then 0 else v x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) next case False have le_Max: "u i / v i \ Max ((\i. u i / v i) ` S)" if "i \ S" for i by (simp add: True that) have "Max ((\i. u i / v i) ` S) \ (\i. u i / v i) ` S" using True v1 by (auto intro: Max_in) then obtain b where "b \ S" and beq: "Max ((\b. u b / v b) ` S) = u b / v b" by blast then have "0 \ u b / v b" using le_Max beq divide_le_0_iff le_numeral_extra(2) sum_nonpos u1 by (metis False eq_iff v0) then have "0 < u b" "0 < v b" using False \b \ S\ u0 v0 by force+ have fin: "finite (insert a (S - {b}))" using sum.infinite v1 by fastforce show ?thesis proof show "x \ convex hull insert a (S - {b})" unfolding convex_hull_finite [OF fin] mem_Collect_eq proof (intro conjI exI ballI) have "(\x \ insert a (S - {b}). if x=a then v b / u b else v x - (v b / u b) * u x) = v b / u b + (\x \ S - {b}. v x - (v b / u b) * u x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = v b / u b + (\x \ S - {b}. v x) - (v b / u b) * (\x \ S - {b}. u x)" by (simp add: Groups_Big.sum_subtractf sum_distrib_left) also have "... = (\x\S. v x)" using \0 < u b\ True by (simp add: Groups_Big.sum_diff1 u1 field_simps) finally show "sum (\x. if x=a then v b / u b else v x - (v b / u b) * u x) (insert a (S - {b})) = 1" by (simp add: v1) show "0 \ (if i = a then v b / u b else v i - v b / u b * u i)" if "i \ insert a (S - {b})" for i using \0 < u b\ \0 < v b\ v0 [of i] le_Max [of i] beq that False by (auto simp: field_simps split: if_split_asm) have "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = (v b / u b) *\<^sub>R a + (\x\S - {b}. (v x - v b / u b * u x) *\<^sub>R x)" using \a \ S\ \b \ S\ True apply simp apply (rule sum.cong, auto) done also have "... = (v b / u b) *\<^sub>R a + (\x \ S - {b}. v x *\<^sub>R x) - (v b / u b) *\<^sub>R (\x \ S - {b}. u x *\<^sub>R x)" by (simp add: Groups_Big.sum_subtractf scaleR_left_diff_distrib sum_distrib_left scale_sum_right) also have "... = (\x\S. v x *\<^sub>R x)" using \0 < u b\ True by (simp add: ua vx Groups_Big.sum_diff1 algebra_simps) finally show "(\x\insert a (S - {b}). (if x=a then v b / u b else v x - v b / u b * u x) *\<^sub>R x) = x" by (simp add: vx) qed qed (rule \b \ S\) qed next case False obtain T where "finite T" "T \ S" and caT: "card T \ Suc (DIM('a))" and xT: "x \ convex hull T" using xS by (auto simp: caratheodory [of S]) with False obtain b where b: "b \ S" "b \ T" by (metis antisym subsetI) show ?thesis proof show "x \ convex hull insert a (S - {b})" using \T \ S\ b by (blast intro: subsetD [OF hull_mono xT]) qed (rule \b \ S\) qed qed lemma convex_hull_exchange_Union: fixes a :: "'a::euclidean_space" assumes "a \ convex hull S" shows "convex hull S = (\b \ S. convex hull (insert a (S - {b})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (blast intro: in_convex_hull_exchange [OF assms]) show "?rhs \ ?lhs" proof clarify fix x b assume"b \ S" "x \ convex hull insert a (S - {b})" then show "x \ convex hull S" if "b \ S" by (metis (no_types) that assms order_refl hull_mono hull_redundant insert_Diff_single insert_subset subsetCE) qed qed lemma Un_closed_segment: fixes a :: "'a::euclidean_space" assumes "b \ closed_segment a c" shows "closed_segment a b \ closed_segment b c = closed_segment a c" proof (cases "c = a") case True with assms show ?thesis by simp next case False with assms have "convex hull {a, b} \ convex hull {b, c} = (\ba\{a, c}. convex hull insert b ({a, c} - {ba}))" by (auto simp: insert_Diff_if insert_commute) then show ?thesis using convex_hull_exchange_Union by (metis assms segment_convex_hull) qed lemma Un_open_segment: fixes a :: "'a::euclidean_space" assumes "b \ open_segment a c" shows "open_segment a b \ {b} \ open_segment b c = open_segment a c" proof - have b: "b \ closed_segment a c" by (simp add: assms open_closed_segment) have *: "open_segment a c \ insert b (open_segment a b \ open_segment b c)" if "{b,c,a} \ open_segment a b \ open_segment b c = {c,a} \ open_segment a c" proof - have "insert a (insert c (insert b (open_segment a b \ open_segment b c))) = insert a (insert c (open_segment a c))" using that by (simp add: insert_commute) then show ?thesis by (metis (no_types) Diff_cancel Diff_eq_empty_iff Diff_insert2 open_segment_def) qed show ?thesis using Un_closed_segment [OF b] apply (simp add: closed_segment_eq_open) apply (rule equalityI) using assms apply (simp add: b subset_open_segment) using * by (simp add: insert_commute) qed subsection\Covering an open set by a countable chain of compact sets\ proposition open_Union_compact_subsets: fixes S :: "'a::euclidean_space set" assumes "open S" obtains C where "\n. compact(C n)" "\n. C n \ S" "\n. C n \ interior(C(Suc n))" "\(range C) = S" "\K. \compact K; K \ S\ \ \N. \n\N. K \ (C n)" proof (cases "S = {}") case True then show ?thesis by (rule_tac C = "\n. {}" in that) auto next case False then obtain a where "a \ S" by auto let ?C = "\n. cball a (real n) - (\x \ -S. \e \ ball 0 (1 / real(Suc n)). {x + e})" have "\N. \n\N. K \ (f n)" if "\n. compact(f n)" and sub_int: "\n. f n \ interior (f(Suc n))" and eq: "\(range f) = S" and "compact K" "K \ S" for f K proof - have *: "\n. f n \ (\n. interior (f n))" by (meson Sup_upper2 UNIV_I \\n. f n \ interior (f (Suc n))\ image_iff) have mono: "\m n. m \ n \f m \ f n" by (meson dual_order.trans interior_subset lift_Suc_mono_le sub_int) obtain I where "finite I" and I: "K \ (\i\I. interior (f i))" proof (rule compactE_image [OF \compact K\]) show "K \ (\n. interior (f n))" using \K \ S\ \\(f ` UNIV) = S\ * by blast qed auto { fix n assume n: "Max I \ n" have "(\i\I. interior (f i)) \ f n" by (rule UN_least) (meson dual_order.trans interior_subset mono I Max_ge [OF \finite I\] n) then have "K \ f n" using I by auto } then show ?thesis by blast qed moreover have "\f. (\n. compact(f n)) \ (\n. (f n) \ S) \ (\n. (f n) \ interior(f(Suc n))) \ ((\(range f) = S))" proof (intro exI conjI allI) show "\n. compact (?C n)" by (auto simp: compact_diff open_sums) show "\n. ?C n \ S" by auto show "?C n \ interior (?C (Suc n))" for n proof (simp add: interior_diff, rule Diff_mono) show "cball a (real n) \ ball a (1 + real n)" by (simp add: cball_subset_ball_iff) have cl: "closed (\x\- S. \e\cball 0 (1 / (2 + real n)). {x + e})" using assms by (auto intro: closed_compact_sums) have "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \e \ cball 0 (1 / (2 + real n)). {x + e})" by (intro closure_minimal UN_mono ball_subset_cball order_refl cl) also have "... \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" apply (intro UN_mono order_refl) apply (simp add: cball_subset_ball_iff field_split_simps) done finally show "closure (\x\- S. \y\ball 0 (1 / (2 + real n)). {x + y}) \ (\x \ -S. \y\ball 0 (1 / (1 + real n)). {x + y})" . qed have "S \ \ (range ?C)" proof fix x assume x: "x \ S" then obtain e where "e > 0" and e: "ball x e \ S" using assms open_contains_ball by blast then obtain N1 where "N1 > 0" and N1: "real N1 > 1/e" using reals_Archimedean2 by (metis divide_less_0_iff less_eq_real_def neq0_conv not_le of_nat_0 of_nat_1 of_nat_less_0_iff) obtain N2 where N2: "norm(x - a) \ real N2" by (meson real_arch_simple) have N12: "inverse((N1 + N2) + 1) \ inverse(N1)" using \N1 > 0\ by (auto simp: field_split_simps) have "x \ y + z" if "y \ S" "norm z < 1 / (1 + (real N1 + real N2))" for y z proof - have "e * real N1 < e * (1 + (real N1 + real N2))" by (simp add: \0 < e\) then have "1 / (1 + (real N1 + real N2)) < e" using N1 \e > 0\ by (metis divide_less_eq less_trans mult.commute of_nat_add of_nat_less_0_iff of_nat_Suc) then have "x - z \ ball x e" using that by simp then have "x - z \ S" using e by blast with that show ?thesis by auto qed with N2 show "x \ \ (range ?C)" by (rule_tac a = "N1+N2" in UN_I) (auto simp: dist_norm norm_minus_commute) qed then show "\ (range ?C) = S" by auto qed ultimately show ?thesis using that by metis qed subsection\Orthogonal complement\ definition\<^marker>\tag important\ orthogonal_comp ("_\<^sup>\" [80] 80) where "orthogonal_comp W \ {x. \y \ W. orthogonal y x}" proposition subspace_orthogonal_comp: "subspace (W\<^sup>\)" unfolding subspace_def orthogonal_comp_def orthogonal_def by (auto simp: inner_right_distrib) lemma orthogonal_comp_anti_mono: assumes "A \ B" shows "B\<^sup>\ \ A\<^sup>\" proof fix x assume x: "x \ B\<^sup>\" show "x \ orthogonal_comp A" using x unfolding orthogonal_comp_def by (simp add: orthogonal_def, metis assms in_mono) qed lemma orthogonal_comp_null [simp]: "{0}\<^sup>\ = UNIV" by (auto simp: orthogonal_comp_def orthogonal_def) lemma orthogonal_comp_UNIV [simp]: "UNIV\<^sup>\ = {0}" unfolding orthogonal_comp_def orthogonal_def by auto (use inner_eq_zero_iff in blast) lemma orthogonal_comp_subset: "U \ U\<^sup>\\<^sup>\" by (auto simp: orthogonal_comp_def orthogonal_def inner_commute) lemma subspace_sum_minimal: assumes "S \ U" "T \ U" "subspace U" shows "S + T \ U" proof fix x assume "x \ S + T" then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" by (meson set_plus_elim) then show "x \ U" by (meson assms subsetCE subspace_add) qed proposition subspace_sum_orthogonal_comp: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U + U\<^sup>\ = UNIV" proof - obtain B where "B \ U" and ortho: "pairwise orthogonal B" "\x. x \ B \ norm x = 1" and "independent B" "card B = dim U" "span B = U" using orthonormal_basis_subspace [OF assms] by metis then have "finite B" by (simp add: indep_card_eq_dim_span) have *: "\x\B. \y\B. x \ y = (if x=y then 1 else 0)" using ortho norm_eq_1 by (auto simp: orthogonal_def pairwise_def) { fix v let ?u = "\b\B. (v \ b) *\<^sub>R b" have "v = ?u + (v - ?u)" by simp moreover have "?u \ U" by (metis (no_types, lifting) \span B = U\ assms subspace_sum span_base span_mul) moreover have "(v - ?u) \ U\<^sup>\" proof (clarsimp simp: orthogonal_comp_def orthogonal_def) fix y assume "y \ U" with \span B = U\ span_finite [OF \finite B\] obtain u where u: "y = (\b\B. u b *\<^sub>R b)" by auto have "b \ (v - ?u) = 0" if "b \ B" for b using that \finite B\ by (simp add: * algebra_simps inner_sum_right if_distrib [of "(*)v" for v] inner_commute cong: if_cong) then show "y \ (v - ?u) = 0" by (simp add: u inner_sum_left) qed ultimately have "v \ U + U\<^sup>\" using set_plus_intro by fastforce } then show ?thesis by auto qed lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}" using orthogonal_comp_def orthogonal_self by (force simp: assms subspace_0 subspace_orthogonal_comp) lemma orthogonal_comp_self: fixes U :: "'a :: euclidean_space set" assumes "subspace U" shows "U\<^sup>\\<^sup>\ = U" proof have ssU': "subspace (U\<^sup>\)" by (simp add: subspace_orthogonal_comp) have "u \ U" if "u \ U\<^sup>\\<^sup>\" for u proof - obtain v w where "u = v+w" "v \ U" "w \ U\<^sup>\" using subspace_sum_orthogonal_comp [OF assms] set_plus_elim by blast then have "u-v \ U\<^sup>\" by simp moreover have "v \ U\<^sup>\\<^sup>\" using \v \ U\ orthogonal_comp_subset by blast then have "u-v \ U\<^sup>\\<^sup>\" by (simp add: subspace_diff subspace_orthogonal_comp that) ultimately have "u-v = 0" using orthogonal_Int_0 ssU' by blast with \v \ U\ show ?thesis by auto qed then show "U\<^sup>\\<^sup>\ \ U" by auto qed (use orthogonal_comp_subset in auto) lemma ker_orthogonal_comp_adjoint: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "f -` {0} = (range (adjoint f))\<^sup>\" apply (auto simp: orthogonal_comp_def orthogonal_def) apply (simp add: adjoint_works assms(1) inner_commute) by (metis adjoint_works all_zero_iff assms(1) inner_commute) subsection\<^marker>\tag unimportant\ \A non-injective linear function maps into a hyperplane.\ lemma linear_surj_adj_imp_inj: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" "surj (adjoint f)" shows "inj f" proof - have "\x. y = adjoint f x" for y using assms by (simp add: surjD) then show "inj f" using assms unfolding inj_on_def image_def by (metis (no_types) adjoint_works euclidean_eqI) qed \ \\<^url>\https://mathonline.wikidot.com/injectivity-and-surjectivity-of-the-adjoint-of-a-linear-map\\ lemma surj_adjoint_iff_inj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "surj (adjoint f) \ inj f" proof assume "surj (adjoint f)" then show "inj f" by (simp add: assms linear_surj_adj_imp_inj) next assume "inj f" have "f -` {0} = {0}" using assms \inj f\ linear_0 linear_injective_0 by fastforce moreover have "f -` {0} = range (adjoint f)\<^sup>\" by (intro ker_orthogonal_comp_adjoint assms) ultimately have "range (adjoint f)\<^sup>\\<^sup>\ = UNIV" by (metis orthogonal_comp_null) then show "surj (adjoint f)" using adjoint_linear \linear f\ by (subst (asm) orthogonal_comp_self) (simp add: adjoint_linear linear_subspace_image) qed lemma inj_adjoint_iff_surj [simp]: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes "linear f" shows "inj (adjoint f) \ surj f" proof assume "inj (adjoint f)" have "(adjoint f) -` {0} = {0}" by (metis \inj (adjoint f)\ adjoint_linear assms surj_adjoint_iff_inj ker_orthogonal_comp_adjoint orthogonal_comp_UNIV) then have "(range(f))\<^sup>\ = {0}" by (metis (no_types, hide_lams) adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint set_zero) then show "surj f" by (metis \inj (adjoint f)\ adjoint_adjoint adjoint_linear assms surj_adjoint_iff_inj) next assume "surj f" then have "range f = (adjoint f -` {0})\<^sup>\" by (simp add: adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint) then have "{0} = adjoint f -` {0}" using \surj f\ adjoint_adjoint adjoint_linear assms ker_orthogonal_comp_adjoint by force then show "inj (adjoint f)" by (simp add: \surj f\ adjoint_adjoint adjoint_linear assms linear_surj_adj_imp_inj) qed lemma linear_singular_into_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" shows "\ inj f \ (\a. a \ 0 \ (\x. a \ f x = 0))" (is "_ = ?rhs") proof assume "\inj f" then show ?rhs using all_zero_iff by (metis (no_types, hide_lams) adjoint_clauses(2) adjoint_linear assms linear_injective_0 linear_injective_imp_surjective linear_surj_adj_imp_inj) next assume ?rhs then show "\inj f" by (metis assms linear_injective_isomorphism all_zero_iff) qed lemma linear_singular_image_hyperplane: fixes f :: "'n::euclidean_space \ 'n" assumes "linear f" "\inj f" obtains a where "a \ 0" "\S. f ` S \ {x. a \ x = 0}" using assms by (fastforce simp add: linear_singular_into_hyperplane) end diff --git a/src/HOL/Analysis/Topology_Euclidean_Space.thy b/src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy @@ -1,2509 +1,2508 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) chapter \Vector Analysis\ theory Topology_Euclidean_Space imports Elementary_Normed_Spaces Linear_Algebra Norm_Arith begin section \Elementary Topology in Euclidean Space\ lemma euclidean_dist_l2: fixes x y :: "'a :: euclidean_space" shows "dist x y = L2_set (\i. dist (x \ i) (y \ i)) Basis" unfolding dist_norm norm_eq_sqrt_inner L2_set_def by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) lemma norm_nth_le: "norm (x \ i) \ norm x" if "i \ Basis" proof - have "(x \ i)\<^sup>2 = (\i\{i}. (x \ i)\<^sup>2)" by simp also have "\ \ (\i\Basis. (x \ i)\<^sup>2)" by (intro sum_mono2) (auto simp: that) finally show ?thesis unfolding norm_conv_dist euclidean_dist_l2[of x] L2_set_def by (auto intro!: real_le_rsqrt) qed subsection\<^marker>\tag unimportant\ \Continuity of the representation WRT an orthogonal basis\ lemma orthogonal_Basis: "pairwise orthogonal Basis" by (simp add: inner_not_same_Basis orthogonal_def pairwise_def) lemma representation_bound: fixes B :: "'N::real_inner set" assumes "finite B" "independent B" "b \ B" and orth: "pairwise orthogonal B" obtains m where "m > 0" "\x. x \ span B \ \representation B x b\ \ m * norm x" proof fix x assume x: "x \ span B" have "b \ 0" using \independent B\ \b \ B\ dependent_zero by blast have [simp]: "b \ b' = (if b' = b then (norm b)\<^sup>2 else 0)" if "b \ B" "b' \ B" for b b' using orth by (simp add: orthogonal_def pairwise_def norm_eq_sqrt_inner that) have "norm x = norm (\b\B. representation B x b *\<^sub>R b)" using real_vector.sum_representation_eq [OF \independent B\ x \finite B\ order_refl] by simp also have "\ = sqrt ((\b\B. representation B x b *\<^sub>R b) \ (\b\B. representation B x b *\<^sub>R b))" by (simp add: norm_eq_sqrt_inner) also have "\ = sqrt (\b\B. (representation B x b *\<^sub>R b) \ (representation B x b *\<^sub>R b))" using \finite B\ by (simp add: inner_sum_left inner_sum_right if_distrib [of "\x. _ * x"] cong: if_cong sum.cong_simp) also have "\ = sqrt (\b\B. (norm (representation B x b *\<^sub>R b))\<^sup>2)" by (simp add: mult.commute mult.left_commute power2_eq_square) also have "\ = sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" by (simp add: norm_mult power_mult_distrib) finally have "norm x = sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" . moreover have "sqrt ((representation B x b)\<^sup>2 * (norm b)\<^sup>2) \ sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" using \b \ B\ \finite B\ by (auto intro: member_le_sum) then have "\representation B x b\ \ (1 / norm b) * sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" using \b \ 0\ by (simp add: field_split_simps real_sqrt_mult del: real_sqrt_le_iff) ultimately show "\representation B x b\ \ (1 / norm b) * norm x" by simp next show "0 < 1 / norm b" using \independent B\ \b \ B\ dependent_zero by auto qed lemma continuous_on_representation: fixes B :: "'N::euclidean_space set" assumes "finite B" "independent B" "b \ B" "pairwise orthogonal B" shows "continuous_on (span B) (\x. representation B x b)" proof show "\d>0. \x'\span B. dist x' x < d \ dist (representation B x' b) (representation B x b) \ e" if "e > 0" "x \ span B" for x e proof - obtain m where "m > 0" and m: "\x. x \ span B \ \representation B x b\ \ m * norm x" using assms representation_bound by blast show ?thesis unfolding dist_norm proof (intro exI conjI ballI impI) show "e/m > 0" by (simp add: \e > 0\ \m > 0\) show "norm (representation B x' b - representation B x b) \ e" if x': "x' \ span B" and less: "norm (x'-x) < e/m" for x' proof - have "\representation B (x'-x) b\ \ m * norm (x'-x)" using m [of "x'-x"] \x \ span B\ span_diff x' by blast also have "\ < e" by (metis \m > 0\ less mult.commute pos_less_divide_eq) finally have "\representation B (x'-x) b\ \ e" by simp then show ?thesis by (simp add: \x \ span B\ \independent B\ representation_diff x') qed qed qed qed subsection\<^marker>\tag unimportant\\Balls in Euclidean Space\ lemma cball_subset_cball_iff: fixes a :: "'a :: euclidean_space" shows "cball a r \ cball a' r' \ dist a a' + r \ r' \ r < 0" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next case False then have [simp]: "r \ 0" by simp have "norm (a - a') + r \ r'" proof (cases "a = a'") case True then show ?thesis using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] by (force simp: SOME_Basis dist_norm) next case False have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" by (simp add: algebra_simps) also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" by (simp add: algebra_simps) also from \a \ a'\ have "... = \- norm (a - a') - r\" by simp (simp add: field_simps) finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \norm (a - a') + r\" by linarith from \a \ a'\ show ?thesis using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] by (simp add: dist_norm scaleR_add_left) qed then show ?rhs by (simp add: dist_norm) qed qed metric lemma cball_subset_ball_iff: "cball a r \ ball a' r' \ dist a a' + r < r' \ r < 0" (is "?lhs \ ?rhs") for a :: "'a::euclidean_space" proof assume ?lhs then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next case False then have [simp]: "r \ 0" by simp have "norm (a - a') + r < r'" proof (cases "a = a'") case True then show ?thesis using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] by (force simp: SOME_Basis dist_norm) next case False have False if "norm (a - a') + r \ r'" proof - from that have "\r' - norm (a - a')\ \ r" by (simp split: abs_split) (metis \0 \ r\ \?lhs\ centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) then show ?thesis using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ apply (simp add: dist_norm) apply (simp add: scaleR_left_diff_distrib) apply (simp add: field_simps) done qed then show ?thesis by force qed then show ?rhs by (simp add: dist_norm) qed next assume ?rhs then show ?lhs by metric qed lemma ball_subset_cball_iff: "ball a r \ cball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") for a :: "'a::euclidean_space" proof (cases "r \ 0") case True then show ?thesis by metric next case False show ?thesis proof assume ?lhs then have "(cball a r \ cball a' r')" by (metis False closed_cball closure_ball closure_closed closure_mono not_less) with False show ?rhs by (fastforce iff: cball_subset_cball_iff) next assume ?rhs with False show ?lhs by metric qed qed lemma ball_subset_ball_iff: fixes a :: "'a :: euclidean_space" shows "ball a r \ ball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True then show ?thesis by metric next case False show ?thesis proof assume ?lhs then have "0 < r'" using False by metric then have "(cball a r \ cball a' r')" by (metis False\?lhs\ closure_ball closure_mono not_less) then show ?rhs using False cball_subset_cball_iff by fastforce qed metric qed lemma ball_eq_ball_iff: fixes x :: "'a :: euclidean_space" shows "ball x d = ball y e \ d \ 0 \ e \ 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (cases "d \ 0 \ e \ 0") case True with \?lhs\ show ?rhs by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) next case False with \?lhs\ show ?rhs apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next assume ?rhs then show ?lhs by (auto simp: set_eq_subset ball_subset_ball_iff) qed lemma cball_eq_cball_iff: fixes x :: "'a :: euclidean_space" shows "cball x d = cball y e \ d < 0 \ e < 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (cases "d < 0 \ e < 0") case True with \?lhs\ show ?rhs by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) next case False with \?lhs\ show ?rhs apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next assume ?rhs then show ?lhs by (auto simp: set_eq_subset cball_subset_cball_iff) qed lemma ball_eq_cball_iff: fixes x :: "'a :: euclidean_space" shows "ball x d = cball y e \ d \ 0 \ e < 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ done next assume ?rhs then show ?lhs by auto qed lemma cball_eq_ball_iff: fixes x :: "'a :: euclidean_space" shows "cball x d = ball y e \ d < 0 \ e \ 0" using ball_eq_cball_iff by blast lemma finite_ball_avoid: fixes S :: "'a :: euclidean_space set" assumes "open S" "finite X" "p \ S" shows "\e>0. \w\ball p e. w\S \ (w\p \ w\X)" proof - obtain e1 where "0 < e1" and e1_b:"ball p e1 \ S" using open_contains_ball_eq[OF \open S\] assms by auto obtain e2 where "0 < e2" and "\x\X. x \ p \ e2 \ dist p x" using finite_set_avoid[OF \finite X\,of p] by auto hence "\w\ball p (min e1 e2). w\S \ (w\p \ w\X)" using e1_b by auto thus "\e>0. \w\ball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ \e1>0\ apply (rule_tac x="min e1 e2" in exI) by auto qed lemma finite_cball_avoid: fixes S :: "'a :: euclidean_space set" assumes "open S" "finite X" "p \ S" shows "\e>0. \w\cball p e. w\S \ (w\p \ w\X)" proof - obtain e1 where "e1>0" and e1: "\w\ball p e1. w\S \ (w\p \ w\X)" using finite_ball_avoid[OF assms] by auto define e2 where "e2 \ e1/2" have "e2>0" and "e2 < e1" unfolding e2_def using \e1>0\ by auto then have "cball p e2 \ ball p e1" by (subst cball_subset_ball_iff,auto) then show "\e>0. \w\cball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ e1 by auto qed lemma dim_cball: assumes "e > 0" shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)" proof - { fix x :: "'n::euclidean_space" define y where "y = (e / norm x) *\<^sub>R x" then have "y \ cball 0 e" using assms by auto moreover have *: "x = (norm x / e) *\<^sub>R y" using y_def assms by simp moreover from * have "x = (norm x/e) *\<^sub>R y" by auto ultimately have "x \ span (cball 0 e)" using span_scale[of y "cball 0 e" "norm x/e"] span_superset[of "cball 0 e"] by (simp add: span_base) } then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" by auto then show ?thesis - using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp: dim_UNIV) + using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto) qed subsection \Boxes\ abbreviation\<^marker>\tag important\ One :: "'a::euclidean_space" where "One \ \Basis" lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False proof - have "dependent (Basis :: 'a set)" apply (simp add: dependent_finite) apply (rule_tac x="\i. 1" in exI) using SOME_Basis apply (auto simp: assms) done with independent_Basis show False by force qed corollary\<^marker>\tag unimportant\ One_neq_0[iff]: "One \ 0" by (metis One_non_0) corollary\<^marker>\tag unimportant\ Zero_neq_One[iff]: "0 \ One" by (metis One_non_0) definition\<^marker>\tag important\ (in euclidean_space) eucl_less (infix " (\i\Basis. a \ i < b \ i)" definition\<^marker>\tag important\ box_eucl_less: "box a b = {x. a x \tag important\ "cbox a b = {x. \i\Basis. a \ i \ x \ i \ x \ i \ b \ i}" lemma box_def: "box a b = {x. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" and in_box_eucl_less: "x \ box a b \ a x box a b \ (\i\Basis. a \ i < x \ i \ x \ i < b \ i)" "x \ cbox a b \ (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i)" by (auto simp: box_eucl_less eucl_less_def cbox_def) lemma cbox_Pair_eq: "cbox (a, c) (b, d) = cbox a b \ cbox c d" by (force simp: cbox_def Basis_prod_def) lemma cbox_Pair_iff [iff]: "(x, y) \ cbox (a, c) (b, d) \ x \ cbox a b \ y \ cbox c d" by (force simp: cbox_Pair_eq) lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (cbox a b \ cbox c d)" apply (auto simp: cbox_def Basis_complex_def) apply (rule_tac x = "(Re x, Im x)" in image_eqI) using complex_eq by auto lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \ cbox a b = {} \ cbox c d = {}" by (force simp: cbox_Pair_eq) lemma swap_cbox_Pair [simp]: "prod.swap ` cbox (c, a) (d, b) = cbox (a,c) (b,d)" by auto lemma mem_box_real[simp]: "(x::real) \ box a b \ a < x \ x < b" "(x::real) \ cbox a b \ a \ x \ x \ b" by (auto simp: mem_box) lemma box_real[simp]: fixes a b:: real shows "box a b = {a <..< b}" "cbox a b = {a .. b}" by auto lemma box_Int_box: fixes a :: "'a::euclidean_space" shows "box a b \ box c d = box (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" unfolding set_eq_iff and Int_iff and mem_box by auto lemma rational_boxes: fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \) \ x \ box a b \ box a b \ ball x e" proof - define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" - using assms by (auto simp: DIM_positive) + using assms by (auto) have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] obtain a where a: "\xa. a xa \ \ \ a xa < x \ xa \ x \ xa - a xa < e'" .. have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] obtain b where b: "\xa. b xa \ \ \ x \ xa < b xa \ b xa - x \ xa < e'" .. let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" show ?thesis proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) fix y :: 'a assume *: "y \ box ?a ?b" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) also have "\ < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule sum_strict_mono) fix i :: "'a" assume i: "i \ Basis" have "a i < y\i \ y\i < b i" using * i by (auto simp: box_def) moreover have "a i < x\i" "x\i - a i < e'" using a by auto moreover have "x\i < b i" "b i - x\i < e'" using b by auto ultimately have "\x\i - y\i\ < 2 * e'" by auto then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) then have "(dist (x \ i) (y \ i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" by (rule power_strict_mono) auto then show "(dist (x \ i) (y \ i))\<^sup>2 < e\<^sup>2 / real DIM('a)" by (simp add: power_divide) qed auto also have "\ = e" using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) qed (insert a b, auto simp: box_def) qed lemma open_UNION_box: fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" defines "I \ {f\Basis \\<^sub>E \ \ \. box (a' f) (b' f) \ M}" shows "M = (\f\I. box (a' f) (b' f))" proof - have "x \ (\f\I. box (a' f) (b' f))" if "x \ M" for x proof - obtain e where e: "e > 0" "ball x e \ M" using openE[OF \open M\ \x \ M\] by auto moreover obtain a b where ab: "x \ box a b" "\i \ Basis. a \ i \ \" "\i\Basis. b \ i \ \" "box a b \ ball x e" using rational_boxes[OF e(1)] by metis ultimately show ?thesis by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) (auto simp: euclidean_representation I_def a'_def b'_def) qed then show ?thesis by (auto simp: I_def) qed corollary open_countable_Union_open_box: fixes S :: "'a :: euclidean_space set" assumes "open S" obtains \ where "countable \" "\ \ Pow S" "\X. X \ \ \ \a b. X = box a b" "\\ = S" proof - let ?a = "\f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" let ?b = "\f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" let ?I = "{f\Basis \\<^sub>E \ \ \. box (?a f) (?b f) \ S}" let ?\ = "(\f. box (?a f) (?b f)) ` ?I" show ?thesis proof have "countable ?I" by (simp add: countable_PiE countable_rat) then show "countable ?\" by blast show "\?\ = S" using open_UNION_box [OF assms] by metis qed auto qed lemma rational_cboxes: fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \) \ x \ cbox a b \ cbox a b \ ball x e" proof - define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" using assms by auto have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] obtain a where a: "\u. a u \ \ \ a u < x \ u \ x \ u - a u < e'" .. have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] obtain b where b: "\u. b u \ \ \ x \ u < b u \ b u - x \ u < e'" .. let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" show ?thesis proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) fix y :: 'a assume *: "y \ cbox ?a ?b" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) also have "\ < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule sum_strict_mono) fix i :: "'a" assume i: "i \ Basis" have "a i \ y\i \ y\i \ b i" using * i by (auto simp: cbox_def) moreover have "a i < x\i" "x\i - a i < e'" using a by auto moreover have "x\i < b i" "b i - x\i < e'" using b by auto ultimately have "\x\i - y\i\ < 2 * e'" by auto then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) then have "(dist (x \ i) (y \ i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" by (rule power_strict_mono) auto then show "(dist (x \ i) (y \ i))\<^sup>2 < e\<^sup>2 / real DIM('a)" by (simp add: power_divide) qed auto also have "\ = e" using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) next show "x \ cbox (\i\Basis. a i *\<^sub>R i) (\i\Basis. b i *\<^sub>R i)" using a b less_imp_le by (auto simp: cbox_def) qed (use a b cbox_def in auto) qed lemma open_UNION_cbox: fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" defines "I \ {f\Basis \\<^sub>E \ \ \. cbox (a' f) (b' f) \ M}" shows "M = (\f\I. cbox (a' f) (b' f))" proof - have "x \ (\f\I. cbox (a' f) (b' f))" if "x \ M" for x proof - obtain e where e: "e > 0" "ball x e \ M" using openE[OF \open M\ \x \ M\] by auto moreover obtain a b where ab: "x \ cbox a b" "\i \ Basis. a \ i \ \" "\i \ Basis. b \ i \ \" "cbox a b \ ball x e" using rational_cboxes[OF e(1)] by metis ultimately show ?thesis by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) (auto simp: euclidean_representation I_def a'_def b'_def) qed then show ?thesis by (auto simp: I_def) qed corollary open_countable_Union_open_cbox: fixes S :: "'a :: euclidean_space set" assumes "open S" obtains \ where "countable \" "\ \ Pow S" "\X. X \ \ \ \a b. X = cbox a b" "\\ = S" proof - let ?a = "\f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" let ?b = "\f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" let ?I = "{f\Basis \\<^sub>E \ \ \. cbox (?a f) (?b f) \ S}" let ?\ = "(\f. cbox (?a f) (?b f)) ` ?I" show ?thesis proof have "countable ?I" by (simp add: countable_PiE countable_rat) then show "countable ?\" by blast show "\?\ = S" using open_UNION_cbox [OF assms] by metis qed auto qed lemma box_eq_empty: fixes a :: "'a::euclidean_space" shows "(box a b = {} \ (\i\Basis. b\i \ a\i))" (is ?th1) and "(cbox a b = {} \ (\i\Basis. b\i < a\i))" (is ?th2) proof - { fix i x assume i: "i\Basis" and as:"b\i \ a\i" and x:"x\box a b" then have "a \ i < x \ i \ x \ i < b \ i" unfolding mem_box by (auto simp: box_def) then have "a\i < b\i" by auto then have False using as by auto } moreover { assume as: "\i\Basis. \ (b\i \ a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i :: 'a assume i: "i \ Basis" have "a\i < b\i" using as[THEN bspec[where x=i]] i by auto then have "a\i < ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i < b\i" by (auto simp: inner_add_left) } then have "box a b \ {}" using mem_box(1)[of "?x" a b] by auto } ultimately show ?th1 by blast { fix i x assume i: "i \ Basis" and as:"b\i < a\i" and x:"x\cbox a b" then have "a \ i \ x \ i \ x \ i \ b \ i" unfolding mem_box by auto then have "a\i \ b\i" by auto then have False using as by auto } moreover { assume as:"\i\Basis. \ (b\i < a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i :: 'a assume i:"i \ Basis" have "a\i \ b\i" using as[THEN bspec[where x=i]] i by auto then have "a\i \ ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i \ b\i" by (auto simp: inner_add_left) } then have "cbox a b \ {}" using mem_box(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed lemma box_ne_empty: fixes a :: "'a::euclidean_space" shows "cbox a b \ {} \ (\i\Basis. a\i \ b\i)" and "box a b \ {} \ (\i\Basis. a\i < b\i)" unfolding box_eq_empty[of a b] by fastforce+ lemma fixes a :: "'a::euclidean_space" shows cbox_sing [simp]: "cbox a a = {a}" and box_sing [simp]: "box a a = {}" unfolding set_eq_iff mem_box eq_iff [symmetric] by (auto intro!: euclidean_eqI[where 'a='a]) (metis all_not_in_conv nonempty_Basis) lemma subset_box_imp: fixes a :: "'a::euclidean_space" shows "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ cbox c d \ cbox a b" and "(\i\Basis. a\i < c\i \ d\i < b\i) \ cbox c d \ box a b" and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ cbox a b" and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ box a b" unfolding subset_eq[unfolded Ball_def] unfolding mem_box by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ lemma box_subset_cbox: fixes a :: "'a::euclidean_space" shows "box a b \ cbox a b" unfolding subset_eq [unfolded Ball_def] mem_box by (fast intro: less_imp_le) lemma subset_box: fixes a :: "'a::euclidean_space" shows "cbox c d \ cbox a b \ (\i\Basis. c\i \ d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th1) and "cbox c d \ box a b \ (\i\Basis. c\i \ d\i) \ (\i\Basis. a\i < c\i \ d\i < b\i)" (is ?th2) and "box c d \ cbox a b \ (\i\Basis. c\i < d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th3) and "box c d \ box a b \ (\i\Basis. c\i < d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th4) proof - let ?lesscd = "\i\Basis. c\i < d\i" let ?lerhs = "\i\Basis. a\i \ c\i \ d\i \ b\i" show ?th1 ?th2 by (fastforce simp: mem_box)+ have acdb: "a\i \ c\i \ d\i \ b\i" if i: "i \ Basis" and box: "box c d \ cbox a b" and cd: "\i. i \ Basis \ c\i < d\i" for i proof - have "box c d \ {}" using that unfolding box_eq_empty by force { let ?x = "(\j\Basis. (if j=i then ((min (a\j) (d\j))+c\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" assume *: "a\i > c\i" then have "c \ j < ?x \ j \ ?x \ j < d \ j" if "j \ Basis" for j using cd that by (fastforce simp add: i *) then have "?x \ box c d" unfolding mem_box by auto moreover have "?x \ cbox a b" using i cd * by (force simp: mem_box) ultimately have False using box by auto } then have "a\i \ c\i" by force moreover { let ?x = "(\j\Basis. (if j=i then ((max (b\j) (c\j))+d\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" assume *: "b\i < d\i" then have "d \ j > ?x \ j \ ?x \ j > c \ j" if "j \ Basis" for j using cd that by (fastforce simp add: i *) then have "?x \ box c d" unfolding mem_box by auto moreover have "?x \ cbox a b" using i cd * by (force simp: mem_box) ultimately have False using box by auto } then have "b\i \ d\i" by (rule ccontr) auto ultimately show ?thesis by auto qed show ?th3 using acdb by (fastforce simp add: mem_box) have acdb': "a\i \ c\i \ d\i \ b\i" if "i \ Basis" "box c d \ box a b" "\i. i \ Basis \ c\i < d\i" for i using box_subset_cbox[of a b] that acdb by auto show ?th4 using acdb' by (fastforce simp add: mem_box) qed lemma eq_cbox: "cbox a b = cbox c d \ cbox a b = {} \ cbox c d = {} \ a = c \ b = d" (is "?lhs = ?rhs") proof assume ?lhs then have "cbox a b \ cbox c d" "cbox c d \ cbox a b" by auto then show ?rhs by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI) next assume ?rhs then show ?lhs by force qed lemma eq_cbox_box [simp]: "cbox a b = box c d \ cbox a b = {} \ box c d = {}" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "cbox a b \ box c d" "box c d \ cbox a b" by auto then show ?rhs apply (simp add: subset_box) using L box_ne_empty box_sing apply (fastforce simp add:) done qed force lemma eq_box_cbox [simp]: "box a b = cbox c d \ box a b = {} \ cbox c d = {}" by (metis eq_cbox_box) lemma eq_box: "box a b = box c d \ box a b = {} \ box c d = {} \ a = c \ b = d" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "box a b \ box c d" "box c d \ box a b" by auto then show ?rhs apply (simp add: subset_box) using box_ne_empty(2) L apply auto apply (meson euclidean_eqI less_eq_real_def not_less)+ done qed force lemma subset_box_complex: "cbox a b \ cbox c d \ (Re a \ Re b \ Im a \ Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" "cbox a b \ box c d \ (Re a \ Re b \ Im a \ Im b) \ Re a > Re c \ Im a > Im c \ Re b < Re d \ Im b < Im d" "box a b \ cbox c d \ (Re a < Re b \ Im a < Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" "box a b \ box c d \ (Re a < Re b \ Im a < Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" by (subst subset_box; force simp: Basis_complex_def)+ lemma in_cbox_complex_iff: "x \ cbox a b \ Re x \ {Re a..Re b} \ Im x \ {Im a..Im b}" by (cases x; cases a; cases b) (auto simp: cbox_Complex_eq) lemma box_Complex_eq: "box (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (box a b \ box c d)" by (auto simp: box_def Basis_complex_def image_iff complex_eq_iff) lemma in_box_complex_iff: "x \ box a b \ Re x \ {Re a<.. Im x \ {Im a<.. cbox c d = cbox (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" unfolding set_eq_iff and Int_iff and mem_box by auto lemma disjoint_interval: fixes a::"'a::euclidean_space" shows "cbox a b \ cbox c d = {} \ (\i\Basis. (b\i < a\i \ d\i < c\i \ b\i < c\i \ d\i < a\i))" (is ?th1) and "cbox a b \ box c d = {} \ (\i\Basis. (b\i < a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th2) and "box a b \ cbox c d = {} \ (\i\Basis. (b\i \ a\i \ d\i < c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th3) and "box a b \ box c d = {} \ (\i\Basis. (b\i \ a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th4) proof - let ?z = "(\i\Basis. (((max (a\i) (c\i)) + (min (b\i) (d\i))) / 2) *\<^sub>R i)::'a" have **: "\P Q. (\i :: 'a. i \ Basis \ Q ?z i \ P i) \ (\i x :: 'a. i \ Basis \ P i \ Q x i) \ (\x. \i\Basis. Q x i) \ (\i\Basis. P i)" by blast note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) show ?th1 unfolding * by (intro **) auto show ?th2 unfolding * by (intro **) auto show ?th3 unfolding * by (intro **) auto show ?th4 unfolding * by (intro **) auto qed lemma UN_box_eq_UNIV: "(\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" proof - have "\x \ b\ < real_of_int (\Max ((\b. \x \ b\)`Basis)\ + 1)" if [simp]: "b \ Basis" for x b :: 'a proof - have "\x \ b\ \ real_of_int \\x \ b\\" by (rule le_of_int_ceiling) also have "\ \ real_of_int \Max ((\b. \x \ b\)`Basis)\" by (auto intro!: ceiling_mono) also have "\ < real_of_int (\Max ((\b. \x \ b\)`Basis)\ + 1)" by simp finally show ?thesis . qed then have "\n::nat. \b\Basis. \x \ b\ < real n" for x :: 'a by (metis order.strict_trans reals_Archimedean2) moreover have "\x b::'a. \n::nat. \x \ b\ < real n \ - real n < x \ b \ x \ b < real n" by auto ultimately show ?thesis by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) qed lemma image_affinity_cbox: fixes m::real fixes a b c :: "'a::euclidean_space" shows "(\x. m *\<^sub>R x + c) ` cbox a b = (if cbox a b = {} then {} else (if 0 \ m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" proof (cases "m = 0") case True { fix x assume "\i\Basis. x \ i \ c \ i" "\i\Basis. c \ i \ x \ i" then have "x = c" by (simp add: dual_order.antisym euclidean_eqI) } moreover have "c \ cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" - unfolding True by (auto simp: cbox_sing) + unfolding True by (auto) ultimately show ?thesis using True by (auto simp: cbox_def) next case False { fix y assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m > 0" then have "\i\Basis. (m *\<^sub>R a + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R b + c) \ i" by (auto simp: inner_distrib) } moreover { fix y assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m < 0" then have "\i\Basis. (m *\<^sub>R b + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R a + c) \ i" by (auto simp: mult_left_mono_neg inner_distrib) } moreover { fix y assume "m > 0" and "\i\Basis. (m *\<^sub>R a + c) \ i \ y \ i" and "\i\Basis. y \ i \ (m *\<^sub>R b + c) \ i" then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) done } moreover { fix y assume "\i\Basis. (m *\<^sub>R b + c) \ i \ y \ i" "\i\Basis. y \ i \ (m *\<^sub>R a + c) \ i" "m < 0" then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) done } ultimately show ?thesis using False by (auto simp: cbox_def) qed lemma image_smult_cbox:"(\x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = (if cbox a b = {} then {} else if 0 \ m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" using image_affinity_cbox[of m 0 a b] by auto lemma swap_continuous: assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "continuous_on (cbox (c,a) (d,b)) (\(x, y). f y x)" proof - have "(\(x, y). f y x) = (\(x, y). f x y) \ prod.swap" by auto then show ?thesis apply (rule ssubst) apply (rule continuous_on_compose) apply (simp add: split_def) apply (rule continuous_intros | simp add: assms)+ done qed subsection \General Intervals\ definition\<^marker>\tag important\ "is_interval (s::('a::euclidean_space) set) \ (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" lemma is_interval_1: "is_interval (s::real set) \ (\a\s. \b\s. \ x. a \ x \ x \ b \ x \ s)" unfolding is_interval_def by auto lemma is_interval_Int: "is_interval X \ is_interval Y \ is_interval (X \ Y)" unfolding is_interval_def by blast lemma is_interval_cbox [simp]: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) and is_interval_box [simp]: "is_interval (box a b)" (is ?th2) unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff by (meson order_trans le_less_trans less_le_trans less_trans)+ lemma is_interval_empty [iff]: "is_interval {}" unfolding is_interval_def by simp lemma is_interval_univ [iff]: "is_interval UNIV" unfolding is_interval_def by simp lemma mem_is_intervalI: assumes "is_interval s" and "a \ s" "b \ s" and "\i. i \ Basis \ a \ i \ x \ i \ x \ i \ b \ i \ b \ i \ x \ i \ x \ i \ a \ i" shows "x \ s" by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) lemma interval_subst: fixes S::"'a::euclidean_space set" assumes "is_interval S" and "x \ S" "y j \ S" and "j \ Basis" shows "(\i\Basis. (if i = j then y i \ i else x \ i) *\<^sub>R i) \ S" by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) lemma mem_box_componentwiseI: fixes S::"'a::euclidean_space set" assumes "is_interval S" assumes "\i. i \ Basis \ x \ i \ ((\x. x \ i) ` S)" shows "x \ S" proof - from assms have "\i \ Basis. \s \ S. x \ i = s \ i" by auto with finite_Basis obtain s and bs::"'a list" where s: "\i. i \ Basis \ x \ i = s i \ i" "\i. i \ Basis \ s i \ S" and bs: "set bs = Basis" "distinct bs" by (metis finite_distinct_list) from nonempty_Basis s obtain j where j: "j \ Basis" "s j \ S" by blast define y where "y = rec_list (s j) (\j _ Y. (\i\Basis. (if i = j then s i \ i else Y \ i) *\<^sub>R i))" have "x = (\i\Basis. (if i \ set bs then s i \ i else s j \ i) *\<^sub>R i)" using bs by (auto simp: s(1)[symmetric] euclidean_representation) also have [symmetric]: "y bs = \" using bs(2) bs(1)[THEN equalityD1] by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) also have "y bs \ S" using bs(1)[THEN equalityD1] apply (induct bs) apply (auto simp: y_def j) apply (rule interval_subst[OF assms(1)]) apply (auto simp: s) done finally show ?thesis . qed lemma cbox01_nonempty [simp]: "cbox 0 One \ {}" by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg) lemma box01_nonempty [simp]: "box 0 One \ {}" by (simp add: box_ne_empty inner_Basis inner_sum_left) lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" using nonempty_Basis box01_nonempty box_eq_empty(1) box_ne_empty(1) by blast lemma interval_subset_is_interval: assumes "is_interval S" shows "cbox a b \ S \ cbox a b = {} \ a \ S \ b \ S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using box_ne_empty(1) mem_box(2) by fastforce next assume ?rhs have "cbox a b \ S" if "a \ S" "b \ S" using assms unfolding is_interval_def apply (clarsimp simp add: mem_box) using that by blast with \?rhs\ show ?lhs by blast qed lemma is_real_interval_union: "is_interval (X \ Y)" if X: "is_interval X" and Y: "is_interval Y" and I: "(X \ {} \ Y \ {} \ X \ Y \ {})" for X Y::"real set" proof - consider "X \ {}" "Y \ {}" | "X = {}" | "Y = {}" by blast then show ?thesis proof cases case 1 then obtain r where "r \ X \ X \ Y = {}" "r \ Y \ X \ Y = {}" by blast then show ?thesis using I 1 X Y unfolding is_interval_1 by (metis (full_types) Un_iff le_cases) qed (use that in auto) qed lemma is_interval_translationI: assumes "is_interval X" shows "is_interval ((+) x ` X)" unfolding is_interval_def proof safe fix b d e assume "b \ X" "d \ X" "\i\Basis. (x + b) \ i \ e \ i \ e \ i \ (x + d) \ i \ (x + d) \ i \ e \ i \ e \ i \ (x + b) \ i" hence "e - x \ X" by (intro mem_is_intervalI[OF assms \b \ X\ \d \ X\, of "e - x"]) (auto simp: algebra_simps) thus "e \ (+) x ` X" by force qed lemma is_interval_uminusI: assumes "is_interval X" shows "is_interval (uminus ` X)" unfolding is_interval_def proof safe fix b d e assume "b \ X" "d \ X" "\i\Basis. (- b) \ i \ e \ i \ e \ i \ (- d) \ i \ (- d) \ i \ e \ i \ e \ i \ (- b) \ i" hence "- e \ X" by (intro mem_is_intervalI[OF assms \b \ X\ \d \ X\, of "- e"]) (auto simp: algebra_simps) thus "e \ uminus ` X" by force qed lemma is_interval_uminus[simp]: "is_interval (uminus ` x) = is_interval x" using is_interval_uminusI[of x] is_interval_uminusI[of "uminus ` x"] by (auto simp: image_image) lemma is_interval_neg_translationI: assumes "is_interval X" shows "is_interval ((-) x ` X)" proof - have "(-) x ` X = (+) x ` uminus ` X" by (force simp: algebra_simps) also have "is_interval \" by (metis is_interval_uminusI is_interval_translationI assms) finally show ?thesis . qed lemma is_interval_translation[simp]: "is_interval ((+) x ` X) = is_interval X" using is_interval_neg_translationI[of "(+) x ` X" x] by (auto intro!: is_interval_translationI simp: image_image) lemma is_interval_minus_translation[simp]: shows "is_interval ((-) x ` X) = is_interval X" proof - have "(-) x ` X = (+) x ` uminus ` X" by (force simp: algebra_simps) also have "is_interval \ = is_interval X" by simp finally show ?thesis . qed lemma is_interval_minus_translation'[simp]: shows "is_interval ((\x. x - c) ` X) = is_interval X" using is_interval_translation[of "-c" X] by (metis image_cong uminus_add_conv_diff) lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real by (simp add: cball_eq_atLeastAtMost is_interval_def) lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real by (simp add: ball_eq_greaterThanLessThan is_interval_def) subsection\<^marker>\tag unimportant\ \Bounded Projections\ lemma bounded_inner_imp_bdd_above: assumes "bounded s" shows "bdd_above ((\x. x \ a) ` s)" by (simp add: assms bounded_imp_bdd_above bounded_linear_image bounded_linear_inner_left) lemma bounded_inner_imp_bdd_below: assumes "bounded s" shows "bdd_below ((\x. x \ a) ` s)" by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left) subsection\<^marker>\tag unimportant\ \Structural rules for pointwise continuity\ lemma continuous_infnorm[continuous_intros]: "continuous F f \ continuous F (\x. infnorm (f x))" unfolding continuous_def by (rule tendsto_infnorm) lemma continuous_inner[continuous_intros]: assumes "continuous F f" and "continuous F g" shows "continuous F (\x. inner (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_inner) subsection\<^marker>\tag unimportant\ \Structural rules for setwise continuity\ lemma continuous_on_infnorm[continuous_intros]: "continuous_on s f \ continuous_on s (\x. infnorm (f x))" unfolding continuous_on by (fast intro: tendsto_infnorm) lemma continuous_on_inner[continuous_intros]: fixes g :: "'a::topological_space \ 'b::real_inner" assumes "continuous_on s f" and "continuous_on s g" shows "continuous_on s (\x. inner (f x) (g x))" using bounded_bilinear_inner assms by (rule bounded_bilinear.continuous_on) subsection\<^marker>\tag unimportant\ \Openness of halfspaces.\ lemma open_halfspace_lt: "open {x. inner a x < b}" - by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: open_Collect_less continuous_on_inner) lemma open_halfspace_gt: "open {x. inner a x > b}" - by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: open_Collect_less continuous_on_inner) lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\i < a}" - by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: open_Collect_less continuous_on_inner) lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\i > a}" - by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: open_Collect_less continuous_on_inner) lemma eucl_less_eq_halfspaces: fixes a :: "'a::euclidean_space" shows "{x. x i\Basis. {x. x \ i < a \ i})" "{x. a i\Basis. {x. a \ i < x \ i})" by (auto simp: eucl_less_def) lemma open_Collect_eucl_less[simp, intro]: fixes a :: "'a::euclidean_space" shows "open {x. x \tag unimportant\ \Closure and Interior of halfspaces and hyperplanes\ lemma continuous_at_inner: "continuous (at x) (inner a)" unfolding continuous_at by (intro tendsto_intros) lemma closed_halfspace_le: "closed {x. inner a x \ b}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: closed_Collect_le continuous_on_inner) lemma closed_halfspace_ge: "closed {x. inner a x \ b}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: closed_Collect_le continuous_on_inner) lemma closed_hyperplane: "closed {x. inner a x = b}" - by (simp add: closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: closed_Collect_eq continuous_on_inner) lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\i \ a}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: closed_Collect_le continuous_on_inner) lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\i \ a}" - by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: closed_Collect_le continuous_on_inner) lemma closed_interval_left: fixes b :: "'a::euclidean_space" shows "closed {x::'a. \i\Basis. x\i \ b\i}" - by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner) lemma closed_interval_right: fixes a :: "'a::euclidean_space" shows "closed {x::'a. \i\Basis. a\i \ x\i}" - by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) + by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner) lemma interior_halfspace_le [simp]: assumes "a \ 0" shows "interior {x. a \ x \ b} = {x. a \ x < b}" proof - have *: "a \ x < b" if x: "x \ S" and S: "S \ {x. a \ x \ b}" and "open S" for S x proof - obtain e where "e>0" and e: "cball x e \ S" using \open S\ open_contains_cball x by blast then have "x + (e / norm a) *\<^sub>R a \ cball x e" by (simp add: dist_norm) then have "x + (e / norm a) *\<^sub>R a \ S" using e by blast then have "x + (e / norm a) *\<^sub>R a \ {x. a \ x \ b}" using S by blast moreover have "e * (a \ a) / norm a > 0" by (simp add: \0 < e\ assms) ultimately show ?thesis by (simp add: algebra_simps) qed show ?thesis by (rule interior_unique) (auto simp: open_halfspace_lt *) qed lemma interior_halfspace_ge [simp]: "a \ 0 \ interior {x. a \ x \ b} = {x. a \ x > b}" using interior_halfspace_le [of "-a" "-b"] by simp lemma closure_halfspace_lt [simp]: assumes "a \ 0" shows "closure {x. a \ x < b} = {x. a \ x \ b}" proof - have [simp]: "-{x. a \ x < b} = {x. a \ x \ b}" by (force simp:) then show ?thesis using interior_halfspace_ge [of a b] assms by (force simp: closure_interior) qed lemma closure_halfspace_gt [simp]: "a \ 0 \ closure {x. a \ x > b} = {x. a \ x \ b}" using closure_halfspace_lt [of "-a" "-b"] by simp lemma interior_hyperplane [simp]: assumes "a \ 0" shows "interior {x. a \ x = b} = {}" proof - have [simp]: "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by (force simp:) then show ?thesis by (auto simp: assms) qed lemma frontier_halfspace_le: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x \ b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def closed_halfspace_le) qed lemma frontier_halfspace_ge: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x \ b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def closed_halfspace_ge) qed lemma frontier_halfspace_lt: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x < b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def interior_open open_halfspace_lt) qed lemma frontier_halfspace_gt: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x > b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def interior_open open_halfspace_gt) qed subsection\<^marker>\tag unimportant\\Some more convenient intermediate-value theorem formulations\ lemma connected_ivt_hyperplane: assumes "connected S" and xy: "x \ S" "y \ S" and b: "inner a x \ b" "b \ inner a y" shows "\z \ S. inner a z = b" proof (rule ccontr) assume as:"\ (\z\S. inner a z = b)" let ?A = "{x. inner a x < b}" let ?B = "{x. inner a x > b}" have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto moreover have "?A \ ?B = {}" by auto moreover have "S \ ?A \ ?B" using as by auto ultimately show False using \connected S\[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] using xy b by auto qed lemma connected_ivt_component: fixes x::"'a::euclidean_space" shows "connected S \ x \ S \ y \ S \ x\k \ a \ a \ y\k \ (\z\S. z\k = a)" using connected_ivt_hyperplane[of S x y "k::'a" a] by (auto simp: inner_commute) subsection \Limit Component Bounds\ lemma Lim_component_le: fixes f :: "'a \ 'b::euclidean_space" assumes "(f \ l) net" and "\ (trivial_limit net)" and "eventually (\x. f(x)\i \ b) net" shows "l\i \ b" by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) lemma Lim_component_ge: fixes f :: "'a \ 'b::euclidean_space" assumes "(f \ l) net" and "\ (trivial_limit net)" and "eventually (\x. b \ (f x)\i) net" shows "b \ l\i" by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) lemma Lim_component_eq: fixes f :: "'a \ 'b::euclidean_space" assumes net: "(f \ l) net" "\ trivial_limit net" and ev:"eventually (\x. f(x)\i = b) net" shows "l\i = b" using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge[OF net, of b i] using Lim_component_le[OF net, of i b] by auto lemma open_box[intro]: "open (box a b)" proof - have "open (\i\Basis. ((\) i) -` {a \ i <..< b \ i})" by (auto intro!: continuous_open_vimage continuous_inner continuous_ident continuous_const) also have "(\i\Basis. ((\) i) -` {a \ i <..< b \ i}) = box a b" by (auto simp: box_def inner_commute) finally show ?thesis . qed lemma closed_cbox[intro]: fixes a b :: "'a::euclidean_space" shows "closed (cbox a b)" proof - have "closed (\i\Basis. (\x. x\i) -` {a\i .. b\i})" by (intro closed_INT ballI continuous_closed_vimage allI linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) also have "(\i\Basis. (\x. x\i) -` {a\i .. b\i}) = cbox a b" by (auto simp: cbox_def) finally show "closed (cbox a b)" . qed lemma interior_cbox [simp]: fixes a b :: "'a::euclidean_space" shows "interior (cbox a b) = box a b" (is "?L = ?R") proof(rule subset_antisym) show "?R \ ?L" using box_subset_cbox open_box by (rule interior_maximal) { fix x assume "x \ interior (cbox a b)" then obtain s where s: "open s" "x \ s" "s \ cbox a b" .. then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ cbox a b" unfolding open_dist and subset_eq by auto { fix i :: 'a assume i: "i \ Basis" have "dist (x - (e / 2) *\<^sub>R i) x < e" and "dist (x + (e / 2) *\<^sub>R i) x < e" unfolding dist_norm apply auto unfolding norm_minus_cancel using norm_Basis[OF i] \e>0\ apply auto done then have "a \ i \ (x - (e / 2) *\<^sub>R i) \ i" and "(x + (e / 2) *\<^sub>R i) \ i \ b \ i" using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] unfolding mem_box using i by blast+ then have "a \ i < x \ i" and "x \ i < b \ i" using \e>0\ i by (auto simp: inner_diff_left inner_Basis inner_add_left) } then have "x \ box a b" unfolding mem_box by auto } then show "?L \ ?R" .. qed lemma bounded_cbox [simp]: fixes a :: "'a::euclidean_space" shows "bounded (cbox a b)" proof - let ?b = "\i\Basis. \a\i\ + \b\i\" { fix x :: "'a" assume "\i. i\Basis \ a \ i \ x \ i \ x \ i \ b \ i" then have "(\i\Basis. \x \ i\) \ ?b" by (force simp: intro!: sum_mono) then have "norm x \ ?b" using norm_le_l1[of x] by auto } then show ?thesis unfolding cbox_def bounded_iff by force qed lemma bounded_box [simp]: fixes a :: "'a::euclidean_space" shows "bounded (box a b)" using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] by simp lemma not_interval_UNIV [simp]: fixes a :: "'a::euclidean_space" shows "cbox a b \ UNIV" "box a b \ UNIV" using bounded_box[of a b] bounded_cbox[of a b] by force+ lemma not_interval_UNIV2 [simp]: fixes a :: "'a::euclidean_space" shows "UNIV \ cbox a b" "UNIV \ box a b" using bounded_box[of a b] bounded_cbox[of a b] by force+ lemma box_midpoint: fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "((1/2) *\<^sub>R (a + b)) \ box a b" proof - have "a \ i < ((1 / 2) *\<^sub>R (a + b)) \ i \ ((1 / 2) *\<^sub>R (a + b)) \ i < b \ i" if "i \ Basis" for i using assms that by (auto simp: inner_add_left box_ne_empty) then show ?thesis unfolding mem_box by auto qed lemma open_cbox_convex: fixes x :: "'a::euclidean_space" assumes x: "x \ box a b" and y: "y \ cbox a b" and e: "0 < e" "e \ 1" shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ box a b" proof - { fix i :: 'a assume i: "i \ Basis" have "a \ i = e * (a \ i) + (1 - e) * (a \ i)" unfolding left_diff_distrib by simp also have "\ < e * (x \ i) + (1 - e) * (y \ i)" proof (rule add_less_le_mono) show "e * (a \ i) < e * (x \ i)" using \0 < e\ i mem_box(1) x by auto show "(1 - e) * (a \ i) \ (1 - e) * (y \ i)" by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) qed finally have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i" unfolding inner_simps by auto moreover { have "b \ i = e * (b\i) + (1 - e) * (b\i)" unfolding left_diff_distrib by simp also have "\ > e * (x \ i) + (1 - e) * (y \ i)" proof (rule add_less_le_mono) show "e * (x \ i) < e * (b \ i)" using \0 < e\ i mem_box(1) x by auto show "(1 - e) * (y \ i) \ (1 - e) * (b \ i)" by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) qed finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" unfolding inner_simps by auto } ultimately have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i \ (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" by auto } then show ?thesis unfolding mem_box by auto qed lemma closure_cbox [simp]: "closure (cbox a b) = cbox a b" by (simp add: closed_cbox) lemma closure_box [simp]: fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "closure (box a b) = cbox a b" proof - have ab: "a cbox a b" define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n { fix n assume fn: "f n a f n = x" and xc: "x \ ?c" have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \ 1" unfolding inverse_le_1_iff by auto have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" by (auto simp: algebra_simps) then have "f n (f \ x) sequentially" { fix e :: real assume "e > 0" then obtain N :: nat where N: "inverse (real (N + 1)) < e" using reals_Archimedean by auto have "inverse (real n + 1) < e" if "N \ n" for n by (auto intro!: that le_less_trans [OF _ N]) then have "\N::nat. \n\N. inverse (real n + 1) < e" by auto } then have "((\n. inverse (real n + 1)) \ 0) sequentially" unfolding lim_sequentially by(auto simp: dist_norm) then have "(f \ x) sequentially" unfolding f_def using tendsto_add[OF tendsto_const, of "\n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] using tendsto_scaleR [OF _ tendsto_const, of "\n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto } ultimately have "x \ closure (box a b)" using as box_midpoint[OF assms] unfolding closure_def islimpt_sequential by (cases "x=?c") (auto simp: in_box_eucl_less) } then show ?thesis using closure_minimal[OF box_subset_cbox, of a b] by blast qed lemma bounded_subset_box_symmetric: fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ box (-a) a" proof - obtain b where "b>0" and b: "\x\S. norm x \ b" using assms[unfolded bounded_pos] by auto define a :: 'a where "a = (\i\Basis. (b + 1) *\<^sub>R i)" have "(-a)\i < x\i" and "x\i < a\i" if "x \ S" and i: "i \ Basis" for x i using b Basis_le_norm[OF i, of x] that by (auto simp: a_def) then have "S \ box (-a) a" by (auto simp: simp add: box_def) then show ?thesis .. qed lemma bounded_subset_cbox_symmetric: fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ cbox (-a) a" proof - obtain a where "S \ box (-a) a" using bounded_subset_box_symmetric[OF assms] by auto then show ?thesis by (meson box_subset_cbox dual_order.trans that) qed lemma frontier_cbox: fixes a b :: "'a::euclidean_space" shows "frontier (cbox a b) = cbox a b - box a b" unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. lemma frontier_box: fixes a b :: "'a::euclidean_space" shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)" proof (cases "box a b = {}") case True then show ?thesis using frontier_empty by auto next case False then show ?thesis unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] by auto qed lemma Int_interval_mixed_eq_empty: fixes a :: "'a::euclidean_space" assumes "box c d \ {}" shows "box a b \ cbox c d = {} \ box a b \ box c d = {}" unfolding closure_box[OF assms, symmetric] unfolding open_Int_closure_eq_empty[OF open_box] .. subsection \Class Instances\ lemma compact_lemma: fixes f :: "nat \ 'a::euclidean_space" assumes "bounded (range f)" shows "\d\Basis. \l::'a. \ r. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) \ i) (l \ i) < e) sequentially)" by (rule compact_lemma_general[where unproj="\e. \i\Basis. e i *\<^sub>R i"]) (auto intro!: assms bounded_linear_inner_left bounded_linear_image simp: euclidean_representation) instance\<^marker>\tag important\ euclidean_space \ heine_borel proof fix f :: "nat \ 'a" assume f: "bounded (range f)" then obtain l::'a and r where r: "strict_mono r" and l: "\e>0. eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e) sequentially" using compact_lemma [OF f] by blast { fix e::real assume "e > 0" - hence "e / real_of_nat DIM('a) > 0" by (simp add: DIM_positive) + hence "e / real_of_nat DIM('a) > 0" by (simp) with l have "eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))) sequentially" by simp moreover { fix n assume n: "\i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))" have "dist (f (r n)) l \ (\i\Basis. dist (f (r n) \ i) (l \ i))" apply (subst euclidean_dist_l2) using zero_le_dist apply (rule L2_set_le_sum) done also have "\ < (\i\(Basis::'a set). e / (real_of_nat DIM('a)))" apply (rule sum_strict_mono) using n apply auto done finally have "dist (f (r n)) l < e" by auto } ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" by (rule eventually_mono) } then have *: "((f \ r) \ l) sequentially" unfolding o_def tendsto_iff by simp with r show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" by auto qed instance\<^marker>\tag important\ euclidean_space \ banach .. instance euclidean_space \ second_countable_topology proof define a where "a f = (\i\Basis. fst (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have a: "\f. (\i\Basis. fst (f i) *\<^sub>R i) = a f" by simp define b where "b f = (\i\Basis. snd (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have b: "\f. (\i\Basis. snd (f i) *\<^sub>R i) = b f" by simp define B where "B = (\f. box (a f) (b f)) ` (Basis \\<^sub>E (\ \ \))" have "Ball B open" by (simp add: B_def open_box) moreover have "(\A. open A \ (\B'\B. \B' = A))" proof safe fix A::"'a set" assume "open A" show "\B'\B. \B' = A" apply (rule exI[of _ "{b\B. b \ A}"]) apply (subst (3) open_UNION_box[OF \open A\]) apply (auto simp: a b B_def) done qed ultimately have "topological_basis B" unfolding topological_basis_def by blast moreover have "countable B" unfolding B_def by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) ultimately show "\B::'a set set. countable B \ open = generate_topology B" by (blast intro: topological_basis_imp_subbasis) qed instance euclidean_space \ polish_space .. subsection \Compact Boxes\ lemma compact_cbox [simp]: fixes a :: "'a::euclidean_space" shows "compact (cbox a b)" using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] by (auto simp: compact_eq_seq_compact_metric) proposition is_interval_compact: "is_interval S \ compact S \ (\a b. S = cbox a b)" (is "?lhs = ?rhs") proof (cases "S = {}") case True with empty_as_interval show ?thesis by auto next case False show ?thesis proof assume L: ?lhs then have "is_interval S" "compact S" by auto define a where "a \ \i\Basis. (INF x\S. x \ i) *\<^sub>R i" define b where "b \ \i\Basis. (SUP x\S. x \ i) *\<^sub>R i" have 1: "\x i. \x \ S; i \ Basis\ \ (INF x\S. x \ i) \ x \ i" by (simp add: cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded L) have 2: "\x i. \x \ S; i \ Basis\ \ x \ i \ (SUP x\S. x \ i)" by (simp add: cSup_upper bounded_inner_imp_bdd_above compact_imp_bounded L) have 3: "x \ S" if inf: "\i. i \ Basis \ (INF x\S. x \ i) \ x \ i" and sup: "\i. i \ Basis \ x \ i \ (SUP x\S. x \ i)" for x proof (rule mem_box_componentwiseI [OF \is_interval S\]) fix i::'a assume i: "i \ Basis" have cont: "continuous_on S (\x. x \ i)" by (intro continuous_intros) obtain a where "a \ S" and a: "\y. y\S \ a \ i \ y \ i" using continuous_attains_inf [OF \compact S\ False cont] by blast obtain b where "b \ S" and b: "\y. y\S \ y \ i \ b \ i" using continuous_attains_sup [OF \compact S\ False cont] by blast have "a \ i \ (INF x\S. x \ i)" by (simp add: False a cINF_greatest) also have "\ \ x \ i" by (simp add: i inf) finally have ai: "a \ i \ x \ i" . have "x \ i \ (SUP x\S. x \ i)" by (simp add: i sup) also have "(SUP x\S. x \ i) \ b \ i" by (simp add: False b cSUP_least) finally have bi: "x \ i \ b \ i" . show "x \ i \ (\x. x \ i) ` S" apply (rule_tac x="\j\Basis. (if j = i then x \ i else a \ j) *\<^sub>R j" in image_eqI) apply (simp add: i) apply (rule mem_is_intervalI [OF \is_interval S\ \a \ S\ \b \ S\]) using i ai bi apply force done qed have "S = cbox a b" by (auto simp: a_def b_def mem_box intro: 1 2 3) then show ?rhs by blast next assume R: ?rhs then show ?lhs using compact_cbox is_interval_cbox by blast qed qed subsection\<^marker>\tag unimportant\\Componentwise limits and continuity\ text\But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\ lemma Euclidean_dist_upper: "i \ Basis \ dist (x \ i) (y \ i) \ dist x y" by (metis (no_types) member_le_L2_set euclidean_dist_l2 finite_Basis) text\But is the premise \<^term>\i \ Basis\ really necessary?\ lemma open_preimage_inner: assumes "open S" "i \ Basis" shows "open {x. x \ i \ S}" proof (rule openI, simp) fix x assume x: "x \ i \ S" with assms obtain e where "0 < e" and e: "ball (x \ i) e \ S" by (auto simp: open_contains_ball_eq) have "\e>0. ball (y \ i) e \ S" if dxy: "dist x y < e / 2" for y proof (intro exI conjI) have "dist (x \ i) (y \ i) < e / 2" by (meson \i \ Basis\ dual_order.trans Euclidean_dist_upper not_le that) then have "dist (x \ i) z < e" if "dist (y \ i) z < e / 2" for z by (metis dist_commute dist_triangle_half_l that) then have "ball (y \ i) (e / 2) \ ball (x \ i) e" using mem_ball by blast with e show "ball (y \ i) (e / 2) \ S" by (metis order_trans) qed (simp add: \0 < e\) then show "\e>0. ball x e \ {s. s \ i \ S}" by (metis (no_types, lifting) \0 < e\ \open S\ half_gt_zero_iff mem_Collect_eq mem_ball open_contains_ball_eq subsetI) qed proposition tendsto_componentwise_iff: fixes f :: "_ \ 'b::euclidean_space" shows "(f \ l) F \ (\i \ Basis. ((\x. (f x \ i)) \ (l \ i)) F)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding tendsto_def apply clarify apply (drule_tac x="{s. s \ i \ S}" in spec) apply (auto simp: open_preimage_inner) done next assume R: ?rhs then have "\e. e > 0 \ \i\Basis. \\<^sub>F x in F. dist (f x \ i) (l \ i) < e" unfolding tendsto_iff by blast then have R': "\e. e > 0 \ \\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e" by (simp add: eventually_ball_finite_distrib [symmetric]) show ?lhs unfolding tendsto_iff proof clarify fix e::real assume "0 < e" have *: "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" if "\i\Basis. dist (f x \ i) (l \ i) < e / real DIM('b)" for x proof - have "L2_set (\i. dist (f x \ i) (l \ i)) Basis \ sum (\i. dist (f x \ i) (l \ i)) Basis" by (simp add: L2_set_le_sum) also have "... < DIM('b) * (e / real DIM('b))" apply (rule sum_bounded_above_strict) using that by auto also have "... = e" by (simp add: field_simps) finally show "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" . qed have "\\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e / DIM('b)" apply (rule R') using \0 < e\ by simp then show "\\<^sub>F x in F. dist (f x) l < e" apply (rule eventually_mono) apply (subst euclidean_dist_l2) using * by blast qed qed corollary continuous_componentwise: "continuous F f \ (\i \ Basis. continuous F (\x. (f x \ i)))" by (simp add: continuous_def tendsto_componentwise_iff [symmetric]) corollary continuous_on_componentwise: fixes S :: "'a :: t2_space set" shows "continuous_on S f \ (\i \ Basis. continuous_on S (\x. (f x \ i)))" apply (simp add: continuous_on_eq_continuous_within) using continuous_componentwise by blast lemma linear_componentwise_iff: "(linear f') \ (\i\Basis. linear (\x. f' x \ i))" apply (auto simp: linear_iff inner_left_distrib) apply (metis inner_left_distrib euclidean_eq_iff) by (metis euclidean_eqI inner_scaleR_left) lemma bounded_linear_componentwise_iff: "(bounded_linear f') \ (\i\Basis. bounded_linear (\x. f' x \ i))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: bounded_linear_inner_left_comp) next assume ?rhs then have "(\i\Basis. \K. \x. \f' x \ i\ \ norm x * K)" "linear f'" by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) then obtain F where F: "\i x. i \ Basis \ \f' x \ i\ \ norm x * F i" by metis have "norm (f' x) \ norm x * sum F Basis" for x proof - have "norm (f' x) \ (\i\Basis. \f' x \ i\)" by (rule norm_le_l1) also have "... \ (\i\Basis. norm x * F i)" by (metis F sum_mono) also have "... = norm x * sum F Basis" by (simp add: sum_distrib_left) finally show ?thesis . qed then show ?lhs by (force simp: bounded_linear_def bounded_linear_axioms_def \linear f'\) qed subsection\<^marker>\tag unimportant\ \Continuous Extension\ definition clamp :: "'a::euclidean_space \ 'a \ 'a \ 'a" where "clamp a b x = (if (\i\Basis. a \ i \ b \ i) then (\i\Basis. (if x\i < a\i then a\i else if x\i \ b\i then x\i else b\i) *\<^sub>R i) else a)" lemma clamp_in_interval[simp]: assumes "\i. i \ Basis \ a \ i \ b \ i" shows "clamp a b x \ cbox a b" unfolding clamp_def using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) lemma clamp_cancel_cbox[simp]: fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "clamp a b x = x" using assms by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a]) lemma clamp_empty_interval: assumes "i \ Basis" "a \ i > b \ i" shows "clamp a b = (\_. a)" using assms by (force simp: clamp_def[abs_def] split: if_splits intro!: ext) lemma dist_clamps_le_dist_args: fixes x :: "'a::euclidean_space" shows "dist (clamp a b y) (clamp a b x) \ dist y x" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" then have "(\i\Basis. (dist (clamp a b y \ i) (clamp a b x \ i))\<^sup>2) \ (\i\Basis. (dist (y \ i) (x \ i))\<^sup>2)" by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) then show ?thesis by (auto intro: real_sqrt_le_mono simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] L2_set_def) qed (auto simp: clamp_def) lemma clamp_continuous_at: fixes f :: "'a::euclidean_space \ 'b::metric_space" and x :: 'a assumes f_cont: "continuous_on (cbox a b) f" shows "continuous (at x) (\x. f (clamp a b x))" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" show ?thesis unfolding continuous_at_eps_delta proof safe fix x :: 'a fix e :: real assume "e > 0" moreover have "clamp a b x \ cbox a b" - by (simp add: clamp_in_interval le) + by (simp add: le) moreover note f_cont[simplified continuous_on_iff] ultimately obtain d where d: "0 < d" "\x'. x' \ cbox a b \ dist x' (clamp a b x) < d \ dist (f x') (f (clamp a b x)) < e" by force show "\d>0. \x'. dist x' x < d \ dist (f (clamp a b x')) (f (clamp a b x)) < e" using le by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) qed qed (auto simp: clamp_empty_interval) lemma clamp_continuous_on: fixes f :: "'a::euclidean_space \ 'b::metric_space" assumes f_cont: "continuous_on (cbox a b) f" shows "continuous_on S (\x. f (clamp a b x))" using assms by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) lemma clamp_bounded: fixes f :: "'a::euclidean_space \ 'b::metric_space" assumes bounded: "bounded (f ` (cbox a b))" shows "bounded (range (\x. f (clamp a b x)))" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" from bounded obtain c where f_bound: "\x\f ` cbox a b. dist undefined x \ c" by (auto simp: bounded_any_center[where a=undefined]) then show ?thesis by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] simp: bounded_any_center[where a=undefined]) qed (auto simp: clamp_empty_interval image_def) definition ext_cont :: "('a::euclidean_space \ 'b::metric_space) \ 'a \ 'a \ 'a \ 'b" where "ext_cont f a b = (\x. f (clamp a b x))" lemma ext_cont_cancel_cbox[simp]: fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "ext_cont f a b x = f x" using assms unfolding ext_cont_def by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) lemma continuous_on_ext_cont[continuous_intros]: "continuous_on (cbox a b) f \ continuous_on S (ext_cont f a b)" by (auto intro!: clamp_continuous_on simp: ext_cont_def) subsection \Separability\ 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}" proof - obtain \ :: "'a set set" where "countable \" and opn: "\C. C \ \ \ open C" and Un: "\S. open S \ \U. U \ \ \ S = \U" using univ_second_countable by blast have *: "infinite (range (\n. ball (0::'a) (inverse(Suc n))))" apply (rule Infinite_Set.range_inj_infinite) apply (simp add: inj_on_def ball_eq_ball_iff) done have "infinite \" proof assume "finite \" then have "finite (Union ` (Pow \))" by simp then have "finite (range (\n. ball (0::'a) (inverse(Suc n))))" apply (rule rev_finite_subset) by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) with * show False by simp qed obtain f :: "nat \ 'a set" where "\ = range f" "inj f" by (blast intro: countable_as_injective_image [OF \countable \\ \infinite \\]) have *: "\k. S = \{f n |n. n \ k}" if "open S" for S using Un [OF that] apply clarify apply (rule_tac x="f-`U" in exI) using \inj f\ \\ = range f\ apply force done show ?thesis apply (rule that [OF \inj f\ _ *]) apply (auto simp: \\ = range f\ opn) done qed proposition separable: fixes S :: "'a::{metric_space, second_countable_topology} set" obtains T where "countable T" "T \ S" "S \ closure T" proof - obtain \ :: "'a set set" where "countable \" and "{} \ \" and ope: "\C. C \ \ \ openin(top_of_set S) C" and if_ope: "\T. openin(top_of_set 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 (top_of_set 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 subsection\<^marker>\tag unimportant\ \Diameter\ lemma diameter_cball [simp]: fixes a :: "'a::euclidean_space" shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" proof - have "diameter(cball a r) = 2*r" if "r \ 0" proof (rule order_antisym) show "diameter (cball a r) \ 2*r" proof (rule diameter_le) fix x y assume "x \ cball a r" "y \ cball a r" then have "norm (x - a) \ r" "norm (a - y) \ r" by (auto simp: dist_norm norm_minus_commute) then have "norm (x - y) \ r+r" using norm_diff_triangle_le by blast then show "norm (x - y) \ 2*r" by simp qed (simp add: that) have "2*r = dist (a + r *\<^sub>R (SOME i. i \ Basis)) (a - r *\<^sub>R (SOME i. i \ Basis))" apply (simp add: dist_norm) by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) also have "... \ diameter (cball a r)" apply (rule diameter_bounded_bound) using that by (auto simp: dist_norm) finally show "2*r \ diameter (cball a r)" . qed then show ?thesis by simp qed lemma diameter_ball [simp]: fixes a :: "'a::euclidean_space" shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" proof - have "diameter(ball a r) = 2*r" if "r > 0" by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) then show ?thesis by (simp add: diameter_def) qed lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" proof - have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm) then show ?thesis by simp qed lemma diameter_open_interval [simp]: "diameter {a<..i \ Basis. a \ i \ b \ i) \ diameter (cbox a b) = dist a b" by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) subsection\<^marker>\tag unimportant\\Relating linear images to open/closed/interior/closure/connected\ proposition open_surjective_linear_image: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "open A" "linear f" "surj f" shows "open(f ` A)" unfolding open_dist proof clarify fix x assume "x \ A" have "bounded (inv f ` Basis)" by (simp add: finite_imp_bounded) with bounded_pos obtain B where "B > 0" and B: "\x. x \ inv f ` Basis \ norm x \ B" by metis obtain e where "e > 0" and e: "\z. dist z x < e \ z \ A" by (metis open_dist \x \ A\ \open A\) define \ where "\ \ e / B / DIM('b)" show "\e>0. \y. dist y (f x) < e \ y \ f ` A" proof (intro exI conjI) show "\ > 0" using \e > 0\ \B > 0\ by (simp add: \_def field_split_simps) have "y \ f ` A" if "dist y (f x) * (B * real DIM('b)) < e" for y proof - define u where "u \ y - f x" show ?thesis proof (rule image_eqI) show "y = f (x + (\i\Basis. (u \ i) *\<^sub>R inv f i))" apply (simp add: linear_add linear_sum linear.scaleR \linear f\ surj_f_inv_f \surj f\) apply (simp add: euclidean_representation u_def) done have "dist (x + (\i\Basis. (u \ i) *\<^sub>R inv f i)) x \ (\i\Basis. norm ((u \ i) *\<^sub>R inv f i))" by (simp add: dist_norm sum_norm_le) also have "... = (\i\Basis. \u \ i\ * norm (inv f i))" by simp also have "... \ (\i\Basis. \u \ i\) * B" by (simp add: B sum_distrib_right sum_mono mult_left_mono) also have "... \ DIM('b) * dist y (f x) * B" apply (rule mult_right_mono [OF sum_bounded_above]) using \0 < B\ by (auto simp: Basis_le_norm dist_norm u_def) also have "... < e" by (metis mult.commute mult.left_commute that) finally show "x + (\i\Basis. (u \ i) *\<^sub>R inv f i) \ A" by (rule e) qed qed then show "\y. dist y (f x) < \ \ y \ f ` A" using \e > 0\ \B > 0\ - by (auto simp: \_def field_split_simps mult_less_0_iff) + by (auto simp: \_def field_split_simps) qed qed corollary open_bijective_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" shows "open(f ` A) \ open A" proof assume "open(f ` A)" then have "open(f -` (f ` A))" using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) then show "open A" by (simp add: assms bij_is_inj inj_vimage_image_eq) next assume "open A" then show "open(f ` A)" by (simp add: assms bij_is_surj open_surjective_linear_image) qed corollary interior_bijective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") proof safe fix x assume x: "x \ ?lhs" then obtain T where "open T" and "x \ T" and "T \ f ` S" by (metis interiorE) then show "x \ ?rhs" by (metis (no_types, hide_lams) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) next fix x assume x: "x \ interior S" then show "f x \ interior (f ` S)" by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) qed lemma interior_injective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "linear f" "inj f" shows "interior(f ` S) = f ` (interior S)" by (simp add: linear_injective_imp_surjective assms bijI interior_bijective_linear_image) lemma interior_surjective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "linear f" "surj f" shows "interior(f ` S) = f ` (interior S)" by (simp add: assms interior_injective_linear_image linear_surjective_imp_injective) lemma interior_negations: fixes S :: "'a::euclidean_space set" shows "interior(uminus ` S) = image uminus (interior S)" by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) lemma connected_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "linear f" and "connected s" shows "connected (f ` s)" using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast subsection\<^marker>\tag unimportant\ \"Isometry" (up to constant bounds) of Injective Linear Map\ proposition injective_imp_isometric: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes s: "closed s" "subspace s" and f: "bounded_linear f" "\x\s. f x = 0 \ x = 0" shows "\e>0. \x\s. norm (f x) \ e * norm x" proof (cases "s \ {0::'a}") case True have "norm x \ norm (f x)" if "x \ s" for x proof - from True that have "x = 0" by auto then show ?thesis by simp qed then show ?thesis by (auto intro!: exI[where x=1]) next case False interpret f: bounded_linear f by fact from False obtain a where a: "a \ 0" "a \ s" by auto from False have "s \ {}" by auto let ?S = "{f x| x. x \ s \ norm x = norm a}" let ?S' = "{x::'a. x\s \ norm x = norm a}" let ?S'' = "{x::'a. norm x = norm a}" have "?S'' = frontier (cball 0 (norm a))" by (simp add: sphere_def dist_norm) then have "compact ?S''" by (metis compact_cball compact_frontier) moreover have "?S' = s \ ?S''" by auto ultimately have "compact ?S'" using closed_Int_compact[of s ?S''] using s(1) by auto moreover have *:"f ` ?S' = ?S" by auto ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto then have "closed ?S" using compact_imp_closed by auto moreover from a have "?S \ {}" by auto ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto then obtain b where "b\s" and ba: "norm b = norm a" and b: "\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" unfolding *[symmetric] unfolding image_iff by auto let ?e = "norm (f b) / norm b" have "norm b > 0" using ba and a and norm_ge_zero by auto moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF \b\s\] using \norm b >0\ by simp ultimately have "0 < norm (f b) / norm b" by simp moreover have "norm (f b) / norm b * norm x \ norm (f x)" if "x\s" for x proof (cases "x = 0") case True then show "norm (f b) / norm b * norm x \ norm (f x)" by auto next case False with \a \ 0\ have *: "0 < norm a / norm x" unfolding zero_less_norm_iff[symmetric] by simp have "\x\s. c *\<^sub>R x \ s" for c using s[unfolded subspace_def] by simp with \x \ s\ \x \ 0\ have "(norm a / norm x) *\<^sub>R x \ {x \ s. norm x = norm a}" by simp with \x \ 0\ \a \ 0\ show "norm (f b) / norm b * norm x \ norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] unfolding f.scaleR and ba by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) qed ultimately show ?thesis by auto qed proposition closed_injective_image_subspace: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 \ x = 0" "closed s" shows "closed(f ` s)" proof - obtain e where "e > 0" and e: "\x\s. e * norm x \ norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto show ?thesis using complete_isometric_image[OF \e>0\ assms(1,2) e] and assms(4) unfolding complete_eq_closed[symmetric] by auto qed subsection\<^marker>\tag unimportant\ \Some properties of a canonical subspace\ lemma closed_substandard: "closed {x::'a::euclidean_space. \i\Basis. P i \ x\i = 0}" (is "closed ?A") proof - let ?D = "{i\Basis. P i}" have "closed (\i\?D. {x::'a. x\i = 0})" - by (simp add: closed_INT closed_Collect_eq continuous_on_inner - continuous_on_const continuous_on_id) + by (simp add: closed_INT closed_Collect_eq continuous_on_inner) also have "(\i\?D. {x::'a. x\i = 0}) = ?A" by auto finally show "closed ?A" . qed lemma closed_subspace: fixes s :: "'a::euclidean_space set" assumes "subspace s" shows "closed s" proof - have "dim s \ card (Basis :: 'a set)" using dim_subset_UNIV by auto with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" by auto let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" using dim_substandard[of d] t d assms by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) then obtain f where f: "linear f" "f ` {x. \i\Basis. i \ d \ x \ i = 0} = s" "inj_on f {x. \i\Basis. i \ d \ x \ i = 0}" by blast interpret f: bounded_linear f using f by (simp add: linear_conv_bounded_linear) have "x \ ?t \ f x = 0 \ x = 0" for x using f.zero d f(3)[THEN inj_onD, of x 0] by auto moreover have "closed ?t" by (rule closed_substandard) moreover have "subspace ?t" by (rule subspace_substandard) ultimately show ?thesis using closed_injective_image_subspace[of ?t f] unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto qed lemma complete_subspace: "subspace s \ complete s" for s :: "'a::euclidean_space set" using complete_eq_closed closed_subspace by auto lemma closed_span [iff]: "closed (span s)" for s :: "'a::euclidean_space set" - by (simp add: closed_subspace subspace_span) + by (simp add: closed_subspace) lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") for s :: "'a::euclidean_space set" proof - have "?dc \ ?d" using closure_minimal[OF span_superset, of s] using closed_subspace[OF subspace_span, of s] using dim_subset[of "closure s" "span s"] by simp then show ?thesis using dim_subset[OF closure_subset, of s] by simp qed subsection \Set Distance\ lemma setdist_compact_closed: fixes A :: "'a::heine_borel set" assumes A: "compact A" and B: "closed B" and "A \ {}" "B \ {}" shows "\x \ A. \y \ B. dist x y = setdist A B" proof - obtain x where "x \ A" "setdist A B = infdist x B" by (metis A assms(3) setdist_attains_inf setdist_sym) moreover obtain y where"y \ B" "infdist x B = dist x y" using B \B \ {}\ infdist_attains_inf by blast ultimately show ?thesis using \x \ A\ \y \ B\ by auto qed lemma setdist_closed_compact: fixes S :: "'a::heine_borel set" assumes S: "closed S" and T: "compact T" and "S \ {}" "T \ {}" shows "\x \ S. \y \ T. dist x y = setdist S T" using setdist_compact_closed [OF T S \T \ {}\ \S \ {}\] by (metis dist_commute setdist_sym) lemma setdist_eq_0_compact_closed: assumes S: "compact S" and T: "closed T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" proof (cases "S = {} \ T = {}") case True then show ?thesis by force next case False then show ?thesis by (metis S T disjoint_iff_not_equal in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym) qed corollary setdist_gt_0_compact_closed: assumes S: "compact S" and T: "closed T" shows "setdist S T > 0 \ (S \ {} \ T \ {} \ S \ T = {})" using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] by linarith lemma setdist_eq_0_closed_compact: assumes S: "closed S" and T: "compact T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" using setdist_eq_0_compact_closed [OF T S] by (metis Int_commute setdist_sym) lemma setdist_eq_0_bounded: fixes S :: "'a::heine_borel set" assumes "bounded S \ bounded T" shows "setdist S T = 0 \ S = {} \ T = {} \ closure S \ closure T \ {}" proof (cases "S = {} \ T = {}") case False then show ?thesis using setdist_eq_0_compact_closed [of "closure S" "closure T"] setdist_eq_0_closed_compact [of "closure S" "closure T"] assms by (force simp: bounded_closure compact_eq_bounded_closed) qed force lemma setdist_eq_0_sing_1: "setdist {x} S = 0 \ S = {} \ x \ closure S" by (metis in_closure_iff_infdist_zero infdist_def infdist_eq_setdist) lemma setdist_eq_0_sing_2: "setdist S {x} = 0 \ S = {} \ x \ closure S" by (metis setdist_eq_0_sing_1 setdist_sym) lemma setdist_neq_0_sing_1: "\setdist {x} S = a; a \ 0\ \ S \ {} \ x \ closure S" by (metis setdist_closure_2 setdist_empty2 setdist_eq_0I singletonI) lemma setdist_neq_0_sing_2: "\setdist S {x} = a; a \ 0\ \ S \ {} \ x \ closure S" by (simp add: setdist_neq_0_sing_1 setdist_sym) lemma setdist_sing_in_set: "x \ S \ setdist {x} S = 0" by (simp add: setdist_eq_0I) lemma setdist_eq_0_closed: "closed S \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (simp add: setdist_eq_0_sing_1) lemma setdist_eq_0_closedin: shows "\closedin (top_of_set U) S; x \ U\ \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) lemma setdist_gt_0_closedin: shows "\closedin (top_of_set U) S; x \ U; S \ {}; x \ S\ \ setdist {x} S > 0" using less_eq_real_def setdist_eq_0_closedin by fastforce no_notation eucl_less (infix "Bernstein-Weierstrass and Stone-Weierstrass\ text\By L C Paulson (2015)\ theory Weierstrass_Theorems imports Uniform_Limit Path_Connected Derivative begin subsection \Bernstein polynomials\ definition\<^marker>\tag important\ Bernstein :: "[nat,nat,real] \ real" where "Bernstein n k x \ of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)" lemma Bernstein_nonneg: "\0 \ x; x \ 1\ \ 0 \ Bernstein n k x" by (simp add: Bernstein_def) lemma Bernstein_pos: "\0 < x; x < 1; k \ n\ \ 0 < Bernstein n k x" by (simp add: Bernstein_def) lemma sum_Bernstein [simp]: "(\k\n. Bernstein n k x) = 1" using binomial_ring [of x "1-x" n] by (simp add: Bernstein_def) lemma binomial_deriv1: "(\k\n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)" apply (rule DERIV_unique [where f = "\a. (a+b)^n" and x=a]) apply (subst binomial_ring) apply (rule derivative_eq_intros sum.cong | simp add: atMost_atLeast0)+ done lemma binomial_deriv2: "(\k\n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) = of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)" apply (rule DERIV_unique [where f = "\a. of_nat n * (a+b::real) ^ (n-1)" and x=a]) apply (subst binomial_deriv1 [symmetric]) apply (rule derivative_eq_intros sum.cong | simp add: Num.numeral_2_eq_2)+ done lemma sum_k_Bernstein [simp]: "(\k\n. real k * Bernstein n k x) = of_nat n * x" apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric]) apply (simp add: sum_distrib_right) apply (auto simp: Bernstein_def algebra_simps power_eq_if intro!: sum.cong) done lemma sum_kk_Bernstein [simp]: "(\k\n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2" proof - have "(\k\n. real k * (real k - 1) * Bernstein n k x) = (\k\n. real k * real (k - Suc 0) * real (n choose k) * x ^ (k - 2) * (1 - x) ^ (n - k) * x\<^sup>2)" proof (rule sum.cong [OF refl], simp) fix k assume "k \ n" then consider "k = 0" | "k = 1" | k' where "k = Suc (Suc k')" by (metis One_nat_def not0_implies_Suc) then show "k = 0 \ (real k - 1) * Bernstein n k x = real (k - Suc 0) * (real (n choose k) * (x ^ (k - 2) * ((1 - x) ^ (n - k) * x\<^sup>2)))" by cases (auto simp add: Bernstein_def power2_eq_square algebra_simps) qed also have "... = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2" by (subst binomial_deriv2 [of n x "1-x", simplified, symmetric]) (simp add: sum_distrib_right) also have "... = n * (n - 1) * x\<^sup>2" by auto finally show ?thesis by auto qed subsection \Explicit Bernstein version of the 1D Weierstrass approximation theorem\ theorem Bernstein_Weierstrass: fixes f :: "real \ real" assumes contf: "continuous_on {0..1} f" and e: "0 < e" shows "\N. \n x. N \ n \ x \ {0..1} \ \f x - (\k\n. f(k/n) * Bernstein n k x)\ < e" proof - have "bounded (f ` {0..1})" using compact_continuous_image compact_imp_bounded contf by blast then obtain M where M: "\x. 0 \ x \ x \ 1 \ \f x\ \ M" by (force simp add: bounded_iff) then have Mge0: "0 \ M" by force have ucontf: "uniformly_continuous_on {0..1} f" using compact_uniformly_continuous contf by blast then obtain d where d: "d>0" "\x x'. \ x \ {0..1}; x' \ {0..1}; \x' - x\ < d\ \ \f x' - f x\ < e/2" apply (rule uniformly_continuous_onE [where e = "e/2"]) using e by (auto simp: dist_norm) { fix n::nat and x::real assume n: "Suc (nat\4*M/(e*d\<^sup>2)\) \ n" and x: "0 \ x" "x \ 1" have "0 < n" using n by simp have ed0: "- (e * d\<^sup>2) < 0" using e \0 by simp also have "... \ M * 4" using \0\M\ by simp finally have [simp]: "real_of_int (nat \4 * M / (e * d\<^sup>2)\) = real_of_int \4 * M / (e * d\<^sup>2)\" using \0\M\ e \0 by (simp add: field_simps) have "4*M/(e*d\<^sup>2) + 1 \ real (Suc (nat\4*M/(e*d\<^sup>2)\))" by (simp add: real_nat_ceiling_ge) also have "... \ real n" using n by (simp add: field_simps) finally have nbig: "4*M/(e*d\<^sup>2) + 1 \ real n" . have sum_bern: "(\k\n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n" proof - have *: "\a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x" by (simp add: algebra_simps power2_eq_square) have "(\k\n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)" apply (simp add: * sum.distrib) apply (simp flip: sum_distrib_left add: mult.assoc) apply (simp add: algebra_simps power2_eq_square) done then have "(\k\n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n" by (simp add: power2_eq_square) then show ?thesis using n by (simp add: sum_divide_distrib field_split_simps mult.commute power2_commute) qed { fix k assume k: "k \ n" then have kn: "0 \ k / n" "k / n \ 1" by (auto simp: field_split_simps) consider (lessd) "\x - k / n\ < d" | (ged) "d \ \x - k / n\" by linarith then have "\(f x - f (k/n))\ \ e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2" proof cases case lessd then have "\(f x - f (k/n))\ < e/2" using d x kn by (simp add: abs_minus_commute) also have "... \ (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)" using Mge0 d by simp finally show ?thesis by simp next case ged then have dle: "d\<^sup>2 \ (x - k/n)\<^sup>2" by (metis d(1) less_eq_real_def power2_abs power_mono) have "\(f x - f (k/n))\ \ \f x\ + \f (k/n)\" by (rule abs_triangle_ineq4) also have "... \ M+M" by (meson M add_mono_thms_linordered_semiring(1) kn x) also have "... \ 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)" apply simp apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified]) using dle \d>0\ \M\0\ by auto also have "... \ e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2" using e by simp finally show ?thesis . qed } note * = this have "\f x - (\k\n. f(k / n) * Bernstein n k x)\ \ \\k\n. (f x - f(k / n)) * Bernstein n k x\" by (simp add: sum_subtractf sum_distrib_left [symmetric] algebra_simps) also have "... \ (\k\n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)" apply (rule order_trans [OF sum_abs sum_mono]) using * apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono) done also have "... \ e/2 + (2 * M) / (d\<^sup>2 * n)" apply (simp only: sum.distrib Rings.semiring_class.distrib_right sum_distrib_left [symmetric] mult.assoc sum_bern) using \d>0\ x apply (simp add: divide_simps Mge0 mult_le_one mult_left_le) done also have "... < e" apply (simp add: field_simps) using \d>0\ nbig e \n>0\ apply (simp add: field_split_simps algebra_simps) using ed0 by linarith finally have "\f x - (\k\n. f (real k / real n) * Bernstein n k x)\ < e" . } then show ?thesis by auto qed subsection \General Stone-Weierstrass theorem\ text\Source: Bruno Brosowski and Frank Deutsch. An Elementary Proof of the Stone-Weierstrass Theorem. Proceedings of the American Mathematical Society Volume 81, Number 1, January 1981. DOI: 10.2307/2043993 \<^url>\https://www.jstor.org/stable/2043993\\ locale function_ring_on = fixes R :: "('a::t2_space \ real) set" and S :: "'a set" assumes compact: "compact S" assumes continuous: "f \ R \ continuous_on S f" assumes add: "f \ R \ g \ R \ (\x. f x + g x) \ R" assumes mult: "f \ R \ g \ R \ (\x. f x * g x) \ R" assumes const: "(\_. c) \ R" assumes separable: "x \ S \ y \ S \ x \ y \ \f\R. f x \ f y" begin lemma minus: "f \ R \ (\x. - f x) \ R" by (frule mult [OF const [of "-1"]]) simp lemma diff: "f \ R \ g \ R \ (\x. f x - g x) \ R" unfolding diff_conv_add_uminus by (metis add minus) lemma power: "f \ R \ (\x. f x ^ n) \ R" by (induct n) (auto simp: const mult) lemma sum: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" by (induct I rule: finite_induct; simp add: const add) lemma prod: "\finite I; \i. i \ I \ f i \ R\ \ (\x. \i \ I. f i x) \ R" by (induct I rule: finite_induct; simp add: const mult) definition\<^marker>\tag important\ normf :: "('a::t2_space \ real) \ real" where "normf f \ SUP x\S. \f x\" lemma normf_upper: "\continuous_on S f; x \ S\ \ \f x\ \ normf f" apply (simp add: normf_def) apply (rule cSUP_upper, assumption) by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs) lemma normf_least: "S \ {} \ (\x. x \ S \ \f x\ \ M) \ normf f \ M" by (simp add: normf_def cSUP_least) end lemma (in function_ring_on) one: assumes U: "open U" and t0: "t0 \ S" "t0 \ U" and t1: "t1 \ S-U" shows "\V. open V \ t0 \ V \ S \ V \ U \ (\e>0. \f \ R. f ` S \ {0..1} \ (\t \ S \ V. f t < e) \ (\t \ S - U. f t > 1 - e))" proof - have "\pt \ R. pt t0 = 0 \ pt t > 0 \ pt ` S \ {0..1}" if t: "t \ S - U" for t proof - have "t \ t0" using t t0 by auto then obtain g where g: "g \ R" "g t \ g t0" using separable t0 by (metis Diff_subset subset_eq t) define h where [abs_def]: "h x = g x - g t0" for x have "h \ R" unfolding h_def by (fast intro: g const diff) then have hsq: "(\w. (h w)\<^sup>2) \ R" by (simp add: power2_eq_square mult) have "h t \ h t0" by (simp add: h_def g) then have "h t \ 0" by (simp add: h_def) then have ht2: "0 < (h t)^2" by simp also have "... \ normf (\w. (h w)\<^sup>2)" using t normf_upper [where x=t] continuous [OF hsq] by force finally have nfp: "0 < normf (\w. (h w)\<^sup>2)" . define p where [abs_def]: "p x = (1 / normf (\w. (h w)\<^sup>2)) * (h x)^2" for x have "p \ R" unfolding p_def by (fast intro: hsq const mult) moreover have "p t0 = 0" by (simp add: p_def h_def) moreover have "p t > 0" using nfp ht2 by (simp add: p_def) moreover have "\x. x \ S \ p x \ {0..1}" using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def) ultimately show "\pt \ R. pt t0 = 0 \ pt t > 0 \ pt ` S \ {0..1}" by auto qed then obtain pf where pf: "\t. t \ S-U \ pf t \ R \ pf t t0 = 0 \ pf t t > 0" and pf01: "\t. t \ S-U \ pf t ` S \ {0..1}" by metis have com_sU: "compact (S-U)" using compact closed_Int_compact U by (simp add: Diff_eq compact_Int_closed open_closed) have "\t. t \ S-U \ \A. open A \ A \ S = {x\S. 0 < pf t x}" apply (rule open_Collect_positive) by (metis pf continuous) then obtain Uf where Uf: "\t. t \ S-U \ open (Uf t) \ (Uf t) \ S = {x\S. 0 < pf t x}" by metis then have open_Uf: "\t. t \ S-U \ open (Uf t)" by blast have tUft: "\t. t \ S-U \ t \ Uf t" using pf Uf by blast then have *: "S-U \ (\x \ S-U. Uf x)" by blast obtain subU where subU: "subU \ S - U" "finite subU" "S - U \ (\x \ subU. Uf x)" by (blast intro: that compactE_image [OF com_sU open_Uf *]) then have [simp]: "subU \ {}" using t1 by auto then have cardp: "card subU > 0" using subU by (simp add: card_gt_0_iff) define p where [abs_def]: "p x = (1 / card subU) * (\t \ subU. pf t x)" for x have pR: "p \ R" unfolding p_def using subU pf by (fast intro: pf const mult sum) have pt0 [simp]: "p t0 = 0" using subU pf by (auto simp: p_def intro: sum.neutral) have pt_pos: "p t > 0" if t: "t \ S-U" for t proof - obtain i where i: "i \ subU" "t \ Uf i" using subU t by blast show ?thesis using subU i t apply (clarsimp simp: p_def field_split_simps) apply (rule sum_pos2 [OF \finite subU\]) using Uf t pf01 apply auto apply (force elim!: subsetCE) done qed have p01: "p x \ {0..1}" if t: "x \ S" for x proof - have "0 \ p x" using subU cardp t apply (simp add: p_def field_split_simps sum_nonneg) apply (rule sum_nonneg) using pf01 by force moreover have "p x \ 1" using subU cardp t apply (simp add: p_def field_split_simps sum_nonneg) apply (rule sum_bounded_above [where 'a=real and K=1, simplified]) using pf01 by force ultimately show ?thesis by auto qed have "compact (p ` (S-U))" by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR) then have "open (- (p ` (S-U)))" by (simp add: compact_imp_closed open_Compl) moreover have "0 \ - (p ` (S-U))" by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos) ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \ - (p ` (S-U))" by (auto simp: elim!: openE) then have pt_delta: "\x. x \ S-U \ p x \ delta0" by (force simp: ball_def dist_norm dest: p01) define \ where "\ = delta0/2" have "delta0 \ 1" using delta0 p01 [of t1] t1 by (force simp: ball_def dist_norm dest: p01) with delta0 have \01: "0 < \" "\ < 1" by (auto simp: \_def) have pt_\: "\x. x \ S-U \ p x \ \" using pt_delta delta0 by (force simp: \_def) have "\A. open A \ A \ S = {x\S. p x < \/2}" by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const]) then obtain V where V: "open V" "V \ S = {x\S. p x < \/2}" by blast define k where "k = nat\1/\\ + 1" have "k>0" by (simp add: k_def) have "k-1 \ 1/\" using \01 by (simp add: k_def) with \01 have "k \ (1+\)/\" by (auto simp: algebra_simps add_divide_distrib) also have "... < 2/\" using \01 by (auto simp: field_split_simps) finally have k2\: "k < 2/\" . have "1/\ < k" using \01 unfolding k_def by linarith with \01 k2\ have k\: "1 < k*\" "k*\ < 2" by (auto simp: field_split_simps) define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t have qR: "q n \ R" for n by (simp add: q_def const diff power pR) have q01: "\n t. t \ S \ q n t \ {0..1}" using p01 by (simp add: q_def power_le_one algebra_simps) have qt0 [simp]: "\n. n>0 \ q n t0 = 1" using t0 pf by (simp add: q_def power_0_left) { fix t and n::nat assume t: "t \ S \ V" with \k>0\ V have "k * p t < k * \ / 2" by force then have "1 - (k * \ / 2)^n \ 1 - (k * p t)^n" using \k>0\ p01 t by (simp add: power_mono) also have "... \ q n t" using Bernoulli_inequality [of "- ((p t)^n)" "k^n"] apply (simp add: q_def) by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t) finally have "1 - (k * \ / 2) ^ n \ q n t" . } note limitV = this { fix t and n::nat assume t: "t \ S - U" with \k>0\ U have "k * \ \ k * p t" by (simp add: pt_\) with k\ have kpt: "1 < k * p t" by (blast intro: less_le_trans) have ptn_pos: "0 < p t ^ n" using pt_pos [OF t] by simp have ptn_le: "p t ^ n \ 1" by (meson DiffE atLeastAtMost_iff p01 power_le_one t) have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n" using pt_pos [OF t] \k>0\ by (simp add: q_def) also have "... \ (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)" using pt_pos [OF t] \k>0\ apply simp apply (simp only: times_divide_eq_right [symmetric]) apply (rule mult_left_mono [of "1::real", simplified]) apply (simp_all add: power_mult_distrib) apply (rule zero_le_power) using ptn_le by linarith also have "... \ (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)" apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]]) using \k>0\ ptn_pos ptn_le apply (auto simp: power_mult_distrib) done also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)" using pt_pos [OF t] \k>0\ by (simp add: algebra_simps power_mult power2_eq_square flip: power_mult_distrib) also have "... \ (1/(k * (p t))^n) * 1" apply (rule mult_left_mono [OF power_le_one]) using pt_pos \k>0\ p01 power_le_one t apply auto done also have "... \ (1 / (k*\))^n" using \k>0\ \01 power_mono pt_\ t by (fastforce simp: field_simps) finally have "q n t \ (1 / (real k * \)) ^ n " . } note limitNonU = this define NN where "NN e = 1 + nat \max (ln e / ln (real k * \ / 2)) (- ln e / ln (real k * \))\" for e have NN: "of_nat (NN e) > ln e / ln (real k * \ / 2)" "of_nat (NN e) > - ln e / ln (real k * \)" if "0e. e>0 \ (k * \ / 2)^NN e < e" apply (subst Transcendental.ln_less_cancel_iff [symmetric]) prefer 3 apply (subst ln_realpow) using \k>0\ \\>0\ NN k\ apply (force simp add: field_simps)+ done have NN0: "(1/(k*\)) ^ (NN e) < e" if "e>0" for e proof - have "0 < ln (real k) + ln \" using \01(1) \0 < k\ k\(1) ln_gt_zero ln_mult by fastforce then have "real (NN e) * ln (1 / (real k * \)) < ln e" using k\(1) NN(2) [of e] that by (simp add: ln_div divide_simps) then have "exp (real (NN e) * ln (1 / (real k * \))) < e" by (metis exp_less_mono exp_ln that) then show ?thesis by (simp add: \01(1) \0 < k\ exp_of_nat_mult) qed { fix t and e::real assume "e>0" have "t \ S \ V \ 1 - q (NN e) t < e" "t \ S - U \ q (NN e) t < e" proof - assume t: "t \ S \ V" show "1 - q (NN e) t < e" by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF \e>0\]]) next assume t: "t \ S - U" show "q (NN e) t < e" using limitNonU [OF t] less_le_trans [OF NN0 [OF \e>0\]] not_le by blast qed } then have "\e. e > 0 \ \f\R. f ` S \ {0..1} \ (\t \ S \ V. f t < e) \ (\t \ S - U. 1 - e < f t)" using q01 by (rule_tac x="\x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR) moreover have t0V: "t0 \ V" "S \ V \ U" using pt_\ t0 U V \01 by fastforce+ ultimately show ?thesis using V t0V by blast qed text\Non-trivial case, with \<^term>\A\ and \<^term>\B\ both non-empty\ lemma (in function_ring_on) two_special: assumes A: "closed A" "A \ S" "a \ A" and B: "closed B" "B \ S" "b \ B" and disj: "A \ B = {}" and e: "0 < e" "e < 1" shows "\f \ R. f ` S \ {0..1} \ (\x \ A. f x < e) \ (\x \ B. f x > 1 - e)" proof - { fix w assume "w \ A" then have "open ( - B)" "b \ S" "w \ B" "w \ S" using assms by auto then have "\V. open V \ w \ V \ S \ V \ -B \ (\e>0. \f \ R. f ` S \ {0..1} \ (\x \ S \ V. f x < e) \ (\x \ S \ B. f x > 1 - e))" using one [of "-B" w b] assms \w \ A\ by simp } then obtain Vf where Vf: "\w. w \ A \ open (Vf w) \ w \ Vf w \ S \ Vf w \ -B \ (\e>0. \f \ R. f ` S \ {0..1} \ (\x \ S \ Vf w. f x < e) \ (\x \ S \ B. f x > 1 - e))" by metis then have open_Vf: "\w. w \ A \ open (Vf w)" by blast have tVft: "\w. w \ A \ w \ Vf w" using Vf by blast then have sum_max_0: "A \ (\x \ A. Vf x)" by blast have com_A: "compact A" using A by (metis compact compact_Int_closed inf.absorb_iff2) obtain subA where subA: "subA \ A" "finite subA" "A \ (\x \ subA. Vf x)" by (blast intro: that compactE_image [OF com_A open_Vf sum_max_0]) then have [simp]: "subA \ {}" using \a \ A\ by auto then have cardp: "card subA > 0" using subA by (simp add: card_gt_0_iff) have "\w. w \ A \ \f \ R. f ` S \ {0..1} \ (\x \ S \ Vf w. f x < e / card subA) \ (\x \ S \ B. f x > 1 - e / card subA)" using Vf e cardp by simp then obtain ff where ff: "\w. w \ A \ ff w \ R \ ff w ` S \ {0..1} \ (\x \ S \ Vf w. ff w x < e / card subA) \ (\x \ S \ B. ff w x > 1 - e / card subA)" by metis define pff where [abs_def]: "pff x = (\w \ subA. ff w x)" for x have pffR: "pff \ R" unfolding pff_def using subA ff by (auto simp: intro: prod) moreover have pff01: "pff x \ {0..1}" if t: "x \ S" for x proof - have "0 \ pff x" using subA cardp t apply (simp add: pff_def field_split_simps sum_nonneg) apply (rule Groups_Big.linordered_semidom_class.prod_nonneg) using ff by fastforce moreover have "pff x \ 1" using subA cardp t apply (simp add: pff_def field_split_simps sum_nonneg) apply (rule prod_mono [where g = "\x. 1", simplified]) using ff by fastforce ultimately show ?thesis by auto qed moreover { fix v x assume v: "v \ subA" and x: "x \ Vf v" "x \ S" from subA v have "pff x = ff v x * (\w \ subA - {v}. ff w x)" unfolding pff_def by (metis prod.remove) also have "... \ ff v x * 1" apply (rule Rings.ordered_semiring_class.mult_left_mono) apply (rule prod_mono [where g = "\x. 1", simplified]) using ff [THEN conjunct2, THEN conjunct1] v subA x apply auto apply (meson atLeastAtMost_iff contra_subsetD imageI) apply (meson atLeastAtMost_iff contra_subsetD image_eqI) using atLeastAtMost_iff by blast also have "... < e / card subA" using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x by auto also have "... \ e" using cardp e by (simp add: field_split_simps) finally have "pff x < e" . } then have "\x. x \ A \ pff x < e" using A Vf subA by (metis UN_E contra_subsetD) moreover { fix x assume x: "x \ B" then have "x \ S" using B by auto have "1 - e \ (1 - e / card subA) ^ card subA" using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp by (auto simp: field_simps) also have "... = (\w \ subA. 1 - e / card subA)" - by (simp add: prod_constant subA(2)) + by (simp add: subA(2)) also have "... < pff x" apply (simp add: pff_def) apply (rule prod_mono_strict [where f = "\x. 1 - e / card subA", simplified]) apply (simp_all add: subA(2)) apply (intro ballI conjI) using e apply (force simp: field_split_simps) using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x apply blast done finally have "1 - e < pff x" . } ultimately show ?thesis by blast qed lemma (in function_ring_on) two: assumes A: "closed A" "A \ S" and B: "closed B" "B \ S" and disj: "A \ B = {}" and e: "0 < e" "e < 1" shows "\f \ R. f ` S \ {0..1} \ (\x \ A. f x < e) \ (\x \ B. f x > 1 - e)" proof (cases "A \ {} \ B \ {}") case True then show ?thesis apply (simp flip: ex_in_conv) using assms apply safe apply (force simp add: intro!: two_special) done next case False with e show ?thesis apply simp apply (erule disjE) apply (rule_tac [2] x="\x. 0" in bexI) apply (rule_tac x="\x. 1" in bexI) apply (auto simp: const) done qed text\The special case where \<^term>\f\ is non-negative and \<^term>\e<1/3\\ lemma (in function_ring_on) Stone_Weierstrass_special: assumes f: "continuous_on S f" and fpos: "\x. x \ S \ f x \ 0" and e: "0 < e" "e < 1/3" shows "\g \ R. \x\S. \f x - g x\ < 2*e" proof - define n where "n = 1 + nat \normf f / e\" define A where "A j = {x \ S. f x \ (j - 1/3)*e}" for j :: nat define B where "B j = {x \ S. f x \ (j + 1/3)*e}" for j :: nat have ngt: "(n-1) * e \ normf f" "n\1" using e apply (simp_all add: n_def field_simps of_nat_Suc) by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq) then have ge_fx: "(n-1) * e \ f x" if "x \ S" for x using f normf_upper that by fastforce { fix j have A: "closed (A j)" "A j \ S" apply (simp_all add: A_def Collect_restrict) apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const]) apply (simp add: compact compact_imp_closed) done have B: "closed (B j)" "B j \ S" apply (simp_all add: B_def Collect_restrict) apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f]) apply (simp add: compact compact_imp_closed) done have disj: "(A j) \ (B j) = {}" using e by (auto simp: A_def B_def field_simps) have "\f \ R. f ` S \ {0..1} \ (\x \ A j. f x < e/n) \ (\x \ B j. f x > 1 - e/n)" apply (rule two) using e A B disj ngt apply simp_all done } then obtain xf where xfR: "\j. xf j \ R" and xf01: "\j. xf j ` S \ {0..1}" and xfA: "\x j. x \ A j \ xf j x < e/n" and xfB: "\x j. x \ B j \ xf j x > 1 - e/n" by metis define g where [abs_def]: "g x = e * (\i\n. xf i x)" for x have gR: "g \ R" unfolding g_def by (fast intro: mult const sum xfR) have gge0: "\x. x \ S \ g x \ 0" using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff sum_nonneg) have A0: "A 0 = {}" using fpos e by (fastforce simp: A_def) have An: "A n = S" using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff) have Asub: "A j \ A i" if "i\j" for i j using e that apply (clarsimp simp: A_def) apply (erule order_trans, simp) done { fix t assume t: "t \ S" define j where "j = (LEAST j. t \ A j)" have jn: "j \ n" using t An by (simp add: Least_le j_def) have Aj: "t \ A j" using t An by (fastforce simp add: j_def intro: LeastI) then have Ai: "t \ A i" if "i\j" for i using Asub [OF that] by blast then have fj1: "f t \ (j - 1/3)*e" by (simp add: A_def) then have Anj: "t \ A i" if "ii apply (simp add: j_def) using not_less_Least by blast have j1: "1 \ j" using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def) then have Anj: "t \ A (j-1)" using Least_le by (fastforce simp add: j_def) then have fj2: "(j - 4/3)*e < f t" using j1 t by (simp add: A_def of_nat_diff) have ***: "xf i t \ e/n" if "i\j" for i using xfA [OF Ai] that by (simp add: less_eq_real_def) { fix i assume "i+2 \ j" then obtain d where "i+2+d = j" using le_Suc_ex that by blast then have "t \ B i" using Anj e ge_fx [OF t] \1 \ n\ fpos [OF t] t apply (simp add: A_def B_def) apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc) apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"]) apply auto done then have "xf i t > 1 - e/n" by (rule xfB) } note **** = this have xf_le1: "\i. xf i t \ 1" using xf01 t by force have "g t = e * (\ii=j..n. xf i t)" using j1 jn e apply (simp add: g_def flip: distrib_left) apply (subst sum.union_disjoint [symmetric]) apply (auto simp: ivl_disj_un) done also have "... \ e*j + e * ((Suc n - j)*e/n)" apply (rule add_mono) apply (simp_all only: mult_le_cancel_left_pos e) apply (rule sum_bounded_above [OF xf_le1, where A = "lessThan j", simplified]) using sum_bounded_above [of "{j..n}" "\i. xf i t", OF ***] apply simp done also have "... \ j*e + e*(n - j + 1)*e/n " using \1 \ n\ e by (simp add: field_simps del: of_nat_Suc) also have "... \ j*e + e*e" using \1 \ n\ e j1 by (simp add: field_simps del: of_nat_Suc) also have "... < (j + 1/3)*e" using e by (auto simp: field_simps) finally have gj1: "g t < (j + 1 / 3) * e" . have gj2: "(j - 4/3)*e < g t" proof (cases "2 \ j") case False then have "j=1" using j1 by simp with t gge0 e show ?thesis by force next case True then have "(j - 4/3)*e < (j-1)*e - e^2" using e by (auto simp: of_nat_diff algebra_simps power2_eq_square) also have "... < (j-1)*e - ((j - 1)/n) * e^2" using e True jn by (simp add: power2_eq_square field_simps) also have "... = e * (j-1) * (1 - e/n)" by (simp add: power2_eq_square field_simps) also have "... \ e * (\i\j-2. xf i t)" using e apply simp apply (rule order_trans [OF _ sum_bounded_below [OF less_imp_le [OF ****]]]) using True apply (simp_all add: of_nat_Suc of_nat_diff) done also have "... \ g t" using jn e using e xf01 t apply (simp add: g_def zero_le_mult_iff image_subset_iff sum_nonneg) apply (rule Groups_Big.sum_mono2, auto) done finally show ?thesis . qed have "\f t - g t\ < 2 * e" using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps) } then show ?thesis by (rule_tac x=g in bexI) (auto intro: gR) qed text\The ``unpretentious'' formulation\ proposition (in function_ring_on) Stone_Weierstrass_basic: assumes f: "continuous_on S f" and e: "e > 0" shows "\g \ R. \x\S. \f x - g x\ < e" proof - have "\g \ R. \x\S. \(f x + normf f) - g x\ < 2 * min (e/2) (1/4)" apply (rule Stone_Weierstrass_special) apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const]) using normf_upper [OF f] apply force apply (simp add: e, linarith) done then obtain g where "g \ R" "\x\S. \g x - (f x + normf f)\ < e" by force then show ?thesis apply (rule_tac x="\x. g x - normf f" in bexI) apply (auto simp: algebra_simps intro: diff const) done qed theorem (in function_ring_on) Stone_Weierstrass: assumes f: "continuous_on S f" shows "\F\UNIV \ R. LIM n sequentially. F n :> uniformly_on S f" proof - { fix e::real assume e: "0 < e" then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e" by (auto simp: real_arch_inverse [of e]) { fix n :: nat and x :: 'a and g :: "'a \ real" assume n: "N \ n" "\x\S. \f x - g x\ < 1 / (1 + real n)" assume x: "x \ S" have "\ real (Suc n) < inverse e" using \N \ n\ N using less_imp_inverse_less by force then have "1 / (1 + real n) \ e" using e by (simp add: field_simps of_nat_Suc) then have "\f x - g x\ < e" using n(2) x by auto } note * = this have "\\<^sub>F n in sequentially. \x\S. \f x - (SOME g. g \ R \ (\x\S. \f x - g x\ < 1 / (1 + real n))) x\ < e" apply (rule eventually_sequentiallyI [of N]) apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *) done } then show ?thesis apply (rule_tac x="\n::nat. SOME g. g \ R \ (\x\S. \f x - g x\ < 1 / (1 + n))" in bexI) prefer 2 apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]]) unfolding uniform_limit_iff apply (auto simp: dist_norm abs_minus_commute) done qed text\A HOL Light formulation\ corollary Stone_Weierstrass_HOL: fixes R :: "('a::t2_space \ real) set" and S :: "'a set" assumes "compact S" "\c. P(\x. c::real)" "\f. P f \ continuous_on S f" "\f g. P(f) \ P(g) \ P(\x. f x + g x)" "\f g. P(f) \ P(g) \ P(\x. f x * g x)" "\x y. x \ S \ y \ S \ x \ y \ \f. P(f) \ f x \ f y" "continuous_on S f" "0 < e" shows "\g. P(g) \ (\x \ S. \f x - g x\ < e)" proof - interpret PR: function_ring_on "Collect P" apply unfold_locales using assms by auto show ?thesis using PR.Stone_Weierstrass_basic [OF \continuous_on S f\ \0 < e\] by blast qed subsection \Polynomial functions\ inductive real_polynomial_function :: "('a::real_normed_vector \ real) \ bool" where linear: "bounded_linear f \ real_polynomial_function f" | const: "real_polynomial_function (\x. c)" | add: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x + g x)" | mult: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x * g x)" declare real_polynomial_function.intros [intro] definition\<^marker>\tag important\ polynomial_function :: "('a::real_normed_vector \ 'b::real_normed_vector) \ bool" where "polynomial_function p \ (\f. bounded_linear f \ real_polynomial_function (f o p))" lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p" unfolding polynomial_function_def proof assume "real_polynomial_function p" then show " \f. bounded_linear f \ real_polynomial_function (f \ p)" proof (induction p rule: real_polynomial_function.induct) case (linear h) then show ?case by (auto simp: bounded_linear_compose real_polynomial_function.linear) next case (const h) then show ?case by (simp add: real_polynomial_function.const) next case (add h) then show ?case by (force simp add: bounded_linear_def linear_add real_polynomial_function.add) next case (mult h) then show ?case by (force simp add: real_bounded_linear const real_polynomial_function.mult) qed next assume [rule_format, OF bounded_linear_ident]: "\f. bounded_linear f \ real_polynomial_function (f \ p)" then show "real_polynomial_function p" by (simp add: o_def) qed lemma polynomial_function_const [iff]: "polynomial_function (\x. c)" by (simp add: polynomial_function_def o_def const) lemma polynomial_function_bounded_linear: "bounded_linear f \ polynomial_function f" by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear) lemma polynomial_function_id [iff]: "polynomial_function(\x. x)" by (simp add: polynomial_function_bounded_linear) lemma polynomial_function_add [intro]: "\polynomial_function f; polynomial_function g\ \ polynomial_function (\x. f x + g x)" by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def) lemma polynomial_function_mult [intro]: assumes f: "polynomial_function f" and g: "polynomial_function g" shows "polynomial_function (\x. f x *\<^sub>R g x)" using g apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR const real_polynomial_function.mult o_def) apply (rule mult) using f apply (auto simp: real_polynomial_function_eq) done lemma polynomial_function_cmul [intro]: assumes f: "polynomial_function f" shows "polynomial_function (\x. c *\<^sub>R f x)" by (rule polynomial_function_mult [OF polynomial_function_const f]) lemma polynomial_function_minus [intro]: assumes f: "polynomial_function f" shows "polynomial_function (\x. - f x)" using polynomial_function_cmul [OF f, of "-1"] by simp lemma polynomial_function_diff [intro]: "\polynomial_function f; polynomial_function g\ \ polynomial_function (\x. f x - g x)" unfolding add_uminus_conv_diff [symmetric] by (metis polynomial_function_add polynomial_function_minus) lemma polynomial_function_sum [intro]: "\finite I; \i. i \ I \ polynomial_function (\x. f x i)\ \ polynomial_function (\x. sum (f x) I)" by (induct I rule: finite_induct) auto lemma real_polynomial_function_minus [intro]: "real_polynomial_function f \ real_polynomial_function (\x. - f x)" using polynomial_function_minus [of f] by (simp add: real_polynomial_function_eq) lemma real_polynomial_function_diff [intro]: "\real_polynomial_function f; real_polynomial_function g\ \ real_polynomial_function (\x. f x - g x)" using polynomial_function_diff [of f] by (simp add: real_polynomial_function_eq) lemma real_polynomial_function_sum [intro]: "\finite I; \i. i \ I \ real_polynomial_function (\x. f x i)\ \ real_polynomial_function (\x. sum (f x) I)" using polynomial_function_sum [of I f] by (simp add: real_polynomial_function_eq) lemma real_polynomial_function_power [intro]: "real_polynomial_function f \ real_polynomial_function (\x. f x ^ n)" by (induct n) (simp_all add: const mult) lemma real_polynomial_function_compose [intro]: assumes f: "polynomial_function f" and g: "real_polynomial_function g" shows "real_polynomial_function (g o f)" using g apply (induction g rule: real_polynomial_function.induct) using f apply (simp_all add: polynomial_function_def o_def const add mult) done lemma polynomial_function_compose [intro]: assumes f: "polynomial_function f" and g: "polynomial_function g" shows "polynomial_function (g o f)" using g real_polynomial_function_compose [OF f] by (auto simp: polynomial_function_def o_def) lemma sum_max_0: fixes x::real (*in fact "'a::comm_ring_1"*) shows "(\i\max m n. x^i * (if i \ m then a i else 0)) = (\i\m. x^i * a i)" proof - have "(\i\max m n. x^i * (if i \ m then a i else 0)) = (\i\max m n. (if i \ m then x^i * a i else 0))" by (auto simp: algebra_simps intro: sum.cong) also have "... = (\i\m. (if i \ m then x^i * a i else 0))" by (rule sum.mono_neutral_right) auto also have "... = (\i\m. x^i * a i)" by (auto simp: algebra_simps intro: sum.cong) finally show ?thesis . qed lemma real_polynomial_function_imp_sum: assumes "real_polynomial_function f" shows "\a n::nat. f = (\x. \i\n. a i * x ^ i)" using assms proof (induct f) case (linear f) then show ?case apply (clarsimp simp add: real_bounded_linear) apply (rule_tac x="\i. if i=0 then 0 else c" in exI) apply (rule_tac x=1 in exI) apply (simp add: mult_ac) done next case (const c) show ?case apply (rule_tac x="\i. c" in exI) apply (rule_tac x=0 in exI) apply (auto simp: mult_ac of_nat_Suc) done case (add f1 f2) then obtain a1 n1 a2 n2 where "f1 = (\x. \i\n1. a1 i * x ^ i)" "f2 = (\x. \i\n2. a2 i * x ^ i)" by auto then show ?case apply (rule_tac x="\i. (if i \ n1 then a1 i else 0) + (if i \ n2 then a2 i else 0)" in exI) apply (rule_tac x="max n1 n2" in exI) using sum_max_0 [where m=n1 and n=n2] sum_max_0 [where m=n2 and n=n1] apply (simp add: sum.distrib algebra_simps max.commute) done case (mult f1 f2) then obtain a1 n1 a2 n2 where "f1 = (\x. \i\n1. a1 i * x ^ i)" "f2 = (\x. \i\n2. a2 i * x ^ i)" by auto then obtain b1 b2 where "f1 = (\x. \i\n1. b1 i * x ^ i)" "f2 = (\x. \i\n2. b2 i * x ^ i)" "b1 = (\i. if i\n1 then a1 i else 0)" "b2 = (\i. if i\n2 then a2 i else 0)" by auto then show ?case apply (rule_tac x="\i. \k\i. b1 k * b2 (i - k)" in exI) apply (rule_tac x="n1+n2" in exI) using polynomial_product [of n1 b1 n2 b2] apply (simp add: Set_Interval.atLeast0AtMost) done qed lemma real_polynomial_function_iff_sum: "real_polynomial_function f \ (\a n::nat. f = (\x. \i\n. a i * x ^ i))" apply (rule iffI) apply (erule real_polynomial_function_imp_sum) apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_sum) done lemma polynomial_function_iff_Basis_inner: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" shows "polynomial_function f \ (\b\Basis. real_polynomial_function (\x. inner (f x) b))" (is "?lhs = ?rhs") unfolding polynomial_function_def proof (intro iffI allI impI) assume "\h. bounded_linear h \ real_polynomial_function (h \ f)" then show ?rhs by (force simp add: bounded_linear_inner_left o_def) next fix h :: "'b \ real" assume rp: "\b\Basis. real_polynomial_function (\x. f x \ b)" and h: "bounded_linear h" have "real_polynomial_function (h \ (\x. \b\Basis. (f x \ b) *\<^sub>R b))" apply (rule real_polynomial_function_compose [OF _ linear [OF h]]) using rp apply (auto simp: real_polynomial_function_eq polynomial_function_mult) done then show "real_polynomial_function (h \ f)" by (simp add: euclidean_representation_sum_fun) qed subsection \Stone-Weierstrass theorem for polynomial functions\ text\First, we need to show that they are continuous, differentiable and separable.\ lemma continuous_real_polymonial_function: assumes "real_polynomial_function f" shows "continuous (at x) f" using assms by (induct f) (auto simp: linear_continuous_at) lemma continuous_polymonial_function: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "polynomial_function f" shows "continuous (at x) f" apply (rule euclidean_isCont) using assms apply (simp add: polynomial_function_iff_Basis_inner) apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR) done lemma continuous_on_polymonial_function: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "polynomial_function f" shows "continuous_on S f" using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on by blast lemma has_real_derivative_polynomial_function: assumes "real_polynomial_function p" shows "\p'. real_polynomial_function p' \ (\x. (p has_real_derivative (p' x)) (at x))" using assms proof (induct p) case (linear p) then show ?case by (force simp: real_bounded_linear const intro!: derivative_eq_intros) next case (const c) show ?case by (rule_tac x="\x. 0" in exI) auto case (add f1 f2) then obtain p1 p2 where "real_polynomial_function p1" "\x. (f1 has_real_derivative p1 x) (at x)" "real_polynomial_function p2" "\x. (f2 has_real_derivative p2 x) (at x)" by auto then show ?case apply (rule_tac x="\x. p1 x + p2 x" in exI) apply (auto intro!: derivative_eq_intros) done case (mult f1 f2) then obtain p1 p2 where "real_polynomial_function p1" "\x. (f1 has_real_derivative p1 x) (at x)" "real_polynomial_function p2" "\x. (f2 has_real_derivative p2 x) (at x)" by auto then show ?case using mult apply (rule_tac x="\x. f1 x * p2 x + f2 x * p1 x" in exI) apply (auto intro!: derivative_eq_intros) done qed lemma has_vector_derivative_polynomial_function: fixes p :: "real \ 'a::euclidean_space" assumes "polynomial_function p" obtains p' where "polynomial_function p'" "\x. (p has_vector_derivative (p' x)) (at x)" proof - { fix b :: 'a assume "b \ Basis" then obtain p' where p': "real_polynomial_function p'" and pd: "\x. ((\x. p x \ b) has_real_derivative p' x) (at x)" using assms [unfolded polynomial_function_iff_Basis_inner, rule_format] \b \ Basis\ has_real_derivative_polynomial_function by blast have "\q. polynomial_function q \ (\x. ((\u. (p u \ b) *\<^sub>R b) has_vector_derivative q x) (at x))" apply (rule_tac x="\x. p' x *\<^sub>R b" in exI) using \b \ Basis\ p' apply (simp add: polynomial_function_iff_Basis_inner inner_Basis) apply (auto intro: derivative_eq_intros pd) done } then obtain qf where qf: "\b. b \ Basis \ polynomial_function (qf b)" "\b x. b \ Basis \ ((\u. (p u \ b) *\<^sub>R b) has_vector_derivative qf b x) (at x)" by metis show ?thesis apply (rule_tac p'="\x. \b\Basis. qf b x" in that) apply (force intro: qf) apply (subst euclidean_representation_sum_fun [of p, symmetric]) apply (auto intro: has_vector_derivative_sum qf) done qed lemma real_polynomial_function_separable: fixes x :: "'a::euclidean_space" assumes "x \ y" shows "\f. real_polynomial_function f \ f x \ f y" proof - have "real_polynomial_function (\u. \b\Basis. (inner (x-u) b)^2)" apply (rule real_polynomial_function_sum) apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff const linear bounded_linear_inner_left) done then show ?thesis apply (intro exI conjI, assumption) using assms apply (force simp add: euclidean_eq_iff [of x y] sum_nonneg_eq_0_iff algebra_simps) done qed lemma Stone_Weierstrass_real_polynomial_function: fixes f :: "'a::euclidean_space \ real" assumes "compact S" "continuous_on S f" "0 < e" obtains g where "real_polynomial_function g" "\x. x \ S \ \f x - g x\ < e" proof - interpret PR: function_ring_on "Collect real_polynomial_function" apply unfold_locales using assms continuous_on_polymonial_function real_polynomial_function_eq apply (auto intro: real_polynomial_function_separable) done show ?thesis using PR.Stone_Weierstrass_basic [OF \continuous_on S f\ \0 < e\] that by blast qed theorem Stone_Weierstrass_polynomial_function: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "compact S" and f: "continuous_on S f" and e: "0 < e" shows "\g. polynomial_function g \ (\x \ S. norm(f x - g x) < e)" proof - { fix b :: 'b assume "b \ Basis" have "\p. real_polynomial_function p \ (\x \ S. \f x \ b - p x\ < e / DIM('b))" apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF S _, of "\x. f x \ b" "e / card Basis"]]) using e f - apply (auto simp: Euclidean_Space.DIM_positive intro: continuous_intros) + apply (auto intro: continuous_intros) done } then obtain pf where pf: "\b. b \ Basis \ real_polynomial_function (pf b) \ (\x \ S. \f x \ b - pf b x\ < e / DIM('b))" apply (rule bchoice [rule_format, THEN exE]) apply assumption apply (force simp add: intro: that) done have "polynomial_function (\x. \b\Basis. pf b x *\<^sub>R b)" using pf by (simp add: polynomial_function_sum polynomial_function_mult real_polynomial_function_eq) moreover { fix x assume "x \ S" have "norm (\b\Basis. (f x \ b) *\<^sub>R b - pf b x *\<^sub>R b) \ (\b\Basis. norm ((f x \ b) *\<^sub>R b - pf b x *\<^sub>R b))" by (rule norm_sum) also have "... < of_nat DIM('b) * (e / DIM('b))" apply (rule sum_bounded_above_strict) apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \x \ S\) apply (rule DIM_positive) done also have "... = e" - using DIM_positive by (simp add: field_simps) + by (simp add: field_simps) finally have "norm (\b\Basis. (f x \ b) *\<^sub>R b - pf b x *\<^sub>R b) < e" . } ultimately show ?thesis apply (subst euclidean_representation_sum_fun [of f, symmetric]) apply (rule_tac x="\x. \b\Basis. pf b x *\<^sub>R b" in exI) apply (auto simp flip: sum_subtractf) done qed proposition Stone_Weierstrass_uniform_limit: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "compact S" and f: "continuous_on S f" obtains g where "uniform_limit S g f sequentially" "\n. polynomial_function (g n)" proof - have pos: "inverse (Suc n) > 0" for n by auto obtain g where g: "\n. polynomial_function (g n)" "\x n. x \ S \ norm(f x - g n x) < inverse (Suc n)" using Stone_Weierstrass_polynomial_function[OF S f pos] by metis have "uniform_limit S g f sequentially" proof (rule uniform_limitI) fix e::real assume "0 < e" with LIMSEQ_inverse_real_of_nat have "\\<^sub>F n in sequentially. inverse (Suc n) < e" by (rule order_tendstoD) moreover have "\\<^sub>F n in sequentially. \x\S. dist (g n x) (f x) < inverse (Suc n)" using g by (simp add: dist_norm norm_minus_commute) ultimately show "\\<^sub>F n in sequentially. \x\S. dist (g n x) (f x) < e" by (eventually_elim) auto qed then show ?thesis using g(1) .. qed subsection\Polynomial functions as paths\ text\One application is to pick a smooth approximation to a path, or just pick a smooth path anyway in an open connected set\ lemma path_polynomial_function: fixes g :: "real \ 'b::euclidean_space" shows "polynomial_function g \ path g" by (simp add: path_def continuous_on_polymonial_function) lemma path_approx_polynomial_function: fixes g :: "real \ 'b::euclidean_space" assumes "path g" "0 < e" shows "\p. polynomial_function p \ pathstart p = pathstart g \ pathfinish p = pathfinish g \ (\t \ {0..1}. norm(p t - g t) < e)" proof - obtain q where poq: "polynomial_function q" and noq: "\x. x \ {0..1} \ norm (g x - q x) < e/4" using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms by (auto simp: path_def) have pf: "polynomial_function (\t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))" by (force simp add: poq) have *: "\t. t \ {0..1} \ norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)" apply (intro Real_Vector_Spaces.norm_add_less) using noq apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1) done show ?thesis apply (intro exI conjI) apply (rule pf) using * apply (auto simp add: pathstart_def pathfinish_def algebra_simps) done qed proposition connected_open_polynomial_connected: fixes S :: "'a::euclidean_space set" assumes S: "open S" "connected S" and "x \ S" "y \ S" shows "\g. polynomial_function g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" proof - have "path_connected S" using assms by (simp add: connected_open_path_connected) with \x \ S\ \y \ S\ obtain p where p: "path p" "path_image p \ S" "pathstart p = x" "pathfinish p = y" by (force simp: path_connected_def) have "\e. 0 < e \ (\x \ path_image p. ball x e \ S)" proof (cases "S = UNIV") case True then show ?thesis by (simp add: gt_ex) next case False then have "- S \ {}" by blast then show ?thesis apply (rule_tac x="setdist (path_image p) (-S)" in exI) using S p apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed) using setdist_le_dist [of _ "path_image p" _ "-S"] by fastforce qed then obtain e where "0 < e"and eb: "\x. x \ path_image p \ ball x e \ S" by auto show ?thesis using path_approx_polynomial_function [OF \path p\ \0 < e\] apply clarify apply (intro exI conjI, assumption) using p apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+ done qed lemma differentiable_componentwise_within: "f differentiable (at a within S) \ (\i \ Basis. (\x. f x \ i) differentiable at a within S)" proof - { assume "\i\Basis. \D. ((\x. f x \ i) has_derivative D) (at a within S)" then obtain f' where f': "\i. i \ Basis \ ((\x. f x \ i) has_derivative f' i) (at a within S)" by metis have eq: "(\x. (\j\Basis. f' j x *\<^sub>R j) \ i) = f' i" if "i \ Basis" for i using that by (simp add: inner_add_left inner_add_right) have "\D. \i\Basis. ((\x. f x \ i) has_derivative (\x. D x \ i)) (at a within S)" apply (rule_tac x="\x::'a. (\j\Basis. f' j x *\<^sub>R j) :: 'b" in exI) apply (simp add: eq f') done } then show ?thesis apply (simp add: differentiable_def) using has_derivative_componentwise_within by blast qed lemma polynomial_function_inner [intro]: fixes i :: "'a::euclidean_space" shows "polynomial_function g \ polynomial_function (\x. g x \ i)" apply (subst euclidean_representation [where x=i, symmetric]) apply (force simp: inner_sum_right polynomial_function_iff_Basis_inner polynomial_function_sum) done text\ Differentiability of real and vector polynomial functions.\ lemma differentiable_at_real_polynomial_function: "real_polynomial_function f \ f differentiable (at a within S)" by (induction f rule: real_polynomial_function.induct) (simp_all add: bounded_linear_imp_differentiable) lemma differentiable_on_real_polynomial_function: "real_polynomial_function p \ p differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on differentiable_at_real_polynomial_function) lemma differentiable_at_polynomial_function: fixes f :: "_ \ 'a::euclidean_space" shows "polynomial_function f \ f differentiable (at a within S)" by (metis differentiable_at_real_polynomial_function polynomial_function_iff_Basis_inner differentiable_componentwise_within) lemma differentiable_on_polynomial_function: fixes f :: "_ \ 'a::euclidean_space" shows "polynomial_function f \ f differentiable_on S" by (simp add: differentiable_at_polynomial_function differentiable_on_def) lemma vector_eq_dot_span: assumes "x \ span B" "y \ span B" and i: "\i. i \ B \ i \ x = i \ y" shows "x = y" proof - have "\i. i \ B \ orthogonal (x - y) i" by (simp add: i inner_commute inner_diff_right orthogonal_def) moreover have "x - y \ span B" by (simp add: assms span_diff) ultimately have "x - y = 0" using orthogonal_to_span orthogonal_self by blast then show ?thesis by simp qed lemma orthonormal_basis_expand: assumes B: "pairwise orthogonal B" and 1: "\i. i \ B \ norm i = 1" and "x \ span B" and "finite B" shows "(\i\B. (x \ i) *\<^sub>R i) = x" proof (rule vector_eq_dot_span [OF _ \x \ span B\]) show "(\i\B. (x \ i) *\<^sub>R i) \ span B" by (simp add: span_clauses span_sum) show "i \ (\i\B. (x \ i) *\<^sub>R i) = i \ x" if "i \ B" for i proof - have [simp]: "i \ j = (if j = i then 1 else 0)" if "j \ B" for j using B 1 that \i \ B\ by (force simp: norm_eq_1 orthogonal_def pairwise_def) have "i \ (\i\B. (x \ i) *\<^sub>R i) = (\j\B. x \ j * (i \ j))" by (simp add: inner_sum_right) also have "... = (\j\B. if j = i then x \ i else 0)" by (rule sum.cong; simp) also have "... = i \ x" by (simp add: \finite B\ that inner_commute sum.delta) finally show ?thesis . qed qed theorem Stone_Weierstrass_polynomial_function_subspace: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" and contf: "continuous_on S f" and "0 < e" and "subspace T" "f ` S \ T" obtains g where "polynomial_function g" "g ` S \ T" "\x. x \ S \ norm(f x - g x) < e" proof - obtain B where "B \ T" and orthB: "pairwise orthogonal B" and B1: "\x. x \ B \ norm x = 1" and "independent B" and cardB: "card B = dim T" and spanB: "span B = T" using orthonormal_basis_subspace \subspace T\ by metis then have "finite B" by (simp add: independent_imp_finite) then obtain n::nat and b where "B = b ` {i. i < n}" "inj_on b {i. i < n}" using finite_imp_nat_seg_image_inj_on by metis with cardB have "n = card B" "dim T = n" by (auto simp: card_image) have fx: "(\i\B. (f x \ i) *\<^sub>R i) = f x" if "x \ S" for x apply (rule orthonormal_basis_expand [OF orthB B1 _ \finite B\]) using \f ` S \ T\ spanB that by auto have cont: "continuous_on S (\x. \i\B. (f x \ i) *\<^sub>R i)" by (intro continuous_intros contf) obtain g where "polynomial_function g" and g: "\x. x \ S \ norm ((\i\B. (f x \ i) *\<^sub>R i) - g x) < e / (n+2)" using Stone_Weierstrass_polynomial_function [OF \compact S\ cont, of "e / real (n + 2)"] \0 < e\ by auto with fx have g: "\x. x \ S \ norm (f x - g x) < e / (n+2)" by auto show ?thesis proof show "polynomial_function (\x. \i\B. (g x \ i) *\<^sub>R i)" apply (rule polynomial_function_sum) apply (simp add: \finite B\) using \polynomial_function g\ by auto show "(\x. \i\B. (g x \ i) *\<^sub>R i) ` S \ T" using \B \ T\ by (blast intro: subspace_sum subspace_mul \subspace T\) show "norm (f x - (\i\B. (g x \ i) *\<^sub>R i)) < e" if "x \ S" for x proof - have orth': "pairwise (\i j. orthogonal ((f x \ i) *\<^sub>R i - (g x \ i) *\<^sub>R i) ((f x \ j) *\<^sub>R j - (g x \ j) *\<^sub>R j)) B" apply (rule pairwise_mono [OF orthB]) apply (auto simp: orthogonal_def inner_diff_right inner_diff_left) done then have "(norm (\i\B. (f x \ i) *\<^sub>R i - (g x \ i) *\<^sub>R i))\<^sup>2 = (\i\B. (norm ((f x \ i) *\<^sub>R i - (g x \ i) *\<^sub>R i))\<^sup>2)" by (simp add: norm_sum_Pythagorean [OF \finite B\ orth']) also have "... = (\i\B. (norm (((f x - g x) \ i) *\<^sub>R i))\<^sup>2)" by (simp add: algebra_simps) also have "... \ (\i\B. (norm (f x - g x))\<^sup>2)" apply (rule sum_mono) apply (simp add: B1) apply (rule order_trans [OF Cauchy_Schwarz_ineq]) by (simp add: B1 dot_square_norm) also have "... = n * norm (f x - g x)^2" by (simp add: \n = card B\) also have "... \ n * (e / (n+2))^2" apply (rule mult_left_mono) apply (meson dual_order.order_iff_strict g norm_ge_zero power_mono that, simp) done also have "... \ e^2 / (n+2)" using \0 < e\ by (simp add: divide_simps power2_eq_square) also have "... < e^2" using \0 < e\ by (simp add: divide_simps) finally have "(norm (\i\B. (f x \ i) *\<^sub>R i - (g x \ i) *\<^sub>R i))\<^sup>2 < e^2" . then have "(norm (\i\B. (f x \ i) *\<^sub>R i - (g x \ i) *\<^sub>R i)) < e" apply (rule power2_less_imp_less) using \0 < e\ by auto then show ?thesis using fx that by (simp add: sum_subtractf) qed qed qed hide_fact linear add mult const end