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 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 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 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) + by (rule Hausdorff_space_subtopology) (metis Hausdorff_space_euclidean Hausdorff_space_product_topology) end diff --git a/src/HOL/Analysis/Analysis.thy b/src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy +++ b/src/HOL/Analysis/Analysis.thy @@ -1,51 +1,52 @@ theory Analysis imports (* Linear Algebra *) Convex Determinants (* Topology *) Connected Abstract_Limits (* Functional Analysis *) Elementary_Normed_Spaces Norm_Arith (* Vector Analysis *) Convex_Euclidean_Space Operator_Norm (* Unsorted *) Line_Segment Derivative Cartesian_Euclidean_Space Weierstrass_Theorems (* Measure and Integration Theory *) Ball_Volume Integral_Test Improper_Integral Equivalence_Measurable_On_Borel Lebesgue_Integral_Substitution Embed_Measure Complete_Measure Radon_Nikodym Fashoda_Theorem Cross3 Homeomorphism Bounded_Continuous_Function Abstract_Topology Product_Topology Lindelof_Spaces Infinite_Products Infinite_Set_Sum Polytope Jordan_Curve Poly_Roots Generalised_Binomial_Theorem Gamma_Function Change_Of_Vars Multivariate_Analysis Simplex_Content FPS_Convergence Smooth_Paths Abstract_Euclidean_Space + Function_Metric begin end diff --git a/src/HOL/Analysis/Function_Metric.thy b/src/HOL/Analysis/Function_Metric.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Function_Metric.thy @@ -0,0 +1,386 @@ +(* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr with additions from LCP + License: BSD +*) + +section \Metrics on product spaces\ + +theory Function_Metric + imports + Function_Topology + Elementary_Metric_Spaces +begin + +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 + +end \ No newline at end of file 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,2061 +1,1688 @@ (* 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 + imports + Elementary_Topology + Abstract_Limits + Connected 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 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" + fixes B :: "'a::real_normed_vector 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) 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/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,4113 +1,4131 @@ (* 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) 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) 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 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 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 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 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 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 linepath_cnj': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) 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 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) 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) 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) 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) 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) next case equal then show ?thesis 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) 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) } 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.\ +lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" +proof - + have "\U V. open U \ open V \ x \ U \ y \ V \ disjnt U V" + if "x \ y" + for x y :: 'a + proof (intro exI conjI) + let ?r = "dist x y / 2" + have [simp]: "?r > 0" + by (simp add: that) + show "open (ball x ?r)" "open (ball y ?r)" "x \ (ball x ?r)" "y \ (ball y ?r)" + by (auto simp add: that) + show "disjnt (ball x ?r) (ball y ?r)" + unfolding disjnt_def by (simp add: disjoint_ballI) + qed + then show ?thesis + by (simp add: Hausdorff_space_def) +qed + 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 subsection\<^marker>\tag unimportant\ \Rectangular paths\ definition\<^marker>\tag unimportant\ rectpath where "rectpath a1 a3 = (let a2 = Complex (Re a3) (Im a1); a4 = Complex (Re a1) (Im a3) in linepath a1 a2 +++ linepath a2 a3 +++ linepath a3 a4 +++ linepath a4 a1)" lemma path_rectpath [simp, intro]: "path (rectpath a b)" by (simp add: Let_def rectpath_def) lemma pathstart_rectpath [simp]: "pathstart (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma pathfinish_rectpath [simp]: "pathfinish (rectpath a1 a3) = a1" by (simp add: rectpath_def Let_def) lemma simple_path_rectpath [simp, intro]: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "simple_path (rectpath a1 a3)" unfolding rectpath_def Let_def using assms by (intro simple_path_join_loop arc_join arc_linepath) (auto simp: complex_eq_iff path_image_join closed_segment_same_Re closed_segment_same_Im) lemma path_image_rectpath: assumes "Re a1 \ Re a3" "Im a1 \ Im a3" shows "path_image (rectpath a1 a3) = {z. Re z \ {Re a1, Re a3} \ Im z \ {Im a1..Im a3}} \ {z. Im z \ {Im a1, Im a3} \ Re z \ {Re a1..Re a3}}" (is "?lhs = ?rhs") proof - define a2 a4 where "a2 = Complex (Re a3) (Im a1)" and "a4 = Complex (Re a1) (Im a3)" have "?lhs = closed_segment a1 a2 \ closed_segment a2 a3 \ closed_segment a4 a3 \ closed_segment a1 a4" by (simp_all add: rectpath_def Let_def path_image_join closed_segment_commute a2_def a4_def Un_assoc) also have "\ = ?rhs" using assms by (auto simp: rectpath_def Let_def path_image_join a2_def a4_def closed_segment_same_Re closed_segment_same_Im closed_segment_eq_real_ivl) finally show ?thesis . qed lemma path_image_rectpath_subset_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ cbox a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff) lemma path_image_rectpath_inter_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) \ box a b = {}" using assms by (auto simp: path_image_rectpath in_box_complex_iff) lemma path_image_rectpath_cbox_minus_box: assumes "Re a \ Re b" "Im a \ Im b" shows "path_image (rectpath a b) = cbox a b - box a b" using assms by (auto simp: path_image_rectpath in_cbox_complex_iff in_box_complex_iff) end diff --git a/src/HOL/Analysis/Product_Topology.thy b/src/HOL/Analysis/Product_Topology.thy --- a/src/HOL/Analysis/Product_Topology.thy +++ b/src/HOL/Analysis/Product_Topology.thy @@ -1,577 +1,577 @@ section\The binary product topology\ theory Product_Topology -imports Function_Topology Abstract_Limits +imports Function_Topology begin -section \Product Topology\ +section \Product Topology\ subsection\Definition\ definition prod_topology :: "'a topology \ 'b topology \ ('a \ 'b) topology" where "prod_topology X Y \ topology (arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T}))" lemma open_product_open: assumes "open A" shows "\\. \ \ {S \ T |S T. open S \ open T} \ \ \ = A" proof - obtain f g where *: "\u. u \ A \ open (f u) \ open (g u) \ u \ (f u) \ (g u) \ (f u) \ (g u) \ A" using open_prod_def [of A] assms by metis let ?\ = "(\u. f u \ g u) ` A" show ?thesis by (rule_tac x="?\" in exI) (auto simp: dest: *) qed lemma open_product_open_eq: "(arbitrary union_of (\U. \S T. U = S \ T \ open S \ open T)) = open" by (force simp: union_of_def arbitrary_def intro: open_product_open open_Times) lemma openin_prod_topology: "openin (prod_topology X Y) = arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T})" unfolding prod_topology_def proof (rule topology_inverse') show "istopology (arbitrary union_of (\U. U \ {S \ T |S T. openin X S \ openin Y T}))" apply (rule istopology_base, simp) by (metis openin_Int Times_Int_Times) qed lemma topspace_prod_topology [simp]: "topspace (prod_topology X Y) = topspace X \ topspace Y" proof - have "topspace(prod_topology X Y) = \ (Collect (openin (prod_topology X Y)))" (is "_ = ?Z") unfolding topspace_def .. also have "\ = topspace X \ topspace Y" proof show "?Z \ topspace X \ topspace Y" apply (auto simp: openin_prod_topology union_of_def arbitrary_def) using openin_subset by force+ next have *: "\A B. topspace X \ topspace Y = A \ B \ openin X A \ openin Y B" by blast show "topspace X \ topspace Y \ ?Z" apply (rule Union_upper) using * by (simp add: openin_prod_topology arbitrary_union_of_inc) qed finally show ?thesis . qed lemma subtopology_Times: shows "subtopology (prod_topology X Y) (S \ T) = prod_topology (subtopology X S) (subtopology Y T)" proof - have "((\U. \S T. U = S \ T \ openin X S \ openin Y T) relative_to S \ T) = (\U. \S' T'. U = S' \ T' \ (openin X relative_to S) S' \ (openin Y relative_to T) T')" by (auto simp: relative_to_def Times_Int_Times fun_eq_iff) metis then show ?thesis by (simp add: topology_eq openin_prod_topology arbitrary_union_of_relative_to flip: openin_relative_to) qed lemma prod_topology_subtopology: "prod_topology (subtopology X S) Y = subtopology (prod_topology X Y) (S \ topspace Y)" "prod_topology X (subtopology Y T) = subtopology (prod_topology X Y) (topspace X \ T)" by (auto simp: subtopology_Times) lemma prod_topology_discrete_topology: "discrete_topology (S \ T) = prod_topology (discrete_topology S) (discrete_topology T)" by (auto simp: discrete_topology_unique openin_prod_topology intro: arbitrary_union_of_inc) lemma prod_topology_euclidean [simp]: "prod_topology euclidean euclidean = euclidean" by (simp add: prod_topology_def open_product_open_eq) lemma prod_topology_subtopology_eu [simp]: "prod_topology (subtopology euclidean S) (subtopology euclidean T) = subtopology euclidean (S \ T)" by (simp add: prod_topology_subtopology subtopology_subtopology Times_Int_Times) lemma openin_prod_topology_alt: "openin (prod_topology X Y) S \ (\x y. (x,y) \ S \ (\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ S))" apply (auto simp: openin_prod_topology arbitrary_union_of_alt, fastforce) by (metis mem_Sigma_iff) lemma open_map_fst: "open_map (prod_topology X Y) X fst" unfolding open_map_def openin_prod_topology_alt by (force simp: openin_subopen [of X "fst ` _"] intro: subset_fst_imageI) lemma open_map_snd: "open_map (prod_topology X Y) Y snd" unfolding open_map_def openin_prod_topology_alt by (force simp: openin_subopen [of Y "snd ` _"] intro: subset_snd_imageI) lemma openin_prod_Times_iff: "openin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ openin X S \ openin Y T" proof (cases "S = {} \ T = {}") case False then show ?thesis apply (simp add: openin_prod_topology_alt openin_subopen [of X S] openin_subopen [of Y T] times_subset_iff, safe) apply (meson|force)+ done qed force lemma closure_of_Times: "(prod_topology X Y) closure_of (S \ T) = (X closure_of S) \ (Y closure_of T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (clarsimp simp: closure_of_def openin_prod_topology_alt) blast show "?rhs \ ?lhs" by (clarsimp simp: closure_of_def openin_prod_topology_alt) (meson SigmaI subsetD) qed lemma closedin_prod_Times_iff: "closedin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ closedin X S \ closedin Y T" by (auto simp: closure_of_Times times_eq_iff simp flip: closure_of_eq) lemma interior_of_Times: "(prod_topology X Y) interior_of (S \ T) = (X interior_of S) \ (Y interior_of T)" proof (rule interior_of_unique) show "(X interior_of S) \ Y interior_of T \ S \ T" by (simp add: Sigma_mono interior_of_subset) show "openin (prod_topology X Y) ((X interior_of S) \ Y interior_of T)" by (simp add: openin_prod_Times_iff) next show "T' \ (X interior_of S) \ Y interior_of T" if "T' \ S \ T" "openin (prod_topology X Y) T'" for T' proof (clarsimp; intro conjI) fix a :: "'a" and b :: "'b" assume "(a, b) \ T'" with that obtain U V where UV: "openin X U" "openin Y V" "a \ U" "b \ V" "U \ V \ T'" by (metis openin_prod_topology_alt) then show "a \ X interior_of S" using interior_of_maximal_eq that(1) by fastforce show "b \ Y interior_of T" using UV interior_of_maximal_eq that(1) by (metis SigmaI mem_Sigma_iff subset_eq) qed qed subsection \Continuity\ lemma continuous_map_pairwise: "continuous_map Z (prod_topology X Y) f \ continuous_map Z X (fst \ f) \ continuous_map Z Y (snd \ f)" (is "?lhs = ?rhs") proof - let ?g = "fst \ f" and ?h = "snd \ f" have f: "f x = (?g x, ?h x)" for x by auto show ?thesis proof (cases "(\x \ topspace Z. ?g x \ topspace X) \ (\x \ topspace Z. ?h x \ topspace Y)") case True show ?thesis proof safe assume "continuous_map Z (prod_topology X Y) f" then have "openin Z {x \ topspace Z. fst (f x) \ U}" if "openin X U" for U unfolding continuous_map_def using True that apply clarify apply (drule_tac x="U \ topspace Y" in spec) by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong) with True show "continuous_map Z X (fst \ f)" by (auto simp: continuous_map_def) next assume "continuous_map Z (prod_topology X Y) f" then have "openin Z {x \ topspace Z. snd (f x) \ V}" if "openin Y V" for V unfolding continuous_map_def using True that apply clarify apply (drule_tac x="topspace X \ V" in spec) by (simp add: openin_prod_Times_iff mem_Times_iff cong: conj_cong) with True show "continuous_map Z Y (snd \ f)" by (auto simp: continuous_map_def) next assume Z: "continuous_map Z X (fst \ f)" "continuous_map Z Y (snd \ f)" have *: "openin Z {x \ topspace Z. f x \ W}" if "\w. w \ W \ \U V. openin X U \ openin Y V \ w \ U \ V \ U \ V \ W" for W proof (subst openin_subopen, clarify) fix x :: "'a" assume "x \ topspace Z" and "f x \ W" with that [OF \f x \ W\] obtain U V where UV: "openin X U" "openin Y V" "f x \ U \ V" "U \ V \ W" by auto with Z UV show "\T. openin Z T \ x \ T \ T \ {x \ topspace Z. f x \ W}" apply (rule_tac x="{x \ topspace Z. ?g x \ U} \ {x \ topspace Z. ?h x \ V}" in exI) apply (auto simp: \x \ topspace Z\ continuous_map_def) done qed show "continuous_map Z (prod_topology X Y) f" using True by (simp add: continuous_map_def openin_prod_topology_alt mem_Times_iff *) qed qed (auto simp: continuous_map_def) qed lemma continuous_map_paired: "continuous_map Z (prod_topology X Y) (\x. (f x,g x)) \ continuous_map Z X f \ continuous_map Z Y g" by (simp add: continuous_map_pairwise o_def) lemma continuous_map_pairedI [continuous_intros]: "\continuous_map Z X f; continuous_map Z Y g\ \ continuous_map Z (prod_topology X Y) (\x. (f x,g x))" by (simp add: continuous_map_pairwise o_def) lemma continuous_map_fst [continuous_intros]: "continuous_map (prod_topology X Y) X fst" using continuous_map_pairwise [of "prod_topology X Y" X Y id] by (simp add: continuous_map_pairwise) lemma continuous_map_snd [continuous_intros]: "continuous_map (prod_topology X Y) Y snd" using continuous_map_pairwise [of "prod_topology X Y" X Y id] by (simp add: continuous_map_pairwise) lemma continuous_map_fst_of [continuous_intros]: "continuous_map Z (prod_topology X Y) f \ continuous_map Z X (fst \ f)" by (simp add: continuous_map_pairwise) lemma continuous_map_snd_of [continuous_intros]: "continuous_map Z (prod_topology X Y) f \ continuous_map Z Y (snd \ f)" by (simp add: continuous_map_pairwise) lemma continuous_map_prod_fst: "i \ I \ continuous_map (prod_topology (product_topology (\i. Y) I) X) Y (\x. fst x i)" using continuous_map_componentwise_UNIV continuous_map_fst by fastforce lemma continuous_map_prod_snd: "i \ I \ continuous_map (prod_topology X (product_topology (\i. Y) I)) Y (\x. snd x i)" using continuous_map_componentwise_UNIV continuous_map_snd by fastforce lemma continuous_map_if_iff [simp]: "continuous_map X Y (\x. if P then f x else g x) \ continuous_map X Y (if P then f else g)" by simp lemma continuous_map_if [continuous_intros]: "\P \ continuous_map X Y f; ~P \ continuous_map X Y g\ \ continuous_map X Y (\x. if P then f x else g x)" by simp lemma continuous_map_subtopology_fst [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) X fst" using continuous_map_from_subtopology continuous_map_fst by force lemma continuous_map_subtopology_snd [continuous_intros]: "continuous_map (subtopology (prod_topology X Y) Z) Y snd" using continuous_map_from_subtopology continuous_map_snd by force lemma quotient_map_fst [simp]: "quotient_map(prod_topology X Y) X fst \ (topspace Y = {} \ topspace X = {})" by (auto simp: continuous_open_quotient_map open_map_fst continuous_map_fst) lemma quotient_map_snd [simp]: "quotient_map(prod_topology X Y) Y snd \ (topspace X = {} \ topspace Y = {})" by (auto simp: continuous_open_quotient_map open_map_snd continuous_map_snd) lemma retraction_map_fst: "retraction_map (prod_topology X Y) X fst \ (topspace Y = {} \ topspace X = {})" proof (cases "topspace Y = {}") case True then show ?thesis using continuous_map_image_subset_topspace by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) next case False have "\g. continuous_map X (prod_topology X Y) g \ (\x\topspace X. fst (g x) = x)" if y: "y \ topspace Y" for y by (rule_tac x="\x. (x,y)" in exI) (auto simp: y continuous_map_paired) with False have "retraction_map (prod_topology X Y) X fst" by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst) with False show ?thesis by simp qed lemma retraction_map_snd: "retraction_map (prod_topology X Y) Y snd \ (topspace X = {} \ topspace Y = {})" proof (cases "topspace X = {}") case True then show ?thesis using continuous_map_image_subset_topspace by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_fst continuous_map_on_empty) next case False have "\g. continuous_map Y (prod_topology X Y) g \ (\y\topspace Y. snd (g y) = y)" if x: "x \ topspace X" for x by (rule_tac x="\y. (x,y)" in exI) (auto simp: x continuous_map_paired) with False have "retraction_map (prod_topology X Y) Y snd" by (fastforce simp: retraction_map_def retraction_maps_def continuous_map_snd) with False show ?thesis by simp qed lemma continuous_map_of_fst: "continuous_map (prod_topology X Y) Z (f \ fst) \ topspace Y = {} \ continuous_map X Z f" proof (cases "topspace Y = {}") case True then show ?thesis by (simp add: continuous_map_on_empty) next case False then show ?thesis by (simp add: continuous_compose_quotient_map_eq quotient_map_fst) qed lemma continuous_map_of_snd: "continuous_map (prod_topology X Y) Z (f \ snd) \ topspace X = {} \ continuous_map Y Z f" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: continuous_map_on_empty) next case False then show ?thesis by (simp add: continuous_compose_quotient_map_eq quotient_map_snd) qed lemma continuous_map_prod_top: "continuous_map (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) \ topspace (prod_topology X Y) = {} \ continuous_map X X' f \ continuous_map Y Y' g" proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis by (simp add: continuous_map_on_empty) next case False then show ?thesis by (simp add: continuous_map_paired case_prod_unfold continuous_map_of_fst [unfolded o_def] continuous_map_of_snd [unfolded o_def]) qed lemma in_prod_topology_closure_of: assumes "z \ (prod_topology X Y) closure_of S" shows "fst z \ X closure_of (fst ` S)" "snd z \ Y closure_of (snd ` S)" using assms continuous_map_eq_image_closure_subset continuous_map_fst apply fastforce using assms continuous_map_eq_image_closure_subset continuous_map_snd apply fastforce done proposition compact_space_prod_topology: "compact_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ compact_space X \ compact_space Y" proof (cases "topspace(prod_topology X Y) = {}") case True then show ?thesis using compact_space_topspace_empty by blast next case False then have non_mt: "topspace X \ {}" "topspace Y \ {}" by auto have "compact_space X" "compact_space Y" if "compact_space(prod_topology X Y)" proof - have "compactin X (fst ` (topspace X \ topspace Y))" by (metis compact_space_def continuous_map_fst image_compactin that topspace_prod_topology) moreover have "compactin Y (snd ` (topspace X \ topspace Y))" by (metis compact_space_def continuous_map_snd image_compactin that topspace_prod_topology) ultimately show "compact_space X" "compact_space Y" by (simp_all add: non_mt compact_space_def) qed moreover define \ where "\ \ (\V. topspace X \ V) ` Collect (openin Y)" define \ where "\ \ (\U. U \ topspace Y) ` Collect (openin X)" have "compact_space(prod_topology X Y)" if "compact_space X" "compact_space Y" proof (rule Alexander_subbase_alt) show toptop: "topspace X \ topspace Y \ \(\ \ \)" unfolding \_def \_def by auto fix \ :: "('a \ 'b) set set" assume \: "\ \ \ \ \" "topspace X \ topspace Y \ \\" then obtain \' \' where XY: "\' \ \" "\' \ \" and \eq: "\ = \' \ \'" using subset_UnE by metis then have sub: "topspace X \ topspace Y \ \(\' \ \')" using \ by simp obtain \ \ where \: "\U. U \ \ \ openin X U" "\' = (\U. U \ topspace Y) ` \" and \: "\V. V \ \ \ openin Y V" "\' = (\V. topspace X \ V) ` \" using XY by (clarsimp simp add: \_def \_def subset_image_iff) (force simp add: subset_iff) have "\\. finite \ \ \ \ \' \ \' \ topspace X \ topspace Y \ \ \" proof - have "topspace X \ \\ \ topspace Y \ \\" using \ \ \ \eq by auto then have *: "\\. finite \ \ (\x \ \. x \ (\V. topspace X \ V) ` \ \ x \ (\U. U \ topspace Y) ` \) \ (topspace X \ topspace Y \ \\)" proof assume "topspace X \ \\" with \compact_space X\ \ obtain \ where "finite \" "\ \ \" "topspace X \ \\" by (meson compact_space_alt) with that show ?thesis by (rule_tac x="(\D. D \ topspace Y) ` \" in exI) auto next assume "topspace Y \ \\" with \compact_space Y\ \ obtain \ where "finite \" "\ \ \" "topspace Y \ \\" by (meson compact_space_alt) with that show ?thesis by (rule_tac x="(\C. topspace X \ C) ` \" in exI) auto qed then show ?thesis using that \ \ by blast qed then show "\\. finite \ \ \ \ \ \ topspace X \ topspace Y \ \ \" using \ \eq by blast next have "(finite intersection_of (\x. x \ \ \ x \ \) relative_to topspace X \ topspace Y) = (\U. \S T. U = S \ T \ openin X S \ openin Y T)" (is "?lhs = ?rhs") proof - have "?rhs U" if "?lhs U" for U proof - have "topspace X \ topspace Y \ \ T \ {A \ B |A B. A \ Collect (openin X) \ B \ Collect (openin Y)}" if "finite T" "T \ \ \ \" for T using that proof induction case (insert B \) then show ?case unfolding \_def \_def apply (simp add: Int_ac subset_eq image_def) apply (metis (no_types) openin_Int openin_topspace Times_Int_Times) done qed auto then show ?thesis using that by (auto simp: subset_eq elim!: relative_toE intersection_ofE) qed moreover have "?lhs Z" if Z: "?rhs Z" for Z proof - obtain U V where "Z = U \ V" "openin X U" "openin Y V" using Z by blast then have UV: "U \ V = (topspace X \ topspace Y) \ (U \ V)" by (simp add: Sigma_mono inf_absorb2 openin_subset) moreover have "?lhs ((topspace X \ topspace Y) \ (U \ V))" proof (rule relative_to_inc) show "(finite intersection_of (\x. x \ \ \ x \ \)) (U \ V)" apply (simp add: intersection_of_def \_def \_def) apply (rule_tac x="{(U \ topspace Y),(topspace X \ V)}" in exI) using \openin X U\ \openin Y V\ openin_subset UV apply (fastforce simp add:) done qed ultimately show ?thesis using \Z = U \ V\ by auto qed ultimately show ?thesis by meson qed then show "topology (arbitrary union_of (finite intersection_of (\x. x \ \ \ \) relative_to (topspace X \ topspace Y))) = prod_topology X Y" by (simp add: prod_topology_def) qed ultimately show ?thesis using False by blast qed lemma compactin_Times: "compactin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ compactin X S \ compactin Y T" by (auto simp: compactin_subspace subtopology_Times compact_space_prod_topology) subsection\Homeomorphic maps\ lemma homeomorphic_maps_prod: "homeomorphic_maps (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) (\(x,y). (f' x, g' y)) \ topspace(prod_topology X Y) = {} \ topspace(prod_topology X' Y') = {} \ homeomorphic_maps X X' f f' \ homeomorphic_maps Y Y' g g'" unfolding homeomorphic_maps_def continuous_map_prod_top by (auto simp: continuous_map_def homeomorphic_maps_def continuous_map_prod_top) lemma homeomorphic_maps_swap: "homeomorphic_maps (prod_topology X Y) (prod_topology Y X) (\(x,y). (y,x)) (\(y,x). (x,y))" by (auto simp: homeomorphic_maps_def case_prod_unfold continuous_map_fst continuous_map_pairedI continuous_map_snd) lemma homeomorphic_map_swap: "homeomorphic_map (prod_topology X Y) (prod_topology Y X) (\(x,y). (y,x))" using homeomorphic_map_maps homeomorphic_maps_swap by metis lemma embedding_map_graph: "embedding_map X (prod_topology X Y) (\x. (x, f x)) \ continuous_map X Y f" (is "?lhs = ?rhs") proof assume L: ?lhs have "snd \ (\x. (x, f x)) = f" by force moreover have "continuous_map X Y (snd \ (\x. (x, f x)))" using L unfolding embedding_map_def by (meson continuous_map_in_subtopology continuous_map_snd_of homeomorphic_imp_continuous_map) ultimately show ?rhs by simp next assume R: ?rhs then show ?lhs unfolding homeomorphic_map_maps embedding_map_def homeomorphic_maps_def by (rule_tac x=fst in exI) (auto simp: continuous_map_in_subtopology continuous_map_paired continuous_map_from_subtopology continuous_map_fst) qed lemma homeomorphic_space_prod_topology: "\X homeomorphic_space X''; Y homeomorphic_space Y'\ \ prod_topology X Y homeomorphic_space prod_topology X'' Y'" using homeomorphic_maps_prod unfolding homeomorphic_space_def by blast lemma prod_topology_homeomorphic_space_left: "topspace Y = {b} \ prod_topology X Y homeomorphic_space X" unfolding homeomorphic_space_def by (rule_tac x=fst in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps) lemma prod_topology_homeomorphic_space_right: "topspace X = {a} \ prod_topology X Y homeomorphic_space Y" unfolding homeomorphic_space_def by (rule_tac x=snd in exI) (simp add: homeomorphic_map_def inj_on_def flip: homeomorphic_map_maps) lemma homeomorphic_space_prod_topology_sing1: "b \ topspace Y \ X homeomorphic_space (prod_topology X (subtopology Y {b}))" by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_left topspace_subtopology) lemma homeomorphic_space_prod_topology_sing2: "a \ topspace X \ Y homeomorphic_space (prod_topology (subtopology X {a}) Y)" by (metis empty_subsetI homeomorphic_space_sym inf.absorb_iff2 insert_subset prod_topology_homeomorphic_space_right topspace_subtopology) lemma topological_property_of_prod_component: assumes major: "P(prod_topology X Y)" and X: "\x. \x \ topspace X; P(prod_topology X Y)\ \ P(subtopology (prod_topology X Y) ({x} \ topspace Y))" and Y: "\y. \y \ topspace Y; P(prod_topology X Y)\ \ P(subtopology (prod_topology X Y) (topspace X \ {y}))" and PQ: "\X X'. X homeomorphic_space X' \ (P X \ Q X')" and PR: "\X X'. X homeomorphic_space X' \ (P X \ R X')" shows "topspace(prod_topology X Y) = {} \ Q X \ R Y" proof - have "Q X \ R Y" if "topspace(prod_topology X Y) \ {}" proof - from that obtain a b where a: "a \ topspace X" and b: "b \ topspace Y" by force show ?thesis using X [OF a major] and Y [OF b major] homeomorphic_space_prod_topology_sing1 [OF b, of X] homeomorphic_space_prod_topology_sing2 [OF a, of Y] by (simp add: subtopology_Times) (meson PQ PR homeomorphic_space_prod_topology_sing2 homeomorphic_space_sym) qed then show ?thesis by metis qed lemma limitin_pairwise: "limitin (prod_topology X Y) f l F \ limitin X (fst \ f) (fst l) F \ limitin Y (snd \ f) (snd l) F" (is "?lhs = ?rhs") proof assume ?lhs then obtain a b where ev: "\U. \(a,b) \ U; openin (prod_topology X Y) U\ \ \\<^sub>F x in F. f x \ U" and a: "a \ topspace X" and b: "b \ topspace Y" and l: "l = (a,b)" by (auto simp: limitin_def) moreover have "\\<^sub>F x in F. fst (f x) \ U" if "openin X U" "a \ U" for U proof - have "\\<^sub>F c in F. f c \ U \ topspace Y" using b that ev [of "U \ topspace Y"] by (auto simp: openin_prod_topology_alt) then show ?thesis by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) qed moreover have "\\<^sub>F x in F. snd (f x) \ U" if "openin Y U" "b \ U" for U proof - have "\\<^sub>F c in F. f c \ topspace X \ U" using a that ev [of "topspace X \ U"] by (auto simp: openin_prod_topology_alt) then show ?thesis by (rule eventually_mono) (metis (mono_tags, lifting) SigmaE2 prod.collapse) qed ultimately show ?rhs by (simp add: limitin_def) next have "limitin (prod_topology X Y) f (a,b) F" if "limitin X (fst \ f) a F" "limitin Y (snd \ f) b F" for a b using that proof (clarsimp simp: limitin_def) fix Z :: "('a \ 'b) set" assume a: "a \ topspace X" "\U. openin X U \ a \ U \ (\\<^sub>F x in F. fst (f x) \ U)" and b: "b \ topspace Y" "\U. openin Y U \ b \ U \ (\\<^sub>F x in F. snd (f x) \ U)" and Z: "openin (prod_topology X Y) Z" "(a, b) \ Z" then obtain U V where "openin X U" "openin Y V" "a \ U" "b \ V" "U \ V \ Z" using Z by (force simp: openin_prod_topology_alt) then have "\\<^sub>F x in F. fst (f x) \ U" "\\<^sub>F x in F. snd (f x) \ V" by (simp_all add: a b) then show "\\<^sub>F x in F. f x \ Z" by (rule eventually_elim2) (use \U \ V \ Z\ subsetD in auto) qed then show "?rhs \ ?lhs" by (metis prod.collapse) qed end diff --git a/src/HOL/Analysis/T1_Spaces.thy b/src/HOL/Analysis/T1_Spaces.thy --- a/src/HOL/Analysis/T1_Spaces.thy +++ b/src/HOL/Analysis/T1_Spaces.thy @@ -1,719 +1,701 @@ section\T1 and Hausdorff spaces\ theory T1_Spaces -imports Product_Topology +imports Product_Topology begin section\T1 spaces with equivalences to many naturally "nice" properties. \ definition t1_space where "t1_space X \ \x \ topspace X. \y \ topspace X. x\y \ (\U. openin X U \ x \ U \ y \ U)" lemma t1_space_expansive: "\topspace Y = topspace X; \U. openin X U \ openin Y U\ \ t1_space X \ t1_space Y" by (metis t1_space_def) lemma t1_space_alt: "t1_space X \ (\x \ topspace X. \y \ topspace X. x\y \ (\U. closedin X U \ x \ U \ y \ U))" by (metis DiffE DiffI closedin_def openin_closedin_eq t1_space_def) lemma t1_space_empty: "topspace X = {} \ t1_space X" by (simp add: t1_space_def) lemma t1_space_derived_set_of_singleton: "t1_space X \ (\x \ topspace X. X derived_set_of {x} = {})" apply (simp add: t1_space_def derived_set_of_def, safe) apply (metis openin_topspace) by force lemma t1_space_derived_set_of_finite: "t1_space X \ (\S. finite S \ X derived_set_of S = {})" proof (intro iffI allI impI) fix S :: "'a set" assume "finite S" then have fin: "finite ((\x. {x}) ` (topspace X \ S))" by blast assume "t1_space X" then have "X derived_set_of (\x \ topspace X \ S. {x}) = {}" unfolding derived_set_of_Union [OF fin] by (auto simp: t1_space_derived_set_of_singleton) then have "X derived_set_of (topspace X \ S) = {}" by simp then show "X derived_set_of S = {}" by simp qed (auto simp: t1_space_derived_set_of_singleton) lemma t1_space_closedin_singleton: "t1_space X \ (\x \ topspace X. closedin X {x})" apply (rule iffI) apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton) using t1_space_alt by auto lemma closedin_t1_singleton: "\t1_space X; a \ topspace X\ \ closedin X {a}" by (simp add: t1_space_closedin_singleton) lemma t1_space_closedin_finite: "t1_space X \ (\S. finite S \ S \ topspace X \ closedin X S)" apply (rule iffI) apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_finite) by (simp add: t1_space_closedin_singleton) lemma closure_of_singleton: "t1_space X \ X closure_of {a} = (if a \ topspace X then {a} else {})" by (simp add: closure_of_eq t1_space_closedin_singleton closure_of_eq_empty_gen) lemma separated_in_singleton: assumes "t1_space X" shows "separatedin X {a} S \ a \ topspace X \ S \ topspace X \ (a \ X closure_of S)" "separatedin X S {a} \ a \ topspace X \ S \ topspace X \ (a \ X closure_of S)" unfolding separatedin_def using assms closure_of closure_of_singleton by fastforce+ lemma t1_space_openin_delete: "t1_space X \ (\U x. openin X U \ x \ U \ openin X (U - {x}))" apply (rule iffI) apply (meson closedin_t1_singleton in_mono openin_diff openin_subset) by (simp add: closedin_def t1_space_closedin_singleton) lemma t1_space_openin_delete_alt: "t1_space X \ (\U x. openin X U \ openin X (U - {x}))" by (metis Diff_empty Diff_insert0 t1_space_openin_delete) lemma t1_space_singleton_Inter_open: "t1_space X \ (\x \ topspace X. \{U. openin X U \ x \ U} = {x})" (is "?P=?Q") and t1_space_Inter_open_supersets: "t1_space X \ (\S. S \ topspace X \ \{U. openin X U \ S \ U} = S)" (is "?P=?R") proof - have "?R \ ?Q" apply clarify apply (drule_tac x="{x}" in spec, simp) done moreover have "?Q \ ?P" apply (clarsimp simp add: t1_space_def) apply (drule_tac x=x in bspec) apply (simp_all add: set_eq_iff) by (metis (no_types, lifting)) moreover have "?P \ ?R" proof (clarsimp simp add: t1_space_closedin_singleton, rule subset_antisym) fix S assume S: "\x\topspace X. closedin X {x}" "S \ topspace X" then show "\ {U. openin X U \ S \ U} \ S" apply clarsimp by (metis Diff_insert_absorb Set.set_insert closedin_def openin_topspace subset_insert) qed force ultimately show "?P=?Q" "?P=?R" by auto qed lemma t1_space_derived_set_of_infinite_openin: "t1_space X \ (\S. X derived_set_of S = {x \ topspace X. \U. x \ U \ openin X U \ infinite(S \ U)})" (is "_ = ?rhs") proof assume "t1_space X" show ?rhs proof safe fix S x U assume "x \ X derived_set_of S" "x \ U" "openin X U" "finite (S \ U)" with \t1_space X\ show "False" apply (simp add: t1_space_derived_set_of_finite) by (metis IntI empty_iff empty_subsetI inf_commute openin_Int_derived_set_of_subset subset_antisym) next fix S x have eq: "(\y. (y \ x) \ y \ S \ y \ T) \ ~((S \ T) \ {x})" for x S T by blast assume "x \ topspace X" "\U. x \ U \ openin X U \ infinite (S \ U)" then show "x \ X derived_set_of S" apply (clarsimp simp add: derived_set_of_def eq) by (meson finite.emptyI finite.insertI finite_subset) qed (auto simp: in_derived_set_of) qed (auto simp: t1_space_derived_set_of_singleton) lemma finite_t1_space_imp_discrete_topology: "\topspace X = U; finite U; t1_space X\ \ X = discrete_topology U" by (metis discrete_topology_unique_derived_set t1_space_derived_set_of_finite) lemma t1_space_subtopology: "t1_space X \ t1_space(subtopology X U)" by (simp add: derived_set_of_subtopology t1_space_derived_set_of_finite) lemma closedin_derived_set_of_gen: "t1_space X \ closedin X (X derived_set_of S)" apply (clarsimp simp add: in_derived_set_of closedin_contains_derived_set derived_set_of_subset_topspace) by (metis DiffD2 insert_Diff insert_iff t1_space_openin_delete) lemma derived_set_of_derived_set_subset_gen: "t1_space X \ X derived_set_of (X derived_set_of S) \ X derived_set_of S" by (meson closedin_contains_derived_set closedin_derived_set_of_gen) lemma subtopology_eq_discrete_topology_gen_finite: "\t1_space X; finite S\ \ subtopology X S = discrete_topology(topspace X \ S)" by (simp add: subtopology_eq_discrete_topology_gen t1_space_derived_set_of_finite) lemma subtopology_eq_discrete_topology_finite: "\t1_space X; S \ topspace X; finite S\ \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq t1_space_derived_set_of_finite) lemma t1_space_closed_map_image: "\closed_map X Y f; f ` (topspace X) = topspace Y; t1_space X\ \ t1_space Y" by (metis closed_map_def finite_subset_image t1_space_closedin_finite) lemma homeomorphic_t1_space: "X homeomorphic_space Y \ (t1_space X \ t1_space Y)" apply (clarsimp simp add: homeomorphic_space_def) by (meson homeomorphic_eq_everything_map homeomorphic_maps_map t1_space_closed_map_image) proposition t1_space_product_topology: "t1_space (product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. t1_space (X i))" proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis using True t1_space_empty by blast next case False then obtain f where f: "f \ (\\<^sub>E i\I. topspace(X i))" by fastforce have "t1_space (product_topology X I) \ (\i\I. t1_space (X i))" proof (intro iffI ballI) show "t1_space (X i)" if "t1_space (product_topology X I)" and "i \ I" for i proof - have clo: "\h. h \ (\\<^sub>E i\I. topspace (X i)) \ closedin (product_topology X I) {h}" using that by (simp add: t1_space_closedin_singleton) show ?thesis unfolding t1_space_closedin_singleton proof clarify show "closedin (X i) {xi}" if "xi \ topspace (X i)" for xi using clo [of "\j \ I. if i=j then xi else f j"] f that \i \ I\ by (fastforce simp add: closedin_product_topology_singleton) qed qed next next show "t1_space (product_topology X I)" if "\i\I. t1_space (X i)" using that by (simp add: t1_space_closedin_singleton Ball_def PiE_iff closedin_product_topology_singleton) qed then show ?thesis using False by blast qed lemma t1_space_prod_topology: "t1_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ t1_space X \ t1_space Y" proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis by (auto simp: t1_space_empty) next case False have eq: "{(x,y)} = {x} \ {y}" for x y by simp have "t1_space (prod_topology X Y) \ (t1_space X \ t1_space Y)" using False by (force simp: t1_space_closedin_singleton closedin_prod_Times_iff eq simp del: insert_Times_insert) with False show ?thesis by simp qed subsection\Hausdorff Spaces\ definition Hausdorff_space where "Hausdorff_space X \ \x y. x \ topspace X \ y \ topspace X \ (x \ y) \ (\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V)" lemma Hausdorff_space_expansive: "\Hausdorff_space X; topspace X = topspace Y; \U. openin X U \ openin Y U\ \ Hausdorff_space Y" by (metis Hausdorff_space_def) lemma Hausdorff_space_topspace_empty: "topspace X = {} \ Hausdorff_space X" by (simp add: Hausdorff_space_def) lemma Hausdorff_imp_t1_space: "Hausdorff_space X \ t1_space X" by (metis Hausdorff_space_def disjnt_iff t1_space_def) lemma closedin_derived_set_of: "Hausdorff_space X \ closedin X (X derived_set_of S)" by (simp add: Hausdorff_imp_t1_space closedin_derived_set_of_gen) lemma t1_or_Hausdorff_space: "t1_space X \ Hausdorff_space X \ t1_space X" using Hausdorff_imp_t1_space by blast lemma Hausdorff_space_sing_Inter_opens: "\Hausdorff_space X; a \ topspace X\ \ \{u. openin X u \ a \ u} = {a}" using Hausdorff_imp_t1_space t1_space_singleton_Inter_open by force lemma Hausdorff_space_subtopology: assumes "Hausdorff_space X" shows "Hausdorff_space(subtopology X S)" proof - have *: "disjnt U V \ disjnt (S \ U) (S \ V)" for U V by (simp add: disjnt_iff) from assms show ?thesis apply (simp add: Hausdorff_space_def openin_subtopology_alt) apply (fast intro: * elim!: all_forward) done qed lemma Hausdorff_space_compact_separation: assumes X: "Hausdorff_space X" and S: "compactin X S" and T: "compactin X T" and "disjnt S T" obtains U V where "openin X U" "openin X V" "S \ U" "T \ V" "disjnt U V" proof (cases "S = {}") case True then show thesis by (metis \compactin X T\ compactin_subset_topspace disjnt_empty1 empty_subsetI openin_empty openin_topspace that) next case False have "\x \ S. \U V. openin X U \ openin X V \ x \ U \ T \ V \ disjnt U V" proof fix a assume "a \ S" then have "a \ T" by (meson assms(4) disjnt_iff) have a: "a \ topspace X" using S \a \ S\ compactin_subset_topspace by blast show "\U V. openin X U \ openin X V \ a \ U \ T \ V \ disjnt U V" proof (cases "T = {}") case True then show ?thesis using a disjnt_empty2 openin_empty by blast next case False have "\x \ topspace X - {a}. \U V. openin X U \ openin X V \ x \ U \ a \ V \ disjnt U V" using X a by (simp add: Hausdorff_space_def) then obtain U V where UV: "\x \ topspace X - {a}. openin X (U x) \ openin X (V x) \ x \ U x \ a \ V x \ disjnt (U x) (V x)" by metis with \a \ T\ compactin_subset_topspace [OF T] have Topen: "\W \ U ` T. openin X W" and Tsub: "T \ \ (U ` T)" by (auto simp: ) then obtain \ where \: "finite \" "\ \ U ` T" and "T \ \\" using T unfolding compactin_def by meson then obtain F where F: "finite F" "F \ T" "\ = U ` F" and SUF: "T \ \(U ` F)" and "a \ F" using finite_subset_image [OF \] \a \ T\ by (metis subsetD) have U: "\x. \x \ topspace X; x \ a\ \ openin X (U x)" and V: "\x. \x \ topspace X; x \ a\ \ openin X (V x)" and disj: "\x. \x \ topspace X; x \ a\ \ disjnt (U x) (V x)" using UV by blast+ show ?thesis proof (intro exI conjI) have "F \ {}" using False SUF by blast with \a \ F\ show "openin X (\(V ` F))" using F compactin_subset_topspace [OF T] by (force intro: V) show "openin X (\(U ` F))" using F Topen Tsub by (force intro: U) show "disjnt (\(V ` F)) (\(U ` F))" using disj apply (auto simp: disjnt_def) using \F \ T\ \a \ F\ compactin_subset_topspace [OF T] by blast show "a \ (\(V ` F))" using \F \ T\ T UV \a \ T\ compactin_subset_topspace by blast qed (auto simp: SUF) qed qed then obtain U V where UV: "\x \ S. openin X (U x) \ openin X (V x) \ x \ U x \ T \ V x \ disjnt (U x) (V x)" by metis then have "S \ \ (U ` S)" by auto moreover have "\W \ U ` S. openin X W" using UV by blast ultimately obtain I where I: "S \ \ (U ` I)" "I \ S" "finite I" by (metis S compactin_def finite_subset_image) show thesis proof show "openin X (\(U ` I))" using \I \ S\ UV by blast show "openin X (\ (V ` I))" using False UV \I \ S\ \S \ \ (U ` I)\ \finite I\ by blast show "disjnt (\(U ` I)) (\ (V ` I))" by simp (meson UV \I \ S\ disjnt_subset2 in_mono le_INF_iff order_refl) qed (use UV I in auto) qed lemma Hausdorff_space_compact_sets: "Hausdorff_space X \ (\S T. compactin X S \ compactin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (meson Hausdorff_space_compact_separation) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: Hausdorff_space_def) fix x y assume "x \ topspace X" "y \ topspace X" "x \ y" then show "\U. openin X U \ (\V. openin X V \ x \ U \ y \ V \ disjnt U V)" using R [of "{x}" "{y}"] by auto qed qed lemma compactin_imp_closedin: assumes X: "Hausdorff_space X" and S: "compactin X S" shows "closedin X S" proof - have "S \ topspace X" by (simp add: assms compactin_subset_topspace) moreover have "\T. openin X T \ x \ T \ T \ topspace X - S" if "x \ topspace X" "x \ S" for x using Hausdorff_space_compact_separation [OF X _ S, of "{x}"] that apply (simp add: disjnt_def) by (metis Diff_mono Diff_triv openin_subset) ultimately show ?thesis using closedin_def openin_subopen by force qed lemma closedin_Hausdorff_singleton: "\Hausdorff_space X; x \ topspace X\ \ closedin X {x}" by (simp add: Hausdorff_imp_t1_space closedin_t1_singleton) lemma closedin_Hausdorff_sing_eq: "Hausdorff_space X \ closedin X {x} \ x \ topspace X" by (meson closedin_Hausdorff_singleton closedin_subset insert_subset) lemma Hausdorff_space_discrete_topology [simp]: "Hausdorff_space (discrete_topology U)" unfolding Hausdorff_space_def apply safe by (metis discrete_topology_unique_alt disjnt_empty2 disjnt_insert2 insert_iff mk_disjoint_insert topspace_discrete_topology) lemma compactin_Int: "\Hausdorff_space X; compactin X S; compactin X T\ \ compactin X (S \ T)" by (simp add: closed_Int_compactin compactin_imp_closedin) lemma finite_topspace_imp_discrete_topology: "\topspace X = U; finite U; Hausdorff_space X\ \ X = discrete_topology U" using Hausdorff_imp_t1_space finite_t1_space_imp_discrete_topology by blast lemma derived_set_of_finite: "\Hausdorff_space X; finite S\ \ X derived_set_of S = {}" using Hausdorff_imp_t1_space t1_space_derived_set_of_finite by auto lemma derived_set_of_singleton: "Hausdorff_space X \ X derived_set_of {x} = {}" by (simp add: derived_set_of_finite) lemma closedin_Hausdorff_finite: "\Hausdorff_space X; S \ topspace X; finite S\ \ closedin X S" by (simp add: compactin_imp_closedin finite_imp_compactin_eq) lemma open_in_Hausdorff_delete: "\Hausdorff_space X; openin X S\ \ openin X (S - {x})" using Hausdorff_imp_t1_space t1_space_openin_delete_alt by auto lemma closedin_Hausdorff_finite_eq: "\Hausdorff_space X; finite S\ \ closedin X S \ S \ topspace X" by (meson closedin_Hausdorff_finite closedin_def) lemma derived_set_of_infinite_openin: "Hausdorff_space X \ X derived_set_of S = {x \ topspace X. \U. x \ U \ openin X U \ infinite(S \ U)}" using Hausdorff_imp_t1_space t1_space_derived_set_of_infinite_openin by fastforce lemma Hausdorff_space_discrete_compactin: "Hausdorff_space X \ S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" using derived_set_of_finite discrete_compactin_eq_finite by fastforce lemma Hausdorff_space_finite_topspace: "Hausdorff_space X \ X derived_set_of (topspace X) = {} \ compact_space X \ finite(topspace X)" using derived_set_of_finite discrete_compact_space_eq_finite by auto lemma derived_set_of_derived_set_subset: "Hausdorff_space X \ X derived_set_of (X derived_set_of S) \ X derived_set_of S" by (simp add: Hausdorff_imp_t1_space derived_set_of_derived_set_subset_gen) lemma Hausdorff_space_injective_preimage: assumes "Hausdorff_space Y" and cmf: "continuous_map X Y f" and "inj_on f (topspace X)" shows "Hausdorff_space X" unfolding Hausdorff_space_def proof clarify fix x y assume x: "x \ topspace X" and y: "y \ topspace X" and "x \ y" then obtain U V where "openin Y U" "openin Y V" "f x \ U" "f y \ V" "disjnt U V" using assms unfolding Hausdorff_space_def continuous_map_def by (meson inj_onD) show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" proof (intro exI conjI) show "openin X {x \ topspace X. f x \ U}" using \openin Y U\ cmf continuous_map by fastforce show "openin X {x \ topspace X. f x \ V}" using \openin Y V\ cmf openin_continuous_map_preimage by blast show "disjnt {x \ topspace X. f x \ U} {x \ topspace X. f x \ V}" using \disjnt U V\ by (auto simp add: disjnt_def) qed (use x \f x \ U\ y \f y \ V\ in auto) qed lemma homeomorphic_Hausdorff_space: "X homeomorphic_space Y \ Hausdorff_space X \ Hausdorff_space Y" unfolding homeomorphic_space_def homeomorphic_maps_map by (auto simp: homeomorphic_eq_everything_map Hausdorff_space_injective_preimage) lemma Hausdorff_space_retraction_map_image: "\retraction_map X Y r; Hausdorff_space X\ \ Hausdorff_space Y" unfolding retraction_map_def using Hausdorff_space_subtopology homeomorphic_Hausdorff_space retraction_maps_section_image2 by blast lemma compact_Hausdorff_space_optimal: assumes eq: "topspace Y = topspace X" and XY: "\U. openin X U \ openin Y U" and "Hausdorff_space X" "compact_space Y" shows "Y = X" proof - have "\U. closedin X U \ closedin Y U" using XY using topology_finer_closedin [OF eq] by metis have "openin Y S = openin X S" for S by (metis XY assms(3) assms(4) closedin_compact_space compactin_contractive compactin_imp_closedin eq openin_closedin_eq) then show ?thesis by (simp add: topology_eq) qed lemma continuous_map_imp_closed_graph: assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y" shows "closedin (prod_topology X Y) ((\x. (x,f x)) ` topspace X)" unfolding closedin_def proof show "(\x. (x, f x)) ` topspace X \ topspace (prod_topology X Y)" using continuous_map_def f by fastforce show "openin (prod_topology X Y) (topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X)" unfolding openin_prod_topology_alt proof (intro allI impI) show "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X" if "(x,y) \ topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X" for x y proof - have "x \ topspace X" "y \ topspace Y" "y \ f x" using that by auto moreover have "f x \ topspace Y" by (meson \x \ topspace X\ continuous_map_def f) ultimately obtain U V where UV: "openin Y U" "openin Y V" "f x \ U" "y \ V" "disjnt U V" using Y Hausdorff_space_def by metis show ?thesis proof (intro exI conjI) show "openin X {x \ topspace X. f x \ U}" using \openin Y U\ f openin_continuous_map_preimage by blast show "{x \ topspace X. f x \ U} \ V \ topspace (prod_topology X Y) - (\x. (x, f x)) ` topspace X" using UV by (auto simp: disjnt_iff dest: openin_subset) qed (use UV \x \ topspace X\ in auto) qed qed qed lemma continuous_imp_closed_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y\ \ closed_map X Y f" by (meson closed_map_def closedin_compact_space compactin_imp_closedin image_compactin) lemma continuous_imp_quotient_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y; f ` (topspace X) = topspace Y\ \ quotient_map X Y f" by (simp add: continuous_imp_closed_map continuous_closed_imp_quotient_map) lemma continuous_imp_homeomorphic_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_imp_closed_map bijective_closed_imp_homeomorphic_map) lemma continuous_imp_embedding_map: "\continuous_map X Y f; compact_space X; Hausdorff_space Y; inj_on f (topspace X)\ \ embedding_map X Y f" by (simp add: continuous_imp_closed_map injective_closed_imp_embedding_map) lemma continuous_inverse_map: assumes "compact_space X" "Hausdorff_space Y" and cmf: "continuous_map X Y f" and gf: "\x. x \ topspace X \ g(f x) = x" and Sf: "S \ f ` (topspace X)" shows "continuous_map (subtopology Y S) X g" proof (rule continuous_map_from_subtopology_mono [OF _ \S \ f ` (topspace X)\]) show "continuous_map (subtopology Y (f ` (topspace X))) X g" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace (subtopology Y (f ` topspace X))" then show "g x \ topspace X" by (auto simp: gf) next fix C assume C: "closedin X C" show "closedin (subtopology Y (f ` topspace X)) {x \ topspace (subtopology Y (f ` topspace X)). g x \ C}" proof (rule compactin_imp_closedin) show "Hausdorff_space (subtopology Y (f ` topspace X))" using Hausdorff_space_subtopology [OF \Hausdorff_space Y\] by blast have "compactin Y (f ` C)" using C cmf image_compactin closedin_compact_space [OF \compact_space X\] by blast moreover have "{x \ topspace Y. x \ f ` topspace X \ g x \ C} = f ` C" using closedin_subset [OF C] cmf by (auto simp: gf continuous_map_def) ultimately have "compactin Y {x \ topspace Y. x \ f ` topspace X \ g x \ C}" by simp then show "compactin (subtopology Y (f ` topspace X)) {x \ topspace (subtopology Y (f ` topspace X)). g x \ C}" by (auto simp add: compactin_subtopology) qed qed qed lemma closed_map_paired_continuous_map_right: "\continuous_map X Y f; Hausdorff_space Y\ \ closed_map X (prod_topology X Y) (\x. (x,f x))" by (simp add: continuous_map_imp_closed_graph embedding_map_graph embedding_imp_closed_map) lemma closed_map_paired_continuous_map_left: assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y" shows "closed_map X (prod_topology Y X) (\x. (f x,x))" proof - have eq: "(\x. (f x,x)) = (\(a,b). (b,a)) \ (\x. (x,f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map X (prod_topology X Y) (\x. (x, f x))" using Y closed_map_paired_continuous_map_right f by blast show "closed_map (prod_topology X Y) (prod_topology Y X) (\(a, b). (b, a))" by (metis homeomorphic_map_swap homeomorphic_imp_closed_map) qed qed lemma proper_map_paired_continuous_map_right: "\continuous_map X Y f; Hausdorff_space Y\ \ proper_map X (prod_topology X Y) (\x. (x,f x))" using closed_injective_imp_proper_map closed_map_paired_continuous_map_right by (metis (mono_tags, lifting) Pair_inject inj_onI) lemma proper_map_paired_continuous_map_left: "\continuous_map X Y f; Hausdorff_space Y\ \ proper_map X (prod_topology Y X) (\x. (f x,x))" using closed_injective_imp_proper_map closed_map_paired_continuous_map_left by (metis (mono_tags, lifting) Pair_inject inj_onI) -lemma Hausdorff_space_euclidean [simp]: "Hausdorff_space (euclidean :: 'a::metric_space topology)" -proof - - have "\U V. open U \ open V \ x \ U \ y \ V \ disjnt U V" - if "x \ y" - for x y :: 'a - proof (intro exI conjI) - let ?r = "dist x y / 2" - have [simp]: "?r > 0" - by (simp add: that) - show "open (ball x ?r)" "open (ball y ?r)" "x \ (ball x ?r)" "y \ (ball y ?r)" - by (auto simp add: that) - show "disjnt (ball x ?r) (ball y ?r)" - unfolding disjnt_def by (simp add: disjoint_ballI) - qed - then show ?thesis - by (simp add: Hausdorff_space_def) -qed - lemma Hausdorff_space_prod_topology: "Hausdorff_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ Hausdorff_space X \ Hausdorff_space Y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (rule topological_property_of_prod_component) (auto simp: Hausdorff_space_subtopology homeomorphic_Hausdorff_space) next assume R: ?rhs show ?lhs proof (cases "(topspace X \ topspace Y) = {}") case False with R have ne: "topspace X \ {}" "topspace Y \ {}" and X: "Hausdorff_space X" and Y: "Hausdorff_space Y" by auto show ?thesis unfolding Hausdorff_space_def proof clarify fix x y x' y' assume xy: "(x, y) \ topspace (prod_topology X Y)" and xy': "(x',y') \ topspace (prod_topology X Y)" and *: "\U V. openin (prod_topology X Y) U \ openin (prod_topology X Y) V \ (x, y) \ U \ (x', y') \ V \ disjnt U V" have False if "x \ x' \ y \ y'" using that proof assume "x \ x'" then obtain U V where "openin X U" "openin X V" "x \ U" "x' \ V" "disjnt U V" by (metis Hausdorff_space_def X mem_Sigma_iff topspace_prod_topology xy xy') let ?U = "U \ topspace Y" let ?V = "V \ topspace Y" have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V" by (simp_all add: openin_prod_Times_iff \openin X U\ \openin X V\) moreover have "disjnt ?U ?V" by (simp add: \disjnt U V\) ultimately show False using * \x \ U\ \x' \ V\ xy xy' by (metis SigmaD2 SigmaI topspace_prod_topology) next assume "y \ y'" then obtain U V where "openin Y U" "openin Y V" "y \ U" "y' \ V" "disjnt U V" by (metis Hausdorff_space_def Y mem_Sigma_iff topspace_prod_topology xy xy') let ?U = "topspace X \ U" let ?V = "topspace X \ V" have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V" by (simp_all add: openin_prod_Times_iff \openin Y U\ \openin Y V\) moreover have "disjnt ?U ?V" by (simp add: \disjnt U V\) ultimately show False using "*" \y \ U\ \y' \ V\ xy xy' by (metis SigmaD1 SigmaI topspace_prod_topology) qed then show "x = x' \ y = y'" by blast qed qed (simp add: Hausdorff_space_topspace_empty) qed lemma Hausdorff_space_product_topology: "Hausdorff_space (product_topology X I) \ (\\<^sub>E i\I. topspace(X i)) = {} \ (\i \ I. Hausdorff_space (X i))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (rule topological_property_of_product_component) apply (blast dest: Hausdorff_space_subtopology homeomorphic_Hausdorff_space)+ done next assume R: ?rhs show ?lhs proof (cases "(\\<^sub>E i\I. topspace(X i)) = {}") case True then show ?thesis by (simp add: Hausdorff_space_topspace_empty) next case False have "\U V. openin (product_topology X I) U \ openin (product_topology X I) V \ f \ U \ g \ V \ disjnt U V" if f: "f \ (\\<^sub>E i\I. topspace (X i))" and g: "g \ (\\<^sub>E i\I. topspace (X i))" and "f \ g" for f g :: "'a \ 'b" proof - obtain m where "f m \ g m" using \f \ g\ by blast then have "m \ I" using f g by fastforce then have "Hausdorff_space (X m)" using False that R by blast then obtain U V where U: "openin (X m) U" and V: "openin (X m) V" and "f m \ U" "g m \ V" "disjnt U V" by (metis Hausdorff_space_def PiE_mem \f m \ g m\ \m \ I\ f g) show ?thesis proof (intro exI conjI) let ?U = "(\\<^sub>E i\I. topspace(X i)) \ {x. x m \ U}" let ?V = "(\\<^sub>E i\I. topspace(X i)) \ {x. x m \ V}" show "openin (product_topology X I) ?U" "openin (product_topology X I) ?V" using \m \ I\ U V by (force simp add: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc)+ show "f \ ?U" using \f m \ U\ f by blast show "g \ ?V" using \g m \ V\ g by blast show "disjnt ?U ?V" using \disjnt U V\ by (auto simp: PiE_def Pi_def disjnt_def) qed qed then show ?thesis by (simp add: Hausdorff_space_def) qed qed end diff --git a/src/HOL/Homology/Simplices.thy b/src/HOL/Homology/Simplices.thy --- a/src/HOL/Homology/Simplices.thy +++ b/src/HOL/Homology/Simplices.thy @@ -1,3259 +1,3261 @@ section\Homology, I: Simplices\ theory "Simplices" - imports "HOL-Analysis.Abstract_Euclidean_Space" "HOL-Algebra.Free_Abelian_Groups" - + imports + "HOL-Analysis.Function_Metric" + "HOL-Analysis.Abstract_Euclidean_Space" + "HOL-Algebra.Free_Abelian_Groups" begin subsection\Standard simplices, all of which are topological subspaces of @{text"R^n"}. \ type_synonym 'a chain = "((nat \ real) \ 'a) \\<^sub>0 int" definition standard_simplex :: "nat \ (nat \ real) set" where "standard_simplex p \ {x. (\i. 0 \ x i \ x i \ 1) \ (\i>p. x i = 0) \ (\i\p. x i) = 1}" lemma topspace_standard_simplex: "topspace(subtopology (powertop_real UNIV) (standard_simplex p)) = standard_simplex p" by simp lemma basis_in_standard_simplex [simp]: "(\j. if j = i then 1 else 0) \ standard_simplex p \ i \ p" by (auto simp: standard_simplex_def) lemma nonempty_standard_simplex: "standard_simplex p \ {}" using basis_in_standard_simplex by blast lemma standard_simplex_0: "standard_simplex 0 = {(\j. if j = 0 then 1 else 0)}" by (auto simp: standard_simplex_def) lemma standard_simplex_mono: assumes "p \ q" shows "standard_simplex p \ standard_simplex q" using assms proof (clarsimp simp: standard_simplex_def) fix x :: "nat \ real" assume "\i. 0 \ x i \ x i \ 1" and "\i>p. x i = 0" and "sum x {..p} = 1" then show "sum x {..q} = 1" using sum.mono_neutral_left [of "{..q}" "{..p}" x] assms by auto qed lemma closedin_standard_simplex: "closedin (powertop_real UNIV) (standard_simplex p)" (is "closedin ?X ?S") proof - have eq: "standard_simplex p = (\i. {x. x \ topspace ?X \ x i \ {0..1}}) \ (\i \ {p<..}. {x \ topspace ?X. x i \ {0}}) \ {x \ topspace ?X. (\i\p. x i) \ {1}}" by (auto simp: standard_simplex_def topspace_product_topology) show ?thesis unfolding eq by (rule closedin_Int closedin_Inter continuous_map_sum continuous_map_product_projection closedin_continuous_map_preimage | force | clarify)+ qed lemma standard_simplex_01: "standard_simplex p \ UNIV \\<^sub>E {0..1}" using standard_simplex_def by auto lemma compactin_standard_simplex: "compactin (powertop_real UNIV) (standard_simplex p)" apply (rule closed_compactin [where K = "PiE UNIV (\i. {0..1})"]) apply (simp_all add: compactin_PiE standard_simplex_01 closedin_standard_simplex) done lemma convex_standard_simplex: "\x \ standard_simplex p; y \ standard_simplex p; 0 \ u; u \ 1\ \ (\i. (1 - u) * x i + u * y i) \ standard_simplex p" by (simp add: standard_simplex_def sum.distrib convex_bound_le flip: sum_distrib_left) lemma path_connectedin_standard_simplex: "path_connectedin (powertop_real UNIV) (standard_simplex p)" proof - define g where "g \ \x y::nat\real. \u i. (1 - u) * x i + u * y i" have 1: "continuous_map (subtopology euclideanreal {0..1}) (powertop_real UNIV) (g x y)" if "x \ standard_simplex p" "y \ standard_simplex p" for x y unfolding g_def continuous_map_componentwise by (force intro: continuous_intros) have 2: "g x y ` {0..1} \ standard_simplex p" "g x y 0 = x" "g x y 1 = y" if "x \ standard_simplex p" "y \ standard_simplex p" for x y using that by (auto simp: convex_standard_simplex g_def) show ?thesis unfolding path_connectedin_def path_connected_space_def pathin_def apply (simp add: topspace_standard_simplex topspace_product_topology continuous_map_in_subtopology) by (metis 1 2) qed lemma connectedin_standard_simplex: "connectedin (powertop_real UNIV) (standard_simplex p)" by (simp add: path_connectedin_imp_connectedin path_connectedin_standard_simplex) subsection\Face map\ definition simplical_face :: "nat \ (nat \ 'a) \ nat \ 'a::comm_monoid_add" where "simplical_face k x \ \i. if i < k then x i else if i = k then 0 else x(i -1)" lemma simplical_face_in_standard_simplex: assumes "1 \ p" "k \ p" "x \ standard_simplex (p - Suc 0)" shows "(simplical_face k x) \ standard_simplex p" proof - have x01: "\i. 0 \ x i \ x i \ 1" and sumx: "sum x {..p - Suc 0} = 1" using assms by (auto simp: standard_simplex_def simplical_face_def) have gg: "\g. sum g {..p} = sum g {..k \ p\ sum.union_disjoint [of "{..i\p. if i < k then x i else if i = k then 0 else x (i -1)) = (\i < k. x i) + (\i \ {k..p}. if i = k then 0 else x (i -1))" by (simp add: gg) consider "k \ p - Suc 0" | "k = p" using \k \ p\ by linarith then have "(\i\p. if i < k then x i else if i = k then 0 else x (i -1)) = 1" proof cases case 1 have [simp]: "Suc (p - Suc 0) = p" using \1 \ p\ by auto have "(\i = k..p. if i = k then 0 else x (i -1)) = (\i = k+1..p. if i = k then 0 else x (i -1))" by (rule sum.mono_neutral_right) auto also have "\ = (\i = k+1..p. x (i -1))" by simp also have "\ = (\i = k..p-1. x i)" using sum.atLeastAtMost_reindex [of Suc k "p-1" "\i. x (i - Suc 0)"] 1 by simp finally have eq2: "(\i = k..p. if i = k then 0 else x (i -1)) = (\i = k..p-1. x i)" . with 1 show ?thesis apply (simp add: eq eq2) by (metis (mono_tags, lifting) One_nat_def assms(3) finite_atLeastAtMost finite_lessThan ivl_disj_int(4) ivl_disj_un(10) mem_Collect_eq standard_simplex_def sum.union_disjoint) next case 2 have [simp]: "({..p} \ {x. x < p}) = {..p - Suc 0}" using assms by auto have "(\i\p. if i < p then x i else if i = k then 0 else x (i -1)) = (\i\p. if i < p then x i else 0)" by (rule sum.cong) (auto simp: 2) also have "\ = sum x {..p-1}" by (simp add: sum.If_cases) also have "\ = 1" by (simp add: sumx) finally show ?thesis using 2 by simp qed then show ?thesis using assms by (auto simp: standard_simplex_def simplical_face_def) qed subsection\Singular simplices, forcing canonicity outside the intended domain\ definition singular_simplex :: "nat \ 'a topology \ ((nat \ real) \ 'a) \ bool" where "singular_simplex p X f \ continuous_map(subtopology (powertop_real UNIV) (standard_simplex p)) X f \ f \ extensional (standard_simplex p)" abbreviation singular_simplex_set :: "nat \ 'a topology \ ((nat \ real) \ 'a) set" where "singular_simplex_set p X \ Collect (singular_simplex p X)" lemma singular_simplex_empty: "topspace X = {} \ \ singular_simplex p X f" by (simp add: singular_simplex_def continuous_map nonempty_standard_simplex) lemma singular_simplex_mono: "\singular_simplex p (subtopology X T) f; T \ S\ \ singular_simplex p (subtopology X S) f" by (auto simp: singular_simplex_def continuous_map_in_subtopology) lemma singular_simplex_subtopology: "singular_simplex p (subtopology X S) f \ singular_simplex p X f \ f ` (standard_simplex p) \ S" by (auto simp: singular_simplex_def continuous_map_in_subtopology) subsubsection\Singular face\ definition singular_face :: "nat \ nat \ ((nat \ real) \ 'a) \ (nat \ real) \ 'a" where "singular_face p k f \ restrict (f \ simplical_face k) (standard_simplex (p - Suc 0))" lemma singular_simplex_singular_face: assumes f: "singular_simplex p X f" and "1 \ p" "k \ p" shows "singular_simplex (p - Suc 0) X (singular_face p k f)" proof - let ?PT = "(powertop_real UNIV)" have 0: "simplical_face k ` standard_simplex (p - Suc 0) \ standard_simplex p" using assms simplical_face_in_standard_simplex by auto have 1: "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) (subtopology ?PT (standard_simplex p)) (simplical_face k)" proof (clarsimp simp add: continuous_map_in_subtopology simplical_face_in_standard_simplex continuous_map_componentwise 0) fix i have "continuous_map ?PT euclideanreal (\x. if i < k then x i else if i = k then 0 else x (i -1))" by (auto intro: continuous_map_product_projection) then show "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) euclideanreal (\x. simplical_face k x i)" by (simp add: simplical_face_def continuous_map_from_subtopology) qed have 2: "continuous_map (subtopology ?PT (standard_simplex p)) X f" using assms(1) singular_simplex_def by blast show ?thesis by (simp add: singular_simplex_def singular_face_def continuous_map_compose [OF 1 2]) qed subsection\Singular chains\ definition singular_chain :: "[nat, 'a topology, 'a chain] \ bool" where "singular_chain p X c \ Poly_Mapping.keys c \ singular_simplex_set p X" abbreviation singular_chain_set :: "[nat, 'a topology] \ ('a chain) set" where "singular_chain_set p X \ Collect (singular_chain p X)" lemma singular_chain_empty: "topspace X = {} \ singular_chain p X c \ c = 0" by (auto simp: singular_chain_def singular_simplex_empty subset_eq poly_mapping_eqI) lemma singular_chain_mono: "\singular_chain p (subtopology X T) c; T \ S\ \ singular_chain p (subtopology X S) c" unfolding singular_chain_def using singular_simplex_mono by blast lemma singular_chain_subtopology: "singular_chain p (subtopology X S) c \ singular_chain p X c \ (\f \ Poly_Mapping.keys c. f ` (standard_simplex p) \ S)" unfolding singular_chain_def by (fastforce simp add: singular_simplex_subtopology subset_eq) lemma singular_chain_0 [iff]: "singular_chain p X 0" by (auto simp: singular_chain_def) lemma singular_chain_of: "singular_chain p X (frag_of c) \ singular_simplex p X c" by (auto simp: singular_chain_def) lemma singular_chain_cmul: "singular_chain p X c \ singular_chain p X (frag_cmul a c)" by (auto simp: singular_chain_def) lemma singular_chain_minus: "singular_chain p X (-c) \ singular_chain p X c" by (auto simp: singular_chain_def) lemma singular_chain_add: "\singular_chain p X a; singular_chain p X b\ \ singular_chain p X (a+b)" unfolding singular_chain_def using keys_add [of a b] by blast lemma singular_chain_diff: "\singular_chain p X a; singular_chain p X b\ \ singular_chain p X (a-b)" unfolding singular_chain_def using keys_diff [of a b] by blast lemma singular_chain_sum: "(\i. i \ I \ singular_chain p X (f i)) \ singular_chain p X (\i\I. f i)" unfolding singular_chain_def using keys_sum [of f I] by blast lemma singular_chain_extend: "(\c. c \ Poly_Mapping.keys x \ singular_chain p X (f c)) \ singular_chain p X (frag_extend f x)" by (simp add: frag_extend_def singular_chain_cmul singular_chain_sum) subsection\Boundary homomorphism for singular chains\ definition chain_boundary :: "nat \ ('a chain) \ 'a chain" where "chain_boundary p c \ (if p = 0 then 0 else frag_extend (\f. (\k\p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f)))) c)" lemma singular_chain_boundary: "singular_chain p X c \ singular_chain (p - Suc 0) X (chain_boundary p c)" unfolding chain_boundary_def apply (clarsimp intro!: singular_chain_extend singular_chain_sum singular_chain_cmul) apply (auto simp: singular_chain_def intro: singular_simplex_singular_face) done lemma singular_chain_boundary_alt: "singular_chain (Suc p) X c \ singular_chain p X (chain_boundary (Suc p) c)" using singular_chain_boundary by force lemma chain_boundary_0 [simp]: "chain_boundary p 0 = 0" by (simp add: chain_boundary_def) lemma chain_boundary_cmul: "chain_boundary p (frag_cmul k c) = frag_cmul k (chain_boundary p c)" by (auto simp: chain_boundary_def frag_extend_cmul) lemma chain_boundary_minus: "chain_boundary p (- c) = - (chain_boundary p c)" by (metis chain_boundary_cmul frag_cmul_minus_one) lemma chain_boundary_add: "chain_boundary p (a+b) = chain_boundary p a + chain_boundary p b" by (simp add: chain_boundary_def frag_extend_add) lemma chain_boundary_diff: "chain_boundary p (a-b) = chain_boundary p a - chain_boundary p b" using chain_boundary_add [of p a "-b"] by (simp add: chain_boundary_minus) lemma chain_boundary_sum: "chain_boundary p (sum g I) = sum (chain_boundary p \ g) I" by (induction I rule: infinite_finite_induct) (simp_all add: chain_boundary_add) lemma chain_boundary_sum': "finite I \ chain_boundary p (sum' g I) = sum' (chain_boundary p \ g) I" by (induction I rule: finite_induct) (simp_all add: chain_boundary_add) lemma chain_boundary_of: "chain_boundary p (frag_of f) = (if p = 0 then 0 else (\k\p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f))))" by (simp add: chain_boundary_def) subsection\Factoring out chains in a subtopology for relative homology\ definition mod_subset where "mod_subset p X \ {(a,b). singular_chain p X (a - b)}" lemma mod_subset_empty [simp]: "(a,b) \ (mod_subset p (subtopology X {})) \ a = b" by (simp add: mod_subset_def singular_chain_empty topspace_subtopology) lemma mod_subset_refl [simp]: "(c,c) \ mod_subset p X" by (auto simp: mod_subset_def) lemma mod_subset_cmul: "(a,b) \ (mod_subset p X) \ (frag_cmul k a, frag_cmul k b) \ (mod_subset p X)" apply (simp add: mod_subset_def) by (metis add_diff_cancel diff_add_cancel frag_cmul_distrib2 singular_chain_cmul) lemma mod_subset_add: "\(c1,c2) \ (mod_subset p X); (d1,d2) \ (mod_subset p X)\ \ (c1+d1, c2+d2) \ (mod_subset p X)" apply (simp add: mod_subset_def) by (simp add: add_diff_add singular_chain_add) subsection\Relative cycles $Z_pX (S)$ where $X$ is a topology and $S$ a subset \ definition singular_relcycle :: "nat \ 'a topology \ 'a set \ ('a chain) \ bool" where "singular_relcycle p X S \ \c. singular_chain p X c \ (chain_boundary p c, 0) \ mod_subset (p-1) (subtopology X S)" abbreviation singular_relcycle_set where "singular_relcycle_set p X S \ Collect (singular_relcycle p X S)" lemma singular_relcycle_restrict [simp]: "singular_relcycle p X (topspace X \ S) = singular_relcycle p X S" proof - have eq: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) show ?thesis by (force simp: singular_relcycle_def eq) qed lemma singular_relcycle: "singular_relcycle p X S c \ singular_chain p X c \ singular_chain (p-1) (subtopology X S) (chain_boundary p c)" by (simp add: singular_relcycle_def mod_subset_def) lemma singular_relcycle_0 [simp]: "singular_relcycle p X S 0" by (auto simp: singular_relcycle_def) lemma singular_relcycle_cmul: "singular_relcycle p X S c \ singular_relcycle p X S (frag_cmul k c)" by (auto simp: singular_relcycle_def chain_boundary_cmul dest: singular_chain_cmul mod_subset_cmul) lemma singular_relcycle_minus: "singular_relcycle p X S (-c) \ singular_relcycle p X S c" by (simp add: chain_boundary_minus singular_chain_minus singular_relcycle) lemma singular_relcycle_add: "\singular_relcycle p X S a; singular_relcycle p X S b\ \ singular_relcycle p X S (a+b)" by (simp add: singular_relcycle_def chain_boundary_add mod_subset_def singular_chain_add) lemma singular_relcycle_sum: "\\i. i \ I \ singular_relcycle p X S (f i)\ \ singular_relcycle p X S (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: singular_relcycle_add) lemma singular_relcycle_diff: "\singular_relcycle p X S a; singular_relcycle p X S b\ \ singular_relcycle p X S (a-b)" by (metis singular_relcycle_add singular_relcycle_minus uminus_add_conv_diff) lemma singular_cycle: "singular_relcycle p X {} c \ singular_chain p X c \ chain_boundary p c = 0" by (simp add: singular_relcycle_def) lemma singular_cycle_mono: "\singular_relcycle p (subtopology X T) {} c; T \ S\ \ singular_relcycle p (subtopology X S) {} c" by (auto simp: singular_cycle elim: singular_chain_mono) subsection\Relative boundaries $B_p X S$, where $X$ is a topology and $S$ a subset.\ definition singular_relboundary :: "nat \ 'a topology \ 'a set \ ('a chain) \ bool" where "singular_relboundary p X S \ \c. \d. singular_chain (Suc p) X d \ (chain_boundary (Suc p) d, c) \ (mod_subset p (subtopology X S))" abbreviation singular_relboundary_set :: "nat \ 'a topology \ 'a set \ ('a chain) set" where "singular_relboundary_set p X S \ Collect (singular_relboundary p X S)" lemma singular_relboundary_restrict [simp]: "singular_relboundary p X (topspace X \ S) = singular_relboundary p X S" unfolding singular_relboundary_def by (metis (no_types, hide_lams) subtopology_subtopology subtopology_topspace) lemma singular_relboundary_alt: "singular_relboundary p X S c \ (\d e. singular_chain (Suc p) X d \ singular_chain p (subtopology X S) e \ chain_boundary (Suc p) d = c + e)" unfolding singular_relboundary_def mod_subset_def by fastforce lemma singular_relboundary: "singular_relboundary p X S c \ (\d e. singular_chain (Suc p) X d \ singular_chain p (subtopology X S) e \ (chain_boundary (Suc p) d) + e = c)" using singular_chain_minus by (fastforce simp add: singular_relboundary_alt) lemma singular_boundary: "singular_relboundary p X {} c \ (\d. singular_chain (Suc p) X d \ chain_boundary (Suc p) d = c)" by (simp add: singular_relboundary_def) lemma singular_boundary_imp_chain: "singular_relboundary p X {} c \ singular_chain p X c" by (auto simp: singular_relboundary singular_chain_boundary_alt singular_chain_empty topspace_subtopology) lemma singular_boundary_mono: "\T \ S; singular_relboundary p (subtopology X T) {} c\ \ singular_relboundary p (subtopology X S) {} c" by (metis mod_subset_empty singular_chain_mono singular_relboundary_def) lemma singular_relboundary_imp_chain: "singular_relboundary p X S c \ singular_chain p X c" unfolding singular_relboundary singular_chain_subtopology by (blast intro: singular_chain_add singular_chain_boundary_alt) lemma singular_chain_imp_relboundary: "singular_chain p (subtopology X S) c \ singular_relboundary p X S c" unfolding singular_relboundary_def apply (rule_tac x=0 in exI) using mod_subset_def singular_chain_diff by fastforce lemma singular_relboundary_0 [simp]: "singular_relboundary p X S 0" unfolding singular_relboundary_def by (rule_tac x=0 in exI) auto lemma singular_relboundary_cmul: "singular_relboundary p X S c \ singular_relboundary p X S (frag_cmul a c)" unfolding singular_relboundary_def by (metis chain_boundary_cmul mod_subset_cmul singular_chain_cmul) lemma singular_relboundary_minus: "singular_relboundary p X S (-c) \ singular_relboundary p X S c" using singular_relboundary_cmul by (metis add.inverse_inverse frag_cmul_minus_one) lemma singular_relboundary_add: "\singular_relboundary p X S a; singular_relboundary p X S b\ \ singular_relboundary p X S (a+b)" unfolding singular_relboundary_def by (metis chain_boundary_add mod_subset_add singular_chain_add) lemma singular_relboundary_diff: "\singular_relboundary p X S a; singular_relboundary p X S b\ \ singular_relboundary p X S (a-b)" by (metis uminus_add_conv_diff singular_relboundary_minus singular_relboundary_add) subsection\The (relative) homology relation\ definition homologous_rel :: "[nat,'a topology,'a set,'a chain,'a chain] \ bool" where "homologous_rel p X S \ \a b. singular_relboundary p X S (a-b)" abbreviation homologous_rel_set where "homologous_rel_set p X S a \ Collect (homologous_rel p X S a)" lemma homologous_rel_restrict [simp]: "homologous_rel p X (topspace X \ S) = homologous_rel p X S" unfolding homologous_rel_def by (metis singular_relboundary_restrict) lemma homologous_rel_refl [simp]: "homologous_rel p X S c c" unfolding homologous_rel_def by auto lemma homologous_rel_sym: "homologous_rel p X S a b = homologous_rel p X S b a" unfolding homologous_rel_def using singular_relboundary_minus by fastforce lemma homologous_rel_trans: assumes "homologous_rel p X S b c" "homologous_rel p X S a b" shows "homologous_rel p X S a c" using homologous_rel_def proof - have "singular_relboundary p X S (b - c)" using assms unfolding homologous_rel_def by blast moreover have "singular_relboundary p X S (b - a)" using assms by (meson homologous_rel_def homologous_rel_sym) ultimately have "singular_relboundary p X S (c - a)" using singular_relboundary_diff by fastforce then show ?thesis by (meson homologous_rel_def homologous_rel_sym) qed lemma homologous_rel_eq: "homologous_rel p X S a = homologous_rel p X S b \ homologous_rel p X S a b" using homologous_rel_sym homologous_rel_trans by fastforce lemma homologous_rel_set_eq: "homologous_rel_set p X S a = homologous_rel_set p X S b \ homologous_rel p X S a b" by (metis homologous_rel_eq mem_Collect_eq) lemma homologous_rel_singular_chain: "homologous_rel p X S a b \ (singular_chain p X a \ singular_chain p X b)" unfolding homologous_rel_def using singular_chain_diff singular_chain_add by (fastforce dest: singular_relboundary_imp_chain) lemma homologous_rel_add: "\homologous_rel p X S a a'; homologous_rel p X S b b'\ \ homologous_rel p X S (a+b) (a'+b')" unfolding homologous_rel_def by (simp add: add_diff_add singular_relboundary_add) lemma homologous_rel_diff: assumes "homologous_rel p X S a a'" "homologous_rel p X S b b'" shows "homologous_rel p X S (a - b) (a' - b')" proof - have "singular_relboundary p X S ((a - a') - (b - b'))" using assms singular_relboundary_diff unfolding homologous_rel_def by blast then show ?thesis by (simp add: homologous_rel_def algebra_simps) qed lemma homologous_rel_sum: assumes f: "finite {i \ I. f i \ 0}" and g: "finite {i \ I. g i \ 0}" and h: "\i. i \ I \ homologous_rel p X S (f i) (g i)" shows "homologous_rel p X S (sum f I) (sum g I)" proof (cases "finite I") case True let ?L = "{i \ I. f i \ 0} \ {i \ I. g i \ 0}" have L: "finite ?L" "?L \ I" using f g by blast+ have "sum f I = sum f ?L" by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto moreover have "sum g I = sum g ?L" by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto moreover have *: "homologous_rel p X S (f i) (g i)" if "i \ ?L" for i using h that by auto have "homologous_rel p X S (sum f ?L) (sum g ?L)" using L proof induction case (insert j J) then show ?case by (simp add: h homologous_rel_add) qed auto ultimately show ?thesis by simp qed auto lemma chain_homotopic_imp_homologous_rel: assumes "\c. singular_chain p X c \ singular_chain (Suc p) X' (h c)" "\c. singular_chain (p -1) (subtopology X S) c \ singular_chain p (subtopology X' T) (h' c)" "\c. singular_chain p X c \ (chain_boundary (Suc p) (h c)) + (h'(chain_boundary p c)) = f c - g c" "singular_relcycle p X S c" shows "homologous_rel p X' T (f c) (g c)" using assms unfolding singular_relcycle_def mod_subset_def homologous_rel_def singular_relboundary_def apply (rule_tac x="h c" in exI, simp) by (metis (no_types, lifting) add_diff_cancel_left' minus_diff_eq singular_chain_minus) subsection\Show that all boundaries are cycles, the key "chain complex" property.\ lemma chain_boundary_boundary: assumes "singular_chain p X c" shows "chain_boundary (p - Suc 0) (chain_boundary p c) = 0" proof (cases "p -1 = 0") case False then have "2 \ p" by auto show ?thesis using assms unfolding singular_chain_def proof (induction rule: frag_induction) case (one g) then have ss: "singular_simplex p X g" by simp have eql: "{..p} \ {..p - Suc 0} \ {(x, y). y < x} = (\(j,i). (Suc i, j)) ` {(i,j). i \ j \ j \ p -1}" using False by (auto simp: image_def) (metis One_nat_def diff_Suc_1 diff_le_mono le_refl lessE less_imp_le_nat) have eqr: "{..p} \ {..p - Suc 0} - {(x, y). y < x} = {(i,j). i \ j \ j \ p -1}" by auto have eqf: "singular_face (p - Suc 0) i (singular_face p (Suc j) g) = singular_face (p - Suc 0) j (singular_face p i g)" if "i \ j" "j \ p - Suc 0" for i j proof (rule ext) fix t show "singular_face (p - Suc 0) i (singular_face p (Suc j) g) t = singular_face (p - Suc 0) j (singular_face p i g) t" proof (cases "t \ standard_simplex (p -1 -1)") case True have fi: "simplical_face i t \ standard_simplex (p - Suc 0)" using False True simplical_face_in_standard_simplex that by force have fj: "simplical_face j t \ standard_simplex (p - Suc 0)" by (metis False One_nat_def True simplical_face_in_standard_simplex less_one not_less that(2)) have eq: "simplical_face (Suc j) (simplical_face i t) = simplical_face i (simplical_face j t)" using True that ss unfolding standard_simplex_def simplical_face_def by fastforce show ?thesis by (simp add: singular_face_def fi fj eq) qed (simp add: singular_face_def) qed show ?case proof (cases "p = 1") case False have eq0: "frag_cmul (-1) a = b \ a + b = 0" for a b by (simp add: neg_eq_iff_add_eq_0) have *: "(\x\p. \i\p - Suc 0. frag_cmul ((-1) ^ (x + i)) (frag_of (singular_face (p - Suc 0) i (singular_face p x g)))) = 0" apply (simp add: sum.cartesian_product sum.Int_Diff [of "_ \ _" _ "{(x,y). y < x}"]) apply (rule eq0) apply (simp only: frag_cmul_sum prod.case_distrib [of "frag_cmul (-1)"] frag_cmul_cmul eql eqr flip: power_Suc) apply (force simp: simp add: inj_on_def sum.reindex add.commute eqf intro: sum.cong) done show ?thesis using False by (simp add: chain_boundary_of chain_boundary_sum chain_boundary_cmul frag_cmul_sum * flip: power_add) qed (simp add: chain_boundary_def) next case (diff a b) then show ?case by (simp add: chain_boundary_diff) qed auto qed (simp add: chain_boundary_def) lemma chain_boundary_boundary_alt: "singular_chain (Suc p) X c \ chain_boundary p (chain_boundary (Suc p) c) = 0" using chain_boundary_boundary by force lemma singular_relboundary_imp_relcycle: assumes "singular_relboundary p X S c" shows "singular_relcycle p X S c" proof - obtain d e where d: "singular_chain (Suc p) X d" and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e" using assms by (auto simp: singular_relboundary singular_relcycle) have 1: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p (chain_boundary (Suc p) d))" using d chain_boundary_boundary_alt by fastforce have 2: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p e)" using \singular_chain p (subtopology X S) e\ singular_chain_boundary by auto have "singular_chain p X c" using assms singular_relboundary_imp_chain by auto moreover have "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p c)" by (simp add: c chain_boundary_add singular_chain_add 1 2) ultimately show ?thesis by (simp add: singular_relcycle) qed lemma homologous_rel_singular_relcycle_1: assumes "homologous_rel p X S c1 c2" "singular_relcycle p X S c1" shows "singular_relcycle p X S c2" using assms by (metis diff_add_cancel homologous_rel_def homologous_rel_sym singular_relboundary_imp_relcycle singular_relcycle_add) lemma homologous_rel_singular_relcycle: assumes "homologous_rel p X S c1 c2" shows "singular_relcycle p X S c1 = singular_relcycle p X S c2" using assms homologous_rel_singular_relcycle_1 using homologous_rel_sym by blast subsection\Operations induced by a continuous map g between topological spaces\ definition simplex_map :: "nat \ ('b \ 'a) \ ((nat \ real) \ 'b) \ (nat \ real) \ 'a" where "simplex_map p g c \ restrict (g \ c) (standard_simplex p)" lemma singular_simplex_simplex_map: "\singular_simplex p X f; continuous_map X X' g\ \ singular_simplex p X' (simplex_map p g f)" unfolding singular_simplex_def simplex_map_def by (auto simp: continuous_map_compose) lemma simplex_map_eq: "\singular_simplex p X c; \x. x \ topspace X \ f x = g x\ \ simplex_map p f c = simplex_map p g c" by (auto simp: singular_simplex_def simplex_map_def continuous_map_def) lemma simplex_map_id_gen: "\singular_simplex p X c; \x. x \ topspace X \ f x = x\ \ simplex_map p f c = c" unfolding singular_simplex_def simplex_map_def continuous_map_def using extensional_arb by fastforce lemma simplex_map_id [simp]: "simplex_map p id = (\c. restrict c (standard_simplex p))" by (auto simp: simplex_map_def) lemma simplex_map_compose: "simplex_map p (h \ g) = simplex_map p h \ simplex_map p g" unfolding simplex_map_def by force lemma singular_face_simplex_map: "\1 \ p; k \ p\ \ singular_face p k (simplex_map p f c) = simplex_map (p - Suc 0) f (c \ simplical_face k)" unfolding simplex_map_def singular_face_def by (force simp: simplical_face_in_standard_simplex) lemma singular_face_restrict [simp]: assumes "p > 0" "i \ p" shows "singular_face p i (restrict f (standard_simplex p)) = singular_face p i f" by (metis assms One_nat_def Suc_leI simplex_map_id singular_face_def singular_face_simplex_map) definition chain_map :: "nat \ ('b \ 'a) \ (((nat \ real) \ 'b) \\<^sub>0 int) \ 'a chain" where "chain_map p g c \ frag_extend (frag_of \ simplex_map p g) c" lemma singular_chain_chain_map: "\singular_chain p X c; continuous_map X X' g\ \ singular_chain p X' (chain_map p g c)" unfolding chain_map_def apply (rule singular_chain_extend) by (metis comp_apply subsetD mem_Collect_eq singular_chain_def singular_chain_of singular_simplex_simplex_map) lemma chain_map_0 [simp]: "chain_map p g 0 = 0" by (auto simp: chain_map_def) lemma chain_map_of [simp]: "chain_map p g (frag_of f) = frag_of (simplex_map p g f)" by (simp add: chain_map_def) lemma chain_map_cmul [simp]: "chain_map p g (frag_cmul a c) = frag_cmul a (chain_map p g c)" by (simp add: frag_extend_cmul chain_map_def) lemma chain_map_minus: "chain_map p g (-c) = - (chain_map p g c)" by (simp add: frag_extend_minus chain_map_def) lemma chain_map_add: "chain_map p g (a+b) = chain_map p g a + chain_map p g b" by (simp add: frag_extend_add chain_map_def) lemma chain_map_diff: "chain_map p g (a-b) = chain_map p g a - chain_map p g b" by (simp add: frag_extend_diff chain_map_def) lemma chain_map_sum: "finite I \ chain_map p g (sum f I) = sum (chain_map p g \ f) I" by (simp add: frag_extend_sum chain_map_def) lemma chain_map_eq: "\singular_chain p X c; \x. x \ topspace X \ f x = g x\ \ chain_map p f c = chain_map p g c" unfolding singular_chain_def apply (erule frag_induction) apply (auto simp: chain_map_diff) apply (metis simplex_map_eq) done lemma chain_map_id_gen: "\singular_chain p X c; \x. x \ topspace X \ f x = x\ \ chain_map p f c = c" unfolding singular_chain_def by (erule frag_induction) (auto simp: chain_map_diff simplex_map_id_gen) lemma chain_map_ident: "singular_chain p X c \ chain_map p id c = c" by (simp add: chain_map_id_gen) lemma chain_map_id: "chain_map p id = frag_extend (frag_of \ (\f. restrict f (standard_simplex p)))" by (auto simp: chain_map_def) lemma chain_map_compose: "chain_map p (h \ g) = chain_map p h \ chain_map p g" proof show "chain_map p (h \ g) c = (chain_map p h \ chain_map p g) c" for c using subset_UNIV proof (induction c rule: frag_induction) case (one x) then show ?case by simp (metis (mono_tags, lifting) comp_eq_dest_lhs restrict_apply simplex_map_def) next case (diff a b) then show ?case by (simp add: chain_map_diff) qed auto qed lemma singular_simplex_chain_map_id: assumes "singular_simplex p X f" shows "chain_map p f (frag_of (restrict id (standard_simplex p))) = frag_of f" proof - have "(restrict (f \ restrict id (standard_simplex p)) (standard_simplex p)) = f" by (rule ext) (metis assms comp_apply extensional_arb id_apply restrict_apply singular_simplex_def) then show ?thesis by (simp add: simplex_map_def) qed lemma chain_boundary_chain_map: assumes "singular_chain p X c" shows "chain_boundary p (chain_map p g c) = chain_map (p - Suc 0) g (chain_boundary p c)" using assms unfolding singular_chain_def proof (induction c rule: frag_induction) case (one x) then have "singular_face p i (simplex_map p g x) = simplex_map (p - Suc 0) g (singular_face p i x)" if "0 \ i" "i \ p" "p \ 0" for i using that by (fastforce simp add: singular_face_def simplex_map_def simplical_face_in_standard_simplex) then show ?case by (auto simp: chain_boundary_of chain_map_sum) next case (diff a b) then show ?case by (simp add: chain_boundary_diff chain_map_diff) qed auto lemma singular_relcycle_chain_map: assumes "singular_relcycle p X S c" "continuous_map X X' g" "g ` S \ T" shows "singular_relcycle p X' T (chain_map p g c)" proof - have "continuous_map (subtopology X S) (subtopology X' T) g" using assms using continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce then show ?thesis using chain_boundary_chain_map [of p X c g] by (metis One_nat_def assms(1) assms(2) singular_chain_chain_map singular_relcycle) qed lemma singular_relboundary_chain_map: assumes "singular_relboundary p X S c" "continuous_map X X' g" "g ` S \ T" shows "singular_relboundary p X' T (chain_map p g c)" proof - obtain d e where d: "singular_chain (Suc p) X d" and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e" using assms by (auto simp: singular_relboundary) have "singular_chain (Suc p) X' (chain_map (Suc p) g d)" using assms(2) d singular_chain_chain_map by blast moreover have "singular_chain p (subtopology X' T) (chain_map p g e)" proof - have "\t. g ` topspace (subtopology t S) \ T" by (metis assms(3) closure_of_subset_subtopology closure_of_topspace dual_order.trans image_mono) then show ?thesis by (meson assms(2) continuous_map_from_subtopology continuous_map_in_subtopology e singular_chain_chain_map) qed moreover have "chain_boundary (Suc p) (chain_map (Suc p) g d) + chain_map p g e = chain_map p g (chain_boundary (Suc p) d + e)" by (metis One_nat_def chain_boundary_chain_map chain_map_add d diff_Suc_1) ultimately show ?thesis unfolding singular_relboundary using c by blast qed subsection\Homology of one-point spaces degenerates except for $p = 0$.\ lemma singular_simplex_singleton: assumes "topspace X = {a}" shows "singular_simplex p X f \ f = restrict (\x. a) (standard_simplex p)" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs proof - have "continuous_map (subtopology (product_topology (\n. euclideanreal) UNIV) (standard_simplex p)) X f" using \singular_simplex p X f\ singular_simplex_def by blast then have "\c. c \ standard_simplex p \ f c = a" by (simp add: assms continuous_map_def) then show ?thesis by (metis (no_types) L extensional_restrict restrict_ext singular_simplex_def) qed next assume ?rhs with assms show ?lhs by (auto simp: singular_simplex_def topspace_subtopology) qed lemma singular_chain_singleton: assumes "topspace X = {a}" shows "singular_chain p X c \ (\b. c = frag_cmul b (frag_of(restrict (\x. a) (standard_simplex p))))" (is "?lhs = ?rhs") proof let ?f = "restrict (\x. a) (standard_simplex p)" assume L: ?lhs with assms have "Poly_Mapping.keys c \ {?f}" by (auto simp: singular_chain_def singular_simplex_singleton) then consider "Poly_Mapping.keys c = {}" | "Poly_Mapping.keys c = {?f}" by blast then show ?rhs proof cases case 1 with L show ?thesis by (metis frag_cmul_zero keys_eq_empty) next case 2 then have "\b. frag_extend frag_of c = frag_cmul b (frag_of (\x\standard_simplex p. a))" by (force simp: frag_extend_def) then show ?thesis by (metis frag_expansion) qed next assume ?rhs with assms show ?lhs by (auto simp: singular_chain_def singular_simplex_singleton) qed lemma chain_boundary_of_singleton: assumes tX: "topspace X = {a}" and sc: "singular_chain p X c" shows "chain_boundary p c = (if p = 0 \ odd p then 0 else frag_extend (\f. frag_of(restrict (\x. a) (standard_simplex (p -1)))) c)" (is "?lhs = ?rhs") proof (cases "p = 0") case False have "?lhs = frag_extend (\f. if odd p then 0 else frag_of(restrict (\x. a) (standard_simplex (p -1)))) c" proof (simp only: chain_boundary_def False if_False, rule frag_extend_eq) fix f assume "f \ Poly_Mapping.keys c" with assms have "singular_simplex p X f" by (auto simp: singular_chain_def) then have *: "\k. k \ p \ singular_face p k f = (\x\standard_simplex (p -1). a)" apply (subst singular_simplex_singleton [OF tX, symmetric]) using False singular_simplex_singular_face by fastforce define c where "c \ frag_of (\x\standard_simplex (p -1). a)" have "(\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f))) = (\k\p. frag_cmul ((-1) ^ k) c)" by (auto simp: c_def * intro: sum.cong) also have "\ = (if odd p then 0 else c)" by (induction p) (auto simp: c_def restrict_def) finally show "(\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f))) = (if odd p then 0 else frag_of (\x\standard_simplex (p -1). a))" unfolding c_def . qed also have "\ = ?rhs" by (auto simp: False frag_extend_eq_0) finally show ?thesis . qed (simp add: chain_boundary_def) lemma singular_cycle_singleton: assumes "topspace X = {a}" shows "singular_relcycle p X {} c \ singular_chain p X c \ (p = 0 \ odd p \ c = 0)" proof - have "c = 0" if "singular_chain p X c" and "chain_boundary p c = 0" and "even p" and "p \ 0" using that assms singular_chain_singleton [of X a p c] chain_boundary_of_singleton [OF assms] by (auto simp: frag_extend_cmul) moreover have "chain_boundary p c = 0" if sc: "singular_chain p X c" and "odd p" by (simp add: chain_boundary_of_singleton [OF assms sc] that) moreover have "chain_boundary 0 c = 0" if "singular_chain 0 X c" and "p = 0" by (simp add: chain_boundary_def) ultimately show ?thesis using assms by (auto simp: singular_cycle) qed lemma singular_boundary_singleton: assumes "topspace X = {a}" shows "singular_relboundary p X {} c \ singular_chain p X c \ (odd p \ c = 0)" proof (cases "singular_chain p X c") case True have eq: "frag_extend (\f. frag_of (\x\standard_simplex p. a)) (frag_of (\x\standard_simplex (Suc p). a)) = frag_of (\x\standard_simplex p. a)" by (simp add: singular_chain_singleton frag_extend_cmul assms) have "\d. singular_chain (Suc p) X d \ chain_boundary (Suc p) d = c" if "singular_chain p X c" and "odd p" using assms that apply (clarsimp simp: singular_chain_singleton) apply (rule_tac x = "frag_cmul b (frag_of (\x\standard_simplex (Suc p). a))" in exI) apply (subst chain_boundary_of_singleton [of X a "Suc p"]) apply (auto simp: singular_chain_singleton frag_extend_cmul eq) done with True assms show ?thesis by (auto simp: singular_boundary chain_boundary_of_singleton) next case False with assms singular_boundary_imp_chain show ?thesis by metis qed lemma singular_boundary_eq_cycle_singleton: assumes "topspace X = {a}" "1 \ p" shows "singular_relboundary p X {} c \ singular_relcycle p X {} c" using assms apply (auto simp: singular_boundary chain_boundary_boundary_alt singular_chain_boundary_alt singular_cycle) by (metis Suc_neq_Zero le_zero_eq singular_boundary singular_boundary_singleton singular_chain_0 singular_cycle_singleton singular_relcycle) lemma singular_boundary_set_eq_cycle_singleton: assumes "topspace X = {a}" "1 \ p" shows "singular_relboundary_set p X {} = singular_relcycle_set p X {}" using singular_boundary_eq_cycle_singleton [OF assms] by blast subsection\Simplicial chains\ text\Simplicial chains, effectively those resulting from linear maps. We still allow the map to be singular, so the name is questionable. These are intended as building-blocks for singular subdivision, rather than as a axis for 1 simplicial homology.\ definition oriented_simplex where "oriented_simplex p l \ (\x\standard_simplex p. \i. (\j\p. l j i * x j))" definition simplicial_simplex where "simplicial_simplex p S f \ singular_simplex p (subtopology (powertop_real UNIV) S) f \ (\l. f = oriented_simplex p l)" lemma simplicial_simplex: "simplicial_simplex p S f \ f ` (standard_simplex p) \ S \ (\l. f = oriented_simplex p l)" (is "?lhs = ?rhs") proof assume R: ?rhs show ?lhs using R apply (clarsimp simp: simplicial_simplex_def singular_simplex_subtopology) apply (simp add: singular_simplex_def oriented_simplex_def) apply (clarsimp simp: continuous_map_componentwise) apply (intro continuous_intros continuous_map_from_subtopology continuous_map_product_projection, auto) done qed (simp add: simplicial_simplex_def singular_simplex_subtopology) lemma simplicial_simplex_empty [simp]: "\ simplicial_simplex p {} f" by (simp add: nonempty_standard_simplex simplicial_simplex) definition simplicial_chain where "simplicial_chain p S c \ Poly_Mapping.keys c \ Collect (simplicial_simplex p S)" lemma simplicial_chain_0 [simp]: "simplicial_chain p S 0" by (simp add: simplicial_chain_def) lemma simplicial_chain_of [simp]: "simplicial_chain p S (frag_of c) \ simplicial_simplex p S c" by (simp add: simplicial_chain_def) lemma simplicial_chain_cmul: "simplicial_chain p S c \ simplicial_chain p S (frag_cmul a c)" by (auto simp: simplicial_chain_def) lemma simplicial_chain_diff: "\simplicial_chain p S c1; simplicial_chain p S c2\ \ simplicial_chain p S (c1 - c2)" unfolding simplicial_chain_def by (meson UnE keys_diff subset_iff) lemma simplicial_chain_sum: "(\i. i \ I \ simplicial_chain p S (f i)) \ simplicial_chain p S (sum f I)" unfolding simplicial_chain_def using order_trans [OF keys_sum [of f I]] by (simp add: UN_least) lemma simplicial_simplex_oriented_simplex: "simplicial_simplex p S (oriented_simplex p l) \ ((\x i. \j\p. l j i * x j) ` standard_simplex p \ S)" by (auto simp: simplicial_simplex oriented_simplex_def) lemma simplicial_imp_singular_simplex: "simplicial_simplex p S f \ singular_simplex p (subtopology (powertop_real UNIV) S) f" by (simp add: simplicial_simplex_def) lemma simplicial_imp_singular_chain: "simplicial_chain p S c \ singular_chain p (subtopology (powertop_real UNIV) S) c" unfolding simplicial_chain_def singular_chain_def by (auto intro: simplicial_imp_singular_simplex) lemma oriented_simplex_eq: "oriented_simplex p l = oriented_simplex p l' \ (\i. i \ p \ l i = l' i)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof clarify fix i assume "i \ p" let ?fi = "(\j. if j = i then 1 else 0)" have "(\j\p. l j k * ?fi j) = (\j\p. l' j k * ?fi j)" for k apply (rule fun_cong [where x=k]) using fun_cong [OF L, of ?fi] apply (simp add: \i \ p\ oriented_simplex_def) done with \i \ p\ show "l i = l' i" by (simp add: if_distrib ext cong: if_cong) qed qed (auto simp: oriented_simplex_def) lemma singular_face_oriented_simplex: assumes "1 \ p" "k \ p" shows "singular_face p k (oriented_simplex p l) = oriented_simplex (p -1) (\j. if j < k then l j else l (Suc j))" proof - have "(\j\p. l j i * simplical_face k x j) = (\j\p - Suc 0. (if j < k then l j else l (Suc j)) i * x j)" if "x \ standard_simplex (p - Suc 0)" for i x proof - show ?thesis unfolding simplical_face_def using sum.zero_middle [OF assms, where 'a=real, symmetric] apply (simp add: if_distrib [of "\x. _ * x"] if_distrib [of "\f. f i * _"] atLeast0AtMost cong: if_cong) done qed then show ?thesis using simplical_face_in_standard_simplex assms by (auto simp: singular_face_def oriented_simplex_def restrict_def) qed lemma simplicial_simplex_singular_face: fixes f :: "(nat \ real) \ nat \ real" assumes ss: "simplicial_simplex p S f" and p: "1 \ p" "k \ p" shows "simplicial_simplex (p - Suc 0) S (singular_face p k f)" proof - let ?X = "subtopology (powertop_real UNIV) S" obtain m where l: "singular_simplex p ?X (oriented_simplex p m)" and feq: "f = oriented_simplex p m" using assms by (force simp: simplicial_simplex_def) moreover have "\l. singular_face p k f = oriented_simplex (p - Suc 0) l" apply (simp add: feq singular_face_def oriented_simplex_def) apply (simp add: simplical_face_in_standard_simplex [OF p] restrict_compose_left subset_eq) apply (rule_tac x="\i. if i < k then m i else m (Suc i)" in exI) using sum.zero_middle [OF p, where 'a=real, symmetric] unfolding simplical_face_def o_def apply (simp add: if_distrib [of "\x. _ * x"] if_distrib [of "\f. f _ * _"] atLeast0AtMost cong: if_cong) done ultimately show ?thesis using assms by (simp add: singular_simplex_singular_face simplicial_simplex_def) qed lemma simplicial_chain_boundary: "simplicial_chain p S c \ simplicial_chain (p -1) S (chain_boundary p c)" unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one f) then have "simplicial_simplex p S f" by simp have "simplicial_chain (p - Suc 0) S (frag_of (singular_face p i f))" if "0 < p" "i \ p" for i using that one apply (simp add: simplicial_simplex_def singular_simplex_singular_face) apply (force simp: singular_face_oriented_simplex) done then have "simplicial_chain (p - Suc 0) S (chain_boundary p (frag_of f))" unfolding chain_boundary_def frag_extend_of by (auto intro!: simplicial_chain_cmul simplicial_chain_sum) then show ?case by (simp add: simplicial_chain_def [symmetric]) next case (diff a b) then show ?case by (metis chain_boundary_diff simplicial_chain_def simplicial_chain_diff) qed auto subsection\The cone construction on simplicial simplices.\ consts simplex_cone :: "[nat, nat \ real, [nat \ real, nat] \ real, nat \ real, nat] \ real" specification (simplex_cone) simplex_cone: "\p v l. simplex_cone p v (oriented_simplex p l) = oriented_simplex (Suc p) (\i. if i = 0 then v else l(i -1))" proof - have *: "\x. \y. \v. (\l. oriented_simplex (Suc x) (\i. if i = 0 then v else l (i -1))) = (y v \ (oriented_simplex x))" apply (subst choice_iff [symmetric]) apply (subst function_factors_left [symmetric]) by (simp add: oriented_simplex_eq) then show ?thesis apply (subst choice_iff [symmetric]) apply (subst fun_eq_iff [symmetric]) unfolding o_def apply (blast intro: sym) done qed lemma simplicial_simplex_simplex_cone: assumes f: "simplicial_simplex p S f" and T: "\x u. \0 \ u; u \ 1; x \ S\ \ (\i. (1 - u) * v i + u * x i) \ T" shows "simplicial_simplex (Suc p) T (simplex_cone p v f)" proof - obtain l where l: "\x. x \ standard_simplex p \ oriented_simplex p l x \ S" and feq: "f = oriented_simplex p l" using f by (auto simp: simplicial_simplex) have "oriented_simplex p l x \ S" if "x \ standard_simplex p" for x using f that by (auto simp: simplicial_simplex feq) then have S: "\x. \\i. 0 \ x i \ x i \ 1; \i. i>p \ x i = 0; sum x {..p} = 1\ \ (\i. \j\p. l j i * x j) \ S" by (simp add: oriented_simplex_def standard_simplex_def) have "oriented_simplex (Suc p) (\i. if i = 0 then v else l (i -1)) x \ T" if "x \ standard_simplex (Suc p)" for x proof (simp add: that oriented_simplex_def sum.atMost_Suc_shift del: sum.atMost_Suc) have x01: "\i. 0 \ x i \ x i \ 1" and x0: "\i. i > Suc p \ x i = 0" and x1: "sum x {..Suc p} = 1" using that by (auto simp: oriented_simplex_def standard_simplex_def) obtain a where "a \ S" using f by force show "(\i. v i * x 0 + (\j\p. l j i * x (Suc j))) \ T" proof (cases "x 0 = 1") case True then have "sum x {Suc 0..Suc p} = 0" using x1 by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) then have [simp]: "x (Suc j) = 0" if "j\p" for j unfolding sum.atLeast_Suc_atMost_Suc_shift using x01 that by (simp add: sum_nonneg_eq_0_iff) then show ?thesis using T [of 0 a] \a \ S\ by (auto simp: True) next case False then have "(\i. v i * x 0 + (\j\p. l j i * x (Suc j))) = (\i. (1 - (1 - x 0)) * v i + (1 - x 0) * (inverse (1 - x 0) * (\j\p. l j i * x (Suc j))))" by (force simp: field_simps) also have "\ \ T" proof (rule T) have "x 0 < 1" by (simp add: False less_le x01) have xle: "x (Suc i) \ (1 - x 0)" for i proof (cases "i \ p") case True have "sum x {0, Suc i} \ sum x {..Suc p}" by (rule sum_mono2) (auto simp: True x01) then show ?thesis using x1 x01 by (simp add: algebra_simps not_less) qed (simp add: x0 x01) have "(\i. (\j\p. l j i * (x (Suc j) * inverse (1 - x 0)))) \ S" proof (rule S) have "x 0 + (\j\p. x (Suc j)) = sum x {..Suc p}" by (metis sum.atMost_Suc_shift) with x1 have "(\j\p. x (Suc j)) = 1 - x 0" by simp with False show "(\j\p. x (Suc j) * inverse (1 - x 0)) = 1" by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero right_inverse sum_distrib_right) qed (use x01 x0 xle \x 0 < 1\ in \auto simp: field_split_simps\) then show "(\i. inverse (1 - x 0) * (\j\p. l j i * x (Suc j))) \ S" by (simp add: field_simps sum_divide_distrib) qed (use x01 in auto) finally show ?thesis . qed qed then show ?thesis by (auto simp: simplicial_simplex feq simplex_cone) qed definition simplicial_cone where "simplicial_cone p v \ frag_extend (frag_of \ simplex_cone p v)" lemma simplicial_chain_simplicial_cone: assumes c: "simplicial_chain p S c" and T: "\x u. \0 \ u; u \ 1; x \ S\ \ (\i. (1 - u) * v i + u * x i) \ T" shows "simplicial_chain (Suc p) T (simplicial_cone p v c)" using c unfolding simplicial_chain_def simplicial_cone_def proof (induction rule: frag_induction) case (one x) then show ?case by (simp add: T simplicial_simplex_simplex_cone) next case (diff a b) then show ?case by (metis frag_extend_diff simplicial_chain_def simplicial_chain_diff) qed auto lemma chain_boundary_simplicial_cone_of': assumes "f = oriented_simplex p l" shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - (if p = 0 then frag_of (\u\standard_simplex p. v) else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))" proof (simp, intro impI conjI) assume "p = 0" have eq: "(oriented_simplex 0 (\j. if j = 0 then v else l j)) = (\u\standard_simplex 0. v)" by (force simp: oriented_simplex_def standard_simplex_def) show "chain_boundary (Suc 0) (simplicial_cone 0 v (frag_of f)) = frag_of f - frag_of (\u\standard_simplex 0. v)" by (simp add: assms simplicial_cone_def chain_boundary_of \p = 0\ simplex_cone singular_face_oriented_simplex eq cong: if_cong) next assume "0 < p" have 0: "simplex_cone (p - Suc 0) v (singular_face p x (oriented_simplex p l)) = oriented_simplex p (\j. if j < Suc x then if j = 0 then v else l (j -1) else if Suc j = 0 then v else l (Suc j -1))" if "x \ p" for x using \0 < p\ that by (auto simp: Suc_leI singular_face_oriented_simplex simplex_cone oriented_simplex_eq) have 1: "frag_extend (frag_of \ simplex_cone (p - Suc 0) v) (\k = 0..p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k (oriented_simplex p l)))) = - (\k = Suc 0..Suc p. frag_cmul ((-1) ^ k) (frag_of (singular_face (Suc p) k (simplex_cone p v (oriented_simplex p l)))))" apply (subst sum.atLeast_Suc_atMost_Suc_shift) apply (simp add: frag_extend_sum frag_extend_cmul flip: sum_negf) apply (auto simp: simplex_cone singular_face_oriented_simplex 0 intro: sum.cong) done moreover have 2: "singular_face (Suc p) 0 (simplex_cone p v (oriented_simplex p l)) = oriented_simplex p l" by (simp add: simplex_cone singular_face_oriented_simplex) show "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - simplicial_cone (p - Suc 0) v (chain_boundary p (frag_of f))" using \p > 0\ apply (simp add: assms simplicial_cone_def chain_boundary_of atMost_atLeast0 del: sum.atMost_Suc) apply (subst sum.atLeast_Suc_atMost [of 0]) apply (simp_all add: 1 2 del: sum.atMost_Suc) done qed lemma chain_boundary_simplicial_cone_of: assumes "simplicial_simplex p S f" shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - (if p = 0 then frag_of (\u\standard_simplex p. v) else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))" using chain_boundary_simplicial_cone_of' assms unfolding simplicial_simplex_def by blast lemma chain_boundary_simplicial_cone: "simplicial_chain p S c \ chain_boundary (Suc p) (simplicial_cone p v c) = c - (if p = 0 then frag_extend (\f. frag_of (\u\standard_simplex p. v)) c else simplicial_cone (p -1) v (chain_boundary p c))" unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then show ?case by (auto simp: chain_boundary_simplicial_cone_of) qed (auto simp: chain_boundary_diff simplicial_cone_def frag_extend_diff) lemma simplex_map_oriented_simplex: assumes l: "simplicial_simplex p (standard_simplex q) (oriented_simplex p l)" and g: "simplicial_simplex r S g" and "q \ r" shows "simplex_map p g (oriented_simplex p l) = oriented_simplex p (g \ l)" proof - obtain m where geq: "g = oriented_simplex r m" using g by (auto simp: simplicial_simplex_def) have "g (\i. \j\p. l j i * x j) i = (\j\p. g (l j) i * x j)" if "x \ standard_simplex p" for x i proof - have ssr: "(\i. \j\p. l j i * x j) \ standard_simplex r" using l that standard_simplex_mono [OF \q \ r\] unfolding simplicial_simplex_oriented_simplex by auto have lss: "l j \ standard_simplex r" if "j\p" for j proof - have q: "(\x i. \j\p. l j i * x j) ` standard_simplex p \ standard_simplex q" using l by (simp add: simplicial_simplex_oriented_simplex) have p: "l j \ (\x i. \j\p. l j i * x j) ` standard_simplex p" apply (rule_tac x="(\i. if i = j then 1 else 0)" in rev_image_eqI) using \j\p\ by (force simp: basis_in_standard_simplex if_distrib cong: if_cong)+ show ?thesis apply (rule subsetD [OF standard_simplex_mono [OF \q \ r\]]) apply (rule subsetD [OF q p]) done qed show ?thesis apply (simp add: geq oriented_simplex_def sum_distrib_left sum_distrib_right mult.assoc ssr lss) by (rule sum.swap) qed then show ?thesis by (force simp: oriented_simplex_def simplex_map_def o_def) qed lemma chain_map_simplicial_cone: assumes g: "simplicial_simplex r S g" and c: "simplicial_chain p (standard_simplex q) c" and v: "v \ standard_simplex q" and "q \ r" shows "chain_map (Suc p) g (simplicial_cone p v c) = simplicial_cone p (g v) (chain_map p g c)" proof - have *: "simplex_map (Suc p) g (simplex_cone p v f) = simplex_cone p (g v) (simplex_map p g f)" if "f \ Poly_Mapping.keys c" for f proof - have "simplicial_simplex p (standard_simplex q) f" using c that by (auto simp: simplicial_chain_def) then obtain m where feq: "f = oriented_simplex p m" by (auto simp: simplicial_simplex) have 0: "simplicial_simplex p (standard_simplex q) (oriented_simplex p m)" using \simplicial_simplex p (standard_simplex q) f\ feq by blast then have 1: "simplicial_simplex (Suc p) (standard_simplex q) (oriented_simplex (Suc p) (\i. if i = 0 then v else m (i -1)))" using convex_standard_simplex v by (simp flip: simplex_cone add: simplicial_simplex_simplex_cone) show ?thesis using simplex_map_oriented_simplex [OF 1 g \q \ r\] simplex_map_oriented_simplex [of p q m r S g, OF 0 g \q \ r\] by (simp add: feq oriented_simplex_eq simplex_cone) qed show ?thesis by (auto simp: chain_map_def simplicial_cone_def frag_extend_compose * intro: frag_extend_eq) qed subsection\Barycentric subdivision of a linear ("simplicial") simplex's image\ definition simplicial_vertex where "simplicial_vertex i f = f(\j. if j = i then 1 else 0)" lemma simplicial_vertex_oriented_simplex: "simplicial_vertex i (oriented_simplex p l) = (if i \ p then l i else undefined)" by (simp add: simplicial_vertex_def oriented_simplex_def if_distrib cong: if_cong) primrec simplicial_subdivision where "simplicial_subdivision 0 = id" | "simplicial_subdivision (Suc p) = frag_extend (\f. simplicial_cone p (\i. (\j\Suc p. simplicial_vertex j f i) / (p + 2)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))" lemma simplicial_subdivision_0 [simp]: "simplicial_subdivision p 0 = 0" by (induction p) auto lemma simplicial_subdivision_diff: "simplicial_subdivision p (c1-c2) = simplicial_subdivision p c1 - simplicial_subdivision p c2" by (induction p) (auto simp: frag_extend_diff) lemma simplicial_subdivision_of: "simplicial_subdivision p (frag_of f) = (if p = 0 then frag_of f else simplicial_cone (p -1) (\i. (\j\p. simplicial_vertex j f i) / (Suc p)) (simplicial_subdivision (p -1) (chain_boundary p (frag_of f))))" by (induction p) (auto simp: add.commute) lemma simplicial_chain_simplicial_subdivision: "simplicial_chain p S c \ simplicial_chain p S (simplicial_subdivision p c)" proof (induction p arbitrary: S c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction c rule: frag_induction) case (one f) then have f: "simplicial_simplex (Suc p) S f" by auto then have "simplicial_chain p (f ` standard_simplex (Suc p)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))" by (metis Suc.IH diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_simplex subsetI) moreover obtain l where l: "\x. x \ standard_simplex (Suc p) \ (\i. (\j\Suc p. l j i * x j)) \ S" and feq: "f = oriented_simplex (Suc p) l" using f by (fastforce simp: simplicial_simplex oriented_simplex_def simp del: sum.atMost_Suc) have "(\i. (1 - u) * ((\j\Suc p. simplicial_vertex j f i) / (real p + 2)) + u * y i) \ S" if "0 \ u" "u \ 1" and y: "y \ f ` standard_simplex (Suc p)" for y u proof - obtain x where x: "x \ standard_simplex (Suc p)" and yeq: "y = oriented_simplex (Suc p) l x" using y feq by blast have "(\i. \j\Suc p. l j i * ((if j \ Suc p then (1 - u) * inverse (p + 2) + u * x j else 0))) \ S" proof (rule l) have i2p: "inverse (2 + real p) \ 1" by (simp add: field_split_simps) show "(\j. if j \ Suc p then (1 - u) * inverse (real (p + 2)) + u * x j else 0) \ standard_simplex (Suc p)" using x \0 \ u\ \u \ 1\ apply (simp add: sum.distrib standard_simplex_def i2p linepath_le_1 flip: sum_distrib_left del: sum.atMost_Suc) apply (simp add: field_split_simps) done qed moreover have "(\i. \j\Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j)) = (\i. (1 - u) * (\j\Suc p. l j i) / (real p + 2) + u * (\j\Suc p. l j i * x j))" proof fix i have "(\j\Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j)) = (\j\Suc p. (1 - u) * l j i / (real p + 2) + u * l j i * x j)" (is "?lhs = _") by (simp add: field_simps cong: sum.cong) also have "\ = (1 - u) * (\j\Suc p. l j i) / (real p + 2) + u * (\j\Suc p. l j i * x j)" (is "_ = ?rhs") by (simp add: sum_distrib_left sum.distrib sum_divide_distrib mult.assoc del: sum.atMost_Suc) finally show "?lhs = ?rhs" . qed ultimately show ?thesis using feq x yeq by (simp add: simplicial_vertex_oriented_simplex) (simp add: oriented_simplex_def) qed ultimately show ?case by (simp add: simplicial_chain_simplicial_cone) next case (diff a b) then show ?case by (metis simplicial_chain_diff simplicial_subdivision_diff) qed auto qed auto lemma chain_boundary_simplicial_subdivision: "simplicial_chain p S c \ chain_boundary p (simplicial_subdivision p c) = simplicial_subdivision (p -1) (chain_boundary p c)" proof (induction p arbitrary: c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction c rule: frag_induction) case (one f) then have f: "simplicial_simplex (Suc p) S f" by simp then have "simplicial_chain p S (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))" by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision) moreover have "simplicial_chain p S (chain_boundary (Suc p) (frag_of f))" using one simplicial_chain_boundary simplicial_chain_of by fastforce moreover have "simplicial_subdivision (p - Suc 0) (chain_boundary p (chain_boundary (Suc p) (frag_of f))) = 0" by (metis f chain_boundary_boundary_alt simplicial_simplex_def simplicial_subdivision_0 singular_chain_of) ultimately show ?case apply (simp add: chain_boundary_simplicial_cone Suc) apply (auto simp: chain_boundary_of frag_extend_diff simplicial_cone_def) done next case (diff a b) then show ?case by (simp add: simplicial_subdivision_diff chain_boundary_diff frag_extend_diff) qed auto qed auto (*A MESS AND USED ONLY ONCE*) lemma simplicial_subdivision_shrinks: "\simplicial_chain p S c; \f x y. \f \ Poly_Mapping.keys c; x \ standard_simplex p; y \ standard_simplex p\ \ \f x k - f y k\ \ d; f \ Poly_Mapping.keys(simplicial_subdivision p c); x \ standard_simplex p; y \ standard_simplex p\ \ \f x k - f y k\ \ (p / (Suc p)) * d" proof (induction p arbitrary: d c f x y) case (Suc p) define Sigp where "Sigp \ \f:: (nat \ real) \ nat \ real. \i. (\j\Suc p. simplicial_vertex j f i) / real (p + 2)" let ?CB = "\f. chain_boundary (Suc p) (frag_of f)" have *: "Poly_Mapping.keys (simplicial_cone p (Sigp f) (simplicial_subdivision p (?CB f))) \ {f. \x\standard_simplex (Suc p). \y\standard_simplex (Suc p). \f x k - f y k\ \ real (Suc p) / real (Suc p + 1) * d}" (is "?lhs \ ?rhs") if f: "f \ Poly_Mapping.keys c" for f proof - have ssf: "simplicial_simplex (Suc p) S f" using Suc.prems(1) simplicial_chain_def that by auto have 2: "\x y. \x \ standard_simplex (Suc p); y \ standard_simplex (Suc p)\ \ \f x k - f y k\ \ d" by (meson Suc.prems(2) f subsetD le_Suc_eq order_refl standard_simplex_mono) have sub: "Poly_Mapping.keys ((frag_of \ simplex_cone p (Sigp f)) g) \ ?rhs" if "g \ Poly_Mapping.keys (simplicial_subdivision p (?CB f))" for g proof - have 1: "simplicial_chain p S (?CB f)" using ssf simplicial_chain_boundary simplicial_chain_of by fastforce have "simplicial_chain (Suc p) (f ` standard_simplex(Suc p)) (frag_of f)" by (metis simplicial_chain_of simplicial_simplex ssf subset_refl) then have sc_sub: "Poly_Mapping.keys (?CB f) \ Collect (simplicial_simplex p (f ` standard_simplex (Suc p)))" by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def) have led: "\h x y. \h \ Poly_Mapping.keys (chain_boundary (Suc p) (frag_of f)); x \ standard_simplex p; y \ standard_simplex p\ \ \h x k - h y k\ \ d" using Suc.prems(2) f sc_sub by (simp add: simplicial_simplex subset_iff image_iff) metis have "\f' x y. \f' \ Poly_Mapping.keys (simplicial_subdivision p (?CB f)); x \ standard_simplex p; y \ standard_simplex p\ \ \f' x k - f' y k\ \ (p / (Suc p)) * d" by (blast intro: led Suc.IH [of "chain_boundary (Suc p) (frag_of f)", OF 1]) then have g: "\x y. \x \ standard_simplex p; y \ standard_simplex p\ \ \g x k - g y k\ \ (p / (Suc p)) * d" using that by blast have "d \ 0" using Suc.prems(2)[OF f] \x \ standard_simplex (Suc p)\ by force have 3: "simplex_cone p (Sigp f) g \ ?rhs" proof - have "simplicial_simplex p (f ` standard_simplex(Suc p)) g" by (metis (mono_tags, hide_lams) sc_sub mem_Collect_eq simplicial_chain_def simplicial_chain_simplicial_subdivision subsetD that) then obtain m where m: "g ` standard_simplex p \ f ` standard_simplex (Suc p)" and geq: "g = oriented_simplex p m" using ssf by (auto simp: simplicial_simplex) have m_in_gim: "\i. i \ p \ m i \ g ` standard_simplex p" apply (rule_tac x = "\j. if j = i then 1 else 0" in image_eqI) apply (simp_all add: geq oriented_simplex_def if_distrib cong: if_cong) done obtain l where l: "f ` standard_simplex (Suc p) \ S" and feq: "f = oriented_simplex (Suc p) l" using ssf by (auto simp: simplicial_simplex) show ?thesis proof (clarsimp simp add: geq simp del: sum.atMost_Suc) fix x y assume x: "x \ standard_simplex (Suc p)" and y: "y \ standard_simplex (Suc p)" then have x': "(\i. 0 \ x i \ x i \ 1) \ (\i>Suc p. x i = 0) \ (\i\Suc p. x i) = 1" and y': "(\i. 0 \ y i \ y i \ 1) \ (\i>Suc p. y i = 0) \ (\i\Suc p. y i) = 1" by (auto simp: standard_simplex_def) have "\(\j\Suc p. (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k * x j) - (\j\Suc p. (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k * y j)\ \ (1 + real p) * d / (2 + real p)" proof - have zero: "\m (s - Suc 0) k - (\j\Suc p. l j k) / (2 + real p)\ \ (1 + real p) * d / (2 + real p)" if "0 < s" and "s \ Suc p" for s proof - have "m (s - Suc 0) \ f ` standard_simplex (Suc p)" using m m_in_gim that(2) by auto then obtain z where eq: "m (s - Suc 0) = (\i. \j\Suc p. l j i * z j)" and z: "z \ standard_simplex (Suc p)" using feq unfolding oriented_simplex_def by auto show ?thesis unfolding eq proof (rule convex_sum_bound_le) fix i assume i: "i \ {..Suc p}" then have [simp]: "card ({..Suc p} - {i}) = Suc p" by (simp add: card_Suc_Diff1) have "(\j\Suc p. \l i k / (p + 2) - l j k / (p + 2)\) = (\j\Suc p. \l i k - l j k\ / (p + 2))" by (rule sum.cong) (simp_all add: flip: diff_divide_distrib) also have "\ = (\j \ {..Suc p} - {i}. \l i k - l j k\ / (p + 2))" by (rule sum.mono_neutral_right) auto also have "\ \ (1 + real p) * d / (p + 2)" proof (rule sum_bounded_above_divide) fix i' :: "nat" assume i': "i' \ {..Suc p} - {i}" have lf: "\r. r \ Suc p \ l r \ f ` standard_simplex(Suc p)" apply (rule_tac x="\j. if j = r then 1 else 0" in image_eqI) apply (auto simp: feq oriented_simplex_def if_distrib [of "\x. _ * x"] cong: if_cong) done show "\l i k - l i' k\ / real (p + 2) \ (1 + real p) * d / real (p + 2) / real (card ({..Suc p} - {i}))" using i i' lf [of i] lf [of i'] 2 by (auto simp: image_iff divide_simps) qed auto finally have "(\j\Suc p. \l i k / (p + 2) - l j k / (p + 2)\) \ (1 + real p) * d / (p + 2)" . then have "\\j\Suc p. l i k / (p + 2) - l j k / (p + 2)\ \ (1 + real p) * d / (p + 2)" by (rule order_trans [OF sum_abs]) then show "\l i k - (\j\Suc p. l j k) / (2 + real p)\ \ (1 + real p) * d / (2 + real p)" by (simp add: sum_subtractf sum_divide_distrib del: sum.atMost_Suc) qed (use standard_simplex_def z in auto) qed have nonz: "\m (s - Suc 0) k - m (r - Suc 0) k\ \ (1 + real p) * d / (2 + real p)" (is "?lhs \ ?rhs") if "r < s" and "0 < r" and "r \ Suc p" and "s \ Suc p" for r s proof - have "?lhs \ (p / (Suc p)) * d" using m_in_gim [of "r - Suc 0"] m_in_gim [of "s - Suc 0"] that g by fastforce also have "\ \ ?rhs" by (simp add: field_simps \0 \ d\) finally show ?thesis . qed have jj: "j \ Suc p \ j' \ Suc p \ \(if j' = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j' -1)) k - (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k\ \ (1 + real p) * d / (2 + real p)" for j j' apply (rule_tac a=j and b = "j'" in linorder_less_wlog) apply (force simp: zero nonz \0 \ d\ simp del: sum.atMost_Suc)+ done show ?thesis apply (rule convex_sum_bound_le) using x' apply blast using x' apply blast apply (subst abs_minus_commute) apply (rule convex_sum_bound_le) using y' apply blast using y' apply blast using jj by blast qed then show "\simplex_cone p (Sigp f) (oriented_simplex p m) x k - simplex_cone p (Sigp f) (oriented_simplex p m) y k\ \ (1 + real p) * d / (2 + real p)" apply (simp add: feq Sigp_def simplicial_vertex_oriented_simplex simplex_cone del: sum.atMost_Suc) apply (simp add: oriented_simplex_def x y del: sum.atMost_Suc) done qed qed show ?thesis using Suc.IH [OF 1, where f=g] 2 3 by simp qed then show ?thesis unfolding simplicial_chain_def simplicial_cone_def by (simp add: order_trans [OF keys_frag_extend] sub UN_subset_iff) qed show ?case using Suc apply (simp del: sum.atMost_Suc) apply (drule subsetD [OF keys_frag_extend]) apply (simp del: sum.atMost_Suc) apply clarify (*OBTAIN?*) apply (rename_tac FFF) using * apply (simp add: add.commute Sigp_def subset_iff) done qed (auto simp: standard_simplex_0) subsection\Singular subdivision\ definition singular_subdivision where "singular_subdivision p \ frag_extend (\f. chain_map p f (simplicial_subdivision p (frag_of(restrict id (standard_simplex p)))))" lemma singular_subdivision_0 [simp]: "singular_subdivision p 0 = 0" by (simp add: singular_subdivision_def) lemma singular_subdivision_add: "singular_subdivision p (a + b) = singular_subdivision p a + singular_subdivision p b" by (simp add: singular_subdivision_def frag_extend_add) lemma singular_subdivision_diff: "singular_subdivision p (a - b) = singular_subdivision p a - singular_subdivision p b" by (simp add: singular_subdivision_def frag_extend_diff) lemma simplicial_simplex_id [simp]: "simplicial_simplex p S (restrict id (standard_simplex p)) \ standard_simplex p \ S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: simplicial_simplex_def singular_simplex_def continuous_map_in_subtopology set_mp) next assume R: ?rhs then have cm: "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p)) (subtopology (powertop_real UNIV) S) id" using continuous_map_from_subtopology_mono continuous_map_id by blast moreover have "\l. restrict id (standard_simplex p) = oriented_simplex p l" apply (rule_tac x="\i j. if i = j then 1 else 0" in exI) apply (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\u. u * _"] cong: if_cong) done ultimately show ?lhs by (simp add: simplicial_simplex_def singular_simplex_def) qed lemma singular_chain_singular_subdivision: "singular_chain p X c \ singular_chain p X (singular_subdivision p c)" unfolding singular_subdivision_def apply (rule singular_chain_extend) apply (rule singular_chain_chain_map [where X = "subtopology (powertop_real UNIV) (standard_simplex p)"]) apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain) by (simp add: singular_chain_def singular_simplex_def subset_iff) lemma naturality_singular_subdivision: "singular_chain p X c \ singular_subdivision p (chain_map p g c) = chain_map p g (singular_subdivision p c)" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have "singular_simplex p X f" by auto have "\simplicial_chain p (standard_simplex p) d\ \ chain_map p (simplex_map p g f) d = chain_map p g (chain_map p f d)" for d unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then have "simplex_map p (simplex_map p g f) x = simplex_map p g (simplex_map p f x)" by (force simp: simplex_map_def restrict_compose_left simplicial_simplex) then show ?case by auto qed (auto simp: chain_map_diff) then show ?case using simplicial_chain_simplicial_subdivision [of p "standard_simplex p" "frag_of (restrict id (standard_simplex p))"] by (simp add: singular_subdivision_def) next case (diff a b) then show ?case by (simp add: chain_map_diff singular_subdivision_diff) qed auto lemma simplicial_chain_chain_map: assumes f: "simplicial_simplex q X f" and c: "simplicial_chain p (standard_simplex q) c" shows "simplicial_chain p X (chain_map p f c)" using c unfolding simplicial_chain_def proof (induction c rule: frag_induction) case (one g) have "\n. simplex_map p (oriented_simplex q l) (oriented_simplex p m) = oriented_simplex p n" if m: "singular_simplex p (subtopology (powertop_real UNIV) (standard_simplex q)) (oriented_simplex p m)" for l m proof - have "(\i. \j\p. m j i * x j) \ standard_simplex q" if "x \ standard_simplex p" for x using that m unfolding oriented_simplex_def singular_simplex_def by (auto simp: continuous_map_in_subtopology image_subset_iff) then show ?thesis unfolding oriented_simplex_def simplex_map_def apply (rule_tac x="\j k. (\i\q. l i k * m j i)" in exI) apply (force simp: sum_distrib_left sum_distrib_right mult.assoc intro: sum.swap) done qed then show ?case using f one apply (auto simp: simplicial_simplex_def) apply (rule singular_simplex_simplex_map [where X = "subtopology (powertop_real UNIV) (standard_simplex q)"]) unfolding singular_simplex_def apply (fastforce simp add:)+ done next case (diff a b) then show ?case by (metis chain_map_diff simplicial_chain_def simplicial_chain_diff) qed auto lemma singular_subdivision_simplicial_simplex: "simplicial_chain p S c \ singular_subdivision p c = simplicial_subdivision p c" proof (induction p arbitrary: S c) case 0 then show ?case unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then show ?case using singular_simplex_chain_map_id simplicial_imp_singular_simplex by (fastforce simp: singular_subdivision_def simplicial_subdivision_def) qed (auto simp: singular_subdivision_diff) next case (Suc p) show ?case using Suc.prems unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one f) then have ssf: "simplicial_simplex (Suc p) S f" by (auto simp: simplicial_simplex) then have 1: "simplicial_chain p (standard_simplex (Suc p)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))" by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_simplex_id) have 2: "(\i. (\j\Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2)) \ standard_simplex (Suc p)" by (simp add: simplicial_vertex_def standard_simplex_def del: sum.atMost_Suc) have ss_Sp: "(\i. (if i \ Suc p then 1 else 0) / (real p + 2)) \ standard_simplex (Suc p)" by (simp add: standard_simplex_def field_split_simps) obtain l where feq: "f = oriented_simplex (Suc p) l" using one unfolding simplicial_simplex by blast then have 3: "f (\i. (\j\Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2)) = (\i. (\j\Suc p. simplicial_vertex j f i) / (real p + 2))" unfolding simplicial_vertex_def oriented_simplex_def by (simp add: ss_Sp if_distrib [of "\x. _ * x"] sum_divide_distrib del: sum.atMost_Suc cong: if_cong) have scp: "singular_chain (Suc p) (subtopology (powertop_real UNIV) (standard_simplex (Suc p))) (frag_of (restrict id (standard_simplex (Suc p))))" by (simp add: simplicial_imp_singular_chain) have scps: "simplicial_chain p (standard_simplex (Suc p)) (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))" by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_simplex_id) have scpf: "simplicial_chain p S (chain_map p f (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))" using scps simplicial_chain_chain_map ssf by blast have 4: "chain_map p f (simplicial_subdivision p (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))) = simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))" apply (simp add: chain_boundary_chain_map [OF scp] del: chain_map_of flip: singular_simplex_chain_map_id [OF simplicial_imp_singular_simplex [OF ssf]]) by (metis (no_types) scp singular_chain_boundary_alt Suc.IH [OF scps] Suc.IH [OF scpf] naturality_singular_subdivision) show ?case apply (simp add: singular_subdivision_def del: sum.atMost_Suc) apply (simp only: ssf 1 2 3 4 chain_map_simplicial_cone [of "Suc p" S _ p "Suc p"]) done qed (auto simp: frag_extend_diff singular_subdivision_diff) qed lemma naturality_simplicial_subdivision: "\simplicial_chain p (standard_simplex q) c; simplicial_simplex q S g\ \ simplicial_subdivision p (chain_map p g c) = chain_map p g (simplicial_subdivision p c)" apply (simp flip: singular_subdivision_simplicial_simplex) by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain singular_subdivision_simplicial_simplex) lemma chain_boundary_singular_subdivision: "singular_chain p X c \ chain_boundary p (singular_subdivision p c) = singular_subdivision (p - Suc 0) (chain_boundary p c)" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have ssf: "singular_simplex p X f" by (auto simp: singular_simplex_def) then have scp: "simplicial_chain p (standard_simplex p) (frag_of (restrict id (standard_simplex p)))" by simp have scp1: "simplicial_chain (p - Suc 0) (standard_simplex p) (chain_boundary p (frag_of (restrict id (standard_simplex p))))" using simplicial_chain_boundary by force have sgp1: "singular_chain (p - Suc 0) (subtopology (powertop_real UNIV) (standard_simplex p)) (chain_boundary p (frag_of (restrict id (standard_simplex p))))" using scp1 simplicial_imp_singular_chain by blast have scpp: "singular_chain p (subtopology (powertop_real UNIV) (standard_simplex p)) (frag_of (restrict id (standard_simplex p)))" using scp simplicial_imp_singular_chain by blast then show ?case unfolding singular_subdivision_def using chain_boundary_chain_map [of p "subtopology (powertop_real UNIV) (standard_simplex p)" _ f] apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain chain_boundary_simplicial_subdivision [OF scp] flip: singular_subdivision_simplicial_simplex [OF scp1] naturality_singular_subdivision [OF sgp1]) by (metis (full_types) singular_subdivision_def chain_boundary_chain_map [OF scpp] singular_simplex_chain_map_id [OF ssf]) qed (auto simp: singular_subdivision_def frag_extend_diff chain_boundary_diff) lemma singular_subdivision_zero: "singular_chain 0 X c \ singular_subdivision 0 c = c" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have "restrict (f \ restrict id (standard_simplex 0)) (standard_simplex 0) = f" by (simp add: extensional_restrict restrict_compose_right singular_simplex_def) then show ?case by (auto simp: singular_subdivision_def simplex_map_def) qed (auto simp: singular_subdivision_def frag_extend_diff) primrec subd where "subd 0 = (\x. 0)" | "subd (Suc p) = frag_extend (\f. simplicial_cone (Suc p) (\i. (\j\Suc p. simplicial_vertex j f i) / real (Suc p + 1)) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - subd p (chain_boundary (Suc p) (frag_of f))))" lemma subd_0 [simp]: "subd p 0 = 0" by (induction p) auto lemma subd_diff [simp]: "subd p (c1 - c2) = subd p c1 - subd p c2" by (induction p) (auto simp: frag_extend_diff) lemma subd_uminus [simp]: "subd p (-c) = - subd p c" by (metis diff_0 subd_0 subd_diff) lemma subd_power_uminus: "subd p (frag_cmul ((-1) ^ k) c) = frag_cmul ((-1) ^ k) (subd p c)" apply (induction k, simp_all) by (metis minus_frag_cmul subd_uminus) lemma subd_power_sum: "subd p (sum f I) = sum (subd p \ f) I" apply (induction I rule: infinite_finite_induct) by auto (metis add_diff_cancel_left' diff_add_cancel subd_diff) lemma subd: "simplicial_chain p (standard_simplex s) c \ (\r g. simplicial_simplex s (standard_simplex r) g \ chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)) \ simplicial_chain (Suc p) (standard_simplex s) (subd p c) \ (chain_boundary (Suc p) (subd p c)) + (subd (p - Suc 0) (chain_boundary p c)) = (simplicial_subdivision p c) - c" proof (induction p arbitrary: c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction rule: frag_induction) case (one f) then obtain l where l: "(\x i. \j\Suc p. l j i * x j) ` standard_simplex (Suc p) \ standard_simplex s" and feq: "f = oriented_simplex (Suc p) l" by (metis (mono_tags) mem_Collect_eq simplicial_simplex simplicial_simplex_oriented_simplex) have scf: "simplicial_chain (Suc p) (standard_simplex s) (frag_of f)" using one by simp have lss: "l i \ standard_simplex s" if "i \ Suc p" for i proof - have "(\i'. \j\Suc p. l j i' * (if j = i then 1 else 0)) \ standard_simplex s" using subsetD [OF l] basis_in_standard_simplex that by blast moreover have "(\i'. \j\Suc p. l j i' * (if j = i then 1 else 0)) = l i" using that by (simp add: if_distrib [of "\x. _ * x"] del: sum.atMost_Suc cong: if_cong) ultimately show ?thesis by simp qed have *: "(\i. i \ n \ l i \ standard_simplex s) \ (\i. (\j\n. l j i) / (Suc n)) \ standard_simplex s" for n proof (induction n) case (Suc n) let ?x = "\i. (1 - inverse (n + 2)) * ((\j\n. l j i) / (Suc n)) + inverse (n + 2) * l (Suc n) i" have "?x \ standard_simplex s" proof (rule convex_standard_simplex) show "(\i. (\j\n. l j i) / real (Suc n)) \ standard_simplex s" using Suc by simp qed (auto simp: lss Suc inverse_le_1_iff) moreover have "?x = (\i. (\j\Suc n. l j i) / real (Suc (Suc n)))" by (force simp: divide_simps) ultimately show ?case by simp qed auto have **: "(\i. (\j\Suc p. simplicial_vertex j f i) / (2 + real p)) \ standard_simplex s" using * [of "Suc p"] lss by (simp add: simplicial_vertex_oriented_simplex feq) show ?case proof (intro conjI impI allI) fix r g assume g: "simplicial_simplex s (standard_simplex r) g" then obtain m where geq: "g = oriented_simplex s m" using simplicial_simplex by blast have 1: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_subdivision (Suc p) (frag_of f))" by (metis mem_Collect_eq one.hyps simplicial_chain_of simplicial_chain_simplicial_subdivision) have 2: "(\j\Suc p. \i\s. m i k * simplicial_vertex j f i) = (\j\Suc p. simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k)" for k proof (rule sum.cong [OF refl]) fix j assume j: "j \ {..Suc p}" have eq: "simplex_map (Suc p) (oriented_simplex s m) (oriented_simplex (Suc p) l) = oriented_simplex (Suc p) (oriented_simplex s m \ l)" proof (rule simplex_map_oriented_simplex) show "simplicial_simplex (Suc p) (standard_simplex s) (oriented_simplex (Suc p) l)" using one by (simp add: feq flip: oriented_simplex_def) show "simplicial_simplex s (standard_simplex r) (oriented_simplex s m)" using g by (simp add: geq) qed auto show "(\i\s. m i k * simplicial_vertex j f i) = simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k" using one j apply (simp add: feq eq simplicial_vertex_oriented_simplex simplicial_simplex_oriented_simplex image_subset_iff) apply (drule_tac x="(\i. if i = j then 1 else 0)" in bspec) apply (auto simp: oriented_simplex_def lss) done qed have 4: "chain_map (Suc p) g (subd p (chain_boundary (Suc p) (frag_of f))) = subd p (chain_boundary (Suc p) (frag_of (simplex_map (Suc p) g f)))" by (metis (no_types) One_nat_def scf Suc.IH chain_boundary_chain_map chain_map_of diff_Suc_Suc diff_zero g simplicial_chain_boundary simplicial_imp_singular_chain) show "chain_map (Suc (Suc p)) g (subd (Suc p) (frag_of f)) = subd (Suc p) (chain_map (Suc p) g (frag_of f))" using g apply (simp only: subd.simps frag_extend_of) apply (subst chain_map_simplicial_cone [of s "standard_simplex r" _ "Suc p" s], assumption) apply (intro simplicial_chain_diff) using "1" apply auto[1] using one.hyps apply auto[1] apply (metis Suc.IH diff_Suc_1 mem_Collect_eq one.hyps simplicial_chain_boundary simplicial_chain_of) using "**" apply auto[1] apply (rule order_refl) apply (simp only: chain_map_of frag_extend_of) apply (rule arg_cong2 [where f = "simplicial_cone (Suc p)"]) apply (simp add: geq sum_distrib_left oriented_simplex_def ** del: sum.atMost_Suc flip: sum_divide_distrib) using 2 apply (simp only: oriented_simplex_def sum.swap [where A = "{..s}"]) using naturality_simplicial_subdivision scf apply (fastforce simp add: 4 chain_map_diff) done next have sc: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_cone p (\i. (\j\Suc p. simplicial_vertex j f i) / (Suc (Suc p))) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))" by (metis diff_Suc_1 nat.simps(3) simplicial_subdivision_of scf simplicial_chain_simplicial_subdivision) have ff: "simplicial_chain (Suc p) (standard_simplex s) (subd p (chain_boundary (Suc p) (frag_of f)))" by (metis (no_types) Suc.IH diff_Suc_1 scf simplicial_chain_boundary) show "simplicial_chain (Suc (Suc p)) (standard_simplex s) (subd (Suc p) (frag_of f))" using one apply (simp only: subd.simps frag_extend_of) apply (rule_tac S="standard_simplex s" in simplicial_chain_simplicial_cone) apply (intro simplicial_chain_diff ff) using sc apply (simp add: algebra_simps) using "**" convex_standard_simplex apply force+ done have "simplicial_chain p (standard_simplex s) (chain_boundary (Suc p) (frag_of f))" using scf simplicial_chain_boundary by fastforce then have "chain_boundary (Suc p) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - subd p (chain_boundary (Suc p) (frag_of f))) = 0" apply (simp only: chain_boundary_diff) using Suc.IH chain_boundary_boundary [of "Suc p" "subtopology (powertop_real UNIV) (standard_simplex s)" "frag_of f"] by (metis One_nat_def add_diff_cancel_left' subd_0 chain_boundary_simplicial_subdivision plus_1_eq_Suc scf simplicial_imp_singular_chain) then show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f)) + subd (Suc p - Suc 0) (chain_boundary (Suc p) (frag_of f)) = simplicial_subdivision (Suc p) (frag_of f) - frag_of f" apply (simp only: subd.simps frag_extend_of) apply (subst chain_boundary_simplicial_cone [of "Suc p" "standard_simplex s"]) apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision) apply (simp add: simplicial_cone_def del: sum.atMost_Suc simplicial_subdivision.simps) done qed next case (diff a b) then show ?case apply safe apply (metis chain_map_diff subd_diff) apply (metis simplicial_chain_diff subd_diff) apply (auto simp: simplicial_subdivision_diff chain_boundary_diff simp del: simplicial_subdivision.simps subd.simps) by (metis (no_types, lifting) add_diff_add add_uminus_conv_diff diff_0 diff_diff_add) qed auto qed simp lemma chain_homotopic_simplicial_subdivision1: "\simplicial_chain p (standard_simplex q) c; simplicial_simplex q (standard_simplex r) g\ \ chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)" by (simp add: subd) lemma chain_homotopic_simplicial_subdivision2: "simplicial_chain p (standard_simplex q) c \ simplicial_chain (Suc p) (standard_simplex q) (subd p c)" by (simp add: subd) lemma chain_homotopic_simplicial_subdivision3: "simplicial_chain p (standard_simplex q) c \ chain_boundary (Suc p) (subd p c) = (simplicial_subdivision p c) - c - subd (p - Suc 0) (chain_boundary p c)" by (simp add: subd algebra_simps) lemma chain_homotopic_simplicial_subdivision: "\h. (\p. h p 0 = 0) \ (\p c1 c2. h p (c1-c2) = h p c1 - h p c2) \ (\p q r g c. simplicial_chain p (standard_simplex q) c \ simplicial_simplex q (standard_simplex r) g \ chain_map (Suc p) g (h p c) = h p (chain_map p g c)) \ (\p q c. simplicial_chain p (standard_simplex q) c \ simplicial_chain (Suc p) (standard_simplex q) (h p c)) \ (\p q c. simplicial_chain p (standard_simplex q) c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (simplicial_subdivision p c) - c)" by (rule_tac x=subd in exI) (fastforce simp: subd) lemma chain_homotopic_singular_subdivision: obtains h where "\p. h p 0 = 0" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" proof - define k where "k \ \p. frag_extend (\f:: (nat \ real) \ 'a. chain_map (Suc p) f (subd p (frag_of(restrict id (standard_simplex p)))))" show ?thesis proof fix p X and c :: "'a chain" assume c: "singular_chain p X c" have "singular_chain (Suc p) X (k p c) \ chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" using c [unfolded singular_chain_def] proof (induction rule: frag_induction) case (one f) let ?X = "subtopology (powertop_real UNIV) (standard_simplex p)" show ?case proof (simp add: k_def, intro conjI) show "singular_chain (Suc p) X (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p)))))" proof (rule singular_chain_chain_map) show "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))" by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain) show "continuous_map ?X X f" using one.hyps singular_simplex_def by auto qed next have scp: "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))" by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain) have feqf: "frag_of (simplex_map p f (restrict id (standard_simplex p))) = frag_of f" using one.hyps singular_simplex_chain_map_id by auto have *: "chain_map p f (subd (p - Suc 0) (\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k id)))) = (\x\p. frag_cmul ((-1) ^ x) (chain_map p (singular_face p x f) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))" (is "?lhs = ?rhs") if "p > 0" proof - have eqc: "subd (p - Suc 0) (frag_of (singular_face p i id)) = chain_map p (singular_face p i id) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))" if "i \ p" for i proof - have 1: "simplicial_chain (p - Suc 0) (standard_simplex (p - Suc 0)) (frag_of (restrict id (standard_simplex (p - Suc 0))))" by simp have 2: "simplicial_simplex (p - Suc 0) (standard_simplex p) (singular_face p i id)" by (metis One_nat_def Suc_leI \0 < p\ simplicial_simplex_id simplicial_simplex_singular_face singular_face_restrict subsetI that) have 3: "simplex_map (p - Suc 0) (singular_face p i id) (restrict id (standard_simplex (p - Suc 0))) = singular_face p i id" by (force simp: simplex_map_def singular_face_def) show ?thesis using chain_homotopic_simplicial_subdivision1 [OF 1 2] that \p > 0\ by (simp add: 3) qed have xx: "simplicial_chain p (standard_simplex(p - Suc 0)) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))" by (metis Suc_pred chain_homotopic_simplicial_subdivision2 order_refl simplicial_chain_of simplicial_simplex_id that) have yy: "\k. k \ p \ chain_map p f (chain_map p (singular_face p k id) h) = chain_map p (singular_face p k f) h" if "simplicial_chain p (standard_simplex(p - Suc 0)) h" for h using that unfolding simplicial_chain_def proof (induction h rule: frag_induction) case (one x) then show ?case using one apply (simp add: chain_map_of singular_simplex_def simplicial_simplex_def, auto) apply (rule_tac f=frag_of in arg_cong, rule) apply (simp add: simplex_map_def) by (simp add: continuous_map_in_subtopology image_subset_iff singular_face_def) qed (auto simp: chain_map_diff) have "?lhs = chain_map p f (\k\p. frag_cmul ((-1) ^ k) (chain_map p (singular_face p k id) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))" by (simp add: subd_power_sum subd_power_uminus eqc) also have "\ = ?rhs" by (simp add: chain_map_sum xx yy) finally show ?thesis . qed have "chain_map p f (simplicial_subdivision p (frag_of (restrict id (standard_simplex p))) - subd (p - Suc 0) (chain_boundary p (frag_of (restrict id (standard_simplex p))))) = singular_subdivision p (frag_of f) - frag_extend (\f. chain_map (Suc (p - Suc 0)) f (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))) (chain_boundary p (frag_of f))" apply (simp add: singular_subdivision_def chain_map_diff) apply (clarsimp simp add: chain_boundary_def) apply (simp add: frag_extend_sum frag_extend_cmul *) done then show "chain_boundary (Suc p) (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p))))) + frag_extend (\f. chain_map (Suc (p - Suc 0)) f (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))) (chain_boundary p (frag_of f)) = singular_subdivision p (frag_of f) - frag_of f" by (simp add: chain_boundary_chain_map [OF scp] chain_homotopic_simplicial_subdivision3 [where q=p] chain_map_diff feqf) qed next case (diff a b) then show ?case apply (simp only: k_def singular_chain_diff chain_boundary_diff frag_extend_diff singular_subdivision_diff) by (metis (no_types, lifting) add_diff_add diff_add_cancel) qed (auto simp: k_def) then show "singular_chain (Suc p) X (k p c)" "chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" by auto qed (auto simp: k_def frag_extend_diff) qed lemma homologous_rel_singular_subdivision: assumes "singular_relcycle p X T c" shows "homologous_rel p X T (singular_subdivision p c) c" proof (cases "p = 0") case True with assms show ?thesis by (auto simp: singular_relcycle_def singular_subdivision_zero) next case False with assms show ?thesis unfolding homologous_rel_def singular_relboundary singular_relcycle by (metis One_nat_def Suc_diff_1 chain_homotopic_singular_subdivision gr_zeroI) qed subsection\Excision argument that we keep doing singular subdivision\ lemma singular_subdivision_power_0 [simp]: "(singular_subdivision p ^^ n) 0 = 0" by (induction n) auto lemma singular_subdivision_power_diff: "(singular_subdivision p ^^ n) (a - b) = (singular_subdivision p ^^ n) a - (singular_subdivision p ^^ n) b" by (induction n) (auto simp: singular_subdivision_diff) lemma iterated_singular_subdivision: "singular_chain p X c \ (singular_subdivision p ^^ n) c = frag_extend (\f. chain_map p f ((simplicial_subdivision p ^^ n) (frag_of(restrict id (standard_simplex p))))) c" proof (induction n arbitrary: c) case 0 then show ?case unfolding singular_chain_def proof (induction c rule: frag_induction) case (one f) then have "restrict f (standard_simplex p) = f" by (simp add: extensional_restrict singular_simplex_def) then show ?case by (auto simp: simplex_map_def cong: restrict_cong) qed (auto simp: frag_extend_diff) next case (Suc n) show ?case using Suc.prems unfolding singular_chain_def proof (induction c rule: frag_induction) case (one f) then have "singular_simplex p X f" by simp have scp: "simplicial_chain p (standard_simplex p) ((simplicial_subdivision p ^^ n) (frag_of (restrict id (standard_simplex p))))" proof (induction n) case 0 then show ?case by (metis funpow_0 order_refl simplicial_chain_of simplicial_simplex_id) next case (Suc n) then show ?case by (simp add: simplicial_chain_simplicial_subdivision) qed have scnp: "simplicial_chain p (standard_simplex p) ((simplicial_subdivision p ^^ n) (frag_of (\x\standard_simplex p. x)))" proof (induction n) case 0 then show ?case by (metis eq_id_iff funpow_0 order_refl simplicial_chain_of simplicial_simplex_id) next case (Suc n) then show ?case by (simp add: simplicial_chain_simplicial_subdivision) qed have sff: "singular_chain p X (frag_of f)" by (simp add: \singular_simplex p X f\ singular_chain_of) then show ?case using Suc.IH [OF sff] naturality_singular_subdivision [OF simplicial_imp_singular_chain [OF scp], of f] singular_subdivision_simplicial_simplex [OF scnp] by (simp add: singular_chain_of id_def del: restrict_apply) qed (auto simp: singular_subdivision_power_diff singular_subdivision_diff frag_extend_diff) qed lemma chain_homotopic_iterated_singular_subdivision: obtains h where "\p. h p 0 = (0 :: 'a chain)" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ n) c - c" proof (induction n arbitrary: thesis) case 0 show ?case by (rule 0 [of "(\p x. 0)"]) auto next case (Suc n) then obtain k where k: "\p. k p 0 = (0 :: 'a chain)" "\p c1 c2. k p (c1-c2) = k p c1 - k p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (k p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ n) c - c" by metis obtain h where h: "\p. h p 0 = (0 :: 'a chain)" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" by (blast intro: chain_homotopic_singular_subdivision) let ?h = "(\p c. singular_subdivision (Suc p) (k p c) + h p c)" show ?case proof (rule Suc.prems) fix p X and c :: "'a chain" assume "singular_chain p X c" then show "singular_chain (Suc p) X (?h p c)" by (simp add: h k singular_chain_add singular_chain_singular_subdivision) next fix p :: "nat" and X :: "'a topology" and c :: "'a chain" assume sc: "singular_chain p X c" have f5: "chain_boundary (Suc p) (singular_subdivision (Suc p) (k p c)) = singular_subdivision p (chain_boundary (Suc p) (k p c))" using chain_boundary_singular_subdivision k(3) sc by fastforce have [simp]: "singular_subdivision (Suc (p - Suc 0)) (k (p - Suc 0) (chain_boundary p c)) = singular_subdivision p (k (p - Suc 0) (chain_boundary p c))" proof (cases p) case 0 then show ?thesis by (simp add: k chain_boundary_def) qed auto show "chain_boundary (Suc p) (?h p c) + ?h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ Suc n) c - c" using chain_boundary_singular_subdivision [of "Suc p" X] apply (simp add: chain_boundary_add f5 h k algebra_simps) apply (erule thin_rl) using h(4) [OF sc] k(4) [OF sc] singular_subdivision_add [of p "chain_boundary (Suc p) (k p c)" "k (p - Suc 0) (chain_boundary p c)"] apply (simp add: algebra_simps) by (smt add.assoc add.left_commute singular_subdivision_add) qed (auto simp: k h singular_subdivision_diff) qed lemma llemma: assumes p: "standard_simplex p \ \\" and \: "\U. U \ \ \ openin (powertop_real UNIV) U" obtains d where "0 < d" "\K. \K \ standard_simplex p; \x y i. \i \ p; x \ K; y \ K\ \ \x i - y i\ \ d\ \ \U. U \ \ \ K \ U" proof - have "\e U. 0 < e \ U \ \ \ x \ U \ (\y. (\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0) \ y \ U)" if x: "x \ standard_simplex p" for x proof- obtain U where U: "U \ \" "x \ U" using x p by blast then obtain V where finV: "finite {i. V i \ UNIV}" and openV: "\i. open (V i)" and xV: "x \ Pi\<^sub>E UNIV V" and UV: "Pi\<^sub>E UNIV V \ U" using \ unfolding openin_product_topology_alt by force have xVi: "x i \ V i" for i using PiE_mem [OF xV] by simp have "\i. \e>0. \x'. \x' - x i\ < e \ x' \ V i" by (rule openV [unfolded open_real, rule_format, OF xVi]) then obtain d where d: "\i. d i > 0" and dV: "\i x'. \x' - x i\ < d i \ x' \ V i" by metis define e where "e \ Inf (insert 1 (d ` {i. V i \ UNIV})) / 3" have ed3: "e \ d i / 3" if "V i \ UNIV" for i using that finV by (auto simp: e_def intro: cInf_le_finite) show "\e U. 0 < e \ U \ \ \ x \ U \ (\y. (\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0) \ y \ U)" proof (intro exI conjI allI impI) show "e > 0" using d finV by (simp add: e_def finite_less_Inf_iff) fix y assume y: "(\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0)" have "y \ Pi\<^sub>E UNIV V" proof show "y i \ V i" for i proof (cases "p < i") case True then show ?thesis by (metis (mono_tags, lifting) y x mem_Collect_eq standard_simplex_def xVi) next case False show ?thesis proof (cases "V i = UNIV") case False show ?thesis proof (rule dV) have "\y i - x i\ \ 2 * e" using y \\ p < i\ by simp also have "\ < d i" using ed3 [OF False] \e > 0\ by simp finally show "\y i - x i\ < d i" . qed qed auto qed qed auto with UV show "y \ U" by blast qed (use U in auto) qed then obtain e U where eU: "\x. x \ standard_simplex p \ 0 < e x \ U x \ \ \ x \ U x" and UI: "\x y. \x \ standard_simplex p; \i. i \ p \ \y i - x i\ \ 2 * e x; \i. i > p \ y i = 0\ \ y \ U x" by metis define F where "F \ \x. Pi\<^sub>E UNIV (\i. if i \ p then {x i - e x<..S \ F ` standard_simplex p. openin (powertop_real UNIV) S" by (simp add: F_def openin_PiE_gen) moreover have pF: "standard_simplex p \ \(F ` standard_simplex p)" by (force simp: F_def PiE_iff eU) ultimately have "\\. finite \ \ \ \ F ` standard_simplex p \ standard_simplex p \ \\" using compactin_standard_simplex [of p] unfolding compactin_def by force then obtain S where "finite S" and ssp: "S \ standard_simplex p" "standard_simplex p \ \(F ` S)" unfolding ex_finite_subset_image by (auto simp: ex_finite_subset_image) then have "S \ {}" by (auto simp: nonempty_standard_simplex) show ?thesis proof show "Inf (e ` S) > 0" using \finite S\ \S \ {}\ ssp eU by (auto simp: finite_less_Inf_iff) fix k :: "(nat \ real) set" assume k: "k \ standard_simplex p" and kle: "\x y i. \i \ p; x \ k; y \ k\ \ \x i - y i\ \ Inf (e ` S)" show "\U. U \ \ \ k \ U" proof (cases "k = {}") case True then show ?thesis using \S \ {}\ eU equals0I ssp(1) subset_eq p by auto next case False with k ssp obtain x a where "x \ k" "x \ standard_simplex p" and a: "a \ S" and Fa: "x \ F a" by blast then have le_ea: "\i. i \ p \ abs (x i - a i) < e a" by (simp add: F_def PiE_iff if_distrib abs_diff_less_iff cong: if_cong) show ?thesis proof (intro exI conjI) show "U a \ \" using a eU ssp(1) by auto show "k \ U a" proof clarify fix y assume "y \ k" with k have y: "y \ standard_simplex p" by blast show "y \ U a" proof (rule UI) show "a \ standard_simplex p" using a ssp(1) by auto fix i :: "nat" assume "i \ p" then have "\x i - y i\ \ e a" by (meson kle [OF \i \ p\] a \finite S\ \x \ k\ \y \ k\ cInf_le_finite finite_imageI imageI order_trans) then show "\y i - a i\ \ 2 * e a" using le_ea [OF \i \ p\] by linarith next fix i assume "p < i" then show "y i = 0" using standard_simplex_def y by auto qed qed qed qed qed qed proposition sufficient_iterated_singular_subdivision_exists: assumes \: "\U. U \ \ \ openin X U" and X: "topspace X \ \\" and p: "singular_chain p X c" obtains n where "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ \. f ` (standard_simplex p) \ V" proof (cases "c = 0") case False then show ?thesis proof (cases "topspace X = {}") case True show ?thesis using p that by (force simp: singular_chain_empty True) next case False show ?thesis proof (cases "\ = {}") case True then show ?thesis using False X by blast next case False have "\e. 0 < e \ (\K. K \ standard_simplex p \ (\x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ e) \ (\V. V \ \ \ f ` K \ V))" if f: "f \ Poly_Mapping.keys c" for f proof - have ssf: "singular_simplex p X f" using f p by (auto simp: singular_chain_def) then have fp: "\x. x \ standard_simplex p \ f x \ topspace X" by (auto simp: singular_simplex_def image_subset_iff dest: continuous_map_image_subset_topspace) have "\T. openin (powertop_real UNIV) T \ standard_simplex p \ f -` V = T \ standard_simplex p" if V: "V \ \" for V proof - have "singular_simplex p X f" using p f unfolding singular_chain_def by blast then have "openin (subtopology (powertop_real UNIV) (standard_simplex p)) {x \ standard_simplex p. f x \ V}" using \ [OF \V \ \\] by (simp add: singular_simplex_def continuous_map_def) moreover have "standard_simplex p \ f -` V = {x \ standard_simplex p. f x \ V}" by blast ultimately show ?thesis by (simp add: openin_subtopology) qed then obtain g where gope: "\V. V \ \ \ openin (powertop_real UNIV) (g V)" and geq: "\V. V \ \ \ standard_simplex p \ f -` V = g V \ standard_simplex p" by metis obtain d where "0 < d" and d: "\K. \K \ standard_simplex p; \x y i. \i \ p; x \ K; y \ K\ \ \x i - y i\ \ d\ \ \U. U \ g ` \ \ K \ U" proof (rule llemma [of p "g ` \"]) show "standard_simplex p \ \(g ` \)" using geq X fp by (fastforce simp add:) show "openin (powertop_real UNIV) U" if "U \ g ` \" for U :: "(nat \ real) set" using gope that by blast qed auto show ?thesis proof (rule exI, intro allI conjI impI) fix K :: "(nat \ real) set" assume K: "K \ standard_simplex p" and Kd: "\x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ d" then have "\U. U \ g ` \ \ K \ U" using d [OF K] by auto then show "\V. V \ \ \ f ` K \ V" using K geq by fastforce qed (rule \d > 0\) qed then obtain \ where epos: "\f \ Poly_Mapping.keys c. 0 < \ f" and e: "\f K. \f \ Poly_Mapping.keys c; K \ standard_simplex p; \x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ \ f\ \ \V. V \ \ \ f ` K \ V" by metis obtain d where "0 < d" and d: "\f K. \f \ Poly_Mapping.keys c; K \ standard_simplex p; \x y i. \x \ K; y \ K; i \ p\ \ \x i - y i\ \ d\ \ \V. V \ \ \ f ` K \ V" proof show "Inf (\ ` Poly_Mapping.keys c) > 0" by (simp add: finite_less_Inf_iff \c \ 0\ epos) fix f K assume fK: "f \ Poly_Mapping.keys c" "K \ standard_simplex p" and le: "\x y i. \x \ K; y \ K; i \ p\ \ \x i - y i\ \ Inf (\ ` Poly_Mapping.keys c)" then have lef: "Inf (\ ` Poly_Mapping.keys c) \ \ f" by (auto intro: cInf_le_finite) show "\V. V \ \ \ f ` K \ V" using le lef by (blast intro: dual_order.trans e [OF fK]) qed let ?d = "\m. (simplicial_subdivision p ^^ m) (frag_of (restrict id (standard_simplex p)))" obtain n where n: "(p / (Suc p)) ^ n < d" using real_arch_pow_inv \0 < d\ by fastforce show ?thesis proof fix m h assume "n \ m" and "h \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)" then obtain f where "f \ Poly_Mapping.keys c" "h \ Poly_Mapping.keys (chain_map p f (?d m))" using subsetD [OF keys_frag_extend] iterated_singular_subdivision [OF p, of m] by force then obtain g where g: "g \ Poly_Mapping.keys (?d m)" and heq: "h = restrict (f \ g) (standard_simplex p)" using keys_frag_extend by (force simp: chain_map_def simplex_map_def) have xx: "simplicial_chain p (standard_simplex p) (?d n) \ (\f \ Poly_Mapping.keys(?d n). \x \ standard_simplex p. \y \ standard_simplex p. \f x i - f y i\ \ (p / (Suc p)) ^ n)" for n i proof (induction n) case 0 have "simplicial_simplex p (standard_simplex p) (\a\standard_simplex p. a)" by (metis eq_id_iff order_refl simplicial_simplex_id) moreover have "(\x\standard_simplex p. \y\standard_simplex p. \x i - y i\ \ 1)" unfolding standard_simplex_def by (auto simp: abs_if dest!: spec [where x=i]) ultimately show ?case unfolding power_0 funpow_0 by simp next case (Suc n) show ?case unfolding power_Suc funpow.simps o_def proof (intro conjI ballI) show "simplicial_chain p (standard_simplex p) (simplicial_subdivision p (?d n))" by (simp add: Suc simplicial_chain_simplicial_subdivision) show "\f x i - f y i\ \ real p / real (Suc p) * (real p / real (Suc p)) ^ n" if "f \ Poly_Mapping.keys (simplicial_subdivision p (?d n))" and "x \ standard_simplex p" and "y \ standard_simplex p" for f x y using Suc that by (blast intro: simplicial_subdivision_shrinks) qed qed have "g ` standard_simplex p \ standard_simplex p" using g xx [of m] unfolding simplicial_chain_def simplicial_simplex by auto moreover have "\g x i - g y i\ \ d" if "i \ p" "x \ standard_simplex p" "y \ standard_simplex p" for x y i proof - have "\g x i - g y i\ \ (p / (Suc p)) ^ m" using g xx [of m] that by blast also have "\ \ (p / (Suc p)) ^ n" by (auto intro: power_decreasing [OF \n \ m\]) finally show ?thesis using n by simp qed then have "\x i - y i\ \ d" if "x \ g ` (standard_simplex p)" "y \ g ` (standard_simplex p)" "i \ p" for i x y using that by blast ultimately show "\V\\. h ` standard_simplex p \ V" using \f \ Poly_Mapping.keys c\ d [of f "g ` standard_simplex p"] by (simp add: Bex_def heq image_image) qed qed qed qed force lemma small_homologous_rel_relcycle_exists: assumes \: "\U. U \ \ \ openin X U" and X: "topspace X \ \\" and p: "singular_relcycle p X S c" obtains c' where "singular_relcycle p X S c'" "homologous_rel p X S c c'" "\f. f \ Poly_Mapping.keys c' \ \V \ \. f ` (standard_simplex p) \ V" proof - have "singular_chain p X c" "(chain_boundary p c, 0) \ (mod_subset (p - Suc 0) (subtopology X S))" using p unfolding singular_relcycle_def by auto then obtain n where n: "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ \. f ` (standard_simplex p) \ V" by (blast intro: sufficient_iterated_singular_subdivision_exists [OF \ X]) let ?c' = "(singular_subdivision p ^^ n) c" show ?thesis proof show "homologous_rel p X S c ?c'" apply (induction n, simp_all) by (metis p homologous_rel_singular_subdivision homologous_rel_singular_relcycle homologous_rel_trans homologous_rel_sym) then show "singular_relcycle p X S ?c'" by (metis homologous_rel_singular_relcycle p) next fix f :: "(nat \ real) \ 'a" assume "f \ Poly_Mapping.keys ?c'" then show "\V\\. f ` standard_simplex p \ V" by (rule n [OF order_refl]) qed qed lemma excised_chain_exists: fixes S :: "'a set" assumes "X closure_of U \ X interior_of T" "T \ S" "singular_chain p (subtopology X S) c" obtains n d e where "singular_chain p (subtopology X (S - U)) d" "singular_chain p (subtopology X T) e" "(singular_subdivision p ^^ n) c = d + e" proof - have *: "\n d e. singular_chain p (subtopology X (S - U)) d \ singular_chain p (subtopology X T) e \ (singular_subdivision p ^^ n) c = d + e" if c: "singular_chain p (subtopology X S) c" and X: "X closure_of U \ X interior_of T" "U \ topspace X" and S: "T \ S" "S \ topspace X" for p X c S and T U :: "'a set" proof - obtain n where n: "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ {S \ X interior_of T, S - X closure_of U}. f ` standard_simplex p \ V" apply (rule sufficient_iterated_singular_subdivision_exists [of "{S \ X interior_of T, S - X closure_of U}"]) using X S c by (auto simp: topspace_subtopology openin_subtopology_Int2 openin_subtopology_diff_closed) let ?c' = "\n. (singular_subdivision p ^^ n) c" have "singular_chain p (subtopology X S) (?c' m)" for m by (induction m) (auto simp: singular_chain_singular_subdivision c) then have scp: "singular_chain p (subtopology X S) (?c' n)" . have SS: "Poly_Mapping.keys (?c' n) \ singular_simplex_set p (subtopology X (S - U)) \ singular_simplex_set p (subtopology X T)" proof (clarsimp) fix f assume f: "f \ Poly_Mapping.keys ((singular_subdivision p ^^ n) c)" and non: "\ singular_simplex p (subtopology X T) f" show "singular_simplex p (subtopology X (S - U)) f" using n [OF order_refl f] scp f non closure_of_subset [OF \U \ topspace X\] interior_of_subset [of X T] by (fastforce simp: image_subset_iff singular_simplex_subtopology singular_chain_def) qed show ?thesis unfolding singular_chain_def using frag_split [OF SS] by metis qed have "(subtopology X (topspace X \ S)) = (subtopology X S)" by (metis subtopology_subtopology subtopology_topspace) with assms have c: "singular_chain p (subtopology X (topspace X \ S)) c" by simp have Xsub: "X closure_of (topspace X \ U) \ X interior_of (topspace X \ T)" using assms closure_of_restrict interior_of_restrict by fastforce obtain n d e where d: "singular_chain p (subtopology X (topspace X \ S - topspace X \ U)) d" and e: "singular_chain p (subtopology X (topspace X \ T)) e" and de: "(singular_subdivision p ^^ n) c = d + e" using *[OF c Xsub, simplified] assms by force show thesis proof show "singular_chain p (subtopology X (S - U)) d" by (metis d Diff_Int_distrib inf.cobounded2 singular_chain_mono) show "singular_chain p (subtopology X T) e" by (metis e inf.cobounded2 singular_chain_mono) show "(singular_subdivision p ^^ n) c = d + e" by (rule de) qed qed lemma excised_relcycle_exists: fixes S :: "'a set" assumes X: "X closure_of U \ X interior_of T" and "T \ S" and c: "singular_relcycle p (subtopology X S) T c" obtains c' where "singular_relcycle p (subtopology X (S - U)) (T - U) c'" "homologous_rel p (subtopology X S) T c c'" proof - have [simp]: "(S - U) \ (T - U) = T - U" "S \ T = T" using \T \ S\ by auto have scc: "singular_chain p (subtopology X S) c" and scp1: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p c)" using c by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology) obtain n d e where d: "singular_chain p (subtopology X (S - U)) d" and e: "singular_chain p (subtopology X T) e" and de: "(singular_subdivision p ^^ n) c = d + e" using excised_chain_exists [OF X \T \ S\ scc] . have scSUd: "singular_chain (p - Suc 0) (subtopology X (S - U)) (chain_boundary p d)" by (simp add: singular_chain_boundary d) have sccn: "singular_chain p (subtopology X S) ((singular_subdivision p ^^ n) c)" for n by (induction n) (auto simp: singular_chain_singular_subdivision scc) have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c))" proof (induction n) case (Suc n) then show ?case by (simp add: singular_chain_singular_subdivision chain_boundary_singular_subdivision [OF sccn]) qed (auto simp: scp1) then have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c - e))" by (simp add: chain_boundary_diff singular_chain_diff singular_chain_boundary e) with de have scTd: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p d)" by simp show thesis proof have "singular_chain (p - Suc 0) X (chain_boundary p d)" using scTd singular_chain_subtopology by blast with scSUd scTd have "singular_chain (p - Suc 0) (subtopology X (T - U)) (chain_boundary p d)" by (fastforce simp add: singular_chain_subtopology) then show "singular_relcycle p (subtopology X (S - U)) (T - U) d" by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology d) have "homologous_rel p (subtopology X S) T (c-0) ((singular_subdivision p ^^ n) c - e)" proof (rule homologous_rel_diff) show "homologous_rel p (subtopology X S) T c ((singular_subdivision p ^^ n) c)" proof (induction n) case (Suc n) then show ?case apply simp apply (rule homologous_rel_trans) using c homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision homologous_rel_sym by blast qed auto show "homologous_rel p (subtopology X S) T 0 e" unfolding homologous_rel_def using e by (intro singular_relboundary_diff singular_chain_imp_relboundary; simp add: subtopology_subtopology) qed with de show "homologous_rel p (subtopology X S) T c d" by simp qed qed subsection\Homotopy invariance\ theorem homotopic_imp_homologous_rel_chain_maps: assumes hom: "homotopic_with (\h. h ` T \ V) S U f g" and c: "singular_relcycle p S T c" shows "homologous_rel p U V (chain_map p f c) (chain_map p g c)" proof - note sum.atMost_Suc [simp del] have contf: "continuous_map S U f" and contg: "continuous_map S U g" using homotopic_with_imp_continuous_maps [OF hom] by metis+ obtain h where conth: "continuous_map (prod_topology (top_of_set {0..1::real}) S) U h" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = g x" and hV: "\t x. \0 \ t; t \ 1; x \ T\ \ h(t,x) \ V" using hom by (fastforce simp: homotopic_with_def) define vv where "vv \ \j i. if i = Suc j then 1 else (0::real)" define ww where "ww \ \j i. if i=0 \ i = Suc j then 1 else (0::real)" define simp where "simp \ \q i. oriented_simplex (Suc q) (\j. if j \ i then vv j else ww(j -1))" define pr where "pr \ \q c. \i\q. frag_cmul ((-1) ^ i) (frag_of (simplex_map (Suc q) (\z. h(z 0, c(z \ Suc))) (simp q i)))" have ss_ss: "simplicial_simplex (Suc q) ({x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}) (simp q i)" if "i \ q" for q i proof - have "(\j\Suc q. (if j \ i then vv j 0 else ww (j -1) 0) * x j) \ {0..1}" if "x \ standard_simplex (Suc q)" for x proof - have "(\j\Suc q. if j \ i then 0 else x j) \ sum x {..Suc q}" using that unfolding standard_simplex_def by (force intro!: sum_mono) with \i \ q\ that show ?thesis by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\u. u * _"] sum_nonneg cong: if_cong) qed moreover have "(\k. \j\Suc q. (if j \ i then vv j k else ww (j -1) k) * x j) \ Suc \ standard_simplex q" if "x \ standard_simplex (Suc q)" for x proof - have card: "({..q} \ {k. Suc k = j}) = {j-1}" if "0 < j" "j \ Suc q" for j using that by auto have eq: "(\j\Suc q. \k\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0) = (\j\Suc q. x j)" by (rule sum.cong [OF refl]) (use \i \ q\ in \simp add: sum.If_cases card\) have "(\j\Suc q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0) \ sum x {..Suc q}" for k using that unfolding standard_simplex_def by (force intro!: sum_mono) then show ?thesis using \i \ q\ that by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\u. u * _"] sum_nonneg sum.swap [where A = "atMost q"] eq cong: if_cong) qed ultimately show ?thesis by (simp add: that simplicial_simplex_oriented_simplex simp_def image_subset_iff if_distribR) qed obtain prism where prism: "\q. prism q 0 = 0" "\q c. singular_chain q S c \ singular_chain (Suc q) U (prism q c)" "\q c. singular_chain q (subtopology S T) c \ singular_chain (Suc q) (subtopology U V) (prism q c)" "\q c. singular_chain q S c \ chain_boundary (Suc q) (prism q c) = chain_map q g c - chain_map q f c - prism (q -1) (chain_boundary q c)" proof show "(frag_extend \ pr) q 0 = 0" for q by (simp add: pr_def) next show "singular_chain (Suc q) U ((frag_extend \ pr) q c)" if "singular_chain q S c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) show ?case proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum) fix i :: "nat" assume "i \ {..q}" define X where "X = subtopology (powertop_real UNIV) {x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}" show "singular_chain (Suc q) U (frag_of (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)))" unfolding singular_chain_of proof (rule singular_simplex_simplex_map) show "singular_simplex (Suc q) X (simp q i)" unfolding X_def using \i \ {..q}\ simplicial_imp_singular_simplex ss_ss by blast have 0: "continuous_map X (top_of_set {0..1}) (\x. x 0)" unfolding continuous_map_in_subtopology topspace_subtopology X_def by (auto intro: continuous_map_product_projection continuous_map_from_subtopology) have 1: "continuous_map X S (m \ (\x j. x (Suc j)))" proof (rule continuous_map_compose) have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\x j. x (Suc j))" by (auto intro: continuous_map_product_projection) then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\x j. x (Suc j))" unfolding X_def o_def by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection) qed (use one in \simp add: singular_simplex_def\) show "continuous_map X U (\z. h (z 0, m (z \ Suc)))" apply (rule continuous_map_compose [unfolded o_def, OF _ conth]) using 0 1 by (simp add: continuous_map_pairwise o_def) qed qed next case (diff a b) then show ?case apply (simp add: frag_extend_diff keys_diff) using singular_chain_def singular_chain_diff by blast qed auto next show "singular_chain (Suc q) (subtopology U V) ((frag_extend \ pr) q c)" if "singular_chain q (subtopology S T) c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) show ?case proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum) fix i :: "nat" assume "i \ {..q}" define X where "X = subtopology (powertop_real UNIV) {x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}" show "singular_chain (Suc q) (subtopology U V) (frag_of (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)))" unfolding singular_chain_of proof (rule singular_simplex_simplex_map) show "singular_simplex (Suc q) X (simp q i)" unfolding X_def using \i \ {..q}\ simplicial_imp_singular_simplex ss_ss by blast have 0: "continuous_map X (top_of_set {0..1}) (\x. x 0)" unfolding continuous_map_in_subtopology topspace_subtopology X_def by (auto intro: continuous_map_product_projection continuous_map_from_subtopology) have 1: "continuous_map X (subtopology S T) (m \ (\x j. x (Suc j)))" proof (rule continuous_map_compose) have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\x j. x (Suc j))" by (auto intro: continuous_map_product_projection) then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\x j. x (Suc j))" unfolding X_def o_def by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection) show "continuous_map (subtopology (powertop_real UNIV) (standard_simplex q)) (subtopology S T) m" using one continuous_map_into_fulltopology by (auto simp: singular_simplex_def) qed have "continuous_map X (subtopology U V) (h \ (\z. (z 0, m (z \ Suc))))" proof (rule continuous_map_compose) show "continuous_map X (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (\z. (z 0, m (z \ Suc)))" using 0 1 by (simp add: continuous_map_pairwise o_def) have "continuous_map (subtopology (prod_topology euclideanreal S) ({0..1} \ T)) U h" by (metis conth continuous_map_from_subtopology subtopology_Times subtopology_topspace) with hV show "continuous_map (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (subtopology U V) h" by (force simp: topspace_subtopology continuous_map_in_subtopology subtopology_restrict subtopology_Times) qed then show "continuous_map X (subtopology U V) (\z. h (z 0, m (z \ Suc)))" by (simp add: o_def) qed qed next case (diff a b) then show ?case by (metis comp_apply frag_extend_diff singular_chain_diff) qed auto next show "chain_boundary (Suc q) ((frag_extend \ pr) q c) = chain_map q g c - chain_map q f c - (frag_extend \ pr) (q -1) (chain_boundary q c)" if "singular_chain q S c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) have eq2: "Sigma S T = (\i. (i,i)) ` {i \ S. i \ T i} \ (Sigma S (\i. T i - {i}))" for S :: "nat set" and T by force have 1: "(\(i,j)\(\i. (i, i)) ` {i. i \ q \ i \ Suc q}. frag_cmul (((-1) ^ i) * (-1) ^ j) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) + (\(i,j)\(\i. (i, i)) ` {i. i \ q}. frag_cmul (- ((-1) ^ i * (-1) ^ j)) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = frag_of (simplex_map q g m) - frag_of (simplex_map q f m)" proof - have "restrict ((\z. h (z 0, m (z \ Suc))) \ (simp q 0 \ simplical_face 0)) (standard_simplex q) = restrict (g \ m) (standard_simplex q)" proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" have "(\j\Suc q. if j = 0 then 0 else x (j - Suc 0)) = (\j\q. x j)" by (simp add: sum.atMost_Suc_shift) with x have "simp q 0 (simplical_face 0 x) 0 = 1" apply (simp add: oriented_simplex_def simp_def simplical_face_in_standard_simplex) apply (simp add: simplical_face_def if_distrib ww_def standard_simplex_def cong: if_cong) done moreover have "(\n. if n \ q then x n else 0) = x" using standard_simplex_def x by auto then have "(\n. simp q 0 (simplical_face 0 x) (Suc n)) = x" unfolding oriented_simplex_def simp_def ww_def using x apply (simp add: simplical_face_in_standard_simplex) apply (simp add: simplical_face_def if_distrib) apply (simp add: if_distribR if_distrib cong: if_cong) done ultimately show "((\z. h (z 0, m (z \ Suc))) \ (simp q 0 \ simplical_face 0)) x = (g \ m) x" by (simp add: o_def h1) qed then have a: "frag_of (singular_face (Suc q) 0 (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q 0))) = frag_of (simplex_map q g m)" by (simp add: singular_face_simplex_map) (simp add: simplex_map_def) have "restrict ((\z. h (z 0, m (z \ Suc))) \ (simp q q \ simplical_face (Suc q))) (standard_simplex q) = restrict (f \ m) (standard_simplex q)" proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" then have "simp q q (simplical_face (Suc q) x) 0 = 0" unfolding oriented_simplex_def simp_def by (simp add: simplical_face_in_standard_simplex sum.atMost_Suc) (simp add: simplical_face_def vv_def) moreover have "(\n. simp q q (simplical_face (Suc q) x) (Suc n)) = x" unfolding oriented_simplex_def simp_def vv_def using x apply (simp add: simplical_face_in_standard_simplex) apply (force simp: standard_simplex_def simplical_face_def if_distribR if_distrib [of "\x. x * _"] sum.atMost_Suc cong: if_cong) done ultimately show "((\z. h (z 0, m (z \ Suc))) \ (simp q q \ simplical_face (Suc q))) x = (f \ m) x" by (simp add: o_def h0) qed then have b: "frag_of (singular_face (Suc q) (Suc q) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q q))) = frag_of (simplex_map q f m)" by (simp add: singular_face_simplex_map) (simp add: simplex_map_def) have sfeq: "simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q (Suc i) \ simplical_face (Suc i)) = simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q i \ simplical_face (Suc i))" if "i < q" for i unfolding simplex_map_def proof (rule restrict_ext) fix x assume "x \ standard_simplex q" then have "(simp q (Suc i) \ simplical_face (Suc i)) x = (simp q i \ simplical_face (Suc i)) x" unfolding oriented_simplex_def simp_def simplical_face_def by (force intro: sum.cong) then show "((\z. h (z 0, m (z \ Suc))) \ (simp q (Suc i) \ simplical_face (Suc i))) x = ((\z. h (z 0, m (z \ Suc))) \ (simp q i \ simplical_face (Suc i))) x" by simp qed have eqq: "{i. i \ q \ i \ Suc q} = {..q}" by force have qeq: "{..q} = insert 0 ((\i. Suc i) ` {i. i < q})" "{i. i \ q} = insert q {i. i < q}" using le_imp_less_Suc less_Suc_eq_0_disj by auto show ?thesis using a b apply (simp add: sum.reindex inj_on_def eqq) apply (simp add: qeq sum.insert_if sum.reindex sum_negf singular_face_simplex_map sfeq) done qed have 2: "(\(i,j)\(SIGMA i:{..q}. {0..min (Suc q) i} - {i}). frag_cmul ((-1) ^ i * (-1) ^ j) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) + (\(i,j)\(SIGMA i:{..q}. {i..q} - {i}). frag_cmul (- ((-1) ^ i * (-1) ^ j)) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = - frag_extend (pr (q - Suc 0)) (chain_boundary q (frag_of m))" proof (cases "q=0") case True then show ?thesis by (simp add: chain_boundary_def flip: sum.Sigma) next case False have eq: "{..q - Suc 0} \ {..q} = Sigma {..q - Suc 0} (\i. {0..min q i}) \ Sigma {..q} (\i. {i<..q})" by force have I: "(\(i,j)\(SIGMA i:{..q}. {0..min (Suc q) i} - {i}). frag_cmul ((-1) ^ (i + j)) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = (\(i,j)\(SIGMA i:{..q - Suc 0}. {0..min q i}). frag_cmul (- ((-1) ^ (j + i))) (frag_of (simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i))))" proof - have seq: "simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) (i - Suc 0)) = simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q i \ simplical_face j)" if ij: "i \ q" "j \ i" "j \ i" for i j unfolding simplex_map_def proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" have "i > 0" using that by force then have iq: "i - Suc 0 \ q - Suc 0" using \i \ q\ False by simp have q0_eq: "{..Suc q} = insert 0 (Suc ` {..q})" by (auto simp: image_def gr0_conv_Suc) have \: "simp (q - Suc 0) (i - Suc 0) x 0 = simp q i (simplical_face j x) 0" using False x ij unfolding oriented_simplex_def simp_def vv_def ww_def apply (simp add: simplical_face_in_standard_simplex) apply (force simp: simplical_face_def q0_eq sum.reindex intro!: sum.cong) done have \: "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \ Suc) = simp q i (simplical_face j x) \ Suc" proof fix k show "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \ Suc) k = (simp q i (simplical_face j x) \ Suc) k" using False x ij unfolding oriented_simplex_def simp_def o_def vv_def ww_def apply (simp add: simplical_face_in_standard_simplex if_distribR) apply (simp add: simplical_face_def if_distrib [of "\u. u * _"] cong: if_cong) apply (intro impI conjI) apply (force simp: sum.atMost_Suc intro: sum.cong) apply (force simp: q0_eq sum.reindex intro!: sum.cong) done qed have "simp (q - Suc 0) (i - Suc 0) x \ Suc \ standard_simplex (q - Suc 0)" using ss_ss [OF iq] \i \ q\ False \i > 0\ apply (simp add: simplicial_simplex image_subset_iff) using \x \ standard_simplex q\ by blast then show "((\z. h (z 0, singular_face q j m (z \ Suc))) \ simp (q - Suc 0) (i - Suc 0)) x = ((\z. h (z 0, m (z \ Suc))) \ (simp q i \ simplical_face j)) x" by (simp add: singular_face_def \ \) qed have [simp]: "(-1::int) ^ (i + j - Suc 0) = - ((-1) ^ (i + j))" if "i \ j" for i j::nat proof - have "i + j > 0" using that by blast then show ?thesis by (metis (no_types, hide_lams) One_nat_def Suc_diff_1 add.inverse_inverse mult.left_neutral mult_minus_left power_Suc) qed show ?thesis apply (rule sum.eq_general_inverses [where h = "\(a,b). (a-1,b)" and k = "\(a,b). (Suc a,b)"]) using False apply (auto simp: singular_face_simplex_map seq add.commute) done qed have *: "singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)) = simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i)" if ij: "i < j" "j \ q" for i j proof - have iq: "i \ q - Suc 0" using that by auto have sf_eqh: "singular_face (Suc q) (Suc j) (\x. if x \ standard_simplex (Suc q) then ((\z. h (z 0, m (z \ Suc))) \ simp q i) x else undefined) x = h (simp (q - Suc 0) i x 0, singular_face q j m (\xa. simp (q - Suc 0) i x (Suc xa)))" if x: "x \ standard_simplex q" for x proof - let ?f = "\k. \j\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0" have fm: "simplical_face (Suc j) x \ standard_simplex (Suc q)" using ss_ss [OF iq] that ij by (simp add: simplical_face_in_standard_simplex) have ss: "?f \ standard_simplex (q - Suc 0)" unfolding standard_simplex_def proof (intro CollectI conjI impI allI) fix k show "0 \ ?f k" using that by (simp add: sum_nonneg standard_simplex_def) show "?f k \ 1" using x sum_le_included [of "{..q}" "{..q}" x "id"] by (simp add: standard_simplex_def) assume k: "q - Suc 0 < k" show "?f k = 0" by (rule sum.neutral) (use that x iq k standard_simplex_def in auto) next have "(\k\q - Suc 0. ?f k) = (\(k,j) \ ({..q - Suc 0} \ {..q}) \ {(k,j). if j \ i then k = j else Suc k = j}. x j)" apply (simp add: sum.Sigma) by (rule sum.mono_neutral_cong) (auto simp: split: if_split_asm) also have "\ = sum x {..q}" apply (rule sum.eq_general_inverses [where h = "\(k,j). if j\i \ k=j \ j>i \ Suc k = j then j else Suc q" and k = "\j. if j \ i then (j,j) else (j - Suc 0, j)"]) using ij by auto also have "\ = 1" using x by (simp add: standard_simplex_def) finally show "(\k\q - Suc 0. ?f k) = 1" by (simp add: standard_simplex_def) qed let ?g = "\k. if k \ i then 0 else if k < Suc j then x k else if k = Suc j then 0 else x (k - Suc 0)" have eq: "{..Suc q} = {..j} \ {Suc j} \ Suc`{j<..q}" "{..q} = {..j} \ {j<..q}" using ij image_iff less_Suc_eq_0_disj less_Suc_eq_le by (force simp: image_iff)+ then have "(\k\Suc q. ?g k) = (\k\{..j} \ {Suc j} \ Suc`{j<..q}. ?g k)" by simp also have "\ = (\k\{..j} \ Suc`{j<..q}. ?g k)" by (rule sum.mono_neutral_right) auto also have "\ = (\k\{..j}. ?g k) + (\k\Suc`{j<..q}. ?g k)" by (rule sum.union_disjoint) auto also have "\ = (\k\{..j}. ?g k) + (\k\{j<..q}. ?g (Suc k))" by (auto simp: sum.reindex) also have "\ = (\k\{..j}. if k \ i then 0 else x k) + (\k\{j<..q}. if k \ i then 0 else x k)" by (intro sum.cong arg_cong2 [of concl: "(+)"]) (use ij in auto) also have "\ = (\k\q. if k \ i then 0 else x k)" unfolding eq by (subst sum.union_disjoint) auto finally have "(\k\Suc q. ?g k) = (\k\q. if k \ i then 0 else x k)" . then have QQ: "(\l\Suc q. if l \ i then 0 else simplical_face (Suc j) x l) = (\j\q. if j \ i then 0 else x j)" by (simp add: simplical_face_def cong: if_cong) have WW: "(\k. \l\Suc q. if l \ i then if k = l then simplical_face (Suc j) x l else 0 else if Suc k = l then simplical_face (Suc j) x l else 0) = simplical_face j (\k. \j\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0)" proof - have *: "(\l\q. if l \ i then 0 else if Suc k = l then x (l - Suc 0) else 0) = (\l\q. if l \ i then if k - Suc 0 = l then x l else 0 else if k = l then x l else 0)" (is "?lhs = ?rhs") if "k \ q" "k > j" for k proof (cases "k \ q") case True have "?lhs = sum (\l. x (l - Suc 0)) {Suc k}" "?rhs = sum x {k}" by (rule sum.mono_neutral_cong_right; use True ij that in auto)+ then show ?thesis by simp next case False have "?lhs = 0" "?rhs = 0" by (rule sum.neutral; use False ij in auto)+ then show ?thesis by simp qed show ?thesis apply (rule ext) unfolding simplical_face_def using ij apply (auto simp: sum.atMost_Suc cong: if_cong) apply (force simp flip: ivl_disj_un(2) intro: sum.neutral) apply (auto simp: *) done qed show ?thesis using False that iq unfolding oriented_simplex_def simp_def vv_def ww_def apply (simp add: if_distribR cong: if_cong) apply (simp add: simplical_face_def if_distrib [of "\u. u * _"] o_def cong: if_cong) apply (simp add: singular_face_def fm ss QQ WW) done qed show ?thesis unfolding simplex_map_def restrict_def apply (rule ext) apply (simp add: simplicial_simplex image_subset_iff o_def sf_eqh) apply (simp add: singular_face_def) done qed have sgeq: "(SIGMA i:{..q}. {i..q} - {i}) = (SIGMA i:{..q}. {i<..q})" by force have II: "(\(i,j)\(SIGMA i:{..q}. {i..q} - {i}). frag_cmul (- ((-1) ^ (i + j))) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = (\(i,j)\(SIGMA i:{..q}. {i<..q}). frag_cmul (- ((-1) ^ (j + i))) (frag_of (simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i))))" by (force simp: * sgeq add.commute intro: sum.cong) show ?thesis using False apply (simp add: chain_boundary_def frag_extend_sum frag_extend_cmul frag_cmul_sum pr_def flip: sum_negf power_add) apply (subst sum.swap [where A = "{..q}"]) apply (simp add: sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal I II) done qed have *: "\a+b = w; c+d = -z\ \ (a + c) + (b+d) = w-z" for a b w c d z :: "'c \\<^sub>0 int" by (auto simp: algebra_simps) have eq: "{..q} \ {..Suc q} = Sigma {..q} (\i. {0..min (Suc q) i}) \ Sigma {..q} (\i. {Suc i..Suc q})" by force show ?case apply (subst pr_def) apply (simp add: chain_boundary_sum chain_boundary_cmul) apply (subst chain_boundary_def) apply (simp add: frag_cmul_sum sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal sum.atLeast_Suc_atMost_Suc_shift del: sum.cl_ivl_Suc flip: comm_monoid_add_class.sum.Sigma) apply (simp add: comm_monoid_add_class.sum.Sigma eq2 [of _ "\i. {_ i.._ i}"]) apply (simp add: sum.union_disjoint disjoint_iff_not_equal * [OF 1 2]) done next case (diff a b) then show ?case by (simp add: chain_boundary_diff frag_extend_diff chain_map_diff) qed auto qed have *: "singular_chain p (subtopology U V) (prism (p - Suc 0) (chain_boundary p c))" if "singular_chain p S c" "singular_chain (p - Suc 0) (subtopology S T) (chain_boundary p c)" proof (cases "p") case 0 then show ?thesis by (simp add: chain_boundary_def prism) next case (Suc p') with prism that show ?thesis by auto qed then show ?thesis using c unfolding singular_relcycle_def homologous_rel_def singular_relboundary_def mod_subset_def apply (rule_tac x="- prism p c" in exI) by (simp add: chain_boundary_minus prism(2) prism(4) singular_chain_minus) qed end