diff --git a/src/HOL/Analysis/Abstract_Metric_Spaces.thy b/src/HOL/Analysis/Abstract_Metric_Spaces.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Abstract_Metric_Spaces.thy @@ -0,0 +1,2634 @@ +section \Abstract Metric Spaces\ + +theory Abstract_Metric_Spaces + imports Elementary_Metric_Spaces Abstract_Limits Abstract_Topological_Spaces +begin + +(*Avoid a clash with the existing metric_space locale (from the type class)*) +locale Metric_space = + fixes M :: "'a set" and d :: "'a \ 'a \ real" + assumes nonneg [simp]: "\x y. 0 \ d x y" + assumes commute: "\x y. d x y = d y x" + assumes zero [simp]: "\x y. \x \ M; y \ M\ \ d x y = 0 \ x=y" + assumes triangle: "\x y z. \x \ M; y \ M; z \ M\ \ d x z \ d x y + d y z" + +text \Link with the type class version\ +interpretation Met_TC: Metric_space UNIV dist + by (simp add: dist_commute dist_triangle Metric_space.intro) + +context Metric_space +begin + +lemma subspace: "M' \ M \ Metric_space M' d" + by (simp add: commute in_mono Metric_space.intro triangle) + +lemma abs_mdist [simp] : "\d x y\ = d x y" + by simp + +lemma mdist_pos_less: "\x \ y; x \ M; y \ M\ \ 0 < d x y" + by (metis less_eq_real_def nonneg zero) + +lemma mdist_zero [simp]: "x \ M \ d x x = 0" + by simp + +lemma mdist_pos_eq [simp]: "\x \ M; y \ M\ \ 0 < d x y \ x \ y" + using mdist_pos_less zero by fastforce + +lemma triangle': "\x \ M; y \ M; z \ M\ \ d x z \ d x y + d z y" + by (simp add: commute triangle) + +lemma triangle'': "\x \ M; y \ M; z \ M\ \ d x z \ d y x + d y z" + by (simp add: commute triangle) + +lemma mdist_reverse_triangle: "\x \ M; y \ M; z \ M\ \ \d x y - d y z\ \ d x z" + by (smt (verit) commute triangle) + +text\ Open and closed balls \ + +definition mball where "mball x r \ {y. x \ M \ y \ M \ d x y < r}" +definition mcball where "mcball x r \ {y. x \ M \ y \ M \ d x y \ r}" + +lemma in_mball [simp]: "y \ mball x r \ x \ M \ y \ M \ d x y < r" + by (simp add: local.Metric_space_axioms Metric_space.mball_def) + +lemma centre_in_mball_iff [iff]: "x \ mball x r \ x \ M \ 0 < r" + using in_mball mdist_zero by force + +lemma mball_subset_mspace: "mball x r \ M" + by auto + +lemma mball_eq_empty: "mball x r = {} \ (x \ M) \ r \ 0" + by (smt (verit, best) Collect_empty_eq centre_in_mball_iff mball_def nonneg) + +lemma mball_subset: "\d x y + a \ b; y \ M\ \ mball x a \ mball y b" + by (smt (verit) commute in_mball subsetI triangle) + +lemma disjoint_mball: "r + r' \ d x x' \ disjnt (mball x r) (mball x' r')" + by (smt (verit) commute disjnt_iff in_mball triangle) + +lemma mball_subset_concentric: "r \ s \ mball x r \ mball x s" + by auto + +lemma in_mcball [simp]: "y \ mcball x r \ x \ M \ y \ M \ d x y \ r" + by (simp add: local.Metric_space_axioms Metric_space.mcball_def) + +lemma centre_in_mcball_iff [iff]: "x \ mcball x r \ x \ M \ 0 \ r" + using mdist_zero by force + +lemma mcball_eq_empty: "mcball x r = {} \ (x \ M) \ r < 0" + by (smt (verit, best) Collect_empty_eq centre_in_mcball_iff empty_iff mcball_def nonneg) + +lemma mcball_subset_mspace: "mcball x r \ M" + by auto + +lemma mball_subset_mcball: "mball x r \ mcball x r" + by auto + +lemma mcball_subset: "\d x y + a \ b; y \ M\ \ mcball x a \ mcball y b" + by (smt (verit) in_mcball mdist_reverse_triangle subsetI) + +lemma mcball_subset_concentric: "r \ s \ mcball x r \ mcball x s" + by force + +lemma mcball_subset_mball: "\d x y + a < b; y \ M\ \ mcball x a \ mball y b" + by (smt (verit) commute in_mball in_mcball subsetI triangle) + +lemma mcball_subset_mball_concentric: "a < b \ mcball x a \ mball x b" + by force + +end + + + +subsection \Metric topology \ + +context Metric_space +begin + +definition mopen where + "mopen U \ U \ M \ (\x. x \ U \ (\r>0. mball x r \ U))" + +definition mtopology :: "'a topology" where + "mtopology \ topology mopen" + +lemma is_topology_metric_topology [iff]: "istopology mopen" +proof - + have "\S T. \mopen S; mopen T\ \ mopen (S \ T)" + by (smt (verit, del_insts) Int_iff in_mball mopen_def subset_eq) + moreover have "\\. (\K\\. mopen K) \ mopen (\\)" + using mopen_def by fastforce + ultimately show ?thesis + by (simp add: istopology_def) +qed + +lemma openin_mtopology: "openin mtopology U \ U \ M \ (\x. x \ U \ (\r>0. mball x r \ U))" + by (simp add: mopen_def mtopology_def) + +lemma topspace_mtopology [simp]: "topspace mtopology = M" + by (meson order.refl mball_subset_mspace openin_mtopology openin_subset openin_topspace subset_antisym zero_less_one) + +lemma subtopology_mspace [simp]: "subtopology mtopology M = mtopology" + by (metis subtopology_topspace topspace_mtopology) + +lemma open_in_mspace [iff]: "openin mtopology M" + by (metis openin_topspace topspace_mtopology) + +lemma closedin_mspace [iff]: "closedin mtopology M" + by (metis closedin_topspace topspace_mtopology) + +lemma openin_mball [iff]: "openin mtopology (mball x r)" +proof - + have "\y. \x \ M; d x y < r\ \ \s>0. mball y s \ mball x r" + by (metis add_diff_cancel_left' add_diff_eq commute less_add_same_cancel1 mball_subset order_refl) + then show ?thesis + by (auto simp: openin_mtopology) +qed + +lemma mcball_eq_cball [simp]: "Met_TC.mcball = cball" + by force + +lemma mball_eq_ball [simp]: "Met_TC.mball = ball" + by force + +lemma mopen_eq_open [simp]: "Met_TC.mopen = open" + by (force simp: open_contains_ball Met_TC.mopen_def) + +lemma limitin_iff_tendsto [iff]: "limitin Met_TC.mtopology \ x F = tendsto \ x F" + by (simp add: Met_TC.mtopology_def) + +lemma mtopology_is_euclideanreal [simp]: "Met_TC.mtopology = euclideanreal" + by (simp add: Met_TC.mtopology_def) + +(* +lemma metric_injective_image: + "\f m s. + f ` s \ M \ + (\x y. x \ s \ y \ s \ f x = f y \ x = y) + \ (mspace(metric(s,\(x,y). d (f x) (f y))) = s) \ + (d(metric(s,\(x,y). d (f x) (f y))) = + \(x,y). d (f x) (f y))" +oops + REWRITE_TAC[\; FORALL_IN_IMAGE; INJECTIVE_ON_ALT] THEN + REPEAT GEN_TAC THEN STRIP_TAC THEN + REWRITE_TAC[mspace; d; GSYM PAIR_EQ] THEN + REWRITE_TAC[GSYM(CONJUNCT2 metric_tybij); is_metric_space] THEN + REWRITE_TAC[GSYM mspace; GSYM d] THEN + ASM_SIMP_TAC[MDIST_POS_LE; MDIST_TRIANGLE; MDIST_0] THEN + ASM_MESON_TAC[MDIST_SYM]);; +*) + +lemma mtopology_base: + "mtopology = topology(arbitrary union_of (\U. \x \ M. \r>0. U = mball x r))" +proof - + have "\S. \x r. x \ M \ 0 < r \ S = mball x r \ openin mtopology S" + using openin_mball by blast + moreover have "\U x. \openin mtopology U; x \ U\ \ \B. (\x r. x \ M \ 0 < r \ B = mball x r) \ x \ B \ B \ U" + by (metis centre_in_mball_iff in_mono openin_mtopology) + ultimately show ?thesis + by (smt (verit) topology_base_unique) +qed + +lemma closedin_metric: + "closedin mtopology C \ C \ M \ (\x. x \ M - C \ (\r>0. disjnt C (mball x r)))" (is "?lhs = ?rhs") +proof + show "?lhs \ ?rhs" + unfolding closedin_def openin_mtopology + by (metis Diff_disjoint disjnt_def disjnt_subset2 topspace_mtopology) + show "?rhs \ ?lhs" + unfolding closedin_def openin_mtopology disjnt_def + by (metis Diff_subset Diff_triv Int_Diff Int_commute inf.absorb_iff2 mball_subset_mspace topspace_mtopology) +qed + +lemma closedin_mcball [iff]: "closedin mtopology (mcball x r)" +proof - + have "\ra>0. disjnt (mcball x r) (mball y ra)" if "x \ M" for y + by (metis disjnt_empty1 gt_ex mcball_eq_empty that) + moreover have "disjnt (mcball x r) (mball y (d x y - r))" if "y \ M" "d x y > r" for y + using that disjnt_iff in_mball in_mcball mdist_reverse_triangle by force + ultimately show ?thesis + using closedin_metric mcball_subset_mspace by fastforce +qed + +lemma mball_iff_mcball: "(\r>0. mball x r \ U) = (\r>0. mcball x r \ U)" + by (meson dense mball_subset_mcball mcball_subset_mball_concentric order_trans) + +lemma openin_mtopology_mcball: + "openin mtopology U \ U \ M \ (\x. x \ U \ (\r. 0 < r \ mcball x r \ U))" + using mball_iff_mcball openin_mtopology by presburger + +lemma metric_derived_set_of: + "mtopology derived_set_of S = {x \ M. \r>0. \y\S. y\x \ y \ mball x r}" (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + unfolding openin_mtopology derived_set_of_def + by clarsimp (metis in_mball openin_mball openin_mtopology zero) + show "?rhs \ ?lhs" + unfolding openin_mtopology derived_set_of_def + by clarify (metis subsetD topspace_mtopology) +qed + +lemma metric_closure_of: + "mtopology closure_of S = {x \ M. \r>0. \y \ S. y \ mball x r}" +proof - + have "\x r. \0 < r; x \ mtopology closure_of S\ \ \y\S. y \ mball x r" + by (metis centre_in_mball_iff in_closure_of openin_mball topspace_mtopology) + moreover have "\x T. \x \ M; \r>0. \y\S. y \ mball x r\ \ x \ mtopology closure_of S" + by (smt (verit) in_closure_of in_mball openin_mtopology subsetD topspace_mtopology) + ultimately show ?thesis + by (auto simp: in_closure_of) +qed + +lemma metric_closure_of_alt: + "mtopology closure_of S = {x \ M. \r>0. \y \ S. y \ mcball x r}" +proof - + have "\x r. \\r>0. x \ M \ (\y\S. y \ mcball x r); 0 < r\ \ \y\S. y \ M \ d x y < r" + by (meson dense in_mcball le_less_trans) + then show ?thesis + by (fastforce simp: metric_closure_of in_closure_of) +qed + +lemma metric_interior_of: + "mtopology interior_of S = {x \ M. \\>0. mball x \ \ S}" (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + using interior_of_maximal_eq openin_mtopology by fastforce + show "?rhs \ ?lhs" + using interior_of_def openin_mball by fastforce +qed + +lemma metric_interior_of_alt: + "mtopology interior_of S = {x \ M. \\>0. mcball x \ \ S}" + by (fastforce simp: mball_iff_mcball metric_interior_of) + +lemma in_interior_of_mball: + "x \ mtopology interior_of S \ x \ M \ (\\>0. mball x \ \ S)" + using metric_interior_of by force + +lemma in_interior_of_mcball: + "x \ mtopology interior_of S \ x \ M \ (\\>0. mcball x \ \ S)" + using metric_interior_of_alt by force + +lemma Hausdorff_space_mtopology: "Hausdorff_space mtopology" + unfolding Hausdorff_space_def +proof clarify + fix x y + assume x: "x \ topspace mtopology" and y: "y \ topspace mtopology" and "x \ y" + then have gt0: "d x y / 2 > 0" + by auto + have "disjnt (mball x (d x y / 2)) (mball y (d x y / 2))" + by (simp add: disjoint_mball) + then show "\U V. openin mtopology U \ openin mtopology V \ x \ U \ y \ V \ disjnt U V" + by (metis centre_in_mball_iff gt0 openin_mball topspace_mtopology x y) +qed + + + +subsection\Bounded sets\ + +definition mbounded where "mbounded S \ (\x B. S \ mcball x B)" + +lemma mbounded_pos: "mbounded S \ (\x B. 0 < B \ S \ mcball x B)" +proof - + have "\x' r'. 0 < r' \ S \ mcball x' r'" if "S \ mcball x r" for x r + by (metis gt_ex less_eq_real_def linorder_not_le mcball_subset_concentric order_trans that) + then show ?thesis + by (auto simp: mbounded_def) +qed + +lemma mbounded_alt: + "mbounded S \ S \ M \ (\B. \x \ S. \y \ S. d x y \ B)" +proof - + have "\x B. S \ mcball x B \ \x\S. \y\S. d x y \ 2 * B" + by (smt (verit, best) commute in_mcball subsetD triangle) + then show ?thesis + apply (auto simp: mbounded_def subset_iff) + apply blast+ + done +qed + + +lemma mbounded_alt_pos: + "mbounded S \ S \ M \ (\B>0. \x \ S. \y \ S. d x y \ B)" + by (smt (verit, del_insts) gt_ex mbounded_alt) + +lemma mbounded_subset: "\mbounded T; S \ T\ \ mbounded S" + by (meson mbounded_def order_trans) + +lemma mbounded_subset_mspace: "mbounded S \ S \ M" + by (simp add: mbounded_alt) + +lemma mbounded: + "mbounded S \ S = {} \ (\x \ S. x \ M) \ (\y B. y \ M \ (\x \ S. d y x \ B))" + by (meson all_not_in_conv in_mcball mbounded_def subset_iff) + +lemma mbounded_empty [iff]: "mbounded {}" + by (simp add: mbounded) + +lemma mbounded_mcball: "mbounded (mcball x r)" + using mbounded_def by auto + +lemma mbounded_mball [iff]: "mbounded (mball x r)" + by (meson mball_subset_mcball mbounded_def) + +lemma mbounded_insert: "mbounded (insert a S) \ a \ M \ mbounded S" +proof - + have "\y B. \y \ M; \x\S. d y x \ B\ + \ \y. y \ M \ (\B \ d y a. \x\S. d y x \ B)" + by (metis order.trans nle_le) + then show ?thesis + by (auto simp: mbounded) +qed + +lemma mbounded_Int: "mbounded S \ mbounded (S \ T)" + by (meson inf_le1 mbounded_subset) + +lemma mbounded_Un: "mbounded (S \ T) \ mbounded S \ mbounded T" (is "?lhs=?rhs") +proof + assume R: ?rhs + show ?lhs + proof (cases "S={} \ T={}") + case True then show ?thesis + using R by auto + next + case False + obtain x y B C where "S \ mcball x B" "T \ mcball y C" "B > 0" "C > 0" "x \ M" "y \ M" + using R mbounded_pos + by (metis False mcball_eq_empty subset_empty) + then have "S \ T \ mcball x (B + C + d x y)" + by (smt (verit) commute dual_order.trans le_supI mcball_subset mdist_pos_eq) + then show ?thesis + using mbounded_def by blast + qed +next + show "?lhs \ ?rhs" + using mbounded_def by auto +qed + +lemma mbounded_Union: + "\finite \; \X. X \ \ \ mbounded X\ \ mbounded (\\)" + by (induction \ rule: finite_induct) (auto simp: mbounded_Un) + +lemma mbounded_closure_of: + "mbounded S \ mbounded (mtopology closure_of S)" + by (meson closedin_mcball closure_of_minimal mbounded_def) + +lemma mbounded_closure_of_eq: + "S \ M \ (mbounded (mtopology closure_of S) \ mbounded S)" + by (metis closure_of_subset mbounded_closure_of mbounded_subset topspace_mtopology) + + +lemma maxdist_thm: + assumes "mbounded S" + and "x \ S" + and "y \ S" + shows "d x y = (SUP z\S. \d x z - d z y\)" +proof - + have "\d x z - d z y\ \ d x y" if "z \ S" for z + by (metis all_not_in_conv assms mbounded mdist_reverse_triangle that) + moreover have "d x y \ r" + if "\z. z \ S \ \d x z - d z y\ \ r" for r :: real + using that assms mbounded_subset_mspace mdist_zero by fastforce + ultimately show ?thesis + by (intro cSup_eq [symmetric]) auto +qed + + +lemma metric_eq_thm: "\S \ M; x \ S; y \ S\ \ (x = y) = (\z\S. d x z = d y z)" + by (metis commute subset_iff zero) + +lemma compactin_imp_mbounded: + assumes "compactin mtopology S" + shows "mbounded S" +proof - + have "S \ M" + and com: "\\. \\U\\. openin mtopology U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" + using assms by (auto simp: compactin_def mbounded_def) + show ?thesis + proof (cases "S = {}") + case False + with \S \ M\ obtain a where "a \ S" "a \ M" + by blast + with \S \ M\ gt_ex have "S \ \(range (mball a))" + by force + moreover have "\U \ range (mball a). openin mtopology U" + by (simp add: openin_mball) + ultimately obtain \ where "finite \" "\ \ range (mball a)" "S \ \\" + by (meson com) + then show ?thesis + using mbounded_Union mbounded_subset by fastforce + qed auto +qed + +end + + +subsection\Subspace of a metric space\ + +locale submetric = Metric_space + + fixes A + assumes subset: "A \ M" + +sublocale submetric \ sub: Metric_space A d + by (simp add: subset subspace) + +context submetric +begin + +lemma mball_submetric_eq: "sub.mball a r = (if a \ A then A \ mball a r else {})" +and mcball_submetric_eq: "sub.mcball a r = (if a \ A then A \ mcball a r else {})" + using subset by force+ + +lemma mtopology_submetric: "sub.mtopology = subtopology mtopology A" + unfolding topology_eq +proof (intro allI iffI) + fix S + assume "openin sub.mtopology S" + then have "\T. openin (subtopology mtopology A) T \ x \ T \ T \ S" if "x \ S" for x + by (metis mball_submetric_eq openin_mball openin_subtopology_Int2 sub.centre_in_mball_iff sub.openin_mtopology subsetD that) + then show "openin (subtopology mtopology A) S" + by (meson openin_subopen) +next + fix S + assume "openin (subtopology mtopology A) S" + then obtain T where "openin mtopology T" "S = T \ A" + by (meson openin_subtopology) + then have "mopen T" + by (simp add: mopen_def openin_mtopology) + then have "sub.mopen (T \ A)" + unfolding sub.mopen_def mopen_def + by (metis inf.coboundedI2 mball_submetric_eq Int_iff \S = T \ A\ inf.bounded_iff subsetI) + then show "openin sub.mtopology S" + using \S = T \ A\ sub.mopen_def sub.openin_mtopology by force +qed + +lemma mbounded_submetric: "sub.mbounded T \ mbounded T \ T \ A" + by (meson mbounded_alt sub.mbounded_alt subset subset_trans) + +end + +lemma (in Metric_space) submetric_empty [iff]: "submetric M d {}" + by (simp add: Metric_space_axioms submetric.intro submetric_axioms_def) + + +subsection\The discrete metric\ + +locale discrete_metric = + fixes M :: "'a set" + +definition (in discrete_metric) dd :: "'a \ 'a \ real" + where "dd \ \x y::'a. if x=y then 0 else 1" + +lemma metric_M_dd: "Metric_space M discrete_metric.dd" + by (simp add: discrete_metric.dd_def Metric_space.intro) + +sublocale discrete_metric \ disc: Metric_space M dd + by (simp add: metric_M_dd) + + +lemma (in discrete_metric) mopen_singleton: + assumes "x \ M" shows "disc.mopen {x}" +proof - + have "disc.mball x (1/2) \ {x}" + by (smt (verit) dd_def disc.in_mball less_divide_eq_1_pos singleton_iff subsetI) + with assms show ?thesis + using disc.mopen_def half_gt_zero_iff zero_less_one by blast +qed + +lemma (in discrete_metric) mtopology_discrete_metric: + "disc.mtopology = discrete_topology M" +proof - + have "\x. x \ M \ openin disc.mtopology {x}" + by (simp add: disc.mtopology_def mopen_singleton) + then show ?thesis + by (metis disc.topspace_mtopology discrete_topology_unique) +qed + +lemma (in discrete_metric) discrete_ultrametric: + "dd x z \ max (dd x y) (dd y z)" + by (simp add: dd_def) + + +lemma (in discrete_metric) dd_le1: "dd x y \ 1" + by (simp add: dd_def) + +lemma (in discrete_metric) mbounded_discrete_metric: "disc.mbounded S \ S \ M" + by (meson dd_le1 disc.mbounded_alt) + + + +subsection\Metrizable spaces\ + +definition metrizable_space where + "metrizable_space X \ \M d. Metric_space M d \ X = Metric_space.mtopology M d" + +lemma (in Metric_space) metrizable_space_mtopology: "metrizable_space mtopology" + using local.Metric_space_axioms metrizable_space_def by blast + +lemma openin_mtopology_eq_open [simp]: "openin Met_TC.mtopology = open" + by (simp add: Met_TC.mtopology_def) + +lemma closedin_mtopology_eq_closed [simp]: "closedin Met_TC.mtopology = closed" +proof - + have "(euclidean::'a topology) = Met_TC.mtopology" + by (simp add: Met_TC.mtopology_def) + then show ?thesis + using closed_closedin by fastforce +qed + +lemma compactin_mtopology_eq_compact [simp]: "compactin Met_TC.mtopology = compact" + by (simp add: compactin_def compact_eq_Heine_Borel fun_eq_iff) meson + +lemma metrizable_space_discrete_topology: + "metrizable_space(discrete_topology U)" + by (metis discrete_metric.mtopology_discrete_metric metric_M_dd metrizable_space_def) + +lemma metrizable_space_subtopology: + assumes "metrizable_space X" + shows "metrizable_space(subtopology X S)" +proof - + obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d" + using assms metrizable_space_def by blast + then interpret submetric M d "M \ S" + by (simp add: submetric.intro submetric_axioms_def) + show ?thesis + unfolding metrizable_space_def + by (metis X mtopology_submetric sub.Metric_space_axioms subtopology_restrict topspace_mtopology) +qed + +lemma homeomorphic_metrizable_space_aux: + assumes "X homeomorphic_space Y" "metrizable_space X" + shows "metrizable_space Y" +proof - + obtain M d where "Metric_space M d" and X: "X = Metric_space.mtopology M d" + using assms by (auto simp: metrizable_space_def) + then interpret m: Metric_space M d + by simp + obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" + and fg: "(\x \ M. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" + using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce + define d' where "d' x y \ d (g x) (g y)" for x y + interpret m': Metric_space "topspace Y" "d'" + unfolding d'_def + proof + show "(d (g x) (g y) = 0) = (x = y)" if "x \ topspace Y" "y \ topspace Y" for x y + by (metis fg X hmg homeomorphic_imp_surjective_map imageI m.topspace_mtopology m.zero that) + show "d (g x) (g z) \ d (g x) (g y) + d (g y) (g z)" + if "x \ topspace Y" and "y \ topspace Y" and "z \ topspace Y" for x y z + by (metis X that hmg homeomorphic_eq_everything_map imageI m.topspace_mtopology m.triangle) + qed (auto simp: m.nonneg m.commute) + have "Y = Metric_space.mtopology (topspace Y) d'" + unfolding topology_eq + proof (intro allI) + fix S + have "openin m'.mtopology S" if S: "S \ topspace Y" and "openin X (g ` S)" + unfolding m'.openin_mtopology + proof (intro conjI that strip) + fix y + assume "y \ S" + then obtain r where "r>0" and r: "m.mball (g y) r \ g ` S" + using X \openin X (g ` S)\ m.openin_mtopology using \y \ S\ by auto + then have "g ` m'.mball y r \ m.mball (g y) r" + using X d'_def hmg homeomorphic_imp_surjective_map by fastforce + with S fg have "m'.mball y r \ S" + by (smt (verit, del_insts) image_iff m'.in_mball r subset_iff) + then show "\r>0. m'.mball y r \ S" + using \0 < r\ by blast + qed + moreover have "openin X (g ` S)" if ope': "openin m'.mtopology S" + proof - + have "\r>0. m.mball (g y) r \ g ` S" if "y \ S" for y + proof - + have y: "y \ topspace Y" + using m'.openin_mtopology ope' that by blast + obtain r where "r > 0" and r: "m'.mball y r \ S" + using ope' by (meson \y \ S\ m'.openin_mtopology) + moreover have "\x. \x \ M; d (g y) x < r\ \ \u. u \ topspace Y \ d' y u < r \ x = g u" + using fg X d'_def hmf homeomorphic_imp_surjective_map by fastforce + ultimately have "m.mball (g y) r \ g ` m'.mball y r" + using y by (force simp: m'.openin_mtopology) + then show ?thesis + using \0 < r\ r by blast + qed + then show ?thesis + using X hmg homeomorphic_imp_surjective_map m.openin_mtopology ope' openin_subset by fastforce + qed + ultimately have "(S \ topspace Y \ openin X (g ` S)) = openin m'.mtopology S" + using m'.topspace_mtopology openin_subset by blast + then show "openin Y S = openin m'.mtopology S" + by (simp add: m'.mopen_def homeomorphic_map_openness_eq [OF hmg]) + qed + then show ?thesis + using m'.metrizable_space_mtopology by force +qed + +lemma homeomorphic_metrizable_space: + assumes "X homeomorphic_space Y" + shows "metrizable_space X \ metrizable_space Y" + using assms homeomorphic_metrizable_space_aux homeomorphic_space_sym by metis + +lemma metrizable_space_retraction_map_image: + "retraction_map X Y r \ metrizable_space X + \ metrizable_space Y" + using hereditary_imp_retractive_property metrizable_space_subtopology homeomorphic_metrizable_space + by blast + + +lemma metrizable_imp_Hausdorff_space: + "metrizable_space X \ Hausdorff_space X" + by (metis Metric_space.Hausdorff_space_mtopology metrizable_space_def) + +(** +lemma metrizable_imp_kc_space: + "metrizable_space X \ kc_space X" +oops + MESON_TAC[METRIZABLE_IMP_HAUSDORFF_SPACE; HAUSDORFF_IMP_KC_SPACE]);; + +lemma kc_space_mtopology: + "kc_space mtopology" +oops + REWRITE_TAC[GSYM FORALL_METRIZABLE_SPACE; METRIZABLE_IMP_KC_SPACE]);; +**) + +lemma metrizable_imp_t1_space: + "metrizable_space X \ t1_space X" + by (simp add: Hausdorff_imp_t1_space metrizable_imp_Hausdorff_space) + +lemma closed_imp_gdelta_in: + assumes X: "metrizable_space X" and S: "closedin X S" + shows "gdelta_in X S" +proof - + obtain M d where "Metric_space M d" and Xeq: "X = Metric_space.mtopology M d" + using X metrizable_space_def by blast + then interpret M: Metric_space M d + by blast + have "S \ M" + using M.closedin_metric \X = M.mtopology\ S by blast + show ?thesis + proof (cases "S = {}") + case True + then show ?thesis + by simp + next + case False + have "\y\S. d x y < inverse (1 + real n)" if "x \ S" for x n + using \S \ M\ M.mdist_zero [of x] that by force + moreover + have "x \ S" if "x \ M" and \
: "\n. \y\S. d x y < inverse(Suc n)" for x + proof - + have *: "\y\S. d x y < \" if "\ > 0" for \ + by (metis \
that not0_implies_Suc order_less_le order_less_le_trans real_arch_inverse) + have "closedin M.mtopology S" + using S by (simp add: Xeq) + then show ?thesis + apply (simp add: M.closedin_metric) + by (metis * \x \ M\ M.in_mball disjnt_insert1 insert_absorb subsetD) + qed + ultimately have Seq: "S = \(range (\n. {x\M. \y\S. d x y < inverse(Suc n)}))" + using \S \ M\ by force + have "openin M.mtopology {xa \ M. \y\S. d xa y < inverse (1 + real n)}" for n + proof (clarsimp simp: M.openin_mtopology) + fix x y + assume "x \ M" "y \ S" and dxy: "d x y < inverse (1 + real n)" + then have "\z. \z \ M; d x z < inverse (1 + real n) - d x y\ \ \y\S. d z y < inverse (1 + real n)" + by (smt (verit) M.commute M.triangle \S \ M\ in_mono) + with dxy show "\r>0. M.mball x r \ {z \ M. \y\S. d z y < inverse (1 + real n)}" + by (rule_tac x="inverse(Suc n) - d x y" in exI) auto + qed + then show ?thesis + apply (subst Seq) + apply (force simp: Xeq intro: gdelta_in_Inter open_imp_gdelta_in) + done + qed +qed + +lemma open_imp_fsigma_in: + "\metrizable_space X; openin X S\ \ fsigma_in X S" + by (meson closed_imp_gdelta_in fsigma_in_gdelta_in openin_closedin openin_subset) + +(*NEEDS first_countable +lemma first_countable_mtopology: + "first_countable mtopology" +oops + GEN_TAC THEN REWRITE_TAC[first_countable; TOPSPACE_MTOPOLOGY] THEN + X_GEN_TAC `x::A` THEN DISCH_TAC THEN + EXISTS_TAC `{ mball m (x::A,r) | rational r \ 0 < r}` THEN + REWRITE_TAC[FORALL_IN_GSPEC; OPEN_IN_MBALL; EXISTS_IN_GSPEC] THEN + ONCE_REWRITE_TAC[SET_RULE + `{f x | S x \ Q x} = f ` {x \ S. Q x}`] THEN + SIMP_TAC[COUNTABLE_IMAGE; COUNTABLE_RATIONAL; COUNTABLE_RESTRICT] THEN + REWRITE_TAC[OPEN_IN_MTOPOLOGY] THEN + X_GEN_TAC `U::A=>bool` THEN STRIP_TAC THEN + FIRST_X_ASSUM(MP_TAC \ SPEC `x::A`) THEN + ASM_REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN + X_GEN_TAC `r::real` THEN STRIP_TAC THEN FIRST_ASSUM + (MP_TAC \ SPEC `r::real` \ MATCH_MP RATIONAL_APPROXIMATION_BELOW) THEN + MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `q::real` THEN + REWRITE_TAC[REAL_SUB_REFL] THEN STRIP_TAC THEN + ASM_SIMP_TAC[CENTRE_IN_MBALL] THEN + TRANS_TAC SUBSET_TRANS `mball m (x::A,r)` THEN + ASM_SIMP_TAC[MBALL_SUBSET_CONCENTRIC; REAL_LT_IMP_LE]);; + +lemma metrizable_imp_first_countable: + "metrizable_space X \ first_countable X" +oops + REWRITE_TAC[FORALL_METRIZABLE_SPACE; FIRST_COUNTABLE_MTOPOLOGY]);; +*) + +lemma mball_eq_ball [simp]: "Met_TC.mball = ball" + by force + +lemma mopen_eq_open [simp]: "Met_TC.mopen = open" + by (force simp: open_contains_ball Met_TC.mopen_def) + +lemma metrizable_space_euclidean: + "metrizable_space (euclidean :: 'a::metric_space topology)" + unfolding metrizable_space_def + by (metis Met_TC.Metric_space_axioms Met_TC.mtopology_def mopen_eq_open) + +lemma (in Metric_space) regular_space_mtopology: + "regular_space mtopology" +unfolding regular_space_def +proof clarify + fix C a + assume C: "closedin mtopology C" and a: "a \ topspace mtopology" and "a \ C" + have "openin mtopology (topspace mtopology - C)" + by (simp add: C openin_diff) + then obtain r where "r>0" and r: "mball a r \ topspace mtopology - C" + unfolding openin_mtopology using \a \ C\ a by auto + show "\U V. openin mtopology U \ openin mtopology V \ a \ U \ C \ V \ disjnt U V" + proof (intro exI conjI) + show "a \ mball a (r/2)" + using \0 < r\ a by force + show "C \ topspace mtopology - mcball a (r/2)" + using C \0 < r\ r by (fastforce simp: closedin_metric) + qed (auto simp: openin_mball closedin_mcball openin_diff disjnt_iff) +qed + +lemma metrizable_imp_regular_space: + "metrizable_space X \ regular_space X" + by (metis Metric_space.regular_space_mtopology metrizable_space_def) + +lemma regular_space_euclidean: + "regular_space (euclidean :: 'a::metric_space topology)" + by (simp add: metrizable_imp_regular_space metrizable_space_euclidean) + + +subsection\Limits at a point in a topological space\ + +lemma (in Metric_space) eventually_atin_metric: + "eventually P (atin mtopology a) \ + (a \ M \ (\\>0. \x. x \ M \ 0 < d x a \ d x a < \ \ P x))" (is "?lhs=?rhs") +proof (cases "a \ M") + case True + show ?thesis + proof + assume L: ?lhs + with True obtain U where "openin mtopology U" "a \ U" and U: "\x\U - {a}. P x" + by (auto simp: eventually_atin) + then obtain r where "r>0" and "mball a r \ U" + by (meson openin_mtopology) + with U show ?rhs + by (smt (verit, ccfv_SIG) commute in_mball insert_Diff_single insert_iff subset_iff) + next + assume ?rhs + then obtain \ where "\>0" and \: "\x. x \ M \ 0 < d x a \ d x a < \ \ P x" + using True by blast + then have "\x \ mball a \ - {a}. P x" + by (simp add: commute) + then show ?lhs + unfolding eventually_atin openin_mtopology + by (metis True \0 < \\ centre_in_mball_iff openin_mball openin_mtopology) + qed +qed auto + +subsection \Normal spaces and metric spaces\ + +lemma (in Metric_space) normal_space_mtopology: + "normal_space mtopology" + unfolding normal_space_def +proof clarify + fix S T + assume "closedin mtopology S" + then have "\x. x \ M - S \ (\r>0. mball x r \ M - S)" + by (simp add: closedin_def openin_mtopology) + then obtain \ where d0: "\x. x \ M - S \ \ x > 0 \ mball x (\ x) \ M - S" + by metis + assume "closedin mtopology T" + then have "\x. x \ M - T \ (\r>0. mball x r \ M - T)" + by (simp add: closedin_def openin_mtopology) + then obtain \ where e: "\x. x \ M - T \ \ x > 0 \ mball x (\ x) \ M - T" + by metis + assume "disjnt S T" + have "S \ M" "T \ M" + using \closedin mtopology S\ \closedin mtopology T\ closedin_metric by blast+ + have \: "\x. x \ T \ \ x > 0 \ mball x (\ x) \ M - S" + by (meson DiffI \T \ M\ \disjnt S T\ d0 disjnt_iff subsetD) + have \: "\x. x \ S \ \ x > 0 \ mball x (\ x) \ M - T" + by (meson Diff_iff \S \ M\ \disjnt S T\ disjnt_iff e subsetD) + show "\U V. openin mtopology U \ openin mtopology V \ S \ U \ T \ V \ disjnt U V" + proof (intro exI conjI) + show "openin mtopology (\x\S. mball x (\ x / 2))" "openin mtopology (\x\T. mball x (\ x / 2))" + by force+ + show "S \ (\x\S. mball x (\ x / 2))" + using \ \S \ M\ by force + show "T \ (\x\T. mball x (\ x / 2))" + using \ \T \ M\ by force + show "disjnt (\x\S. mball x (\ x / 2)) (\x\T. mball x (\ x / 2))" + using \ \ + apply (clarsimp simp: disjnt_iff subset_iff) + by (smt (verit, ccfv_SIG) field_sum_of_halves triangle') + qed +qed + +lemma metrizable_imp_normal_space: + "metrizable_space X \ normal_space X" + by (metis Metric_space.normal_space_mtopology metrizable_space_def) + +subsection\Topological limitin in metric spaces\ + + +lemma (in Metric_space) limitin_mspace: + "limitin mtopology f l F \ l \ M" + using limitin_topspace by fastforce + +lemma (in Metric_space) limitin_metric_unique: + "\limitin mtopology f l1 F; limitin mtopology f l2 F; F \ bot\ \ l1 = l2" + by (meson Hausdorff_space_mtopology limitin_Hausdorff_unique) + +lemma (in Metric_space) limitin_metric: + "limitin mtopology f l F \ l \ M \ (\\>0. eventually (\x. f x \ M \ d (f x) l < \) F)" + (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding limitin_def + proof (intro conjI strip) + show "l \ M" + using L limitin_mspace by blast + fix \::real + assume "\>0" + then have "\\<^sub>F x in F. f x \ mball l \" + using L openin_mball by (fastforce simp: limitin_def) + then show "\\<^sub>F x in F. f x \ M \ d (f x) l < \" + using commute eventually_mono by fastforce + qed +next + assume R: ?rhs + then show ?lhs + by (force simp: limitin_def commute openin_mtopology subset_eq elim: eventually_mono) +qed + +lemma (in Metric_space) limit_metric_sequentially: + "limitin mtopology f l sequentially \ + l \ M \ (\\>0. \N. \n\N. f n \ M \ d (f n) l < \)" + by (auto simp: limitin_metric eventually_sequentially) + +lemma (in submetric) limitin_submetric_iff: + "limitin sub.mtopology f l F \ + l \ A \ eventually (\x. f x \ A) F \ limitin mtopology f l F" (is "?lhs=?rhs") + by (simp add: limitin_subtopology mtopology_submetric) + +lemma (in Metric_space) metric_closedin_iff_sequentially_closed: + "closedin mtopology S \ + S \ M \ (\\ l. range \ \ S \ limitin mtopology \ l sequentially \ l \ S)" (is "?lhs=?rhs") +proof + assume ?lhs then show ?rhs + by (force simp: closedin_metric limitin_closedin range_subsetD) +next + assume R: ?rhs + show ?lhs + unfolding closedin_metric + proof (intro conjI strip) + show "S \ M" + using R by blast + fix x + assume "x \ M - S" + have False if "\r>0. \y. y \ M \ y \ S \ d x y < r" + proof - + have "\n. \y. y \ M \ y \ S \ d x y < inverse(Suc n)" + using that by auto + then obtain \ where \: "\n. \ n \ M \ \ n \ S \ d x (\ n) < inverse(Suc n)" + by metis + then have "range \ \ M" + by blast + have "\N. \n\N. d x (\ n) < \" if "\>0" for \ + proof - + have "real (Suc (nat \inverse \\)) \ inverse \" + by linarith + then have "\n \ nat \inverse \\. d x (\ n) < \" + by (metis \ inverse_inverse_eq inverse_le_imp_le nat_ceiling_le_eq nle_le not_less_eq_eq order.strict_trans2 that) + then show ?thesis .. + qed + with \ have "limitin mtopology \ x sequentially" + using \x \ M - S\ commute limit_metric_sequentially by auto + then show ?thesis + by (metis R DiffD2 \ image_subset_iff \x \ M - S\) + qed + then show "\r>0. disjnt S (mball x r)" + by (meson disjnt_iff in_mball) + qed +qed + +lemma (in Metric_space) limit_atin_metric: + "limitin X f y (atin mtopology x) \ + y \ topspace X \ + (x \ M + \ (\V. openin X V \ y \ V + \ (\\>0. \x'. x' \ M \ 0 < d x' x \ d x' x < \ \ f x' \ V)))" + by (force simp: limitin_def eventually_atin_metric) + +lemma (in Metric_space) limitin_metric_dist_null: + "limitin mtopology f l F \ l \ M \ eventually (\x. f x \ M) F \ ((\x. d (f x) l) \ 0) F" + by (simp add: limitin_metric tendsto_iff eventually_conj_iff all_conj_distrib imp_conjR gt_ex) + + +subsection\Cauchy sequences and complete metric spaces\ + +context Metric_space +begin + +definition MCauchy :: "(nat \ 'a) \ bool" + where "MCauchy \ \ range \ \ M \ (\\>0. \N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \)" + +definition mcomplete + where "mcomplete \ (\\. MCauchy \ \ (\x. limitin mtopology \ x sequentially))" + +lemma mcomplete_empty [iff]: "Metric_space.mcomplete {} d" + by (simp add: Metric_space.MCauchy_def Metric_space.mcomplete_def subspace) + +lemma MCauchy_imp_MCauchy_suffix: "MCauchy \ \ MCauchy (\ \ (+)n)" + unfolding MCauchy_def image_subset_iff comp_apply + by (metis UNIV_I add.commute trans_le_add1) + +lemma mcomplete: + "mcomplete \ + (\\. (\\<^sub>F n in sequentially. \ n \ M) \ + (\\>0. \N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \) \ + (\x. limitin mtopology \ x sequentially))" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof clarify + fix \ + assume "\\<^sub>F n in sequentially. \ n \ M" + and \: "\\>0. \N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" + then obtain N where "\n. n\N \ \ n \ M" + by (auto simp: eventually_sequentially) + with \ have "MCauchy (\ \ (+)N)" + unfolding MCauchy_def image_subset_iff comp_apply by (meson le_add1 trans_le_add2) + then obtain x where "limitin mtopology (\ \ (+)N) x sequentially" + using L MCauchy_imp_MCauchy_suffix mcomplete_def by blast + then have "limitin mtopology \ x sequentially" + unfolding o_def by (auto simp: add.commute limitin_sequentially_offset_rev) + then show "\x. limitin mtopology \ x sequentially" .. + qed +qed (simp add: mcomplete_def MCauchy_def image_subset_iff) + +lemma mcomplete_empty_mspace: "M = {} \ mcomplete" + using MCauchy_def mcomplete_def by blast + +lemma MCauchy_const [simp]: "MCauchy (\n. a) \ a \ M" + using MCauchy_def mdist_zero by auto + +lemma convergent_imp_MCauchy: + assumes "range \ \ M" and lim: "limitin mtopology \ l sequentially" + shows "MCauchy \" + unfolding MCauchy_def image_subset_iff +proof (intro conjI strip) + fix \::real + assume "\ > 0" + then have "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) l < \/2" + using half_gt_zero lim limitin_metric by blast + then obtain N where "\n. n\N \ \ n \ M \ d (\ n) l < \/2" + by (force simp: eventually_sequentially) + then show "\N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" + by (smt (verit) Metric_space.limitin_mspace Metric_space.mdist_reverse_triangle Metric_space_axioms field_sum_of_halves lim) +qed (use assms in blast) + + +lemma mcomplete_alt: + "mcomplete \ (\\. MCauchy \ \ range \ \ M \ (\x. limitin mtopology \ x sequentially))" + using MCauchy_def convergent_imp_MCauchy mcomplete_def by blast + +lemma MCauchy_subsequence: + assumes "strict_mono r" "MCauchy \" + shows "MCauchy (\ \ r)" +proof - + have "d (\ (r n)) (\ (r n')) < \" + if "N \ n" "N \ n'" "strict_mono r" "\n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" + for \ N n n' + using that by (meson le_trans strict_mono_imp_increasing) + moreover have "range (\x. \ (r x)) \ M" + using MCauchy_def assms by blast + ultimately show ?thesis + using assms by (simp add: MCauchy_def) metis +qed + +lemma MCauchy_offset: + assumes cau: "MCauchy (\ \ (+)k)" and \: "\n. n < k \ \ n \ M" + shows "MCauchy \" + unfolding MCauchy_def image_subset_iff +proof (intro conjI strip) + fix n + show "\ n \ M" + using assms + unfolding MCauchy_def image_subset_iff + by (metis UNIV_I comp_apply le_iff_add linorder_not_le) +next + fix \ :: real + assume "\ > 0" + obtain N where "\n n'. N \ n \ N \ n' \ d ((\ \ (+)k) n) ((\ \ (+)k) n') < \" + using cau \\ > 0\ by (fastforce simp: MCauchy_def) + then show "\N. \n n'. N \ n \ N \ n' \ d (\ n) (\ n') < \" + unfolding o_def + apply (rule_tac x="k+N" in exI) + by (smt (verit, del_insts) add.assoc le_add1 less_eqE) +qed + +lemma MCauchy_convergent_subsequence: + assumes cau: "MCauchy \" and "strict_mono r" + and lim: "limitin mtopology (\ \ r) a sequentially" + shows "limitin mtopology \ a sequentially" + unfolding limitin_metric +proof (intro conjI strip) + show "a \ M" + by (meson assms limitin_mspace) + fix \ :: real + assume "\ > 0" + then obtain N1 where N1: "\n n'. \n\N1; n'\N1\ \ d (\ n) (\ n') < \/2" + using cau unfolding MCauchy_def by (meson half_gt_zero) + obtain N2 where N2: "\n. n \ N2 \ (\ \ r) n \ M \ d ((\ \ r) n) a < \/2" + by (metis (no_types, lifting) lim \\ > 0\ half_gt_zero limit_metric_sequentially) + have "\ n \ M \ d (\ n) a < \" if "n \ max N1 N2" for n + proof (intro conjI) + show "\ n \ M" + using MCauchy_def cau by blast + have "N1 \ r n" + by (meson \strict_mono r\ le_trans max.cobounded1 strict_mono_imp_increasing that) + then show "d (\ n) a < \" + using N1[of n "r n"] N2[of n] \\ n \ M\ \a \ M\ triangle that by fastforce + qed + then show "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) a < \" + using eventually_sequentially by blast +qed + +lemma MCauchy_interleaving_gen: + "MCauchy (\n. if even n then x(n div 2) else y(n div 2)) \ + (MCauchy x \ MCauchy y \ (\n. d (x n) (y n)) \ 0)" (is "?lhs=?rhs") +proof + assume L: ?lhs + have evens: "strict_mono (\n::nat. 2 * n)" and odds: "strict_mono (\n::nat. Suc (2 * n))" + by (auto simp: strict_mono_def) + show ?rhs + proof (intro conjI) + show "MCauchy x" "MCauchy y" + using MCauchy_subsequence [OF evens L] MCauchy_subsequence [OF odds L] by (auto simp: o_def) + show "(\n. d (x n) (y n)) \ 0" + unfolding LIMSEQ_iff + proof (intro strip) + fix \ :: real + assume "\ > 0" + then obtain N where N: + "\n n'. \n\N; n'\N\ \ d (if even n then x (n div 2) else y (n div 2)) + (if even n' then x (n' div 2) else y (n' div 2)) < \" + using L MCauchy_def by fastforce + have "d (x n) (y n) < \" if "n\N" for n + using N [of "2*n" "Suc(2*n)"] that by auto + then show "\N. \n\N. norm (d (x n) (y n) - 0) < \" + by auto + qed + qed +next + assume R: ?rhs + show ?lhs + unfolding MCauchy_def + proof (intro conjI strip) + show "range (\n. if even n then x (n div 2) else y (n div 2)) \ M" + using R by (auto simp: MCauchy_def) + fix \ :: real + assume "\ > 0" + obtain Nx where Nx: "\n n'. \n\Nx; n'\Nx\ \ d (x n) (x n') < \/2" + by (meson half_gt_zero MCauchy_def R \\ > 0\) + obtain Ny where Ny: "\n n'. \n\Ny; n'\Ny\ \ d (y n) (y n') < \/2" + by (meson half_gt_zero MCauchy_def R \\ > 0\) + obtain Nxy where Nxy: "\n. n\Nxy \ d (x n) (y n) < \/2" + using R \\ > 0\ half_gt_zero unfolding LIMSEQ_iff + by (metis abs_mdist diff_zero real_norm_def) + define N where "N \ 2 * Max{Nx,Ny,Nxy}" + show "\N. \n n'. N \ n \ N \ n' \ d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \" + proof (intro exI strip) + fix n n' + assume "N \ n" and "N \ n'" + then have "n div 2 \ Nx" "n div 2 \ Ny" "n div 2 \ Nxy" "n' div 2 \ Nx" "n' div 2 \ Ny" + by (auto simp: N_def) + then have dxyn: "d (x (n div 2)) (y (n div 2)) < \/2" + and dxnn': "d (x (n div 2)) (x (n' div 2)) < \/2" + and dynn': "d (y (n div 2)) (y (n' div 2)) < \/2" + using Nx Ny Nxy by blast+ + have inM: "x (n div 2) \ M" "x (n' div 2) \ M""y (n div 2) \ M" "y (n' div 2) \ M" + using Metric_space.MCauchy_def Metric_space_axioms R by blast+ + show "d (if even n then x (n div 2) else y (n div 2)) (if even n' then x (n' div 2) else y (n' div 2)) < \" + proof (cases "even n") + case nt: True + show ?thesis + proof (cases "even n'") + case True + with \\ > 0\ nt dxnn' show ?thesis by auto + next + case False + with nt dxyn dynn' inM triangle show ?thesis + by fastforce + qed + next + case nf: False + show ?thesis + proof (cases "even n'") + case True + then show ?thesis + by (smt (verit) \\ > 0\ dxyn dxnn' triangle commute inM field_sum_of_halves) + next + case False + with \\ > 0\ nf dynn' show ?thesis by auto + qed + qed + qed + qed +qed + +lemma MCauchy_interleaving: + "MCauchy (\n. if even n then \(n div 2) else a) \ + range \ \ M \ limitin mtopology \ a sequentially" (is "?lhs=?rhs") +proof - + have "?lhs \ (MCauchy \ \ a \ M \ (\n. d (\ n) a) \ 0)" + by (simp add: MCauchy_interleaving_gen [where y = "\n. a"]) + also have "... = ?rhs" + by (metis MCauchy_def always_eventually convergent_imp_MCauchy limitin_metric_dist_null range_subsetD) + finally show ?thesis . +qed + +lemma mcomplete_nest: + "mcomplete \ + (\C::nat \'a set. (\n. closedin mtopology (C n)) \ + (\n. C n \ {}) \ decseq C \ (\\>0. \n a. C n \ mcball a \) + \ \ (range C) \ {})" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding imp_conjL + proof (intro strip) + fix C :: "nat \ 'a set" + assume clo: "\n. closedin mtopology (C n)" + and ne: "\n. C n \ ({}::'a set)" + and dec: "decseq C" + and cover [rule_format]: "\\>0. \n a. C n \ mcball a \" + obtain \ where \: "\n. \ n \ C n" + by (meson ne empty_iff set_eq_iff) + have "MCauchy \" + unfolding MCauchy_def + proof (intro conjI strip) + show "range \ \ M" + using \ clo metric_closedin_iff_sequentially_closed by auto + fix \ :: real + assume "\ > 0" + then obtain N a where N: "C N \ mcball a (\/3)" + using cover by fastforce + have "d (\ m) (\ n) < \" if "N \ m" "N \ n" for m n + proof - + have "d a (\ m) \ \/3" "d a (\ n) \ \/3" + using dec N \ that by (fastforce simp: decseq_def)+ + then have "d (\ m) (\ n) \ \/3 + \/3" + using triangle \ commute dec decseq_def subsetD that N + by (smt (verit, ccfv_threshold) in_mcball) + also have "... < \" + using \\ > 0\ by auto + finally show ?thesis . + qed + then show "\N. \m n. N \ m \ N \ n \ d (\ m) (\ n) < \" + by blast + qed + then obtain x where x: "limitin mtopology \ x sequentially" + using L mcomplete_def by blast + have "x \ C n" for n + proof (rule limitin_closedin [OF x]) + show "closedin mtopology (C n)" + by (simp add: clo) + show "\\<^sub>F x in sequentially. \ x \ C n" + by (metis \ dec decseq_def eventually_sequentiallyI subsetD) + qed auto + then show "\ (range C) \ {}" + by blast +qed +next + assume R: ?rhs + show ?lhs + unfolding mcomplete_def + proof (intro strip) + fix \ + assume "MCauchy \" + then have "range \ \ M" + using MCauchy_def by blast + define C where "C \ \n. mtopology closure_of (\ ` {n..})" + have "\n. closedin mtopology (C n)" + by (auto simp: C_def) + moreover + have ne: "\n. C n \ {}" + using \MCauchy \\ by (auto simp: C_def MCauchy_def disjnt_iff closure_of_eq_empty_gen) + moreover + have dec: "decseq C" + unfolding monotone_on_def + proof (intro strip) + fix m n::nat + assume "m \ n" + then have "{n..} \ {m..}" + by auto + then show "C n \ C m" + unfolding C_def by (meson closure_of_mono image_mono) + qed + moreover + have C: "\N u. C N \ mcball u \" if "\>0" for \ + proof - + obtain N where "\m n. N \ m \ N \ n \ d (\ m) (\ n) < \" + by (meson MCauchy_def \0 < \\ \MCauchy \\) + then have "\ ` {N..} \ mcball (\ N) \" + using MCauchy_def \MCauchy \\ by (force simp: less_eq_real_def) + then have "C N \ mcball (\ N) \" + by (simp add: C_def closure_of_minimal) + then show ?thesis + by blast + qed + ultimately obtain l where x: "l \ \ (range C)" + by (metis R ex_in_conv) + then have *: "\\ N. 0 < \ \ \n'. N \ n' \ l \ M \ \ n' \ M \ d l (\ n') < \" + by (force simp: C_def metric_closure_of) + then have "l \ M" + using gt_ex by blast + show "\l. limitin mtopology \ l sequentially" + unfolding limitin_metric + proof (intro conjI strip exI) + show "l \ M" + using \\n. closedin mtopology (C n)\ closedin_subset x by fastforce + fix \::real + assume "\ > 0" + obtain N where N: "\m n. N \ m \ N \ n \ d (\ m) (\ n) < \/2" + by (meson MCauchy_def \0 < \\ \MCauchy \\ half_gt_zero) + with * [of "\/2" N] + have "\n\N. \ n \ M \ d (\ n) l < \" + by (smt (verit) \range \ \ M\ commute field_sum_of_halves range_subsetD triangle) + then show "\\<^sub>F n in sequentially. \ n \ M \ d (\ n) l < \" + using eventually_sequentially by blast + qed + qed +qed + + +lemma mcomplete_nest_sing: + "mcomplete \ + (\C. (\n. closedin mtopology (C n)) \ + (\n. C n \ {}) \ decseq C \ (\e>0. \n a. C n \ mcball a e) + \ (\l. l \ M \ \ (range C) = {l}))" +proof - + have *: False + if clo: "\n. closedin mtopology (C n)" + and cover: "\\>0. \n a. C n \ mcball a \" + and no_sing: "\y. y \ M \ \ (range C) \ {y}" + and l: "\n. l \ C n" + for C :: "nat \ 'a set" and l + proof - + have inM: "\x. x \ \ (range C) \ x \ M" + using closedin_metric clo by fastforce + then have "l \ M" + by (simp add: l) + have False if l': "l' \ \ (range C)" and "l' \ l" for l' + proof - + have "l' \ M" + using inM l' by blast + obtain n a where na: "C n \ mcball a (d l l' / 3)" + using inM \l \ M\ l' \l' \ l\ cover by force + then have "d a l \ (d l l' / 3)" "d a l' \ (d l l' / 3)" "a \ M" + using l l' na in_mcball by auto + then have "d l l' \ (d l l' / 3) + (d l l' / 3)" + using \l \ M\ \l' \ M\ mdist_reverse_triangle by fastforce + then show False + using nonneg [of l l'] \l' \ l\ \l \ M\ \l' \ M\ zero by force + qed + then show False + by (metis l \l \ M\ no_sing INT_I empty_iff insertI1 is_singletonE is_singletonI') + qed + show ?thesis + unfolding mcomplete_nest imp_conjL + apply (intro all_cong1 imp_cong refl) + using * + by (smt (verit) Inter_iff ex_in_conv range_constant range_eqI) +qed + +lemma mcomplete_fip: + "mcomplete \ + (\\. (\C \ \. closedin mtopology C) \ + (\e>0. \C a. C \ \ \ C \ mcball a e) \ (\\. finite \ \ \ \ \ \ \ \ \ {}) + \ \ \ \ {})" + (is "?lhs = ?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding mcomplete_nest_sing imp_conjL + proof (intro strip) + fix \ :: "'a set set" + assume clo: "\C\\. closedin mtopology C" + and cover: "\e>0. \C a. C \ \ \ C \ mcball a e" + and fip: "\\. finite \ \ \ \ \ \ \ \ \ {}" + then have "\n. \C. C \ \ \ (\a. C \ mcball a (inverse (Suc n)))" + by simp + then obtain C where C: "\n. C n \ \" + and coverC: "\n. \a. C n \ mcball a (inverse (Suc n))" + by metis + define D where "D \ \n. \ (C ` {..n})" + have cloD: "closedin mtopology (D n)" for n + unfolding D_def using clo C by blast + have neD: "D n \ {}" for n + using fip C by (simp add: D_def image_subset_iff) + have decD: "decseq D" + by (force simp: D_def decseq_def) + have coverD: "\n a. D n \ mcball a \" if " \ >0" for \ + proof - + obtain n where "inverse (Suc n) < \" + using \0 < \\ reals_Archimedean by blast + then obtain a where "C n \ mcball a \" + by (meson coverC less_eq_real_def mcball_subset_concentric order_trans) + then show ?thesis + unfolding D_def by blast + qed + have *: "a \ \\" if a: "\ (range D) = {a}" and "a \ M" for a + proof - + have aC: "a \ C n" for n + using that by (auto simp: D_def) + have eqa: "\u. (\n. u \ C n) \ a = u" + using that by (auto simp: D_def) + have "a \ T" if "T \ \" for T + proof - + have cloT: "closedin mtopology (T \ D n)" for n + using clo cloD that by blast + have "\ (insert T (C ` {..n})) \ {}" for n + using that C by (intro fip [rule_format]) auto + then have neT: "T \ D n \ {}" for n + by (simp add: D_def) + have decT: "decseq (\n. T \ D n)" + by (force simp: D_def decseq_def) + have coverT: "\n a. T \ D n \ mcball a \" if " \ >0" for \ + by (meson coverD le_infI2 that) + show ?thesis + using L [unfolded mcomplete_nest_sing, rule_format, of "\n. T \ D n"] a + by (force simp: cloT neT decT coverT) + qed + then show ?thesis by auto + qed + show "\ \ \ {}" + by (metis L cloD neD decD coverD * empty_iff mcomplete_nest_sing) + qed +next + assume R [rule_format]: ?rhs + show ?lhs + unfolding mcomplete_nest imp_conjL + proof (intro strip) + fix C :: "nat \ 'a set" + assume clo: "\n. closedin mtopology (C n)" + and ne: "\n. C n \ {}" + and dec: "decseq C" + and cover: "\\>0. \n a. C n \ mcball a \" + have "\(C ` N) \ {}" if "finite N" for N + proof - + obtain k where "N \ {..k}" + using \finite N\ finite_nat_iff_bounded_le by auto + with dec have "C k \ \(C ` N)" by (auto simp: decseq_def) + then show ?thesis + using ne by force + qed + with clo cover R [of "range C"] show "\ (range C) \ {}" + by (metis (no_types, opaque_lifting) finite_subset_image image_iff UNIV_I) + qed +qed + + +lemma mcomplete_fip_sing: + "mcomplete \ + (\\. (\C\\. closedin mtopology C) \ + (\e>0. \c a. c \ \ \ c \ mcball a e) \ + (\\. finite \ \ \ \ \ \ \ \ \ {}) \ + (\l. l \ M \ \ \ = {l}))" + (is "?lhs = ?rhs") +proof + have *: "l \ M" "\ \ = {l}" + if clo: "Ball \ (closedin mtopology)" + and cover: "\e>0. \C a. C \ \ \ C \ mcball a e" + and fin: "\\. finite \ \ \ \ \ \ \ \ \ {}" + and l: "l \ \ \" + for \ :: "'a set set" and l + proof - + show "l \ M" + by (meson Inf_lower2 clo cover gt_ex metric_closedin_iff_sequentially_closed subsetD that(4)) + show "\ \ = {l}" + proof (cases "\ = {}") + case True + then show ?thesis + using cover mbounded_pos by auto + next + case False + have CM: "\a. a \ \\ \ a \ M" + using False clo closedin_subset by fastforce + have "l' \ \ \" if "l' \ l" for l' + proof + assume l': "l' \ \ \" + with CM have "l' \ M" by blast + with that \l \ M\ have gt0: "0 < d l l'" + by simp + then obtain C a where "C \ \" and C: "C \ mcball a (d l l' / 3)" + using cover [rule_format, of "d l l' / 3"] by auto + then have "d a l \ (d l l' / 3)" "d a l' \ (d l l' / 3)" "a \ M" + using l l' in_mcball by auto + then have "d l l' \ (d l l' / 3) + (d l l' / 3)" + using \l \ M\ \l' \ M\ mdist_reverse_triangle by fastforce + with gt0 show False by auto + qed + then show ?thesis + using l by fastforce + qed + qed + assume L: ?lhs + with * show ?rhs + unfolding mcomplete_fip imp_conjL ex_in_conv [symmetric] + by (elim all_forward imp_forward2 asm_rl) (blast intro: elim: ) +next + assume ?rhs then show ?lhs + unfolding mcomplete_fip by (force elim!: all_forward) +qed + +end + +lemma MCauchy_iff_Cauchy [iff]: "Met_TC.MCauchy = Cauchy" + by (force simp: Cauchy_def Met_TC.MCauchy_def) + +lemma mcomplete_iff_complete [iff]: + "Met_TC.mcomplete (Pure.type ::'a::metric_space itself) \ complete (UNIV::'a set)" + by (auto simp: Met_TC.mcomplete_def complete_def) + +lemma euclidean_metric: "Met_TC.mcomplete (Pure.type ::'a::euclidean_space itself)" + using complete_UNIV mcomplete_iff_complete by blast + +context submetric +begin + +lemma MCauchy_submetric: + "sub.MCauchy \ \ range \ \ A \ MCauchy \" + using MCauchy_def sub.MCauchy_def subset by force + +lemma closedin_mcomplete_imp_mcomplete: + assumes clo: "closedin mtopology A" and "mcomplete" + shows "sub.mcomplete" + unfolding sub.mcomplete_def +proof (intro strip) + fix \ + assume "sub.MCauchy \" + then have \: "MCauchy \" "range \ \ A" + using MCauchy_submetric by blast+ + then obtain x where x: "limitin mtopology \ x sequentially" + using \mcomplete\ unfolding mcomplete_def by blast + then have "x \ A" + using \ clo metric_closedin_iff_sequentially_closed by force + with \ x show "\x. limitin sub.mtopology \ x sequentially" + using limitin_submetric_iff range_subsetD by fastforce +qed + + +lemma sequentially_closedin_mcomplete_imp_mcomplete: + assumes "mcomplete" and "\\ l. range \ \ A \ limitin mtopology \ l sequentially \ l \ A" + shows "sub.mcomplete" + using assms closedin_mcomplete_imp_mcomplete metric_closedin_iff_sequentially_closed subset by blast + +end + + +context Metric_space +begin + +lemma mcomplete_Un: + assumes A: "submetric M d A" "Metric_space.mcomplete A d" + and B: "submetric M d B" "Metric_space.mcomplete B d" + shows "submetric M d (A \ B)" "Metric_space.mcomplete (A \ B) d" +proof - + show "submetric M d (A \ B)" + by (meson assms le_sup_iff submetric_axioms_def submetric_def) + then interpret MAB: Metric_space "A \ B" d + by (meson submetric.subset subspace) + interpret MA: Metric_space A d + by (meson A submetric.subset subspace) + interpret MB: Metric_space B d + by (meson B submetric.subset subspace) + show "Metric_space.mcomplete (A \ B) d" + unfolding MAB.mcomplete_def + proof (intro strip) + fix \ + assume "MAB.MCauchy \" + then have "range \ \ A \ B" + using MAB.MCauchy_def by blast + then have "UNIV \ \ -` A \ \ -` B" + by blast + then consider "infinite (\ -` A)" | "infinite (\ -` B)" + using finite_subset by auto + then show "\x. limitin MAB.mtopology \ x sequentially" + proof cases + case 1 + then obtain r where "strict_mono r" and r: "\n::nat. r n \ \ -` A" + using infinite_enumerate by blast + then have "MA.MCauchy (\ \ r)" + using MA.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence \MAB.MCauchy \\ by auto + with A obtain x where "limitin MA.mtopology (\ \ r) x sequentially" + using MA.mcomplete_def by blast + then have "limitin MAB.mtopology (\ \ r) x sequentially" + by (metis MA.limit_metric_sequentially MAB.limit_metric_sequentially UnCI) + then show ?thesis + using MAB.MCauchy_convergent_subsequence \MAB.MCauchy \\ \strict_mono r\ by blast + next + case 2 + then obtain r where "strict_mono r" and r: "\n::nat. r n \ \ -` B" + using infinite_enumerate by blast + then have "MB.MCauchy (\ \ r)" + using MB.MCauchy_def MAB.MCauchy_def MAB.MCauchy_subsequence \MAB.MCauchy \\ by auto + with B obtain x where "limitin MB.mtopology (\ \ r) x sequentially" + using MB.mcomplete_def by blast + then have "limitin MAB.mtopology (\ \ r) x sequentially" + by (metis MB.limit_metric_sequentially MAB.limit_metric_sequentially UnCI) + then show ?thesis + using MAB.MCauchy_convergent_subsequence \MAB.MCauchy \\ \strict_mono r\ by blast + qed + qed +qed + +lemma mcomplete_Union: + assumes "finite \" + and "\A. A \ \ \ submetric M d A" "\A. A \ \ \ Metric_space.mcomplete A d" + shows "submetric M d (\\)" "Metric_space.mcomplete (\\) d" + using assms + by (induction rule: finite_induct) (auto simp: mcomplete_Un) + +lemma mcomplete_Inter: + assumes "finite \" "\ \ {}" + and sub: "\A. A \ \ \ submetric M d A" + and comp: "\A. A \ \ \ Metric_space.mcomplete A d" + shows "submetric M d (\\)" "Metric_space.mcomplete (\\) d" +proof - + show "submetric M d (\\)" + using assms unfolding submetric_def submetric_axioms_def + by (metis Inter_lower equals0I inf.orderE le_inf_iff) + then interpret MS: submetric M d "\\" + by (meson submetric.subset subspace) + show "Metric_space.mcomplete (\\) d" + unfolding MS.sub.mcomplete_def + proof (intro strip) + fix \ + assume "MS.sub.MCauchy \" + then have "range \ \ \\" + using MS.MCauchy_submetric by blast + obtain A where "A \ \" and A: "Metric_space.mcomplete A d" + using assms by blast + then have "range \ \ A" + using \range \ \ \\\ by blast + interpret SA: submetric M d A + by (meson \A \ \\ sub submetric.subset subspace) + have "MCauchy \" + using MS.MCauchy_submetric \MS.sub.MCauchy \\ by blast + then obtain x where x: "limitin SA.sub.mtopology \ x sequentially" + by (metis A SA.sub.MCauchy_def SA.sub.mcomplete_alt MCauchy_def \range \ \ A\) + show "\x. limitin MS.sub.mtopology \ x sequentially" + apply (rule_tac x="x" in exI) + unfolding MS.limitin_submetric_iff + proof (intro conjI) + show "x \ \ \" + proof clarsimp + fix U + assume "U \ \" + interpret SU: submetric M d U + by (meson \U \ \\ sub submetric.subset subspace) + have "range \ \ U" + using \U \ \\ \range \ \ \ \\ by blast + moreover have "Metric_space.mcomplete U d" + by (simp add: \U \ \\ comp) + ultimately obtain x' where x': "limitin SU.sub.mtopology \ x' sequentially" + using MCauchy_def SU.sub.MCauchy_def SU.sub.mcomplete_alt \MCauchy \\ by meson + have "x' = x" + proof (intro limitin_metric_unique) + show "limitin mtopology \ x' sequentially" + by (meson SU.submetric_axioms submetric.limitin_submetric_iff x') + show "limitin mtopology \ x sequentially" + by (meson SA.submetric_axioms submetric.limitin_submetric_iff x) + qed auto + then show "x \ U" + using SU.sub.limitin_mspace x' by blast + qed + show "\\<^sub>F n in sequentially. \ n \ \\" + by (meson \range \ \ \ \\ always_eventually range_subsetD) + show "limitin mtopology \ x sequentially" + by (meson SA.submetric_axioms submetric.limitin_submetric_iff x) + qed + qed +qed + + +lemma mcomplete_Int: + assumes A: "submetric M d A" "Metric_space.mcomplete A d" + and B: "submetric M d B" "Metric_space.mcomplete B d" + shows "submetric M d (A \ B)" "Metric_space.mcomplete (A \ B) d" + using mcomplete_Inter [of "{A,B}"] assms by force+ + +subsection\Totally bounded subsets of metric spaces\ + +definition mtotally_bounded + where "mtotally_bounded S \ \\>0. \K. finite K \ K \ S \ S \ (\x\K. mball x \)" + +lemma mtotally_bounded_empty [iff]: "mtotally_bounded {}" +by (simp add: mtotally_bounded_def) + +lemma finite_imp_mtotally_bounded: + "\finite S; S \ M\ \ mtotally_bounded S" + by (auto simp: mtotally_bounded_def) + +lemma mtotally_bounded_imp_subset: "mtotally_bounded S \ S \ M" + by (force simp: mtotally_bounded_def intro!: zero_less_one) + +lemma mtotally_bounded_sing [simp]: + "mtotally_bounded {x} \ x \ M" + by (meson empty_subsetI finite.simps finite_imp_mtotally_bounded insert_subset mtotally_bounded_imp_subset) + +lemma mtotally_bounded_Un: + assumes "mtotally_bounded S" "mtotally_bounded T" + shows "mtotally_bounded (S \ T)" +proof - + have "\K. finite K \ K \ S \ T \ S \ T \ (\x\K. mball x e)" + if "e>0" and K: "finite K \ K \ S \ S \ (\x\K. mball x e)" + and L: "finite L \ L \ T \ T \ (\x\L. mball x e)" for K L e + using that by (rule_tac x="K \ L" in exI) auto + with assms show ?thesis + unfolding mtotally_bounded_def by presburger +qed + +lemma mtotally_bounded_Union: + assumes "finite f" "\S. S \ f \ mtotally_bounded S" + shows "mtotally_bounded (\f)" + using assms by (induction f) (auto simp: mtotally_bounded_Un) + +lemma mtotally_bounded_imp_mbounded: + assumes "mtotally_bounded S" + shows "mbounded S" +proof - + obtain K where "finite K \ K \ S \ S \ (\x\K. mball x 1)" + using assms by (force simp: mtotally_bounded_def) + then show ?thesis + by (smt (verit) finite_imageI image_iff mbounded_Union mbounded_mball mbounded_subset) +qed + + +lemma mtotally_bounded_sequentially: + "mtotally_bounded S \ + S \ M \ (\\::nat \ 'a. range \ \ S \ (\r. strict_mono r \ MCauchy (\ \ r)))" + (is "_ \ _ \ ?rhs") +proof (cases "S \ M") + case True + show ?thesis + proof - + { fix \ :: "nat \ 'a" + assume L: "mtotally_bounded S" and \: "range \ \ S" + have "\j > i. d (\ i) (\ j) < 3*\/2 \ infinite (\ -` mball (\ j) (\/2))" + if inf: "infinite (\ -` mball (\ i) \)" and "\ > 0" for i \ + proof - + obtain K where "finite K" "K \ S" and K: "S \ (\x\K. mball x (\/4))" + by (metis L mtotally_bounded_def \\ > 0\ zero_less_divide_iff zero_less_numeral) + then have K_imp_ex: "\y. y \ S \ \x\K. d x y < \/4" + by fastforce + have False if "\x\K. d x (\ i) < \ + \/4 \ finite (\ -` mball x (\/4))" + proof - + have "\w. w \ K \ d w (\ i) < 5 * \/4 \ d w (\ j) < \/4" + if "d (\ i) (\ j) < \" for j + proof - + obtain w where w: "d w (\ j) < \/4" "w \ K" + using K_imp_ex \ by blast + then have "d w (\ i) < \ + \/4" + by (smt (verit, ccfv_SIG) True \K \ S\ \ rangeI subset_eq that triangle') + with w show ?thesis + using in_mball by auto + qed + then have "(\ -` mball (\ i) \) \ (\x\K. if d x (\ i) < \ + \/4 then \ -` mball x (\/4) else {})" + using True \K \ S\ by force + then show False + using finite_subset inf \finite K\ that by fastforce + qed + then obtain x where "x \ K" and dxi: "d x (\ i) < \ + \/4" and infx: "infinite (\ -` mball x (\/4))" + by blast + then obtain j where "j \ (\ -` mball x (\/4)) - {..i}" + using bounded_nat_set_is_finite by (meson Diff_infinite_finite finite_atMost) + then have "j > i" and dxj: "d x (\ j) < \/4" + by auto + have "(\ -` mball x (\/4)) \ (\ -` mball y (\/2))" if "d x y < \/4" "y \ M" for y + using that by (simp add: mball_subset Metric_space_axioms vimage_mono) + then have infj: "infinite (\ -` mball (\ j) (\/2))" + by (meson True \d x (\ j) < \/4\ \ in_mono infx rangeI finite_subset) + have "\ i \ M" "\ j \ M" "x \ M" + using True \K \ S\ \x \ K\ \ by force+ + then have "d (\ i) (\ j) \ d x (\ i) + d x (\ j)" + using triangle'' by blast + also have "\ < 3*\/2" + using dxi dxj by auto + finally have "d (\ i) (\ j) < 3*\/2" . + with \i < j\ infj show ?thesis by blast + qed + then obtain nxt where nxt: "\i \. \\ > 0; infinite (\ -` mball (\ i) \)\ \ + nxt i \ > i \ d (\ i) (\ (nxt i \)) < 3*\/2 \ infinite (\ -` mball (\ (nxt i \)) (\/2))" + by metis + have "mbounded S" + using L by (simp add: mtotally_bounded_imp_mbounded) + then obtain B where B: "\y \ S. d (\ 0) y \ B" and "B > 0" + by (meson \ mbounded_alt_pos range_subsetD) + define eps where "eps \ \n. (B+1) / 2^n" + have [simp]: "eps (Suc n) = eps n / 2" "eps n > 0" for n + using \B > 0\ by (auto simp: eps_def) + have "UNIV \ \ -` mball (\ 0) (B+1)" + using B True \ unfolding image_iff subset_iff + by (smt (verit, best) UNIV_I in_mball vimageI) + then have inf0: "infinite (\ -` mball (\ 0) (eps 0))" + using finite_subset by (auto simp: eps_def) + define r where "r \ rec_nat 0 (\n rec. nxt rec (eps n))" + have [simp]: "r 0 = 0" "r (Suc n) = nxt (r n) (eps n)" for n + by (auto simp: r_def) + have \rM[simp]: "\ (r n) \ M" for n + using True \ by blast + have inf: "infinite (\ -` mball (\ (r n)) (eps n))" for n + proof (induction n) + case 0 then show ?case + by (simp add: inf0) + next + case (Suc n) then show ?case + using nxt [of "eps n" "r n"] by simp + qed + then have "r (Suc n) > r n" for n + by (simp add: nxt) + then have "strict_mono r" + by (simp add: strict_mono_Suc_iff) + have d_less: "d (\ (r n)) (\ (r (Suc n))) < 3 * eps n / 2" for n + using nxt [OF _ inf] by simp + have eps_plus: "eps (k + n) = eps n * (1/2)^k" for k n + by (simp add: eps_def power_add field_simps) + have *: "d (\ (r n)) (\ (r (k + n))) < 3 * eps n" for n k + proof - + have "d (\ (r n)) (\ (r (k+n))) \ 3/2 * eps n * (\i (r n)) (\ (r (Suc k + n))) \ d (\ (r n)) (\ (r (k + n))) + d (\ (r (k + n))) (\ (r (Suc (k + n))))" + by (metis \rM add.commute add_Suc_right triangle) + with d_less[of "k+n"] Suc show ?case + by (simp add: algebra_simps eps_plus) + qed + also have "\ < 3/2 * eps n * 2" + using geometric_sum [of "1/2::real" k] by simp + finally show ?thesis by simp + qed + have "\N. \n\N. \n'\N. d (\ (r n)) (\ (r n')) < \" if "\ > 0" for \ + proof - + define N where "N \ nat \(log 2 (6*(B+1) / \))\" + have \
: "b \ 2 ^ nat \log 2 b\" for b + by (smt (verit) less_log_of_power real_nat_ceiling_ge) + have N: "6 * eps N \ \" + using \
[of "(6*(B+1) / \)"] that by (auto simp: N_def eps_def field_simps) + have "d (\ (r N)) (\ (r n)) < 3 * eps N" if "n \ N" for n + by (metis * add.commute nat_le_iff_add that) + then have "\n\N. \n'\N. d (\ (r n)) (\ (r n')) < 3 * eps N + 3 * eps N" + by (smt (verit, best) \rM triangle'') + with N show ?thesis + by fastforce + qed + then have "MCauchy (\ \ r)" + unfolding MCauchy_def using True \ by auto + then have "\r. strict_mono r \ MCauchy (\ \ r)" + using \strict_mono r\ by blast + } + moreover + { assume R: ?rhs + have "mtotally_bounded S" + unfolding mtotally_bounded_def + proof (intro strip) + fix \ :: real + assume "\ > 0" + have False if \
: "\K. \finite K; K \ S\ \ \s\S. s \ (\x\K. mball x \)" + proof - + obtain f where f: "\K. \finite K; K \ S\ \ f K \ S \ f K \ (\x\K. mball x \)" + using \
by metis + define \ where "\ \ wfrec less_than (\seq n. f (seq ` {.._eq: "\ n = f (\ ` {.._def]) + have [simp]: "\ n \ S" for n + using wf_less_than + proof (induction n rule: wf_induct_rule) + case (less n) with f show ?case + by (auto simp: \_eq [of n]) + qed + then have "range \ \ S" by blast + have \: "p < n \ \ \ d (\ p) (\ n)" for n p + using f[of "\ ` {.._eq [of n] Ball_def) + then obtain r where "strict_mono r" "MCauchy (\ \ r)" + by (meson R \range \ \ S\) + with \0 < \\ obtain N + where N: "\n n'. \n\N; n'\N\ \ d (\ (r n)) (\ (r n')) < \" + by (force simp: MCauchy_def) + show ?thesis + using N [of N "Suc (r N)"] \strict_mono r\ + by (smt (verit) Suc_le_eq \ le_SucI order_refl strict_mono_imp_increasing) + qed + then show "\K. finite K \ K \ S \ S \ (\x\K. mball x \)" + by blast + qed + } + ultimately show ?thesis + using True by blast + qed +qed (use mtotally_bounded_imp_subset in auto) + + +lemma mtotally_bounded_subset: + "\mtotally_bounded S; T \ S\ \ mtotally_bounded T" + by (meson mtotally_bounded_sequentially order_trans) + +lemma mtotally_bounded_submetric: + assumes "mtotally_bounded S" "S \ T" "T \ M" + shows "Metric_space.mtotally_bounded T d S" +proof - + interpret submetric M d T + by (simp add: Metric_space_axioms assms submetric.intro submetric_axioms.intro) + show ?thesis + using assms + unfolding sub.mtotally_bounded_def mtotally_bounded_def + by (force simp: subset_iff elim!: all_forward ex_forward) +qed + +lemma mtotally_bounded_absolute: + "mtotally_bounded S \ S \ M \ Metric_space.mtotally_bounded S d S " +proof - + have "mtotally_bounded S" if "S \ M" "Metric_space.mtotally_bounded S d S" + proof - + interpret submetric M d S + by (simp add: Metric_space_axioms submetric_axioms.intro submetric_def \S \ M\) + show ?thesis + using that + by (metis MCauchy_submetric Metric_space.mtotally_bounded_sequentially Metric_space_axioms subspace) + qed + moreover have "mtotally_bounded S \ Metric_space.mtotally_bounded S d S" + by (simp add: mtotally_bounded_imp_subset mtotally_bounded_submetric) + ultimately show ?thesis + using mtotally_bounded_imp_subset by blast +qed + +lemma mtotally_bounded_closure_of: + assumes "mtotally_bounded S" + shows "mtotally_bounded (mtopology closure_of S)" +proof - + have "S \ M" + by (simp add: assms mtotally_bounded_imp_subset) + have "mtotally_bounded(mtopology closure_of S)" + unfolding mtotally_bounded_def + proof (intro strip) + fix \::real + assume "\ > 0" + then obtain K where "finite K" "K \ S" and K: "S \ (\x\K. mball x (\/2))" + by (metis assms mtotally_bounded_def half_gt_zero) + have "mtopology closure_of S \ (\x\K. mball x \)" + unfolding metric_closure_of + proof clarsimp + fix x + assume "x \ M" and x: "\r>0. \y\S. y \ M \ d x y < r" + then obtain y where "y \ S" and y: "d x y < \/2" + using \0 < \\ half_gt_zero by blast + then obtain x' where "x' \ K" "y \ mball x' (\/2)" + using K by auto + then have "d x' x < \/2 + \/2" + using triangle y \x \ M\ commute by fastforce + then show "\x'\K. x' \ M \ d x' x < \" + using \K \ S\ \S \ M\ \x' \ K\ by force + qed + then show "\K. finite K \ K \ mtopology closure_of S \ mtopology closure_of S \ (\x\K. mball x \)" + using closure_of_subset_Int \K \ S\ \finite K\ K by fastforce + qed + then show ?thesis + by (simp add: assms inf.absorb2 mtotally_bounded_imp_subset) +qed + +lemma mtotally_bounded_closure_of_eq: + "S \ M \ mtotally_bounded (mtopology closure_of S) \ mtotally_bounded S" + by (metis closure_of_subset mtotally_bounded_closure_of mtotally_bounded_subset topspace_mtopology) + +lemma mtotally_bounded_cauchy_sequence: + assumes "MCauchy \" + shows "mtotally_bounded (range \)" + unfolding MCauchy_def mtotally_bounded_def +proof (intro strip) + fix \::real + assume "\ > 0" + then obtain N where "\n. N \ n \ d (\ N) (\ n) < \" + using assms by (force simp: MCauchy_def) + then have "\m. \n\N. \ n \ M \ \ m \ M \ d (\ n) (\ m) < \" + by (metis MCauchy_def assms mdist_zero nle_le range_subsetD) + then + show "\K. finite K \ K \ range \ \ range \ \ (\x\K. mball x \)" + by (rule_tac x="\ ` {0..N}" in exI) force +qed + +lemma MCauchy_imp_mbounded: + "MCauchy \ \ mbounded (range \)" + by (simp add: mtotally_bounded_cauchy_sequence mtotally_bounded_imp_mbounded) + +subsection\Compactness in metric spaces\ + +lemma Bolzano_Weierstrass_property: + assumes "S \ U" "S \ M" + shows + "(\\::nat\'a. range \ \ S + \ (\l r. l \ U \ strict_mono r \ limitin mtopology (\ \ r) l sequentially)) \ + (\T. T \ S \ infinite T \ U \ mtopology derived_set_of T \ {})" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof clarify + fix T + assume "T \ S" and "infinite T" + and T: "U \ mtopology derived_set_of T = {}" + then obtain \ :: "nat\'a" where "inj \" "range \ \ T" + by (meson infinite_countable_subset) + with L obtain l r where "l \ U" "strict_mono r" + and lr: "limitin mtopology (\ \ r) l sequentially" + by (meson \T \ S\ subset_trans) + then obtain \ where "\ > 0" and \: "\y. y \ T \ y = l \ \ d l y < \" + using T \T \ S\ \S \ M\ + by (force simp: metric_derived_set_of limitin_metric disjoint_iff) + with lr have "\\<^sub>F n in sequentially. \ (r n) \ M \ d (\ (r n)) l < \" + by (auto simp: limitin_metric) + then obtain N where N: "d (\ (r N)) l < \" "d (\ (r (Suc N))) l < \" + using less_le_not_le by (auto simp: eventually_sequentially) + moreover have "\ (r N) \ l \ \ (r (Suc N)) \ l" + by (meson \inj \\ \strict_mono r\ injD n_not_Suc_n strict_mono_eq) + ultimately + show False + using \ \range \ \ T\ commute by fastforce + qed +next + assume R: ?rhs + show ?lhs + proof (intro strip) + fix \ :: "nat \ 'a" + assume "range \ \ S" + show "\l r. l \ U \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" + proof (cases "finite (range \)") + case True + then obtain m where "infinite (\ -` {\ m})" + by (metis image_iff inf_img_fin_dom nat_not_finite) + then obtain r where [iff]: "strict_mono r" and r: "\n::nat. r n \ \ -` {\ m}" + using infinite_enumerate by blast + have [iff]: "\ m \ U" "\ m \ M" + using \range \ \ S\ assms by blast+ + show ?thesis + proof (intro conjI exI) + show "limitin mtopology (\ \ r) (\ m) sequentially" + using r by (simp add: limitin_metric) + qed auto + next + case False + then obtain l where "l \ U" and l: "l \ mtopology derived_set_of (range \)" + by (meson R \range \ \ S\ disjoint_iff) + then obtain g where g: "\\. \>0 \ \ (g \) \ l \ d l (\ (g \)) < \" + by (simp add: metric_derived_set_of) metis + have "range \ \ M" + using \range \ \ S\ assms by auto + have "l \ M" + using l metric_derived_set_of by auto + define E where \\a construction to ensure monotonicity\ + "E \ \rec n. insert (inverse (Suc n)) ((\i. d l (\ i)) ` (\k wfrec less_than (\rec n. g (Min (E rec n)))" + have "(\kk (r n)) > 0" for n + using wf_less_than + proof (induction n rule: wf_induct_rule) + case (less n) + then have *: "Min (E r n) > 0" + using \l \ M\ \range \ \ M\ by (auto simp: E_def image_subset_iff) + show ?case + using g [OF *] r_eq [of n] + by (metis \l \ M\ \range \ \ M\ mdist_pos_less range_subsetD) + qed + then have non_l: "\ (r n) \ l" for n + using \range \ \ M\ mdist_pos_eq by blast + have Min_pos: "Min (E r n) > 0" for n + using dl_pos \l \ M\ \range \ \ M\ by (auto simp: E_def image_subset_iff) + have d_small: "d (\(r n)) l < inverse(Suc n)" for n + proof - + have "d (\(r n)) l < Min (E r n)" + by (simp add: \0 < Min (E r n)\ commute g r_eq) + also have "... \ inverse(Suc n)" + by (simp add: E_def) + finally show ?thesis . + qed + have d_lt_d: "d l (\ (r n)) < d l (\ i)" if \
: "p < n" "i \ r p" "\ i \ l" for i p n + proof - + have 1: "d l (\ i) \ E r n" + using \
\l \ M\ \range \ \ M\ + by (force simp: E_def image_subset_iff image_iff) + have "d l (\ (g (Min (E r n)))) < Min (E r n)" + by (rule conjunct2 [OF g [OF Min_pos]]) + also have "Min (E r n) \ d l (\ i)" + using 1 unfolding E_def by (force intro!: Min.coboundedI) + finally show ?thesis + by (simp add: r_eq) + qed + have r: "r p < r n" if "p < n" for p n + using d_lt_d [OF that] non_l by (meson linorder_not_le order_less_irrefl) + show ?thesis + proof (intro exI conjI) + show "strict_mono r" + by (simp add: r strict_monoI) + show "limitin mtopology (\ \ r) l sequentially" + unfolding limitin_metric + proof (intro conjI strip \l \ M\) + fix \ :: real + assume "\ > 0" + then have "\\<^sub>F n in sequentially. inverse(Suc n) < \" + using Archimedean_eventually_inverse by auto + then show "\\<^sub>F n in sequentially. (\ \ r) n \ M \ d ((\ \ r) n) l < \" + by (smt (verit) \range \ \ M\ commute comp_apply d_small eventually_mono range_subsetD) + qed + qed (use \l \ U\ in auto) + qed + qed +qed + +subsubsection \More on Bolzano Weierstrass\ + +lemma Bolzano_Weierstrass_A: + assumes "compactin mtopology S" "T \ S" "infinite T" + shows "S \ mtopology derived_set_of T \ {}" + by (simp add: assms compactin_imp_Bolzano_Weierstrass) + +lemma Bolzano_Weierstrass_B: + fixes \ :: "nat \ 'a" + assumes "S \ M" "range \ \ S" + and "\T. \T \ S \ infinite T\ \ S \ mtopology derived_set_of T \ {}" + shows "\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" + using Bolzano_Weierstrass_property assms by blast + +lemma Bolzano_Weierstrass_C: + assumes "S \ M" + assumes "\\:: nat \ 'a. range \ \ S \ + (\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially)" + shows "mtotally_bounded S" + unfolding mtotally_bounded_sequentially + by (metis convergent_imp_MCauchy assms image_comp image_mono subset_UNIV subset_trans) + +lemma Bolzano_Weierstrass_D: + assumes "S \ M" "S \ \\" and opeU: "\U. U \ \ \ openin mtopology U" + assumes \
: "(\\::nat\'a. range \ \ S + \ (\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially))" + shows "\\>0. \x \ S. \U \ \. mball x \ \ U" +proof (rule ccontr) + assume "\ (\\>0. \x \ S. \U \ \. mball x \ \ U)" + then have "\n. \x\S. \U\\. \ mball x (inverse (Suc n)) \ U" + by simp + then obtain \ where "\n. \ n \ S" + and \: "\n U. U \ \ \ \ mball (\ n) (inverse (Suc n)) \ U" + by metis + then obtain l r where "l \ S" "strict_mono r" + and lr: "limitin mtopology (\ \ r) l sequentially" + by (meson \
image_subsetI) + with \S \ \\\ obtain B where "l \ B" "B \ \" + by auto + then obtain \ where "\ > 0" and \: "\z. \z \ M; d z l < \\ \ z \ B" + by (metis opeU [OF \B \ \\] commute in_mball openin_mtopology subset_iff) + then have "\\<^sub>F n in sequentially. \ (r n) \ M \ d (\ (r n)) l < \/2" + using lr half_gt_zero unfolding limitin_metric o_def by blast + moreover have "\\<^sub>F n in sequentially. inverse (real (Suc n)) < \/2" + using Archimedean_eventually_inverse \0 < \\ half_gt_zero by blast + ultimately obtain n where n: "d (\ (r n)) l < \/2" "inverse (real (Suc n)) < \/2" + by (smt (verit, del_insts) eventually_sequentially le_add1 le_add2) + have "x \ B" if "d (\ (r n)) x < inverse (Suc(r n))" "x \ M" for x + proof - + have rle: "inverse (real (Suc (r n))) \ inverse (real (Suc n))" + using \strict_mono r\ strict_mono_imp_increasing by auto + have "d x l \ d (\ (r n)) x + d (\ (r n)) l" + using that by (metis triangle \\n. \ n \ S\ \l \ S\ \S \ M\ commute subsetD) + also have "... < \" + using that n rle by linarith + finally show ?thesis + by (simp add: \ that) + qed + then show False + using \ [of B "r n"] by (simp add: \B \ \\ subset_iff) +qed + + +lemma Bolzano_Weierstrass_E: + assumes "mtotally_bounded S" "S \ M" + and S: "\\. \\U. U \ \ \ openin mtopology U; S \ \\\ \ \\>0. \x \ S. \U \ \. mball x \ \ U" + shows "compactin mtopology S" +proof (clarsimp simp: compactin_def assms) + fix \ :: "'a set set" + assume \: "\x\\. openin mtopology x" and "S \ \\" + then obtain \ where "\>0" and \: "\x. x \ S \ \U \ \. mball x \ \ U" + by (metis S) + then obtain f where f: "\x. x \ S \ f x \ \ \ mball x \ \ f x" + by metis + then obtain K where "finite K" "K \ S" and K: "S \ (\x\K. mball x \)" + by (metis \0 < \\ \mtotally_bounded S\ mtotally_bounded_def) + show "\\. finite \ \ \ \ \ \ S \ \\" + proof (intro conjI exI) + show "finite (f ` K)" + by (simp add: \finite K\) + show "f ` K \ \" + using \K \ S\ f by blast + show "S \ \(f ` K)" + using K \K \ S\ by (force dest: f) + qed +qed + + +lemma compactin_eq_Bolzano_Weierstrass: + "compactin mtopology S \ + S \ M \ (\T. T \ S \ infinite T \ S \ mtopology derived_set_of T \ {})" + using Bolzano_Weierstrass_C Bolzano_Weierstrass_D Bolzano_Weierstrass_E + by (smt (verit, del_insts) Bolzano_Weierstrass_property compactin_imp_Bolzano_Weierstrass compactin_subspace subset_refl topspace_mtopology) + +lemma compactin_sequentially: + shows "compactin mtopology S \ + S \ M \ + ((\\::nat\'a. range \ \ S + \ (\l r. l \ S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially)))" + by (metis Bolzano_Weierstrass_property compactin_eq_Bolzano_Weierstrass subset_refl) + +lemma compactin_imp_mtotally_bounded: + "compactin mtopology S \ mtotally_bounded S" + by (simp add: Bolzano_Weierstrass_C compactin_sequentially) + +lemma lebesgue_number: + "\compactin mtopology S; S \ \\; \U. U \ \ \ openin mtopology U\ + \ \\>0. \x \ S. \U \ \. mball x \ \ U" + by (simp add: Bolzano_Weierstrass_D compactin_sequentially) + +lemma compact_space_sequentially: + "compact_space mtopology \ + (\\::nat\'a. range \ \ M + \ (\l r. l \ M \ strict_mono r \ limitin mtopology (\ \ r) l sequentially))" + by (simp add: compact_space_def compactin_sequentially) + +lemma compact_space_eq_Bolzano_Weierstrass: + "compact_space mtopology \ + (\S. S \ M \ infinite S \ mtopology derived_set_of S \ {})" + using Int_absorb1 [OF derived_set_of_subset_topspace [of mtopology]] + by (force simp: compact_space_def compactin_eq_Bolzano_Weierstrass) + +lemma compact_space_nest: + "compact_space mtopology \ + (\C. (\n::nat. closedin mtopology (C n)) \ (\n. C n \ {}) \ decseq C \ \(range C) \ {})" + (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof clarify + fix C :: "nat \ 'a set" + assume "\n. closedin mtopology (C n)" + and "\n. C n \ {}" + and "decseq C" + and "\ (range C) = {}" + then obtain K where K: "finite K" "\(C ` K) = {}" + by (metis L compact_space_imp_nest) + then obtain k where "K \ {..k}" + using finite_nat_iff_bounded_le by auto + then have "C k \ \(C ` K)" + using \decseq C\ by (auto simp:decseq_def) + then show False + by (simp add: K \\n. C n \ {}\) + qed +next + assume R [rule_format]: ?rhs + show ?lhs + unfolding compact_space_sequentially + proof (intro strip) + fix \ :: "nat \ 'a" + assume \: "range \ \ M" + have "mtopology closure_of \ ` {n..} \ {}" for n + using \range \ \ M\ by (auto simp: closure_of_eq_empty image_subset_iff) + moreover have "decseq (\n. mtopology closure_of \ ` {n..})" + using closure_of_mono image_mono by (smt (verit) atLeast_subset_iff decseq_def) + ultimately obtain l where l: "\n. l \ mtopology closure_of \ ` {n..}" + using R [of "\n. mtopology closure_of (\ ` {n..})"] by auto + then have "l \ M" and "\n. \r>0. \k\n. \ k \ M \ d l (\ k) < r" + using metric_closure_of by fastforce+ + then obtain f where f: "\n r. r>0 \ f n r \ n \ \ (f n r) \ M \ d l (\ (f n r)) < r" + by metis + define r where "r = rec_nat (f 0 1) (\n rec. (f (Suc rec) (inverse (Suc (Suc n)))))" + have r: "d l (\(r n)) < inverse(Suc n)" for n + by (induction n) (auto simp: rec_nat_0_imp [OF r_def] rec_nat_Suc_imp [OF r_def] f) + have "r n < r(Suc n)" for n + by (simp add: Suc_le_lessD f r_def) + then have "strict_mono r" + by (simp add: strict_mono_Suc_iff) + moreover have "limitin mtopology (\ \ r) l sequentially" + proof (clarsimp simp: limitin_metric \l \ M\) + fix \ :: real + assume "\ > 0" + then have "(\\<^sub>F n in sequentially. inverse (real (Suc n)) < \)" + using Archimedean_eventually_inverse by blast + then show "\\<^sub>F n in sequentially. \ (r n) \ M \ d (\ (r n)) l < \" + by eventually_elim (metis commute \range \ \ M\ order_less_trans r range_subsetD) + qed + ultimately show "\l r. l \ M \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" + using \l \ M\ by blast + qed +qed + + +lemma (in discrete_metric) mcomplete_discrete_metric: "disc.mcomplete" +proof (clarsimp simp: disc.mcomplete_def) + fix \ :: "nat \ 'a" + assume "disc.MCauchy \" + then obtain N where "\n. N \ n \ \ N = \ n" + unfolding disc.MCauchy_def by (metis dd_def dual_order.refl order_less_irrefl zero_less_one) + moreover have "range \ \ M" + using \disc.MCauchy \\ disc.MCauchy_def by blast + ultimately have "limitin disc.mtopology \ (\ N) sequentially" + by (metis disc.limit_metric_sequentially disc.zero range_subsetD) + then show "\x. limitin disc.mtopology \ x sequentially" .. +qed + +lemma compact_space_imp_mcomplete: "compact_space mtopology \ mcomplete" + by (simp add: compact_space_nest mcomplete_nest) + +lemma (in submetric) compactin_imp_mcomplete: + "compactin mtopology A \ sub.mcomplete" + by (simp add: compactin_subspace mtopology_submetric sub.compact_space_imp_mcomplete) + +lemma (in submetric) mcomplete_imp_closedin: + assumes "sub.mcomplete" + shows "closedin mtopology A" +proof - + have "l \ A" + if "range \ \ A" and l: "limitin mtopology \ l sequentially" + for \ :: "nat \ 'a" and l + proof - + have "sub.MCauchy \" + using convergent_imp_MCauchy subset that by (force simp: MCauchy_submetric) + then have "limitin sub.mtopology \ l sequentially" + using assms unfolding sub.mcomplete_def + using l limitin_metric_unique limitin_submetric_iff trivial_limit_sequentially by blast + then show ?thesis + using limitin_submetric_iff by blast + qed + then show ?thesis + using metric_closedin_iff_sequentially_closed subset by auto +qed + +lemma (in submetric) closedin_eq_mcomplete: + "mcomplete \ (closedin mtopology A \ sub.mcomplete)" + using closedin_mcomplete_imp_mcomplete mcomplete_imp_closedin by blast + +lemma compact_space_eq_mcomplete_mtotally_bounded: + "compact_space mtopology \ mcomplete \ mtotally_bounded M" + by (meson Bolzano_Weierstrass_C compact_space_imp_mcomplete compact_space_sequentially limitin_mspace + mcomplete_alt mtotally_bounded_sequentially subset_refl) + + +lemma compact_closure_of_imp_mtotally_bounded: + "\compactin mtopology (mtopology closure_of S); S \ M\ + \ mtotally_bounded S" + using compactin_imp_mtotally_bounded mtotally_bounded_closure_of_eq by blast + +lemma mtotally_bounded_eq_compact_closure_of: + assumes "mcomplete" + shows "mtotally_bounded S \ S \ M \ compactin mtopology (mtopology closure_of S)" + (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + unfolding compactin_subspace + proof (intro conjI) + show "S \ M" + using L by (simp add: mtotally_bounded_imp_subset) + show "mtopology closure_of S \ topspace mtopology" + by (simp add: \S \ M\ closure_of_minimal) + then have MSM: "mtopology closure_of S \ M" + by auto + interpret S: submetric M d "mtopology closure_of S" + proof qed (use MSM in auto) + have "S.sub.mtotally_bounded (mtopology closure_of S)" + using L mtotally_bounded_absolute mtotally_bounded_closure_of by blast + then + show "compact_space (subtopology mtopology (mtopology closure_of S))" + using S.closedin_mcomplete_imp_mcomplete S.mtopology_submetric S.sub.compact_space_eq_mcomplete_mtotally_bounded assms by force + qed +qed (auto simp: compact_closure_of_imp_mtotally_bounded) + + + +lemma compact_closure_of_eq_Bolzano_Weierstrass: + "compactin mtopology (mtopology closure_of S) \ + (\T. infinite T \ T \ S \ T \ M \ mtopology derived_set_of T \ {})" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof (intro strip) + fix T + assume T: "infinite T \ T \ S \ T \ M" + show "mtopology derived_set_of T \ {}" + proof (intro compact_closure_of_imp_Bolzano_Weierstrass) + show "compactin mtopology (mtopology closure_of S)" + by (simp add: L) + qed (use T in auto) + qed +next + have "compactin mtopology (mtopology closure_of S)" + if \
: "\T. \infinite T; T \ S\ \ mtopology derived_set_of T \ {}" and "S \ M" for S + unfolding compactin_sequentially + proof (intro conjI strip) + show MSM: "mtopology closure_of S \ M" + using closure_of_subset_topspace by fastforce + fix \ :: "nat \ 'a" + assume \: "range \ \ mtopology closure_of S" + then have "\y \ S. d (\ n) y < inverse(Suc n)" for n + by (simp add: metric_closure_of image_subset_iff) (metis inverse_Suc of_nat_Suc) + then obtain \ where \: "\n. \ n \ S \ d (\ n) (\ n) < inverse(Suc n)" + by metis + then have "range \ \ S" + by blast + moreover + have *: "\T. T \ S \ infinite T \ mtopology closure_of S \ mtopology derived_set_of T \ {}" + using "\
"(1) derived_set_of_mono derived_set_of_subset_closure_of by fastforce + moreover have "S \ mtopology closure_of S" + by (simp add: \S \ M\ closure_of_subset) + ultimately obtain l r where lr: + "l \ mtopology closure_of S" "strict_mono r" "limitin mtopology (\ \ r) l sequentially" + using Bolzano_Weierstrass_property \S \ M\ by metis + then have "l \ M" + using limitin_mspace by blast + have dr_less: "d ((\ \ r) n) ((\ \ r) n) < inverse(Suc n)" for n + proof - + have "d ((\ \ r) n) ((\ \ r) n) < inverse(Suc (r n))" + using \ by auto + also have "... \ inverse(Suc n)" + using lr strict_mono_imp_increasing by auto + finally show ?thesis . + qed + have "limitin mtopology (\ \ r) l sequentially" + unfolding limitin_metric + proof (intro conjI strip) + show "l \ M" + using limitin_mspace lr by blast + fix \ :: real + assume "\ > 0" + then have "\\<^sub>F n in sequentially. (\ \ r) n \ M \ d ((\ \ r) n) l < \/2" + using lr half_gt_zero limitin_metric by blast + moreover have "\\<^sub>F n in sequentially. inverse (real (Suc n)) < \/2" + using Archimedean_eventually_inverse \0 < \\ half_gt_zero by blast + then have "\\<^sub>F n in sequentially. d ((\ \ r) n) ((\ \ r) n) < \/2" + by eventually_elim (smt (verit, del_insts) dr_less) + ultimately have "\\<^sub>F n in sequentially. d ((\ \ r) n) l < \/2 + \/2" + by eventually_elim (smt (verit) triangle \l \ M\ MSM \ comp_apply order_trans range_subsetD) + then show "\\<^sub>F n in sequentially. (\ \ r) n \ M \ d ((\ \ r) n) l < \" + apply eventually_elim + using \mtopology closure_of S \ M\ \ by auto + qed + with lr show "\l r. l \ mtopology closure_of S \ strict_mono r \ limitin mtopology (\ \ r) l sequentially" + by blast + qed + then show "?rhs \ ?lhs" + by (metis Int_subset_iff closure_of_restrict inf_le1 topspace_mtopology) +qed + +end + +lemma (in discrete_metric) mtotally_bounded_discrete_metric: + "disc.mtotally_bounded S \ finite S \ S \ M" (is "?lhs=?rhs") +proof + assume L: ?lhs + show ?rhs + proof + show "finite S" + by (metis (no_types) L closure_of_subset_Int compactin_discrete_topology disc.mtotally_bounded_eq_compact_closure_of + disc.topspace_mtopology discrete_metric.mcomplete_discrete_metric inf.absorb_iff2 mtopology_discrete_metric finite_subset) + show "S \ M" + by (simp add: L disc.mtotally_bounded_imp_subset) + qed +qed (simp add: disc.finite_imp_mtotally_bounded) + + +context Metric_space +begin + +lemma derived_set_of_infinite_openin_metric: + "mtopology derived_set_of S = + {x \ M. \U. x \ U \ openin mtopology U \ infinite(S \ U)}" + by (simp add: derived_set_of_infinite_openin Hausdorff_space_mtopology) + +lemma derived_set_of_infinite_1: + assumes "infinite (S \ mball x \)" + shows "infinite (S \ mcball x \)" + by (meson Int_mono assms finite_subset mball_subset_mcball subset_refl) + +lemma derived_set_of_infinite_2: + assumes "openin mtopology U" "\\. 0 < \ \ infinite (S \ mcball x \)" and "x \ U" + shows "infinite (S \ U)" + by (metis assms openin_mtopology_mcball finite_Int inf.absorb_iff2 inf_assoc) + +lemma derived_set_of_infinite_mball: + "mtopology derived_set_of S = {x \ M. \e>0. infinite(S \ mball x e)}" + unfolding derived_set_of_infinite_openin_metric + by (meson centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2) + +lemma derived_set_of_infinite_mcball: + "mtopology derived_set_of S = {x \ M. \e>0. infinite(S \ mcball x e)}" + unfolding derived_set_of_infinite_openin_metric + by (meson centre_in_mball_iff openin_mball derived_set_of_infinite_1 derived_set_of_infinite_2) + +end + +subsection\Continuous functions on metric spaces\ + +context Metric_space +begin + +lemma continuous_map_to_metric: + "continuous_map X mtopology f \ + (\x \ topspace X. \\>0. \U. openin X U \ x \ U \ (\y\U. f y \ mball (f x) \))" + (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + unfolding continuous_map_eq_topcontinuous_at topcontinuous_at_def + by (metis centre_in_mball_iff openin_mball topspace_mtopology) +next + assume R: ?rhs + then have "\x\topspace X. f x \ M" + by (meson gt_ex in_mball) + moreover + have "\x V. \x \ topspace X; openin mtopology V; f x \ V\ \ \U. openin X U \ x \ U \ (\y\U. f y \ V)" + unfolding openin_mtopology by (metis Int_iff R inf.orderE) + ultimately + show ?lhs + by (simp add: continuous_map_eq_topcontinuous_at topcontinuous_at_def) +qed + +lemma continuous_map_from_metric: + "continuous_map mtopology X f \ + f ` M \ topspace X \ + (\a \ M. \U. openin X U \ f a \ U \ (\r>0. \x. x \ M \ d a x < r \ f x \ U))" +proof (cases "f ` M \ topspace X") + case True + then show ?thesis + by (fastforce simp: continuous_map openin_mtopology subset_eq) +next + case False + then show ?thesis + by (simp add: continuous_map_eq_image_closure_subset_gen) +qed + +text \An abstract formulation, since the limits do not have to be sequential\ +lemma continuous_map_uniform_limit: + assumes contf: "\\<^sub>F \ in F. continuous_map X mtopology (f \)" + and dfg: "\\. 0 < \ \ \\<^sub>F \ in F. \x \ topspace X. g x \ M \ d (f \ x) (g x) < \" + and nontriv: "\ trivial_limit F" + shows "continuous_map X mtopology g" + unfolding continuous_map_to_metric +proof (intro strip) + fix x and \::real + assume "x \ topspace X" and "\ > 0" + then obtain \ where k: "continuous_map X mtopology (f \)" + and gM: "\x \ topspace X. g x \ M" + and third: "\x \ topspace X. d (f \ x) (g x) < \/3" + using eventually_conj [OF contf] contf dfg [of "\/3"] eventually_happens' [OF nontriv] + by (smt (verit, ccfv_SIG) zero_less_divide_iff) + then obtain U where U: "openin X U" "x \ U" and Uthird: "\y\U. d (f \ y) (f \ x) < \/3" + unfolding continuous_map_to_metric + by (metis \0 < \\ \x \ topspace X\ commute divide_pos_pos in_mball zero_less_numeral) + have f_inM: "f \ y \ M" if "y\U" for y + using U k openin_subset that by (fastforce simp: continuous_map_def) + have "d (g y) (g x) < \" if "y\U" for y + proof - + have "g y \ M" + using U gM openin_subset that by blast + have "d (g y) (g x) \ d (g y) (f \ x) + d (f \ x) (g x)" + by (simp add: U \g y \ M\ \x \ topspace X\ f_inM gM triangle) + also have "\ \ d (g y) (f \ y) + d (f \ y) (f \ x) + d (f \ x) (g x)" + by (simp add: U \g y \ M\ commute f_inM that triangle') + also have "\ < \/3 + \/3 + \/3" + by (smt (verit) U(1) Uthird \x \ topspace X\ commute openin_subset subsetD that third) + finally show ?thesis by simp + qed + with U gM show "\U. openin X U \ x \ U \ (\y\U. g y \ mball (g x) \)" + by (metis commute in_mball in_mono openin_subset) +qed + + +lemma continuous_map_uniform_limit_alt: + assumes contf: "\\<^sub>F \ in F. continuous_map X mtopology (f \)" + and gim: "g ` (topspace X) \ M" + and dfg: "\\. 0 < \ \ \\<^sub>F \ in F. \x \ topspace X. d (f \ x) (g x) < \" + and nontriv: "\ trivial_limit F" + shows "continuous_map X mtopology g" +proof (rule continuous_map_uniform_limit [OF contf]) + fix \ :: real + assume "\ > 0" + with gim dfg show "\\<^sub>F \ in F. \x\topspace X. g x \ M \ d (f \ x) (g x) < \" + by (simp add: image_subset_iff) +qed (use nontriv in auto) + + +lemma continuous_map_uniformly_Cauchy_limit: + assumes "mcomplete" + assumes contf: "\\<^sub>F n in sequentially. continuous_map X mtopology (f n)" + and Cauchy': "\\. \ > 0 \ \N. \m n x. N \ m \ N \ n \ x \ topspace X \ d (f m x) (f n x) < \" + obtains g where + "continuous_map X mtopology g" + "\\. 0 < \ \ \\<^sub>F n in sequentially. \x\topspace X. d (f n x) (g x) < \" +proof - + have "\x. x \ topspace X \ \l. limitin mtopology (\n. f n x) l sequentially" + using \mcomplete\ [unfolded mcomplete, rule_format] assms + by (smt (verit) contf continuous_map_def eventually_mono topspace_mtopology) + then obtain g where g: "\x. x \ topspace X \ limitin mtopology (\n. f n x) (g x) sequentially" + by metis + show thesis + proof + show "\\<^sub>F n in sequentially. \x\topspace X. d (f n x) (g x) < \" + if "\ > 0" for \ :: real + proof - + obtain N where N: "\m n x. \N \ m; N \ n; x \ topspace X\ \ d (f m x) (f n x) < \/2" + by (meson Cauchy' \0 < \\ half_gt_zero) + obtain P where P: "\n x. \n \ P; x \ topspace X\ \ f n x \ M" + using contf by (auto simp: eventually_sequentially continuous_map_def) + show ?thesis + proof (intro eventually_sequentiallyI strip) + fix n x + assume "max N P \ n" and x: "x \ topspace X" + obtain L where "g x \ M" and L: "\n\L. f n x \ M \ d (f n x) (g x) < \/2" + using g [OF x] \\ > 0\ unfolding limitin_metric + by (metis (no_types, lifting) eventually_sequentially half_gt_zero) + define n' where "n' \ Max{L,N,P}" + have L': "\m \ n'. f m x \ M \ d (f m x) (g x) < \/2" + using L by (simp add: n'_def) + moreover + have "d (f n x) (f n' x) < \/2" + using N [of n n' x] \max N P \ n\ n'_def x by fastforce + ultimately have "d (f n x) (g x) < \/2 + \/2" + by (smt (verit, ccfv_SIG) P \g x \ M\ \max N P \ n\ le_refl max.bounded_iff mdist_zero triangle' x) + then show "d (f n x) (g x) < \" by simp + qed + qed + then show "continuous_map X mtopology g" + by (smt (verit, del_insts) eventually_mono g limitin_mspace trivial_limit_sequentially continuous_map_uniform_limit [OF contf]) + qed +qed + +lemma metric_continuous_map: + assumes "Metric_space M' d'" + shows + "continuous_map mtopology (Metric_space.mtopology M' d') f \ + f ` M \ M' \ (\a \ M. \\>0. \\>0. (\x. x \ M \ d a x < \ \ d' (f a) (f x) < \))" + (is "?lhs = ?rhs") +proof - + interpret M': Metric_space M' d' + by (simp add: assms) + show ?thesis + proof + assume L: ?lhs + show ?rhs + proof (intro conjI strip) + show "f ` M \ M'" + using L by (auto simp: continuous_map_def) + fix a and \ :: real + assume "a \ M" and "\ > 0" + then have "openin mtopology {x \ M. f x \ M'.mball (f a) \}" "f a \ M'" + using L unfolding continuous_map_def by fastforce+ + then obtain \ where "\ > 0" "mball a \ \ {x \ M. f x \ M' \ d' (f a) (f x) < \}" + using \0 < \\ \a \ M\ openin_mtopology by auto + then show "\\>0. \x. x \ M \ d a x < \ \ d' (f a) (f x) < \" + using \a \ M\ in_mball by blast + qed + next + assume R: ?rhs + show ?lhs + unfolding continuous_map_def + proof (intro conjI strip) + fix U + assume "openin M'.mtopology U" + then show "openin mtopology {x \ topspace mtopology. f x \ U}" + apply (simp add: continuous_map_def openin_mtopology M'.openin_mtopology subset_iff) + by (metis R image_subset_iff) + qed (use R in auto) + qed +qed + +end (*Metric_space*) + +end + diff --git a/src/HOL/Analysis/Abstract_Topological_Spaces.thy b/src/HOL/Analysis/Abstract_Topological_Spaces.thy --- a/src/HOL/Analysis/Abstract_Topological_Spaces.thy +++ b/src/HOL/Analysis/Abstract_Topological_Spaces.thy @@ -1,4437 +1,5309 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Various Forms of Topological Spaces\ theory Abstract_Topological_Spaces imports Lindelof_Spaces Locally Sum_Topology FSigma begin subsection\Connected topological spaces\ lemma connected_space_eq_frontier_eq_empty: "connected_space X \ (\S. S \ topspace X \ X frontier_of S = {} \ S = {} \ S = topspace X)" by (meson clopenin_eq_frontier_of connected_space_clopen_in) lemma connected_space_frontier_eq_empty: "connected_space X \ S \ topspace X \ (X frontier_of S = {} \ S = {} \ S = topspace X)" by (meson connected_space_eq_frontier_eq_empty frontier_of_empty frontier_of_topspace) lemma connectedin_eq_subset_separated_union: "connectedin X C \ C \ topspace X \ (\S T. separatedin X S T \ C \ S \ T \ C \ S \ C \ T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using connectedin_subset_topspace connectedin_subset_separated_union by blast next assume ?rhs then show ?lhs by (metis closure_of_subset connectedin_separation dual_order.eq_iff inf.orderE separatedin_def sup.boundedE) qed lemma connectedin_clopen_cases: "\connectedin X C; closedin X T; openin X T\ \ C \ T \ disjnt C T" by (metis Diff_eq_empty_iff Int_empty_right clopenin_eq_frontier_of connectedin_Int_frontier_of disjnt_def) lemma connected_space_quotient_map_image: "\quotient_map X X' q; connected_space X\ \ connected_space X'" by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) lemma connected_space_retraction_map_image: "\retraction_map X X' r; connected_space X\ \ connected_space X'" using connected_space_quotient_map_image retraction_imp_quotient_map by blast lemma connectedin_imp_perfect_gen: assumes X: "t1_space X" and S: "connectedin X S" and nontriv: "\a. S = {a}" shows "S \ X derived_set_of S" unfolding derived_set_of_def proof (intro subsetI CollectI conjI strip) show XS: "x \ topspace X" if "x \ S" for x using that S connectedin by fastforce show "\y. y \ x \ y \ S \ y \ T" if "x \ S" and "x \ T \ openin X T" for x T proof - have opeXx: "openin X (topspace X - {x})" by (meson X openin_topspace t1_space_openin_delete_alt) moreover have "S \ T \ (topspace X - {x})" using XS that(2) by auto moreover have "(topspace X - {x}) \ S \ {}" by (metis Diff_triv S connectedin double_diff empty_subsetI inf_commute insert_subsetI nontriv that(1)) ultimately show ?thesis using that connectedinD [OF S, of T "topspace X - {x}"] by blast qed qed lemma connectedin_imp_perfect: "\Hausdorff_space X; connectedin X S; \a. S = {a}\ \ S \ X derived_set_of S" by (simp add: Hausdorff_imp_t1_space connectedin_imp_perfect_gen) subsection\The notion of "separated between" (complement of "connected between)"\ definition separated_between where "separated_between X S T \ \U V. openin X U \ openin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V" lemma separated_between_alt: "separated_between X S T \ (\U V. closedin X U \ closedin X V \ U \ V = topspace X \ disjnt U V \ S \ U \ T \ V)" unfolding separated_between_def by (metis separatedin_open_sets separation_closedin_Un_gen subtopology_topspace separatedin_closed_sets separation_openin_Un_gen) lemma separated_between: "separated_between X S T \ (\U. closedin X U \ openin X U \ S \ U \ T \ topspace X - U)" unfolding separated_between_def closedin_def disjnt_def by (smt (verit, del_insts) Diff_cancel Diff_disjoint Diff_partition Un_Diff Un_Diff_Int openin_subset) lemma separated_between_mono: "\separated_between X S T; S' \ S; T' \ T\ \ separated_between X S' T'" by (meson order.trans separated_between) lemma separated_between_refl: "separated_between X S S \ S = {}" unfolding separated_between_def by (metis Un_empty_right disjnt_def disjnt_empty2 disjnt_subset2 disjnt_sym le_iff_inf openin_empty openin_topspace) lemma separated_between_sym: "separated_between X S T \ separated_between X T S" by (metis disjnt_sym separated_between_alt sup_commute) lemma separated_between_imp_subset: "separated_between X S T \ S \ topspace X \ T \ topspace X" by (metis le_supI1 le_supI2 separated_between_def) lemma separated_between_empty: "(separated_between X {} S \ S \ topspace X) \ (separated_between X S {} \ S \ topspace X)" by (metis Diff_empty bot.extremum closedin_empty openin_empty separated_between separated_between_imp_subset separated_between_sym) lemma separated_between_Un: "separated_between X S (T \ U) \ separated_between X S T \ separated_between X S U" by (auto simp: separated_between) lemma separated_between_Un': "separated_between X (S \ T) U \ separated_between X S U \ separated_between X T U" by (simp add: separated_between_Un separated_between_sym) lemma separated_between_imp_disjoint: "separated_between X S T \ disjnt S T" by (meson disjnt_iff separated_between_def subsetD) lemma separated_between_imp_separatedin: "separated_between X S T \ separatedin X S T" by (meson separated_between_def separatedin_mono separatedin_open_sets) lemma separated_between_full: assumes "S \ T = topspace X" shows "separated_between X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" proof - have "separated_between X S T \ separatedin X S T" by (simp add: separated_between_imp_separatedin) then show ?thesis unfolding separated_between_def by (metis assms separation_closedin_Un_gen separation_openin_Un_gen subset_refl subtopology_topspace) qed lemma separated_between_eq_separatedin: "S \ T = topspace X \ (separated_between X S T \ separatedin X S T)" by (simp add: separated_between_full separatedin_full) lemma separated_between_pointwise_left: assumes "compactin X S" shows "separated_between X S T \ (S = {} \ T \ topspace X) \ (\x \ S. separated_between X {x} T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using separated_between_imp_subset separated_between_mono by fastforce next assume R: ?rhs then have "T \ topspace X" by (meson equals0I separated_between_imp_subset) show ?lhs proof - obtain U where U: "\x \ S. openin X (U x)" "\x \ S. \V. openin X V \ U x \ V = topspace X \ disjnt (U x) V \ {x} \ U x \ T \ V" using R unfolding separated_between_def by metis then have "S \ \(U ` S)" by blast then obtain K where "finite K" "K \ S" and K: "S \ (\i\K. U i)" using assms U unfolding compactin_def by (smt (verit) finite_subset_image imageE) show ?thesis unfolding separated_between proof (intro conjI exI) have "\x. x \ K \ closedin X (U x)" by (smt (verit) \K \ S\ Diff_cancel U(2) Un_Diff Un_Diff_Int disjnt_def openin_closedin_eq subsetD) then show "closedin X (\ (U ` K))" by (metis (mono_tags, lifting) \finite K\ closedin_Union finite_imageI image_iff) show "openin X (\ (U ` K))" using U(1) \K \ S\ by blast show "S \ \ (U ` K)" by (simp add: K) have "\x i. \x \ T; i \ K; x \ U i\ \ False" by (meson U(2) \K \ S\ disjnt_iff subsetD) then show "T \ topspace X - \ (U ` K)" using \T \ topspace X\ by auto qed qed qed lemma separated_between_pointwise_right: "compactin X T \ separated_between X S T \ (T = {} \ S \ topspace X) \ (\y \ T. separated_between X S {y})" by (meson separated_between_pointwise_left separated_between_sym) lemma separated_between_closure_of: "S \ topspace X \ separated_between X (X closure_of S) T \ separated_between X S T" by (meson closure_of_minimal_eq separated_between_alt) lemma separated_between_closure_of': "T \ topspace X \ separated_between X S (X closure_of T) \ separated_between X S T" by (meson separated_between_closure_of separated_between_sym) lemma separated_between_closure_of_eq: "separated_between X S T \ S \ topspace X \ separated_between X (X closure_of S) T" by (metis separated_between_closure_of separated_between_imp_subset) lemma separated_between_closure_of_eq': "separated_between X S T \ T \ topspace X \ separated_between X S (X closure_of T)" by (metis separated_between_closure_of' separated_between_imp_subset) lemma separated_between_frontier_of_eq': "separated_between X S T \ T \ topspace X \ disjnt S T \ separated_between X S (X frontier_of T)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis interior_of_union_frontier_of separated_between_Un separated_between_closure_of_eq' separated_between_imp_disjoint) next assume R: ?rhs then obtain U where U: "closedin X U" "openin X U" "S \ U" "X closure_of T - X interior_of T \ topspace X - U" by (metis frontier_of_def separated_between) show ?lhs proof (rule separated_between_mono [of _ S "X closure_of T"]) have "separated_between X S T" unfolding separated_between proof (intro conjI exI) show "S \ U - T" "T \ topspace X - (U - T)" using R U(3) by (force simp: disjnt_iff)+ have "T \ X closure_of T" by (simp add: R closure_of_subset) then have *: "U - T = U - X interior_of T" using U(4) interior_of_subset by fastforce then show "closedin X (U - T)" by (simp add: U(1) closedin_diff) have "U \ X frontier_of T = {}" using U(4) frontier_of_def by fastforce then show "openin X (U - T)" by (metis * Diff_Un U(2) Un_Diff_Int Un_Int_eq(1) closedin_closure_of interior_of_union_frontier_of openin_diff sup_bot_right) qed then show "separated_between X S (X closure_of T)" by (simp add: R separated_between_closure_of') qed (auto simp: R closure_of_subset) qed lemma separated_between_frontier_of_eq: "separated_between X S T \ S \ topspace X \ disjnt S T \ separated_between X (X frontier_of S) T" by (metis disjnt_sym separated_between_frontier_of_eq' separated_between_sym) lemma separated_between_frontier_of: "\S \ topspace X; disjnt S T\ \ (separated_between X (X frontier_of S) T \ separated_between X S T)" using separated_between_frontier_of_eq by blast lemma separated_between_frontier_of': "\T \ topspace X; disjnt S T\ \ (separated_between X S (X frontier_of T) \ separated_between X S T)" using separated_between_frontier_of_eq' by auto lemma connected_space_separated_between: "connected_space X \ (\S T. separated_between X S T \ S = {} \ T = {})" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Diff_cancel connected_space_clopen_in separated_between subset_empty) next assume ?rhs then show ?lhs by (meson connected_space_eq_not_separated separated_between_eq_separatedin) qed lemma connected_space_imp_separated_between_trivial: "connected_space X \ (separated_between X S T \ S = {} \ T \ topspace X \ S \ topspace X \ T = {})" by (metis connected_space_separated_between separated_between_empty) subsection\Connected components\ lemma connected_component_of_subtopology_eq: "connected_component_of (subtopology X U) a = connected_component_of X a \ connected_component_of_set X a \ U" by (force simp: connected_component_of_set connectedin_subtopology connected_component_of_def fun_eq_iff subset_iff) lemma connected_components_of_subtopology: assumes "C \ connected_components_of X" "C \ U" shows "C \ connected_components_of (subtopology X U)" proof - obtain a where a: "connected_component_of_set X a \ U" and "a \ topspace X" and Ceq: "C = connected_component_of_set X a" using assms by (force simp: connected_components_of_def) then have "a \ U" by (simp add: connected_component_of_refl in_mono) then have "connected_component_of_set X a = connected_component_of_set (subtopology X U) a" by (metis a connected_component_of_subtopology_eq) then show ?thesis by (simp add: Ceq \a \ U\ \a \ topspace X\ connected_component_in_connected_components_of) qed thm connected_space_iff_components_eq lemma open_in_finite_connected_components: assumes "finite(connected_components_of X)" "C \ connected_components_of X" shows "openin X C" proof - have "closedin X (topspace X - C)" by (metis DiffD1 assms closedin_Union closedin_connected_components_of complement_connected_components_of_Union finite_Diff) then show ?thesis by (simp add: assms connected_components_of_subset openin_closedin) qed thm connected_component_of_eq_overlap lemma connected_components_of_disjoint: assumes "C \ connected_components_of X" "C' \ connected_components_of X" shows "(disjnt C C' \ (C \ C'))" proof - have "C \ {}" using nonempty_connected_components_of assms by blast with assms show ?thesis by (metis disjnt_self_iff_empty pairwiseD pairwise_disjoint_connected_components_of) qed lemma connected_components_of_overlap: "\C \ connected_components_of X; C' \ connected_components_of X\ \ C \ C' \ {} \ C = C'" by (meson connected_components_of_disjoint disjnt_def) lemma pairwise_separated_connected_components_of: "pairwise (separatedin X) (connected_components_of X)" by (simp add: closedin_connected_components_of connected_components_of_disjoint pairwiseI separatedin_closed_sets) lemma finite_connected_components_of_finite: "finite(topspace X) \ finite(connected_components_of X)" by (simp add: Union_connected_components_of finite_UnionD) lemma connected_component_of_unique: "\x \ C; connectedin X C; \C'. x \ C' \ connectedin X C' \ C' \ C\ \ connected_component_of_set X x = C" by (meson connected_component_of_maximal connectedin_connected_component_of subsetD subset_antisym) lemma closedin_connected_component_of_subtopology: "\C \ connected_components_of (subtopology X s); X closure_of C \ s\ \ closedin X C" by (metis closedin_Int_closure_of closedin_connected_components_of closure_of_eq inf.absorb_iff2) lemma connected_component_of_discrete_topology: "connected_component_of_set (discrete_topology U) x = (if x \ U then {x} else {})" by (simp add: locally_path_connected_space_discrete_topology flip: path_component_eq_connected_component_of) lemma connected_components_of_discrete_topology: "connected_components_of (discrete_topology U) = (\x. {x}) ` U" by (simp add: connected_component_of_discrete_topology connected_components_of_def) lemma connected_component_of_continuous_image: "\continuous_map X Y f; connected_component_of X x y\ \ connected_component_of Y (f x) (f y)" by (meson connected_component_of_def connectedin_continuous_map_image image_eqI) lemma homeomorphic_map_connected_component_of: assumes "homeomorphic_map X Y f" and x: "x \ topspace X" shows "connected_component_of_set Y (f x) = f ` (connected_component_of_set X x)" proof - obtain g where g: "continuous_map X Y f" "continuous_map Y X g " "\x. x \ topspace X \ g (f x) = x" "\y. y \ topspace Y \ f (g y) = y" using assms(1) homeomorphic_map_maps homeomorphic_maps_def by fastforce show ?thesis using connected_component_in_topspace [of Y] x g connected_component_of_continuous_image [of X Y f] connected_component_of_continuous_image [of Y X g] by force qed lemma homeomorphic_map_connected_components_of: assumes "homeomorphic_map X Y f" shows "connected_components_of Y = (image f) ` (connected_components_of X)" proof - have "topspace Y = f ` topspace X" by (metis assms homeomorphic_imp_surjective_map) with homeomorphic_map_connected_component_of [OF assms] show ?thesis by (auto simp: connected_components_of_def image_iff) qed lemma connected_component_of_pair: "connected_component_of_set (prod_topology X Y) (x,y) = connected_component_of_set X x \ connected_component_of_set Y y" proof (cases "x \ topspace X \ y \ topspace Y") case True show ?thesis proof (rule connected_component_of_unique) show "(x, y) \ connected_component_of_set X x \ connected_component_of_set Y y" using True by (simp add: connected_component_of_refl) show "connectedin (prod_topology X Y) (connected_component_of_set X x \ connected_component_of_set Y y)" by (metis connectedin_Times connectedin_connected_component_of) show "C \ connected_component_of_set X x \ connected_component_of_set Y y" if "(x, y) \ C \ connectedin (prod_topology X Y) C" for C using that unfolding connected_component_of_def apply clarsimp by (metis (no_types) connectedin_continuous_map_image continuous_map_fst continuous_map_snd fst_conv imageI snd_conv) qed next case False then show ?thesis by (metis Sigma_empty1 Sigma_empty2 connected_component_of_eq_empty mem_Sigma_iff topspace_prod_topology) qed lemma connected_components_of_prod_topology: "connected_components_of (prod_topology X Y) = {C \ D |C D. C \ connected_components_of X \ D \ connected_components_of Y}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" apply (clarsimp simp: connected_components_of_def) by (metis (no_types) connected_component_of_pair imageI) next show "?rhs \ ?lhs" using connected_component_of_pair by (fastforce simp: connected_components_of_def) qed lemma connected_component_of_product_topology: "connected_component_of_set (product_topology X I) x = (if x \ extensional I then PiE I (\i. connected_component_of_set (X i) (x i)) else {})" (is "?lhs = If _ ?R _") proof (cases "x \ topspace(product_topology X I)") case True have "?lhs = (\\<^sub>E i\I. connected_component_of_set (X i) (x i))" if xX: "\i. i\I \ x i \ topspace (X i)" and ext: "x \ extensional I" proof (rule connected_component_of_unique) show "x \ ?R" by (simp add: PiE_iff connected_component_of_refl local.ext xX) show "connectedin (product_topology X I) ?R" by (simp add: connectedin_PiE connectedin_connected_component_of) show "C \ ?R" if "x \ C \ connectedin (product_topology X I) C" for C proof - have "C \ extensional I" using PiE_def connectedin_subset_topspace that by fastforce have "\y. y \ C \ y \ (\ i\I. connected_component_of_set (X i) (x i))" apply (simp add: connected_component_of_def Pi_def) by (metis connectedin_continuous_map_image continuous_map_product_projection imageI that) then show ?thesis using PiE_def \C \ extensional I\ by fastforce qed qed with True show ?thesis by (simp add: PiE_iff) next case False then show ?thesis apply (simp add: PiE_iff) by (smt (verit) Collect_empty_eq False PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty) qed lemma connected_components_of_product_topology: "connected_components_of (product_topology X I) = {PiE I B |B. \i \ I. B i \ connected_components_of(X i)}" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (auto simp: connected_components_of_def connected_component_of_product_topology PiE_iff) show "?rhs \ ?lhs" proof fix F assume "F \ ?rhs" then obtain B where Feq: "F = Pi\<^sub>E I B" and "\i\I. \x\topspace (X i). B i = connected_component_of_set (X i) x" by (force simp: connected_components_of_def connected_component_of_product_topology image_iff) then obtain f where f: "\i. i \ I \ f i \ topspace (X i) \ B i = connected_component_of_set (X i) (f i)" by metis then have "(\i\I. f i) \ ((\\<^sub>E i\I. topspace (X i)) \ extensional I)" by simp with f show "F \ ?lhs" unfolding Feq connected_components_of_def connected_component_of_product_topology image_iff by (smt (verit, del_insts) PiE_cong restrict_PiE_iff restrict_apply' restrict_extensional topspace_product_topology) qed qed subsection \Monotone maps (in the general topological sense)\ definition monotone_map where "monotone_map X Y f == f ` (topspace X) \ topspace Y \ (\y \ topspace Y. connectedin X {x \ topspace X. f x = y})" lemma monotone_map: "monotone_map X Y f \ f ` (topspace X) \ topspace Y \ (\y. connectedin X {x \ topspace X. f x = y})" apply (simp add: monotone_map_def) by (metis (mono_tags, lifting) connectedin_empty [of X] Collect_empty_eq image_subset_iff) lemma monotone_map_in_subtopology: "monotone_map X (subtopology Y S) f \ monotone_map X Y f \ f ` (topspace X) \ S" by (smt (verit, del_insts) le_inf_iff monotone_map topspace_subtopology) lemma monotone_map_from_subtopology: assumes "monotone_map X Y f" "\x y. x \ topspace X \ y \ topspace X \ x \ S \ f x = f y \ y \ S" shows "monotone_map (subtopology X S) Y f" using assms unfolding monotone_map_def connectedin_subtopology by (smt (verit, del_insts) Collect_cong Collect_empty_eq IntE IntI connectedin_empty image_subset_iff mem_Collect_eq subsetI topspace_subtopology) lemma monotone_map_restriction: "monotone_map X Y f \ {x \ topspace X. f x \ v} = u \ monotone_map (subtopology X u) (subtopology Y v) f" by (smt (verit, best) IntI Int_Collect image_subset_iff mem_Collect_eq monotone_map monotone_map_from_subtopology topspace_subtopology) lemma injective_imp_monotone_map: assumes "f ` topspace X \ topspace Y" "inj_on f (topspace X)" shows "monotone_map X Y f" unfolding monotone_map_def proof (intro conjI assms strip) fix y assume "y \ topspace Y" then have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using assms(2) unfolding inj_on_def by blast then show "connectedin X {x \ topspace X. f x = y}" by (metis (no_types, lifting) connectedin_empty connectedin_sing) qed lemma embedding_imp_monotone_map: "embedding_map X Y f \ monotone_map X Y f" by (metis (no_types) embedding_map_def homeomorphic_eq_everything_map inf.absorb_iff2 injective_imp_monotone_map topspace_subtopology) lemma section_imp_monotone_map: "section_map X Y f \ monotone_map X Y f" by (simp add: embedding_imp_monotone_map section_imp_embedding_map) lemma homeomorphic_imp_monotone_map: "homeomorphic_map X Y f \ monotone_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_monotone_map) proposition connected_space_monotone_quotient_map_preimage: assumes f: "monotone_map X Y f" "quotient_map X Y f" and "connected_space Y" shows "connected_space X" proof (rule ccontr) assume "\ connected_space X" then obtain U V where "openin X U" "openin X V" "U \ V = {}" "U \ {}" "V \ {}" and topUV: "topspace X \ U \ V" by (auto simp: connected_space_def) then have UVsub: "U \ topspace X" "V \ topspace X" by (auto simp: openin_subset) have "\ connected_space Y" unfolding connected_space_def not_not proof (intro exI conjI) show "topspace Y \ f`U \ f`V" by (metis f(2) image_Un quotient_imp_surjective_map subset_Un_eq topUV) show "f`U \ {}" by (simp add: \U \ {}\) show "(f`V) \ {}" by (simp add: \V \ {}\) have *: "y \ f ` V" if "y \ f ` U" for y proof - have \
: "connectedin X {x \ topspace X. f x = y}" using f(1) monotone_map by fastforce show ?thesis using connectedinD [OF \
\openin X U\ \openin X V\] UVsub topUV \U \ V = {}\ that by (force simp: disjoint_iff) qed then show "f`U \ f`V = {}" by blast show "openin Y (f`U)" using f \openin X U\ topUV * unfolding quotient_map_saturated_open by force show "openin Y (f`V)" using f \openin X V\ topUV * unfolding quotient_map_saturated_open by force qed then show False by (simp add: assms) qed lemma connectedin_monotone_quotient_map_preimage: assumes "monotone_map X Y f" "quotient_map X Y f" "connectedin Y C" "openin Y C \ closedin Y C" shows "connectedin X {x \ topspace X. f x \ C}" proof - have "connected_space (subtopology X {x \ topspace X. f x \ C})" proof - have "connected_space (subtopology Y C)" using \connectedin Y C\ connectedin_def by blast moreover have "quotient_map (subtopology X {a \ topspace X. f a \ C}) (subtopology Y C) f" by (simp add: assms quotient_map_restriction) ultimately show ?thesis using \monotone_map X Y f\ connected_space_monotone_quotient_map_preimage monotone_map_restriction by blast qed then show ?thesis by (simp add: connectedin_def) qed lemma monotone_open_map: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding connectedin_def proof (intro strip conjI) fix C assume C: "C \ topspace Y \ connected_space (subtopology Y C)" show "connected_space (subtopology X {x \ topspace X. f x \ C})" proof (rule connected_space_monotone_quotient_map_preimage) show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" by (simp add: L monotone_map_restriction) show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" proof (rule continuous_open_imp_quotient_map) show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce qed (use open_map_restriction assms in fastforce)+ qed (simp add: C) qed auto next assume ?rhs then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" by (smt (verit) Collect_cong singletonD singletonI) then show ?lhs by (simp add: fim monotone_map_def) qed lemma monotone_closed_map: assumes "continuous_map X Y f" "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "monotone_map X Y f \ (\C. connectedin Y C \ connectedin X {x \ topspace X. f x \ C})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding connectedin_def proof (intro strip conjI) fix C assume C: "C \ topspace Y \ connected_space (subtopology Y C)" show "connected_space (subtopology X {x \ topspace X. f x \ C})" proof (rule connected_space_monotone_quotient_map_preimage) show "monotone_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" by (simp add: L monotone_map_restriction) show "quotient_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" proof (rule continuous_closed_imp_quotient_map) show "continuous_map (subtopology X {x \ topspace X. f x \ C}) (subtopology Y C) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce qed (use closed_map_restriction assms in fastforce)+ qed (simp add: C) qed auto next assume ?rhs then have "\y. connectedin Y {y} \ connectedin X {x \ topspace X. f x = y}" by (smt (verit) Collect_cong singletonD singletonI) then show ?lhs by (simp add: fim monotone_map_def) qed subsection\Other countability properties\ definition second_countable where "second_countable X \ \\. countable \ \ (\V \ \. openin X V) \ (\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" definition first_countable where "first_countable X \ \x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ (\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U))" definition separable_space where "separable_space X \ \C. countable C \ C \ topspace X \ X closure_of C = topspace X" lemma second_countable: "second_countable X \ (\\. countable \ \ openin X = arbitrary union_of (\x. x \ \))" by (smt (verit) openin_topology_base_unique second_countable_def) lemma second_countable_subtopology: assumes "second_countable X" shows "second_countable (subtopology X S)" proof - obtain \ where \: "countable \" "\V. V \ \ \ openin X V" "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) show ?thesis unfolding second_countable_def proof (intro exI conjI) show "\V\((\)S) ` \. openin (subtopology X S) V" using openin_subtopology_Int2 \ by blast show "\U x. openin (subtopology X S) U \ x \ U \ (\V\((\)S) ` \. x \ V \ V \ U)" using \ unfolding openin_subtopology by (smt (verit, del_insts) IntI image_iff inf_commute inf_le1 subset_iff) qed (use \ in auto) qed lemma second_countable_discrete_topology: "second_countable(discrete_topology U) \ countable U" (is "?lhs=?rhs") proof assume L: ?lhs then obtain \ where \: "countable \" "\V. V \ \ \ V \ U" "\W x. W \ U \ x \ W \ (\V \ \. x \ V \ V \ W)" by (auto simp: second_countable_def) then have "{x} \ \" if "x \ U" for x by (metis empty_subsetI insertCI insert_subset subset_antisym that) then show ?rhs by (smt (verit) countable_subset image_subsetI \countable \\ countable_image_inj_on [OF _ inj_singleton]) next assume ?rhs then show ?lhs unfolding second_countable_def by (rule_tac x="(\x. {x}) ` U" in exI) auto qed lemma second_countable_open_map_image: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "second_countable X" shows "second_countable Y" proof - have openXYf: "\U. openin X U \ openin Y (f ` U)" using assms by (auto simp: open_map_def) obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) show ?thesis unfolding second_countable_def proof (intro exI conjI strip) fix V y assume V: "openin Y V \ y \ V" then obtain x where "x \ topspace X" and x: "f x = y" by (metis fim image_iff openin_subset subsetD) then obtain W where "W\\" "x \ W" "W \ {x \ topspace X. f x \ V}" using * [of "{x \ topspace X. f x \ V}" x] V assms openin_continuous_map_preimage by force then show "\W \ (image f) ` \. y \ W \ W \ V" using x by auto qed (use \ openXYf in auto) qed lemma homeomorphic_space_second_countability: "X homeomorphic_space Y \ (second_countable X \ second_countable Y)" by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym second_countable_open_map_image) lemma second_countable_retraction_map_image: "\retraction_map X Y r; second_countable X\ \ second_countable Y" using hereditary_imp_retractive_property homeomorphic_space_second_countability second_countable_subtopology by blast lemma second_countable_imp_first_countable: "second_countable X \ first_countable X" by (metis first_countable_def second_countable_def) lemma first_countable_subtopology: assumes "first_countable X" shows "first_countable (subtopology X S)" unfolding first_countable_def proof fix x assume "x \ topspace (subtopology X S)" then obtain \ where "countable \" and \: "\V. V \ \ \ openin X V" "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms first_countable_def by force show "\\. countable \ \ (\V\\. openin (subtopology X S) V) \ (\U. openin (subtopology X S) U \ x \ U \ (\V\\. x \ V \ V \ U))" proof (intro exI conjI strip) show "countable (((\)S) ` \)" using \countable \\ by blast show "openin (subtopology X S) V" if "V \ ((\)S) ` \" for V using \ openin_subtopology_Int2 that by fastforce show "\V\((\)S) ` \. x \ V \ V \ U" if "openin (subtopology X S) U \ x \ U" for U using that \(2) by (clarsimp simp: openin_subtopology) (meson le_infI2) qed qed lemma first_countable_discrete_topology: "first_countable (discrete_topology U)" unfolding first_countable_def topspace_discrete_topology openin_discrete_topology proof fix x assume "x \ U" show "\\. countable \ \ (\V\\. V \ U) \ (\Ua. Ua \ U \ x \ Ua \ (\V\\. x \ V \ V \ Ua))" using \x \ U\ by (rule_tac x="{{x}}" in exI) auto qed lemma first_countable_open_map_image: assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "first_countable X" shows "first_countable Y" unfolding first_countable_def proof fix y assume "y \ topspace Y" have openXYf: "\U. openin X U \ openin Y (f ` U)" using assms by (auto simp: open_map_def) then obtain x where x: "x \ topspace X" "f x = y" by (metis \y \ topspace Y\ fim imageE) obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms x first_countable_def by force show "\\. countable \ \ (\V\\. openin Y V) \ (\U. openin Y U \ y \ U \ (\V\\. y \ V \ V \ U))" proof (intro exI conjI strip) fix V assume "openin Y V \ y \ V" then have "\W\\. x \ W \ W \ {x \ topspace X. f x \ V}" using * [of "{x \ topspace X. f x \ V}"] assms openin_continuous_map_preimage x by fastforce then show "\V' \ (image f) ` \. y \ V' \ V' \ V" using image_mono x by auto qed (use \ openXYf in force)+ qed lemma homeomorphic_space_first_countability: "X homeomorphic_space Y \ first_countable X \ first_countable Y" by (meson first_countable_open_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) lemma first_countable_retraction_map_image: "\retraction_map X Y r; first_countable X\ \ first_countable Y" using first_countable_subtopology hereditary_imp_retractive_property homeomorphic_space_first_countability by blast lemma separable_space_open_subset: assumes "separable_space X" "openin X S" shows "separable_space (subtopology X S)" proof - obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" by (meson assms separable_space_def) then have "\x T. \x \ topspace X; x \ T; openin (subtopology X S) T\ \ \y. y \ S \ y \ C \ y \ T" by (smt (verit) \openin X S\ in_closure_of openin_open_subtopology subsetD) with C \openin X S\ show ?thesis unfolding separable_space_def by (rule_tac x="S \ C" in exI) (force simp: in_closure_of) qed lemma separable_space_continuous_map_image: assumes "separable_space X" "continuous_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "separable_space Y" proof - have cont: "\S. f ` (X closure_of S) \ Y closure_of f ` S" by (simp add: assms continuous_map_image_closure_subset) obtain C where C: "countable C" "C \ topspace X" "X closure_of C = topspace X" by (meson assms separable_space_def) then show ?thesis unfolding separable_space_def by (metis cont fim closure_of_subset_topspace countable_image image_mono subset_antisym) qed lemma separable_space_quotient_map_image: "\quotient_map X Y q; separable_space X\ \ separable_space Y" by (meson quotient_imp_continuous_map quotient_imp_surjective_map separable_space_continuous_map_image) lemma separable_space_retraction_map_image: "\retraction_map X Y r; separable_space X\ \ separable_space Y" using retraction_imp_quotient_map separable_space_quotient_map_image by blast lemma homeomorphic_separable_space: "X homeomorphic_space Y \ (separable_space X \ separable_space Y)" by (meson homeomorphic_eq_everything_map homeomorphic_maps_map homeomorphic_space_def separable_space_continuous_map_image) lemma separable_space_discrete_topology: "separable_space(discrete_topology U) \ countable U" by (metis countable_Int2 discrete_topology_closure_of dual_order.refl inf.orderE separable_space_def topspace_discrete_topology) lemma second_countable_imp_separable_space: assumes "second_countable X" shows "separable_space X" proof - obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) obtain c where c: "\V. \V \ \; V \ {}\ \ c V \ V" by (metis all_not_in_conv) then have **: "\x. x \ topspace X \ x \ X closure_of c ` (\ - {{}})" using * by (force simp: closure_of_def) show ?thesis unfolding separable_space_def proof (intro exI conjI) show "countable (c ` (\-{{}}))" using \(1) by blast show "(c ` (\-{{}})) \ topspace X" using \(2) c openin_subset by fastforce show "X closure_of (c ` (\-{{}})) = topspace X" by (meson ** closure_of_subset_topspace subsetI subset_antisym) qed qed lemma second_countable_imp_Lindelof_space: assumes "second_countable X" shows "Lindelof_space X" unfolding Lindelof_space_def proof clarify fix \ assume "\U \ \. openin X U" and UU: "\\ = topspace X" obtain \ where \: "countable \" "\V. V \ \ \ openin X V" and *: "\U x. openin X U \ x \ U \ (\V \ \. x \ V \ V \ U)" using assms by (auto simp: second_countable_def) define \' where "\' = {B \ \. \U. U \ \ \ B \ U}" have \': "countable \'" "\\' = \\" using \ using "*" \\U\\. openin X U\ by (fastforce simp: \'_def)+ have "\b. \U. b \ \' \ U \ \ \ b \ U" by (simp add: \'_def) then obtain G where G: "\b. b \ \' \ G b \ \ \ b \ G b" by metis with \' UU show "\\. countable \ \ \ \ \ \ \\ = topspace X" by (rule_tac x="G ` \'" in exI) fastforce qed subsection \Neigbourhood bases EXTRAS\ (* Neigbourhood bases (useful for "local" properties of various kind). *) lemma openin_topology_neighbourhood_base_unique: "openin X = arbitrary union_of P \ (\u. P u \ openin X u) \ neighbourhood_base_of P X" by (smt (verit, best) open_neighbourhood_base_of openin_topology_base_unique) lemma neighbourhood_base_at_topology_base: " openin X = arbitrary union_of b \ (neighbourhood_base_at x P X \ (\w. b w \ x \ w \ (\u v. openin X u \ P v \ x \ u \ u \ v \ v \ w)))" apply (simp add: neighbourhood_base_at_def) by (smt (verit, del_insts) openin_topology_base_unique subset_trans) lemma neighbourhood_base_of_unlocalized: assumes "\S t. P S \ openin X t \ (t \ {}) \ t \ S \ P t" shows "neighbourhood_base_of P X \ (\x \ topspace X. \u v. openin X u \ P v \ x \ u \ u \ v \ v \ topspace X)" apply (simp add: neighbourhood_base_of_def) by (smt (verit, ccfv_SIG) assms empty_iff neighbourhood_base_at_unlocalized) lemma neighbourhood_base_at_discrete_topology: "neighbourhood_base_at x P (discrete_topology u) \ x \ u \ P {x}" apply (simp add: neighbourhood_base_at_def) by (smt (verit) empty_iff empty_subsetI insert_subset singletonI subsetD subset_singletonD) lemma neighbourhood_base_of_discrete_topology: "neighbourhood_base_of P (discrete_topology u) \ (\x \ u. P {x})" apply (simp add: neighbourhood_base_of_def) using neighbourhood_base_at_discrete_topology[of _ P u] by (metis empty_subsetI insert_subset neighbourhood_base_at_def openin_discrete_topology singletonI) lemma second_countable_neighbourhood_base_alt: "second_countable X \ (\\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_of (\A. A\\) X)" by (metis (full_types) openin_topology_neighbourhood_base_unique second_countable) lemma first_countable_neighbourhood_base_alt: "first_countable X \ (\x \ topspace X. \\. countable \ \ (\V \ \. openin X V) \ neighbourhood_base_at x (\V. V \ \) X)" unfolding first_countable_def apply (intro ball_cong refl ex_cong conj_cong) by (metis (mono_tags, lifting) open_neighbourhood_base_at) lemma second_countable_neighbourhood_base: "second_countable X \ (\\. countable \ \ neighbourhood_base_of (\V. V \ \) X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using second_countable_neighbourhood_base_alt by blast next assume ?rhs then obtain \ where "countable \" and \: "\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. V \ \ \ x \ U \ U \ V \ V \ W))" by (metis neighbourhood_base_of) then show ?lhs unfolding second_countable_neighbourhood_base_alt neighbourhood_base_of apply (rule_tac x="(\u. X interior_of u) ` \" in exI) by (smt (verit, best) interior_of_eq interior_of_mono countable_image image_iff openin_interior_of) qed lemma first_countable_neighbourhood_base: "first_countable X \ (\x \ topspace X. \\. countable \ \ neighbourhood_base_at x (\V. V \ \) X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis first_countable_neighbourhood_base_alt) next assume R: ?rhs show ?lhs unfolding first_countable_neighbourhood_base_alt proof fix x assume "x \ topspace X" with R obtain \ where "countable \" and \: "neighbourhood_base_at x (\V. V \ \) X" by blast then show "\\. countable \ \ Ball \ (openin X) \ neighbourhood_base_at x (\V. V \ \) X" unfolding neighbourhood_base_at_def apply (rule_tac x="(\u. X interior_of u) ` \" in exI) by (smt (verit, best) countable_image image_iff interior_of_eq interior_of_mono openin_interior_of) qed qed subsection\$T_0$ spaces and the Kolmogorov quotient\ definition t0_space where "t0_space X \ \x \ topspace X. \y \ topspace X. x \ y \ (\U. openin X U \ (x \ U \ y \ U))" lemma t0_space_expansive: "\topspace Y = topspace X; \U. openin X U \ openin Y U\ \ t0_space X \ t0_space Y" by (metis t0_space_def) lemma t1_imp_t0_space: "t1_space X \ t0_space X" by (metis t0_space_def t1_space_def) lemma t1_eq_symmetric_t0_space_alt: "t1_space X \ t0_space X \ (\x \ topspace X. \y \ topspace X. x \ X closure_of {y} \ y \ X closure_of {x})" apply (simp add: t0_space_def t1_space_def closure_of_def) by (smt (verit, best) openin_topspace) lemma t1_eq_symmetric_t0_space: "t1_space X \ t0_space X \ (\x y. x \ X closure_of {y} \ y \ X closure_of {x})" by (auto simp: t1_eq_symmetric_t0_space_alt in_closure_of) lemma Hausdorff_imp_t0_space: "Hausdorff_space X \ t0_space X" by (simp add: Hausdorff_imp_t1_space t1_imp_t0_space) lemma t0_space: "t0_space X \ (\x \ topspace X. \y \ topspace X. x \ y \ (\C. closedin X C \ (x \ C \ y \ C)))" unfolding t0_space_def by (metis Diff_iff closedin_def openin_closedin_eq) lemma homeomorphic_t0_space: assumes "X homeomorphic_space Y" shows "t0_space X \ t0_space Y" proof - obtain f where f: "homeomorphic_map X Y f" and F: "inj_on f (topspace X)" and "topspace Y = f ` topspace X" by (metis assms homeomorphic_imp_injective_map homeomorphic_imp_surjective_map homeomorphic_space) with inj_on_image_mem_iff [OF F] show ?thesis apply (simp add: t0_space_def homeomorphic_eq_everything_map continuous_map_def open_map_def inj_on_def) by (smt (verit) mem_Collect_eq openin_subset) qed lemma t0_space_closure_of_sing: "t0_space X \ (\x \ topspace X. \y \ topspace X. X closure_of {x} = X closure_of {y} \ x = y)" by (simp add: t0_space_def closure_of_def set_eq_iff) (smt (verit)) lemma t0_space_discrete_topology: "t0_space (discrete_topology S)" by (simp add: Hausdorff_imp_t0_space) lemma t0_space_subtopology: "t0_space X \ t0_space (subtopology X U)" by (simp add: t0_space_def openin_subtopology) (metis Int_iff) lemma t0_space_retraction_map_image: "\retraction_map X Y r; t0_space X\ \ t0_space Y" using hereditary_imp_retractive_property homeomorphic_t0_space t0_space_subtopology by blast lemma XY: "{x}\{y} = {(x,y)}" by simp lemma t0_space_prod_topologyI: "\t0_space X; t0_space Y\ \ t0_space (prod_topology X Y)" by (simp add: t0_space_closure_of_sing closure_of_Times closure_of_eq_empty_gen times_eq_iff flip: XY insert_Times_insert) lemma t0_space_prod_topology_iff: "t0_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ t0_space X \ t0_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Sigma_empty1 Sigma_empty2 retraction_map_fst retraction_map_snd t0_space_retraction_map_image topspace_prod_topology) qed (metis empty_iff t0_space_def t0_space_prod_topologyI) proposition t0_space_product_topology: "t0_space (product_topology X I) \ topspace(product_topology X I) = {} \ (\i \ I. t0_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (meson retraction_map_product_projection t0_space_retraction_map_image) next assume R: ?rhs show ?lhs proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis by (simp add: t0_space_def) next case False show ?thesis unfolding t0_space proof (intro strip) fix x y assume x: "x \ topspace (product_topology X I)" and y: "y \ topspace (product_topology X I)" and "x \ y" then obtain i where "i \ I" "x i \ y i" by (metis PiE_ext topspace_product_topology) then have "t0_space (X i)" using False R by blast then obtain U where "closedin (X i) U" "(x i \ U \ y i \ U)" by (metis t0_space PiE_mem \i \ I\ \x i \ y i\ topspace_product_topology x y) with \i \ I\ x y show "\U. closedin (product_topology X I) U \ (x \ U) = (y \ U)" by (rule_tac x="PiE I (\j. if j = i then U else topspace(X j))" in exI) (simp add: closedin_product_topology PiE_iff) qed qed qed subsection \Kolmogorov quotients\ definition Kolmogorov_quotient where "Kolmogorov_quotient X \ \x. @y. \U. openin X U \ (y \ U \ x \ U)" lemma Kolmogorov_quotient_in_open: "openin X U \ (Kolmogorov_quotient X x \ U \ x \ U)" by (smt (verit, ccfv_SIG) Kolmogorov_quotient_def someI_ex) lemma Kolmogorov_quotient_in_topspace: "Kolmogorov_quotient X x \ topspace X \ x \ topspace X" by (simp add: Kolmogorov_quotient_in_open) lemma Kolmogorov_quotient_in_closed: "closedin X C \ (Kolmogorov_quotient X x \ C \ x \ C)" unfolding closedin_def by (meson DiffD2 DiffI Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace in_mono) lemma continuous_map_Kolmogorov_quotient: "continuous_map X X (Kolmogorov_quotient X)" using Kolmogorov_quotient_in_open openin_subopen openin_subset by (fastforce simp: continuous_map_def Kolmogorov_quotient_in_topspace) lemma open_map_Kolmogorov_quotient_explicit: "openin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" using Kolmogorov_quotient_in_open openin_subset by fastforce lemma open_map_Kolmogorov_quotient_gen: "open_map (subtopology X S) (subtopology X (image (Kolmogorov_quotient X) S)) (Kolmogorov_quotient X)" proof (clarsimp simp: open_map_def openin_subtopology_alt image_iff) fix U assume "openin X U" then have "Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ U" using Kolmogorov_quotient_in_open [of X U] by auto then show "\V. openin X V \ Kolmogorov_quotient X ` (S \ U) = Kolmogorov_quotient X ` S \ V" using \openin X U\ by blast qed lemma open_map_Kolmogorov_quotient: "open_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis open_map_Kolmogorov_quotient_gen subtopology_topspace) lemma closed_map_Kolmogorov_quotient_explicit: "closedin X U \ Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \ U" using closedin_subset by (fastforce simp: Kolmogorov_quotient_in_closed) lemma closed_map_Kolmogorov_quotient_gen: "closed_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" using Kolmogorov_quotient_in_closed by (force simp: closed_map_def closedin_subtopology_alt image_iff) lemma closed_map_Kolmogorov_quotient: "closed_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis closed_map_Kolmogorov_quotient_gen subtopology_topspace) lemma quotient_map_Kolmogorov_quotient_gen: "quotient_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" proof (intro continuous_open_imp_quotient_map) show "continuous_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" by (simp add: continuous_map_Kolmogorov_quotient continuous_map_from_subtopology continuous_map_in_subtopology image_mono) show "open_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" using open_map_Kolmogorov_quotient_gen by blast show "Kolmogorov_quotient X ` topspace (subtopology X S) = topspace (subtopology X (Kolmogorov_quotient X ` S))" by (force simp: Kolmogorov_quotient_in_open) qed lemma quotient_map_Kolmogorov_quotient: "quotient_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" by (metis quotient_map_Kolmogorov_quotient_gen subtopology_topspace) lemma Kolmogorov_quotient_eq: "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ (\U. openin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Kolmogorov_quotient_in_open) next assume ?rhs then show ?lhs by (simp add: Kolmogorov_quotient_def) qed lemma Kolmogorov_quotient_eq_alt: "Kolmogorov_quotient X x = Kolmogorov_quotient X y \ (\U. closedin X U \ (x \ U \ y \ U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis Kolmogorov_quotient_in_closed) next assume ?rhs then show ?lhs by (smt (verit) Diff_iff Kolmogorov_quotient_eq closedin_topspace in_mono openin_closedin_eq) qed lemma Kolmogorov_quotient_continuous_map: assumes "continuous_map X Y f" "t0_space Y" and x: "x \ topspace X" shows "f (Kolmogorov_quotient X x) = f x" using assms unfolding continuous_map_def t0_space_def by (smt (verit, ccfv_SIG) Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace x mem_Collect_eq) lemma t0_space_Kolmogorov_quotient: "t0_space (subtopology X (Kolmogorov_quotient X ` topspace X))" apply (clarsimp simp: t0_space_def ) by (smt (verit, best) Kolmogorov_quotient_eq imageE image_eqI open_map_Kolmogorov_quotient open_map_def) lemma Kolmogorov_quotient_id: "t0_space X \ x \ topspace X \ Kolmogorov_quotient X x = x" by (metis Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace t0_space_def) lemma Kolmogorov_quotient_idemp: "Kolmogorov_quotient X (Kolmogorov_quotient X x) = Kolmogorov_quotient X x" by (simp add: Kolmogorov_quotient_eq Kolmogorov_quotient_in_open) lemma retraction_maps_Kolmogorov_quotient: "retraction_maps X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X) id" unfolding retraction_maps_def continuous_map_in_subtopology using Kolmogorov_quotient_idemp continuous_map_Kolmogorov_quotient by force lemma retraction_map_Kolmogorov_quotient: "retraction_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" using retraction_map_def retraction_maps_Kolmogorov_quotient by blast lemma retract_of_space_Kolmogorov_quotient_image: "Kolmogorov_quotient X ` topspace X retract_of_space X" proof - have "continuous_map X X (Kolmogorov_quotient X)" by (simp add: continuous_map_Kolmogorov_quotient) then have "Kolmogorov_quotient X ` topspace X \ topspace X" by (simp add: continuous_map_image_subset_topspace) then show ?thesis by (meson retract_of_space_retraction_maps retraction_maps_Kolmogorov_quotient) qed lemma Kolmogorov_quotient_lift_exists: assumes "S \ topspace X" "t0_space Y" and f: "continuous_map (subtopology X S) Y f" obtains g where "continuous_map (subtopology X (image (Kolmogorov_quotient X) S)) Y g" "\x. x \ S \ g(Kolmogorov_quotient X x) = f x" proof - have "\x y. \x \ S; y \ S; Kolmogorov_quotient X x = Kolmogorov_quotient X y\ \ f x = f y" using assms apply (simp add: Kolmogorov_quotient_eq t0_space_def continuous_map_def Int_absorb1 openin_subtopology) by (smt (verit, del_insts) Int_iff mem_Collect_eq) then obtain g where g: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) Y g" "g ` (topspace X \ Kolmogorov_quotient X ` S) = f ` S" "\x. x \ S \ g (Kolmogorov_quotient X x) = f x" using quotient_map_lift_exists [OF quotient_map_Kolmogorov_quotient_gen [of X S] f] by (metis assms(1) topspace_subtopology topspace_subtopology_subset) show ?thesis proof qed (use g in auto) qed subsection\Closed diagonals and graphs\ lemma Hausdorff_space_closedin_diagonal: "Hausdorff_space X \ closedin (prod_topology X X) ((\x. (x,x)) ` topspace X)" proof - have \
: "((\x. (x, x)) ` topspace X) \ topspace X \ topspace X" by auto show ?thesis apply (simp add: closedin_def openin_prod_topology_alt Hausdorff_space_def disjnt_iff \
) apply (intro all_cong1 imp_cong ex_cong1 conj_cong refl) by (force dest!: openin_subset)+ qed lemma closed_map_diag_eq: "closed_map X (prod_topology X X) (\x. (x,x)) \ Hausdorff_space X" proof - have "section_map X (prod_topology X X) (\x. (x, x))" unfolding section_map_def retraction_maps_def by (smt (verit) continuous_map_fst continuous_map_of_fst continuous_map_on_empty continuous_map_pairwise fst_conv fst_diag_fst snd_diag_fst) then have "embedding_map X (prod_topology X X) (\x. (x, x))" by (rule section_imp_embedding_map) then show ?thesis using Hausdorff_space_closedin_diagonal embedding_imp_closed_map_eq by blast qed +lemma proper_map_diag_eq [simp]: + "proper_map X (prod_topology X X) (\x. (x,x)) \ Hausdorff_space X" + by (simp add: closed_map_diag_eq inj_on_convol_ident injective_imp_proper_eq_closed_map) + lemma closedin_continuous_maps_eq: assumes "Hausdorff_space Y" and f: "continuous_map X Y f" and g: "continuous_map X Y g" shows "closedin X {x \ topspace X. f x = g x}" proof - have \
:"{x \ topspace X. f x = g x} = {x \ topspace X. (f x,g x) \ ((\y.(y,y)) ` topspace Y)}" using f continuous_map_image_subset_topspace by fastforce show ?thesis unfolding \
proof (intro closedin_continuous_map_preimage) show "continuous_map X (prod_topology Y Y) (\x. (f x, g x))" by (simp add: continuous_map_pairedI f g) show "closedin (prod_topology Y Y) ((\y. (y, y)) ` topspace Y)" using Hausdorff_space_closedin_diagonal assms by blast qed qed lemma retract_of_space_imp_closedin: assumes "Hausdorff_space X" and S: "S retract_of_space X" shows "closedin X S" proof - obtain r where r: "continuous_map X (subtopology X S) r" "\x\S. r x = x" using assms by (meson retract_of_space_def) then have \
: "S = {x \ topspace X. r x = x}" using S retract_of_space_imp_subset by (force simp: continuous_map_def) show ?thesis unfolding \
using r continuous_map_into_fulltopology assms by (force intro: closedin_continuous_maps_eq) qed lemma homeomorphic_maps_graph: "homeomorphic_maps X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` (topspace X))) (\x. (x, f x)) fst \ continuous_map X Y f" (is "?lhs=?rhs") proof assume ?lhs then have h: "homeomorphic_map X (subtopology (prod_topology X Y) ((\x. (x, f x)) ` topspace X)) (\x. (x, f x))" by (auto simp: homeomorphic_maps_map) have "f = snd \ (\x. (x, f x))" by force then show ?rhs by (metis (no_types, lifting) h continuous_map_in_subtopology continuous_map_snd_of homeomorphic_eq_everything_map) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by (smt (verit, ccfv_threshold) continuous_map_eq continuous_map_subtopology_fst embedding_map_def embedding_map_graph homeomorphic_eq_everything_map image_cong image_iff prod.collapse prod.inject) qed subsection \ KC spaces, those where all compact sets are closed.\ definition kc_space where "kc_space X \ \S. compactin X S \ closedin X S" lemma kc_space_euclidean: "kc_space (euclidean :: 'a::metric_space topology)" by (simp add: compact_imp_closed kc_space_def) lemma kc_space_expansive: "\kc_space X; topspace Y = topspace X; \U. openin X U \ openin Y U\ \ kc_space Y" by (meson compactin_contractive kc_space_def topology_finer_closedin) lemma compactin_imp_closedin_gen: "\kc_space X; compactin X S\ \ closedin X S" using kc_space_def by blast lemma Hausdorff_imp_kc_space: "Hausdorff_space X \ kc_space X" by (simp add: compactin_imp_closedin kc_space_def) lemma kc_imp_t1_space: "kc_space X \ t1_space X" by (simp add: finite_imp_compactin kc_space_def t1_space_closedin_finite) lemma kc_space_subtopology: "kc_space X \ kc_space(subtopology X S)" by (metis closedin_Int_closure_of closure_of_eq compactin_subtopology inf.absorb2 kc_space_def) lemma kc_space_discrete_topology: "kc_space(discrete_topology U)" using Hausdorff_space_discrete_topology compactin_imp_closedin kc_space_def by blast lemma kc_space_continuous_injective_map_preimage: assumes "kc_space Y" "continuous_map X Y f" and injf: "inj_on f (topspace X)" shows "kc_space X" unfolding kc_space_def proof (intro strip) fix S assume S: "compactin X S" have "S = {x \ topspace X. f x \ f ` S}" using S compactin_subset_topspace inj_onD [OF injf] by blast with assms S show "closedin X S" by (metis (no_types, lifting) Collect_cong closedin_continuous_map_preimage compactin_imp_closedin_gen image_compactin) qed lemma kc_space_retraction_map_image: assumes "retraction_map X Y r" "kc_space X" shows "kc_space Y" proof - obtain s where s: "continuous_map X Y r" "continuous_map Y X s" "\x. x \ topspace Y \ r (s x) = x" using assms by (force simp: retraction_map_def retraction_maps_def) then have inj: "inj_on s (topspace Y)" by (metis inj_on_def) show ?thesis unfolding kc_space_def proof (intro strip) fix S assume S: "compactin Y S" have "S = {x \ topspace Y. s x \ s ` S}" using S compactin_subset_topspace inj_onD [OF inj] by blast with assms S show "closedin Y S" by (meson compactin_imp_closedin_gen inj kc_space_continuous_injective_map_preimage s(2)) qed qed lemma homeomorphic_kc_space: "X homeomorphic_space Y \ kc_space X \ kc_space Y" by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym kc_space_continuous_injective_map_preimage) lemma compact_kc_eq_maximal_compact_space: assumes "compact_space X" shows "kc_space X \ (\Y. topspace Y = topspace X \ (\S. openin X S \ openin Y S) \ compact_space Y \ Y = X)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis closedin_compact_space compactin_contractive kc_space_def topology_eq topology_finer_closedin) next assume R: ?rhs show ?lhs unfolding kc_space_def proof (intro strip) fix S assume S: "compactin X S" define Y where "Y \ topology (arbitrary union_of (finite intersection_of (\A. A = topspace X - S \ openin X A) relative_to (topspace X)))" have "topspace Y = topspace X" by (auto simp: Y_def) have "openin X T \ openin Y T" for T by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase openin_subset relative_to_subset) have "compact_space Y" proof (rule Alexander_subbase_alt) show "\\'. finite \' \ \' \ \ \ topspace X \ \ \'" if \: "\ \ insert (topspace X - S) (Collect (openin X))" and sub: "topspace X \ \\" for \ proof - consider "\ \ Collect (openin X)" | \ where "\ = insert (topspace X - S) \" "\ \ Collect (openin X)" using \ by (metis insert_Diff subset_insert_iff) then show ?thesis proof cases case 1 then show ?thesis by (metis assms compact_space_alt mem_Collect_eq subsetD that(2)) next case 2 then have "S \ \\" using S sub compactin_subset_topspace by blast with 2 obtain \ where "finite \ \ \ \ \ \ S \ \\" using S unfolding compactin_def by (metis Ball_Collect) with 2 show ?thesis by (rule_tac x="insert (topspace X - S) \" in exI) blast qed qed qed (auto simp: Y_def) have "Y = X" using R \\S. openin X S \ openin Y S\ \compact_space Y\ \topspace Y = topspace X\ by blast moreover have "openin Y (topspace X - S)" by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase relative_to_subset) ultimately show "closedin X S" unfolding closedin_def using S compactin_subset_topspace by blast qed qed lemma continuous_imp_closed_map_gen: "\compact_space X; kc_space Y; continuous_map X Y f\ \ closed_map X Y f" by (meson closed_map_def closedin_compact_space compactin_imp_closedin_gen image_compactin) lemma kc_space_compact_subtopologies: "kc_space X \ (\K. compactin X K \ kc_space(subtopology X K))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (auto simp: kc_space_def closedin_closed_subtopology compactin_subtopology) next assume R: ?rhs show ?lhs unfolding kc_space_def proof (intro strip) fix K assume K: "compactin X K" then have "K \ topspace X" by (simp add: compactin_subset_topspace) moreover have "X closure_of K \ K" proof fix x assume x: "x \ X closure_of K" have "kc_space (subtopology X K)" by (simp add: R \compactin X K\) have "compactin X (insert x K)" by (metis K x compactin_Un compactin_sing in_closure_of insert_is_Un) then show "x \ K" by (metis R x K Int_insert_left_if1 closedin_Int_closure_of compact_imp_compactin_subtopology insertCI kc_space_def subset_insertI) qed ultimately show "closedin X K" using closure_of_subset_eq by blast qed qed lemma kc_space_compact_prod_topology: assumes "compact_space X" shows "kc_space(prod_topology X X) \ Hausdorff_space X" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs unfolding closed_map_diag_eq [symmetric] proof (intro continuous_imp_closed_map_gen) show "continuous_map X (prod_topology X X) (\x. (x, x))" by (intro continuous_intros) qed (use L assms in auto) next assume ?rhs then show ?lhs by (simp add: Hausdorff_imp_kc_space Hausdorff_space_prod_topology) qed lemma kc_space_prod_topology: "kc_space(prod_topology X X) \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis compactin_subspace kc_space_compact_prod_topology kc_space_subtopology subtopology_Times) next assume R: ?rhs have "kc_space (subtopology (prod_topology X X) L)" if "compactin (prod_topology X X) L" for L proof - define K where "K \ fst ` L \ snd ` L" have "L \ K \ K" by (force simp: K_def) have "compactin X K" by (metis K_def compactin_Un continuous_map_fst continuous_map_snd image_compactin that) then have "Hausdorff_space (subtopology X K)" by (simp add: R) then have "kc_space (prod_topology (subtopology X K) (subtopology X K))" by (simp add: \compactin X K\ compact_space_subtopology kc_space_compact_prod_topology) then have "kc_space (subtopology (prod_topology (subtopology X K) (subtopology X K)) L)" using kc_space_subtopology by blast then show ?thesis using \L \ K \ K\ subtopology_Times subtopology_subtopology by (metis (no_types, lifting) Sigma_cong inf.absorb_iff2) qed then show ?lhs using kc_space_compact_subtopologies by blast qed lemma kc_space_prod_topology_alt: "kc_space(prod_topology X X) \ kc_space X \ (\K. compactin X K \ Hausdorff_space(subtopology X K))" using Hausdorff_imp_kc_space kc_space_compact_subtopologies kc_space_prod_topology by blast proposition kc_space_prod_topology_left: assumes X: "kc_space X" and Y: "Hausdorff_space Y" shows "kc_space (prod_topology X Y)" unfolding kc_space_def proof (intro strip) fix K assume K: "compactin (prod_topology X Y) K" then have "K \ topspace X \ topspace Y" using compactin_subset_topspace topspace_prod_topology by blast moreover have "\T. openin (prod_topology X Y) T \ (a,b) \ T \ T \ (topspace X \ topspace Y) - K" if ab: "(a,b) \ (topspace X \ topspace Y) - K" for a b proof - have "compactin Y {b}" using that by force moreover have "compactin Y {y \ topspace Y. (a,y) \ K}" proof - have "compactin (prod_topology X Y) (K \ {a} \ topspace Y)" using that compact_Int_closedin [OF K] by (simp add: X closedin_prod_Times_iff compactin_imp_closedin_gen) moreover have "subtopology (prod_topology X Y) (K \ {a} \ topspace Y) homeomorphic_space subtopology Y {y \ topspace Y. (a, y) \ K}" unfolding homeomorphic_space_def homeomorphic_maps_def using that apply (rule_tac x="snd" in exI) apply (rule_tac x="Pair a" in exI) by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology continuous_map_subtopology_snd continuous_map_paired) ultimately show ?thesis by (simp add: compactin_subspace homeomorphic_compact_space) qed moreover have "disjnt {b} {y \ topspace Y. (a,y) \ K}" using ab by force ultimately obtain V U where VU: "openin Y V" "openin Y U" "{b} \ V" "{y \ topspace Y. (a,y) \ K} \ U" "disjnt V U" using Hausdorff_space_compact_separation [OF Y] by blast define V' where "V' \ topspace Y - U" have W: "closedin Y V'" "{y \ topspace Y. (a,y) \ K} \ topspace Y - V'" "disjnt V (topspace Y - V')" using VU by (auto simp: V'_def disjnt_iff) with VU obtain "V \ topspace Y" "V' \ topspace Y" by (meson closedin_subset openin_closedin_eq) then obtain "b \ V" "disjnt {y \ topspace Y. (a,y) \ K} V'" "V \ V'" using VU unfolding disjnt_iff V'_def by force define C where "C \ image fst (K \ {z \ topspace(prod_topology X Y). snd z \ V'})" have "closedin (prod_topology X Y) {z \ topspace (prod_topology X Y). snd z \ V'}" using closedin_continuous_map_preimage \closedin Y V'\ continuous_map_snd by blast then have "compactin X C" unfolding C_def by (meson K compact_Int_closedin continuous_map_fst image_compactin) then have "closedin X C" using assms by (auto simp: kc_space_def) show ?thesis proof (intro exI conjI) show "openin (prod_topology X Y) ((topspace X - C) \ V)" by (simp add: VU \closedin X C\ openin_diff openin_prod_Times_iff) have "a \ C" using VU by (auto simp: C_def V'_def) then show "(a, b) \ (topspace X - C) \ V" using \a \ C\ \b \ V\ that by blast show "(topspace X - C) \ V \ topspace X \ topspace Y - K" using \V \ V'\ \V \ topspace Y\ apply (simp add: C_def ) by (smt (verit, ccfv_threshold) DiffE DiffI IntI SigmaE SigmaI image_eqI mem_Collect_eq prod.sel(1) snd_conv subset_iff) qed qed ultimately show "closedin (prod_topology X Y) K" by (metis surj_pair closedin_def openin_subopen topspace_prod_topology) qed lemma kc_space_prod_topology_right: "\Hausdorff_space X; kc_space Y\ \ kc_space (prod_topology X Y)" using kc_space_prod_topology_left homeomorphic_kc_space homeomorphic_space_prod_topology_swap by blast +subsection \Technical results about proper maps, perfect maps, etc\ + +lemma compact_imp_proper_map_gen: + assumes Y: "\S. \S \ topspace Y; \K. compactin Y K \ compactin Y (S \ K)\ + \ closedin Y S" + and fim: "f ` (topspace X) \ topspace Y" + and f: "continuous_map X Y f \ kc_space X" + and YX: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" + shows "proper_map X Y f" + unfolding proper_map_alt closed_map_def +proof (intro conjI strip) + fix C + assume C: "closedin X C" + show "closedin Y (f ` C)" + proof (intro Y) + show "f ` C \ topspace Y" + using C closedin_subset fim by blast + fix K + assume K: "compactin Y K" + define A where "A \ {x \ topspace X. f x \ K}" + have eq: "f ` C \ K = f ` ({x \ topspace X. f x \ K} \ C)" + using C closedin_subset by auto + show "compactin Y (f ` C \ K)" + unfolding eq + proof (rule image_compactin) + show "compactin (subtopology X A) ({x \ topspace X. f x \ K} \ C)" + proof (rule closedin_compact_space) + show "compact_space (subtopology X A)" + by (simp add: A_def K YX compact_space_subtopology) + show "closedin (subtopology X A) ({x \ topspace X. f x \ K} \ C)" + using A_def C closedin_subtopology by blast + qed + have "continuous_map (subtopology X A) (subtopology Y K) f" if "kc_space X" + unfolding continuous_map_closedin + proof (intro conjI strip) + show "f x \ topspace (subtopology Y K)" + if "x \ topspace (subtopology X A)" for x + using that A_def K compactin_subset_topspace by auto + next + fix C + assume C: "closedin (subtopology Y K) C" + show "closedin (subtopology X A) {x \ topspace (subtopology X A). f x \ C}" + proof (rule compactin_imp_closedin_gen) + show "kc_space (subtopology X A)" + by (simp add: kc_space_subtopology that) + have [simp]: "{x \ topspace X. f x \ K \ f x \ C} = {x \ topspace X. f x \ C}" + using C closedin_imp_subset by auto + have "compactin (subtopology Y K) C" + by (simp add: C K closedin_compact_space compact_space_subtopology) + then have "compactin X {x \ topspace X. x \ A \ f x \ C}" + by (auto simp: A_def compactin_subtopology dest: YX) + then show "compactin (subtopology X A) {x \ topspace (subtopology X A). f x \ C}" + by (auto simp add: compactin_subtopology) + qed + qed + with f show "continuous_map (subtopology X A) Y f" + using continuous_map_from_subtopology continuous_map_in_subtopology by blast + qed + qed +qed (simp add: YX) + +lemma tube_lemma_left: + assumes W: "openin (prod_topology X Y) W" and C: "compactin X C" + and y: "y \ topspace Y" and subW: "C \ {y} \ W" + shows "\U V. openin X U \ openin Y V \ C \ U \ y \ V \ U \ V \ W" +proof (cases "C = {}") + case True + with y show ?thesis by auto +next + case False + have "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ W" + if "x \ C" for x + using W openin_prod_topology_alt subW subsetD that by fastforce + then obtain U V where UV: "\x. x \ C \ openin X (U x) \ openin Y (V x) \ x \ U x \ y \ V x \ U x \ V x \ W" + by metis + then obtain D where D: "finite D" "D \ C" "C \ \ (U ` D)" + using compactinD [OF C, of "U`C"] + by (smt (verit) UN_I finite_subset_image imageE subsetI) + show ?thesis + proof (intro exI conjI) + show "openin X (\ (U ` D))" "openin Y (\ (V ` D))" + using D False UV by blast+ + show "y \ \ (V ` D)" "C \ \ (U ` D)" "\(U ` D) \ \(V ` D) \ W" + using D UV by force+ + qed +qed + +lemma Wallace_theorem_prod_topology: + assumes "compactin X K" "compactin Y L" + and W: "openin (prod_topology X Y) W" and subW: "K \ L \ W" + obtains U V where "openin X U" "openin Y V" "K \ U" "L \ V" "U \ V \ W" +proof - + have "\y. y \ L \ \U V. openin X U \ openin Y V \ K \ U \ y \ V \ U \ V \ W" + proof (intro tube_lemma_left assms) + fix y assume "y \ L" + show "y \ topspace Y" + using assms \y \ L\ compactin_subset_topspace by blast + show "K \ {y} \ W" + using \y \ L\ subW by force + qed + then obtain U V where UV: + "\y. y \ L \ openin X (U y) \ openin Y (V y) \ K \ U y \ y \ V y \ U y \ V y \ W" + by metis + then obtain M where "finite M" "M \ L" and M: "L \ \ (V ` M)" + using \compactin Y L\ unfolding compactin_def + by (smt (verit) UN_iff finite_subset_image imageE subset_iff) + show thesis + proof (cases "M={}") + case True + with M have "L={}" + by blast + then show ?thesis + using \compactin X K\ compactin_subset_topspace that by fastforce + next + case False + show ?thesis + proof + show "openin X (\(U`M))" + using False UV \M \ L\ \finite M\ by blast + show "openin Y (\(V`M))" + using UV \M \ L\ by blast + show "K \ \(U`M)" + by (meson INF_greatest UV \M \ L\ subsetD) + show "L \ \(V`M)" + by (simp add: M) + show "\(U`M) \ \(V`M) \ W" + using UV \M \ L\ by fastforce + qed + qed +qed + +lemma proper_map_prod: + "proper_map (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y)) \ + topspace(prod_topology X Y) = {} \ proper_map X X' f \ proper_map Y Y' g" + (is "?lhs \ _ \ ?rhs") +proof (cases "topspace(prod_topology X Y) = {}") + case True + then show ?thesis + by (simp add: proper_map_on_empty) +next + case False + then have ne: "topspace X \ {}" "topspace Y \ {}" + by auto + define h where "h \ \(x,y). (f x, g y)" + have "proper_map X X' f" "proper_map Y Y' g" if ?lhs + proof - + have cm: "closed_map X X' f" "closed_map Y Y' g" + using that False closed_map_prod proper_imp_closed_map by blast+ + show "proper_map X X' f" + proof (clarsimp simp add: proper_map_def cm) + fix y + assume y: "y \ topspace X'" + obtain z where z: "z \ topspace Y" + using ne by blast + then have eq: "{x \ topspace X. f x = y} = + fst ` {u \ topspace X \ topspace Y. h u = (y,g z)}" + by (force simp: h_def) + show "compactin X {x \ topspace X. f x = y}" + unfolding eq + proof (intro image_compactin) + have "g z \ topspace Y'" + by (meson closed_map_def closedin_subset closedin_topspace cm image_subset_iff z) + with y show "compactin (prod_topology X Y) {u \ topspace X \ topspace Y. (h u) = (y, g z)}" + using that by (simp add: h_def proper_map_def) + show "continuous_map (prod_topology X Y) X fst" + by (simp add: continuous_map_fst) + qed + qed + show "proper_map Y Y' g" + proof (clarsimp simp add: proper_map_def cm) + fix y + assume y: "y \ topspace Y'" + obtain z where z: "z \ topspace X" + using ne by blast + then have eq: "{x \ topspace Y. g x = y} = + snd ` {u \ topspace X \ topspace Y. h u = (f z,y)}" + by (force simp: h_def) + show "compactin Y {x \ topspace Y. g x = y}" + unfolding eq + proof (intro image_compactin) + have "f z \ topspace X'" + by (meson closed_map_def closedin_subset closedin_topspace cm image_subset_iff z) + with y show "compactin (prod_topology X Y) {u \ topspace X \ topspace Y. (h u) = (f z, y)}" + using that by (simp add: proper_map_def h_def) + show "continuous_map (prod_topology X Y) Y snd" + by (simp add: continuous_map_snd) + qed + qed + qed + moreover + { assume R: ?rhs + then have fgim: "f ` topspace X \ topspace X'" "g ` topspace Y \ topspace Y'" + and cm: "closed_map X X' f" "closed_map Y Y' g" + by (auto simp: proper_map_def closed_map_imp_subset_topspace) + have "closed_map (prod_topology X Y) (prod_topology X' Y') h" + unfolding closed_map_fibre_neighbourhood imp_conjL + proof (intro conjI strip) + show "h ` topspace (prod_topology X Y) \ topspace (prod_topology X' Y')" + unfolding h_def using fgim by auto + fix W w + assume W: "openin (prod_topology X Y) W" + and w: "w \ topspace (prod_topology X' Y')" + and subW: "{x \ topspace (prod_topology X Y). h x = w} \ W" + then obtain x' y' where weq: "w = (x',y')" "x' \ topspace X'" "y' \ topspace Y'" + by auto + have eq: "{u \ topspace X \ topspace Y. h u = (x',y')} = {x \ topspace X. f x = x'} \ {y \ topspace Y. g y = y'}" + by (auto simp: h_def) + obtain U V where "openin X U" "openin Y V" "U \ V \ W" + and U: "{x \ topspace X. f x = x'} \ U" + and V: "{x \ topspace Y. g x = y'} \ V" + proof (rule Wallace_theorem_prod_topology) + show "compactin X {x \ topspace X. f x = x'}" "compactin Y {x \ topspace Y. g x = y'}" + using R weq unfolding proper_map_def closed_map_fibre_neighbourhood by fastforce+ + show "{x \ topspace X. f x = x'} \ {x \ topspace Y. g x = y'} \ W" + using weq subW by (auto simp: h_def) + qed (use W in auto) + obtain U' where "openin X' U'" "x' \ U'" and U': "{x \ topspace X. f x \ U'} \ U" + using cm U \openin X U\ weq unfolding closed_map_fibre_neighbourhood by meson + obtain V' where "openin Y' V'" "y' \ V'" and V': "{x \ topspace Y. g x \ V'} \ V" + using cm V \openin Y V\ weq unfolding closed_map_fibre_neighbourhood by meson + show "\V. openin (prod_topology X' Y') V \ w \ V \ {x \ topspace (prod_topology X Y). h x \ V} \ W" + proof (intro conjI exI) + show "openin (prod_topology X' Y') (U' \ V')" + by (simp add: \openin X' U'\ \openin Y' V'\ openin_prod_Times_iff) + show "w \ U' \ V'" + using \x' \ U'\ \y' \ V'\ weq by blast + show "{x \ topspace (prod_topology X Y). h x \ U' \ V'} \ W" + using \U \ V \ W\ U' V' h_def by auto + qed + qed + moreover + have "compactin (prod_topology X Y) {u \ topspace X \ topspace Y. h u = (w, z)}" + if "w \ topspace X'" and "z \ topspace Y'" for w z + proof - + have eq: "{u \ topspace X \ topspace Y. h u = (w,z)} = + {u \ topspace X. f u = w} \ {y. y \ topspace Y \ g y = z}" + by (auto simp: h_def) + show ?thesis + using R that by (simp add: eq compactin_Times proper_map_def) + qed + ultimately have ?lhs + by (auto simp: h_def proper_map_def) + } + ultimately show ?thesis using False by metis +qed + +lemma proper_map_paired: + assumes "Hausdorff_space X \ proper_map X Y f \ proper_map X Z g \ + Hausdorff_space Y \ continuous_map X Y f \ proper_map X Z g \ + Hausdorff_space Z \ proper_map X Y f \ continuous_map X Z g" + shows "proper_map X (prod_topology Y Z) (\x. (f x,g x))" + using assms +proof (elim disjE conjE) + assume \
: "Hausdorff_space X" "proper_map X Y f" "proper_map X Z g" + have eq: "(\x. (f x, g x)) = (\(x, y). (f x, g y)) \ (\x. (x, x))" + by auto + show "proper_map X (prod_topology Y Z) (\x. (f x, g x))" + unfolding eq + proof (rule proper_map_compose) + show "proper_map X (prod_topology X X) (\x. (x,x))" + by (simp add: \
) + show "proper_map (prod_topology X X) (prod_topology Y Z) (\(x,y). (f x, g y))" + by (simp add: \
proper_map_prod) + qed +next + assume \
: "Hausdorff_space Y" "continuous_map X Y f" "proper_map X Z g" + have eq: "(\x. (f x, g x)) = (\(x,y). (x,g y)) \ (\x. (f x,x))" + by auto + show "proper_map X (prod_topology Y Z) (\x. (f x, g x))" + unfolding eq + proof (rule proper_map_compose) + show "proper_map X (prod_topology Y X) (\x. (f x,x))" + by (simp add: \
proper_map_paired_continuous_map_left) + show "proper_map (prod_topology Y X) (prod_topology Y Z) (\(x,y). (x,g y))" + by (simp add: \
proper_map_prod proper_map_id [unfolded id_def]) + qed +next + assume \
: "Hausdorff_space Z" "proper_map X Y f" "continuous_map X Z g" + have eq: "(\x. (f x, g x)) = (\(x,y). (f x,y)) \ (\x. (x,g x))" + by auto + show "proper_map X (prod_topology Y Z) (\x. (f x, g x))" + unfolding eq + proof (rule proper_map_compose) + show "proper_map X (prod_topology X Z) (\x. (x, g x))" + using \
proper_map_paired_continuous_map_right by auto + show "proper_map (prod_topology X Z) (prod_topology Y Z) (\(x,y). (f x,y))" + by (simp add: \
proper_map_prod proper_map_id [unfolded id_def]) + qed +qed + +lemma proper_map_pairwise: + assumes + "Hausdorff_space X \ proper_map X Y (fst \ f) \ proper_map X Z (snd \ f) \ + Hausdorff_space Y \ continuous_map X Y (fst \ f) \ proper_map X Z (snd \ f) \ + Hausdorff_space Z \ proper_map X Y (fst \ f) \ continuous_map X Z (snd \ f)" + shows "proper_map X (prod_topology Y Z) f" + using proper_map_paired [OF assms] by (simp add: o_def) + +lemma proper_map_from_composition_right: + assumes "Hausdorff_space Y" "proper_map X Z (g \ f)" and "continuous_map X Y f" + and contg: "continuous_map Y Z g" + shows "proper_map X Y f" +proof - + define YZ where "YZ \ subtopology (prod_topology Y Z) ((\x. (x, g x)) ` topspace Y)" + have "proper_map X Y (fst \ (\x. (f x, (g \ f) x)))" + proof (rule proper_map_compose) + have [simp]: "x \ topspace X \ f x \ topspace Y" for x + by (meson assms(3) continuous_map_def) + show "proper_map X YZ (\x. (f x, (g \ f) x))" + unfolding YZ_def + using assms + by (force intro!: proper_map_into_subtopology proper_map_paired simp: o_def image_iff)+ + show "proper_map YZ Y fst" + using contg + by (simp flip: homeomorphic_maps_graph add: YZ_def homeomorphic_maps_map homeomorphic_imp_proper_map) + qed + moreover have "fst \ (\x. (f x, (g \ f) x)) = f" + by auto + ultimately show ?thesis + by auto +qed + +lemma perfect_map_from_composition_right: + "\Hausdorff_space Y; perfect_map X Z (g \ f); + continuous_map X Y f; continuous_map Y Z g; f ` topspace X = topspace Y\ + \ perfect_map X Y f" + by (meson perfect_map_def proper_map_from_composition_right) + +lemma perfect_map_from_composition_right_inj: + "\perfect_map X Z (g \ f); f ` topspace X = topspace Y; + continuous_map X Y f; continuous_map Y Z g; inj_on g (topspace Y)\ + \ perfect_map X Y f" + by (meson continuous_map_image_subset_topspace perfect_map_def proper_map_from_composition_right_inj) + subsection \Regular spaces\ text \Regular spaces are *not* a priori assumed to be Hausdorff or $T_1$\ definition regular_space where "regular_space X \ \C a. closedin X C \ a \ topspace X - C \ (\U V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V)" lemma homeomorphic_regular_space_aux: assumes hom: "X homeomorphic_space Y" and X: "regular_space X" shows "regular_space Y" proof - obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" and fg: "(\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce show ?thesis unfolding regular_space_def proof clarify fix C a assume Y: "closedin Y C" "a \ topspace Y" and "a \ C" then obtain "closedin X (g ` C)" "g a \ topspace X" "g a \ g ` C" using \closedin Y C\ hmg homeomorphic_map_closedness_eq by (smt (verit, ccfv_SIG) fg homeomorphic_imp_surjective_map image_iff in_mono) then obtain S T where ST: "openin X S" "openin X T" "g a \ S" "g`C \ T" "disjnt S T" using X unfolding regular_space_def by (metis DiffI) then have "openin Y (f`S)" "openin Y (f`T)" by (meson hmf homeomorphic_map_openness_eq)+ moreover have "a \ f`S \ C \ f`T" using ST by (smt (verit, best) Y closedin_subset fg image_eqI subset_iff) moreover have "disjnt (f`S) (f`T)" using ST by (smt (verit, ccfv_SIG) disjnt_iff fg image_iff openin_subset subsetD) ultimately show "\U V. openin Y U \ openin Y V \ a \ U \ C \ V \ disjnt U V" by metis qed qed lemma homeomorphic_regular_space: "X homeomorphic_space Y \ (regular_space X \ regular_space Y)" by (meson homeomorphic_regular_space_aux homeomorphic_space_sym) lemma regular_space: "regular_space X \ (\C a. closedin X C \ a \ topspace X - C \ (\U. openin X U \ a \ U \ disjnt C (X closure_of U)))" unfolding regular_space_def proof (intro all_cong1 imp_cong refl ex_cong1) fix C a U assume C: "closedin X C \ a \ topspace X - C" show "(\V. openin X U \ openin X V \ a \ U \ C \ V \ disjnt U V) \ (openin X U \ a \ U \ disjnt C (X closure_of U))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (smt (verit, best) disjnt_iff in_closure_of subsetD) next assume R: ?rhs then have "disjnt U (topspace X - X closure_of U)" by (meson DiffD2 closure_of_subset disjnt_iff openin_subset subsetD) moreover have "C \ topspace X - X closure_of U" by (meson C DiffI R closedin_subset disjnt_iff subset_eq) ultimately show ?lhs using R by (rule_tac x="topspace X - X closure_of U" in exI) auto qed qed lemma neighbourhood_base_of_closedin: "neighbourhood_base_of (closedin X) X \ regular_space X" (is "?lhs=?rhs") proof - have "?lhs \ (\W x. openin X W \ x \ W \ (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ W)))" by (simp add: neighbourhood_base_of) also have "\ \ (\W x. closedin X W \ x \ topspace X - W \ (\U. openin X U \ (\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W)))" by (smt (verit) Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) also have "\ \ ?rhs" proof - have \
: "(\V. closedin X V \ x \ U \ U \ V \ V \ topspace X - W) \ (\V. openin X V \ x \ U \ W \ V \ disjnt U V)" (is "?lhs=?rhs") if "openin X U" "closedin X W" "x \ topspace X" "x \ W" for W U x proof assume ?lhs with \closedin X W\ show ?rhs unfolding closedin_def by (smt (verit) Diff_mono disjnt_Diff1 double_diff subset_eq) next assume ?rhs with \openin X U\ show ?lhs unfolding openin_closedin_eq disjnt_def by (smt (verit) Diff_Diff_Int Diff_disjoint Diff_eq_empty_iff Int_Diff inf.orderE) qed show ?thesis unfolding regular_space_def by (intro all_cong1 ex_cong1 imp_cong refl) (metis \
DiffE) qed finally show ?thesis . qed lemma regular_space_discrete_topology: "regular_space (discrete_topology S)" using neighbourhood_base_of_closedin neighbourhood_base_of_discrete_topology by fastforce lemma regular_space_subtopology: "regular_space X \ regular_space (subtopology X S)" unfolding regular_space_def openin_subtopology_alt closedin_subtopology_alt disjnt_iff by clarsimp (smt (verit, best) inf.orderE inf_le1 le_inf_iff) lemma regular_space_retraction_map_image: "\retraction_map X Y r; regular_space X\ \ regular_space Y" using hereditary_imp_retractive_property homeomorphic_regular_space regular_space_subtopology by blast lemma regular_t0_imp_Hausdorff_space: "\regular_space X; t0_space X\ \ Hausdorff_space X" apply (clarsimp simp: regular_space_def t0_space Hausdorff_space_def) by (metis disjnt_sym subsetD) lemma regular_t0_eq_Hausdorff_space: "regular_space X \ (t0_space X \ Hausdorff_space X)" using Hausdorff_imp_t0_space regular_t0_imp_Hausdorff_space by blast lemma regular_t1_imp_Hausdorff_space: "\regular_space X; t1_space X\ \ Hausdorff_space X" by (simp add: regular_t0_imp_Hausdorff_space t1_imp_t0_space) lemma regular_t1_eq_Hausdorff_space: "regular_space X \ t1_space X \ Hausdorff_space X" using regular_t0_imp_Hausdorff_space t1_imp_t0_space t1_or_Hausdorff_space by blast lemma compact_Hausdorff_imp_regular_space: assumes "compact_space X" "Hausdorff_space X" shows "regular_space X" unfolding regular_space_def proof clarify fix S a assume "closedin X S" and "a \ topspace X" and "a \ S" then show "\U V. openin X U \ openin X V \ a \ U \ S \ V \ disjnt U V" using assms unfolding Hausdorff_space_compact_sets by (metis closedin_compact_space compactin_sing disjnt_empty1 insert_subset disjnt_insert1) qed lemma regular_space_topspace_empty: "topspace X = {} \ regular_space X" by (simp add: Hausdorff_space_topspace_empty compact_Hausdorff_imp_regular_space compact_space_topspace_empty) lemma neighbourhood_base_of_closed_Hausdorff_space: "regular_space X \ Hausdorff_space X \ neighbourhood_base_of (\C. closedin X C \ Hausdorff_space(subtopology X C)) X" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (simp add: Hausdorff_space_subtopology neighbourhood_base_of_closedin) next assume ?rhs then show ?lhs by (metis (mono_tags, lifting) Hausdorff_space_closed_neighbourhood neighbourhood_base_of neighbourhood_base_of_closedin openin_topspace) qed lemma locally_compact_imp_kc_eq_Hausdorff_space: "neighbourhood_base_of (compactin X) X \ kc_space X \ Hausdorff_space X" by (metis Hausdorff_imp_kc_space kc_imp_t1_space kc_space_def neighbourhood_base_of_closedin neighbourhood_base_of_mono regular_t1_imp_Hausdorff_space) lemma regular_space_compact_closed_separation: assumes X: "regular_space X" and S: "compactin X S" and T: "closedin X T" and "disjnt S T" shows "\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V" proof (cases "S={}") case True then show ?thesis by (meson T closedin_def disjnt_empty1 empty_subsetI openin_empty openin_topspace) next case False then have "\U V. x \ S \ openin X U \ openin X V \ x \ U \ T \ V \ disjnt U V" for x using assms unfolding regular_space_def by (smt (verit) Diff_iff compactin_subset_topspace disjnt_iff subsetD) then obtain U V where UV: "\x. x \ S \ openin X (U x) \ openin X (V x) \ x \ (U x) \ T \ (V x) \ disjnt (U x) (V x)" by metis then obtain \ where "finite \" "\ \ U ` S" "S \ \ \" using S unfolding compactin_def by (smt (verit) UN_iff image_iff subsetI) then obtain K where "finite K" "K \ S" and K: "S \ \(U ` K)" by (metis finite_subset_image) show ?thesis proof (intro exI conjI) show "openin X (\(U ` K))" using \K \ S\ UV by blast show "openin X (\(V ` K))" using False K UV \K \ S\ \finite K\ by blast show "S \ \(U ` K)" by (simp add: K) show "T \ \(V ` K)" using UV \K \ S\ by blast show "disjnt (\(U ` K)) (\(V ` K))" by (smt (verit) Inter_iff UN_E UV \K \ S\ disjnt_iff image_eqI subset_iff) qed qed lemma regular_space_compact_closed_sets: "regular_space X \ (\S T. compactin X S \ closedin 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 using regular_space_compact_closed_separation by fastforce next assume R: ?rhs show ?lhs unfolding regular_space proof clarify fix S x assume "closedin X S" and "x \ topspace X" and "x \ S" then obtain U V where "openin X U \ openin X V \ {x} \ U \ S \ V \ disjnt U V" by (metis R compactin_sing disjnt_empty1 disjnt_insert1) then show "\U. openin X U \ x \ U \ disjnt S (X closure_of U)" by (smt (verit, best) disjnt_iff in_closure_of insert_subset subsetD) qed qed lemma regular_space_prod_topology: "regular_space (prod_topology X Y) \ topspace X = {} \ topspace Y = {} \ regular_space X \ regular_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis regular_space_retraction_map_image retraction_map_fst retraction_map_snd) next assume R: ?rhs show ?lhs proof (cases "topspace X = {} \ topspace Y = {}") case True then show ?thesis by (simp add: regular_space_topspace_empty) next case False then have "regular_space X" "regular_space Y" using R by auto show ?thesis unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of proof clarify fix W x y assume W: "openin (prod_topology X Y) W" and "(x,y) \ W" then obtain U V where U: "openin X U" "x \ U" and V: "openin Y V" "y \ V" and "U \ V \ W" by (metis openin_prod_topology_alt) obtain D1 C1 where 1: "openin X D1" "closedin X C1" "x \ D1" "D1 \ C1" "C1 \ U" by (metis \regular_space X\ U neighbourhood_base_of neighbourhood_base_of_closedin) obtain D2 C2 where 2: "openin Y D2" "closedin Y C2" "y \ D2" "D2 \ C2" "C2 \ V" by (metis \regular_space Y\ V neighbourhood_base_of neighbourhood_base_of_closedin) show "\U V. openin (prod_topology X Y) U \ closedin (prod_topology X Y) V \ (x,y) \ U \ U \ V \ V \ W" proof (intro conjI exI) show "openin (prod_topology X Y) (D1 \ D2)" by (simp add: 1 2 openin_prod_Times_iff) show "closedin (prod_topology X Y) (C1 \ C2)" by (simp add: 1 2 closedin_prod_Times_iff) qed (use 1 2 \U \ V \ W\ in auto) qed qed qed lemma regular_space_product_topology: "regular_space (product_topology X I) \ topspace (product_topology X I) = {} \ (\i \ I. regular_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (meson regular_space_retraction_map_image retraction_map_product_projection) next assume R: ?rhs show ?lhs proof (cases "topspace(product_topology X I) = {}") case True then show ?thesis by (simp add: regular_space_topspace_empty) next case False then obtain x where x: "x \ topspace (product_topology X I)" by blast define \ where "\ \ {Pi\<^sub>E I U |U. finite {i \ I. U i \ topspace (X i)} \ (\i\I. openin (X i) (U i))}" have oo: "openin (product_topology X I) = arbitrary union_of (\W. W \ \)" by (simp add: \_def openin_product_topology product_topology_base_alt) have "\U V. openin (product_topology X I) U \ closedin (product_topology X I) V \ x \ U \ U \ V \ V \ Pi\<^sub>E I W" if fin: "finite {i \ I. W i \ topspace (X i)}" and opeW: "\k. k \ I \ openin (X k) (W k)" and x: "x \ PiE I W" for W x proof - have "\i. i \ I \ \U V. openin (X i) U \ closedin (X i) V \ x i \ U \ U \ V \ V \ W i" by (metis False PiE_iff R neighbourhood_base_of neighbourhood_base_of_closedin opeW x) then obtain U C where UC: "\i. i \ I \ openin (X i) (U i) \ closedin (X i) (C i) \ x i \ U i \ U i \ C i \ C i \ W i" by metis define PI where "PI \ \V. PiE I (\i. if W i = topspace(X i) then topspace(X i) else V i)" show ?thesis proof (intro conjI exI) have "\i\I. W i \ topspace (X i) \ openin (X i) (U i)" using UC by force with fin show "openin (product_topology X I) (PI U)" by (simp add: Collect_mono_iff PI_def openin_PiE_gen rev_finite_subset) show "closedin (product_topology X I) (PI C)" by (simp add: UC closedin_product_topology PI_def) show "x \ PI U" using UC x by (fastforce simp: PI_def) show "PI U \ PI C" by (smt (verit) UC Orderings.order_eq_iff PiE_mono PI_def) show "PI C \ Pi\<^sub>E I W" by (simp add: UC PI_def subset_PiE) qed qed then have "neighbourhood_base_of (closedin (product_topology X I)) (product_topology X I)" unfolding neighbourhood_base_of_topology_base [OF oo] by (force simp: \_def) then show ?thesis by (simp add: neighbourhood_base_of_closedin) qed qed lemma closed_map_paired_gen_aux: assumes "regular_space Y" and f: "closed_map Z X f" and g: "closed_map Z Y g" and clo: "\y. y \ topspace X \ closedin Z {x \ topspace Z. f x = y}" and contg: "continuous_map Z Y g" shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" unfolding closed_map_def proof (intro strip) fix C assume "closedin Z C" then have "C \ topspace Z" by (simp add: closedin_subset) have "f ` topspace Z \ topspace X" "g ` topspace Z \ topspace Y" by (simp_all add: assms closed_map_imp_subset_topspace) show "closedin (prod_topology X Y) ((\x. (f x, g x)) ` C)" unfolding closedin_def topspace_prod_topology proof (intro conjI) have "closedin Y (g ` C)" using \closedin Z C\ assms(3) closed_map_def by blast with assms show "(\x. (f x, g x)) ` C \ topspace X \ topspace Y" using \C \ topspace Z\ \f ` topspace Z \ topspace X\ continuous_map_closedin subsetD by fastforce have *: "\T. openin (prod_topology X Y) T \ (y1,y2) \ T \ T \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" if "(y1,y2) \ (\x. (f x, g x)) ` C" and y1: "y1 \ topspace X" and y2: "y2 \ topspace Y" for y1 y2 proof - define A where "A \ topspace Z - (C \ {x \ topspace Z. f x = y1})" have A: "openin Z A" "{x \ topspace Z. g x = y2} \ A" using that \closedin Z C\ clo that(2) by (auto simp: A_def) obtain V0 where "openin Y V0 \ y2 \ V0" and UA: "{x \ topspace Z. g x \ V0} \ A" using g A y2 unfolding closed_map_fibre_neighbourhood by blast then obtain V V' where VV: "openin Y V \ closedin Y V' \ y2 \ V \ V \ V'" and "V' \ V0" by (metis (no_types, lifting) \regular_space Y\ neighbourhood_base_of neighbourhood_base_of_closedin) with UA have subA: "{x \ topspace Z. g x \ V'} \ A" by blast show ?thesis proof - define B where "B \ topspace Z - (C \ {x \ topspace Z. g x \ V'})" have "openin Z B" using VV \closedin Z C\ contg by (fastforce simp: B_def continuous_map_closedin) have "{x \ topspace Z. f x = y1} \ B" using A_def subA by (auto simp: A_def B_def) then obtain U where "openin X U" "y1 \ U" and subB: "{x \ topspace Z. f x \ U} \ B" using \openin Z B\ y1 f unfolding closed_map_fibre_neighbourhood by meson show ?thesis proof (intro conjI exI) show "openin (prod_topology X Y) (U \ V)" by (metis VV \openin X U\ openin_prod_Times_iff) show "(y1, y2) \ U \ V" by (simp add: VV \y1 \ U\) show "U \ V \ topspace X \ topspace Y - (\x. (f x, g x)) ` C" using VV \C \ topspace Z\ \openin X U\ subB by (force simp: image_iff B_def subset_iff dest: openin_subset) qed qed qed then show "openin (prod_topology X Y) (topspace X \ topspace Y - (\x. (f x, g x)) ` C)" by (smt (verit, ccfv_threshold) Diff_iff SigmaE openin_subopen) qed qed lemma closed_map_paired_gen: assumes f: "closed_map Z X f" and g: "closed_map Z Y g" and D: "(regular_space X \ continuous_map Z X f \ (\z \ topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y \ topspace X. closedin Z {x \ topspace Z. f x = y}))" (is "?RSX \ ?RSY") shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" using D proof assume RSX: ?RSX have eq: "(\x. (f x, g x)) = (\(x,y). (y,x)) \ (\x. (g x, f x))" by auto show ?thesis unfolding eq proof (rule closed_map_compose) show "closed_map Z (prod_topology Y X) (\x. (g x, f x))" using RSX closed_map_paired_gen_aux f g by fastforce show "closed_map (prod_topology Y X) (prod_topology X Y) (\(x, y). (y, x))" using homeomorphic_imp_closed_map homeomorphic_map_swap by blast qed qed (blast intro: assms closed_map_paired_gen_aux) lemma closed_map_paired: assumes "closed_map Z X f" and contf: "continuous_map Z X f" "closed_map Z Y g" and contg: "continuous_map Z Y g" and D: "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" shows "closed_map Z (prod_topology X Y) (\x. (f x, g x))" proof (rule closed_map_paired_gen) show "regular_space X \ continuous_map Z X f \ (\z\topspace Y. closedin Z {x \ topspace Z. g x = z}) \ regular_space Y \ continuous_map Z Y g \ (\y\topspace X. closedin Z {x \ topspace Z. f x = y})" using D contf contg by (smt (verit, del_insts) Collect_cong closedin_continuous_map_preimage t1_space_closedin_singleton singleton_iff) qed (use assms in auto) lemma closed_map_pairwise: assumes "closed_map Z X (fst \ f)" "continuous_map Z X (fst \ f)" "closed_map Z Y (snd \ f)" "continuous_map Z Y (snd \ f)" "t1_space X \ regular_space Y \ regular_space X \ t1_space Y" shows "closed_map Z (prod_topology X Y) f" proof - have "closed_map Z (prod_topology X Y) (\a. ((fst \ f) a, (snd \ f) a))" using assms closed_map_paired by blast then show ?thesis by auto qed +lemma continuous_imp_proper_map: + "\compact_space X; kc_space Y; continuous_map X Y f\ \ proper_map X Y f" + by (simp add: continuous_closed_imp_proper_map continuous_imp_closed_map_gen kc_imp_t1_space) + lemma tube_lemma_right: assumes W: "openin (prod_topology X Y) W" and C: "compactin Y C" and x: "x \ topspace X" and subW: "{x} \ C \ W" shows "\U V. openin X U \ openin Y V \ x \ U \ C \ V \ U \ V \ W" proof (cases "C = {}") case True with x show ?thesis by auto next case False have "\U V. openin X U \ openin Y V \ x \ U \ y \ V \ U \ V \ W" if "y \ C" for y using W openin_prod_topology_alt subW subsetD that by fastforce then obtain U V where UV: "\y. y \ C \ openin X (U y) \ openin Y (V y) \ x \ U y \ y \ V y \ U y \ V y \ W" by metis then obtain D where D: "finite D" "D \ C" "C \ \ (V ` D)" using compactinD [OF C, of "V`C"] by (smt (verit) UN_I finite_subset_image imageE subsetI) show ?thesis proof (intro exI conjI) show "openin X (\ (U ` D))" "openin Y (\ (V ` D))" using D False UV by blast+ show "x \ \ (U ` D)" "C \ \ (V ` D)" "\ (U ` D) \ \ (V ` D) \ W" using D UV by force+ qed qed lemma closed_map_fst: assumes "compact_space Y" shows "closed_map (prod_topology X Y) X fst" proof - have *: "{x \ topspace X \ topspace Y. fst x \ U} = U \ topspace Y" if "U \ topspace X" for U using that by force have **: "\U y. \openin (prod_topology X Y) U; y \ topspace X; {x \ topspace X \ topspace Y. fst x = y} \ U\ \ \V. openin X V \ y \ V \ V \ topspace Y \ U" - using tube_lemma_right[of X Y _ "topspace Y"] assms compact_space_def - by force + using tube_lemma_right[of X Y _ "topspace Y"] assms by (fastforce simp: compact_space_def) show ?thesis unfolding closed_map_fibre_neighbourhood by (force simp: * openin_subset cong: conj_cong intro: **) qed lemma closed_map_snd: assumes "compact_space X" shows "closed_map (prod_topology X Y) Y snd" proof - have "snd = fst o prod.swap" by force moreover have "closed_map (prod_topology X Y) Y (fst o prod.swap)" proof (rule closed_map_compose) show "closed_map (prod_topology X Y) (prod_topology Y X) prod.swap" by (metis (no_types, lifting) homeomorphic_imp_closed_map homeomorphic_map_eq homeomorphic_map_swap prod.swap_def split_beta) show "closed_map (prod_topology Y X) Y fst" by (simp add: closed_map_fst assms) qed ultimately show ?thesis by metis qed lemma closed_map_paired_closed_map_right: "\closed_map X Y f; regular_space X; \y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}\ \ closed_map X (prod_topology X Y) (\x. (x, f x))" by (rule closed_map_paired_gen [OF closed_map_id, unfolded id_def]) auto lemma closed_map_paired_closed_map_left: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "closed_map X (prod_topology Y X) (\x. (f x, x))" proof - have eq: "(\x. (f x, x)) = (\(x,y). (y,x)) \ (\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))" by (simp add: assms closed_map_paired_closed_map_right) show "closed_map (prod_topology X Y) (prod_topology Y X) (\(x, y). (y, x))" using homeomorphic_imp_closed_map homeomorphic_map_swap by blast qed qed lemma closed_map_imp_closed_graph: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "closedin (prod_topology X Y) ((\x. (x, f x)) ` topspace X)" using assms closed_map_def closed_map_paired_closed_map_right by blast lemma proper_map_paired_closed_map_right: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "proper_map X (prod_topology X Y) (\x. (x, f x))" by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_right) lemma proper_map_paired_closed_map_left: assumes "closed_map X Y f" "regular_space X" "\y. y \ topspace Y \ closedin X {x \ topspace X. f x = y}" shows "proper_map X (prod_topology Y X) (\x. (f x, x))" by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_left) proposition regular_space_continuous_proper_map_image: assumes "regular_space X" and contf: "continuous_map X Y f" and pmapf: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "regular_space Y" unfolding regular_space_def proof clarify fix C y assume "closedin Y C" and "y \ topspace Y" and "y \ C" have "closed_map X Y f" "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" using pmapf proper_map_def by force+ moreover have "closedin X {z \ topspace X. f z \ C}" using \closedin Y C\ contf continuous_map_closedin by fastforce moreover have "disjnt {z \ topspace X. f z = y} {z \ topspace X. f z \ C}" using \y \ C\ disjnt_iff by blast ultimately obtain U V where UV: "openin X U" "openin X V" "{z \ topspace X. f z = y} \ U" "{z \ topspace X. f z \ C} \ V" and dUV: "disjnt U V" using \y \ topspace Y\ \regular_space X\ unfolding regular_space_compact_closed_sets by meson have *: "\U T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)" using \closed_map X Y f\ unfolding closed_map_preimage_neighbourhood by blast obtain V1 where V1: "openin Y V1 \ y \ V1" and sub1: "{x \ topspace X. f x \ V1} \ U" using * [of U "{y}"] UV \y \ topspace Y\ by auto moreover obtain V2 where "openin Y V2 \ C \ V2" and sub2: "{x \ topspace X. f x \ V2} \ V" by (smt (verit, ccfv_SIG) * UV \closedin Y C\ closedin_subset mem_Collect_eq subset_iff) moreover have "disjnt V1 V2" proof - have "\x. \\x. x \ U \ x \ V; x \ V1; x \ V2\ \ False" by (smt (verit) V1 fim image_iff mem_Collect_eq openin_subset sub1 sub2 subsetD) with dUV show ?thesis by (auto simp: disjnt_iff) qed ultimately show "\U V. openin Y U \ openin Y V \ y \ U \ C \ V \ disjnt U V" by meson qed lemma regular_space_perfect_map_image: "\regular_space X; perfect_map X Y f\ \ regular_space Y" by (meson perfect_map_def regular_space_continuous_proper_map_image) proposition regular_space_perfect_map_image_eq: assumes "Hausdorff_space X" and perf: "perfect_map X Y f" shows "regular_space X \ regular_space Y" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using perf regular_space_perfect_map_image by blast next assume R: ?rhs have "continuous_map X Y f" "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" using perf by (auto simp: perfect_map_def) then have "closed_map X Y f" and preYf: "(\y \ topspace Y. compactin X {x \ topspace X. f x = y})" by (simp_all add: proper_map_def) show ?lhs unfolding regular_space_def proof clarify fix C x assume "closedin X C" and "x \ topspace X" and "x \ C" obtain U1 U2 where "openin X U1" "openin X U2" "{x} \ U1" and "disjnt U1 U2" and subV: "C \ {z \ topspace X. f z = f x} \ U2" proof (rule Hausdorff_space_compact_separation [of X "{x}" "C \ {z \ topspace X. f z = f x}", OF \Hausdorff_space X\]) show "compactin X {x}" by (simp add: \x \ topspace X\) show "compactin X (C \ {z \ topspace X. f z = f x})" using \closedin X C\ fim \x \ topspace X\ closed_Int_compactin preYf by fastforce show "disjnt {x} (C \ {z \ topspace X. f z = f x})" using \x \ C\ by force qed have "closedin Y (f ` (C - U2))" using \closed_map X Y f\ \closedin X C\ \openin X U2\ closed_map_def by blast moreover have "f x \ topspace Y - f ` (C - U2)" using \closedin X C\ \continuous_map X Y f\ \x \ topspace X\ closedin_subset continuous_map_def subV by fastforce ultimately obtain V1 V2 where VV: "openin Y V1" "openin Y V2" "f x \ V1" and subV2: "f ` (C - U2) \ V2" and "disjnt V1 V2" by (meson R regular_space_def) show "\U U'. openin X U \ openin X U' \ x \ U \ C \ U' \ disjnt U U'" proof (intro exI conjI) show "openin X (U1 \ {x \ topspace X. f x \ V1})" using VV(1) \continuous_map X Y f\ \openin X U1\ continuous_map by fastforce show "openin X (U2 \ {x \ topspace X. f x \ V2})" using VV(2) \continuous_map X Y f\ \openin X U2\ continuous_map by fastforce show "x \ U1 \ {x \ topspace X. f x \ V1}" using VV(3) \x \ topspace X\ \{x} \ U1\ by auto show "C \ U2 \ {x \ topspace X. f x \ V2}" using \closedin X C\ closedin_subset subV2 by auto show "disjnt (U1 \ {x \ topspace X. f x \ V1}) (U2 \ {x \ topspace X. f x \ V2})" using \disjnt U1 U2\ \disjnt V1 V2\ by (auto simp: disjnt_iff) qed qed qed subsection\Locally compact spaces\ definition locally_compact_space where "locally_compact_space X \ \x \ topspace X. \U K. openin X U \ compactin X K \ x \ U \ U \ K" lemma homeomorphic_locally_compact_spaceD: assumes X: "locally_compact_space X" and "X homeomorphic_space Y" shows "locally_compact_space Y" proof - obtain f where hmf: "homeomorphic_map X Y f" using assms homeomorphic_space by blast then have eq: "topspace Y = f ` (topspace X)" by (simp add: homeomorphic_imp_surjective_map) have "\V K. openin Y V \ compactin Y K \ f x \ V \ V \ K" if "x \ topspace X" "openin X U" "compactin X K" "x \ U" "U \ K" for x U K using that by (meson hmf homeomorphic_map_compactness_eq homeomorphic_map_openness_eq image_mono image_eqI) with X show ?thesis by (smt (verit) eq image_iff locally_compact_space_def) qed lemma homeomorphic_locally_compact_space: assumes "X homeomorphic_space Y" shows "locally_compact_space X \ locally_compact_space Y" by (meson assms homeomorphic_locally_compact_spaceD homeomorphic_space_sym) lemma locally_compact_space_retraction_map_image: assumes "retraction_map X Y r" and X: "locally_compact_space X" shows "locally_compact_space Y" proof - obtain s where s: "retraction_maps X Y r s" using assms retraction_map_def by blast obtain T where "T retract_of_space X" and Teq: "T = s ` topspace Y" using retraction_maps_section_image1 s by blast then obtain r where r: "continuous_map X (subtopology X T) r" "\x\T. r x = x" by (meson retract_of_space_def) have "locally_compact_space (subtopology X T)" unfolding locally_compact_space_def openin_subtopology_alt proof clarsimp fix x assume "x \ topspace X" "x \ T" obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" by (meson X \x \ topspace X\ locally_compact_space_def) then have "compactin (subtopology X T) (r ` K) \ T \ U \ r ` K" by (smt (verit) IntD1 image_compactin image_iff inf_le2 r subset_iff) then show "\U. openin X U \ (\K. compactin (subtopology X T) K \ x \ U \ T \ U \ K)" using UK by auto qed with Teq show ?thesis using homeomorphic_locally_compact_space retraction_maps_section_image2 s by blast qed lemma compact_imp_locally_compact_space: "compact_space X \ locally_compact_space X" using compact_space_def locally_compact_space_def by blast lemma neighbourhood_base_imp_locally_compact_space: "neighbourhood_base_of (compactin X) X \ locally_compact_space X" by (metis locally_compact_space_def neighbourhood_base_of openin_topspace) lemma locally_compact_imp_neighbourhood_base: assumes loc: "locally_compact_space X" and reg: "regular_space X" shows "neighbourhood_base_of (compactin X) X" unfolding neighbourhood_base_of proof clarify fix W x assume "openin X W" and "x \ W" then obtain U K where "openin X U" "compactin X K" "x \ U" "U \ K" by (metis loc locally_compact_space_def openin_subset subsetD) moreover have "openin X (U \ W) \ x \ U \ W" using \openin X W\ \x \ W\ \openin X U\ \x \ U\ by blast then have "\u' v. openin X u' \ closedin X v \ x \ u' \ u' \ v \ v \ U \ v \ W" using reg by (metis le_infE neighbourhood_base_of neighbourhood_base_of_closedin) then show "\U V. openin X U \ compactin X V \ x \ U \ U \ V \ V \ W" by (meson \U \ K\ \compactin X K\ closed_compactin subset_trans) qed lemma Hausdorff_regular: "\Hausdorff_space X; neighbourhood_base_of (compactin X) X\ \ regular_space X" by (metis compactin_imp_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono) lemma locally_compact_Hausdorff_imp_regular_space: assumes loc: "locally_compact_space X" and "Hausdorff_space X" shows "regular_space X" unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of proof clarify fix W x assume "openin X W" and "x \ W" then have "x \ topspace X" using openin_subset by blast then obtain U K where "openin X U" "compactin X K" and UK: "x \ U" "U \ K" by (meson loc locally_compact_space_def) with \Hausdorff_space X\ have "regular_space (subtopology X K)" using Hausdorff_space_subtopology compact_Hausdorff_imp_regular_space compact_space_subtopology by blast then have "\U' V'. openin (subtopology X K) U' \ closedin (subtopology X K) V' \ x \ U' \ U' \ V' \ V' \ K \ W" unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of by (meson IntI \U \ K\ \openin X W\ \x \ U\ \x \ W\ openin_subtopology_Int2 subsetD) then obtain V C where "openin X V" "closedin X C" and VC: "x \ K \ V" "K \ V \ K \ C" "K \ C \ K \ W" by (metis Int_commute closedin_subtopology openin_subtopology) show "\U V. openin X U \ closedin X V \ x \ U \ U \ V \ V \ W" proof (intro conjI exI) show "openin X (U \ V)" using \openin X U\ \openin X V\ by blast show "closedin X (K \ C)" using \closedin X C\ \compactin X K\ compactin_imp_closedin \Hausdorff_space X\ by blast qed (use UK VC in auto) qed lemma locally_compact_space_neighbourhood_base: "Hausdorff_space X \ regular_space X \ locally_compact_space X \ neighbourhood_base_of (compactin X) X" by (metis locally_compact_imp_neighbourhood_base locally_compact_Hausdorff_imp_regular_space neighbourhood_base_imp_locally_compact_space) lemma locally_compact_Hausdorff_or_regular: "locally_compact_space X \ (Hausdorff_space X \ regular_space X) \ locally_compact_space X \ regular_space X" using locally_compact_Hausdorff_imp_regular_space by blast lemma locally_compact_space_compact_closedin: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\x \ topspace X. \U K. openin X U \ compactin X K \ closedin X K \ x \ U \ U \ K)" using locally_compact_Hausdorff_or_regular unfolding locally_compact_space_def by (metis assms closed_compactin inf.absorb_iff2 le_infE neighbourhood_base_of neighbourhood_base_of_closedin) lemma locally_compact_space_compact_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\x \ topspace X. \U. openin X U \ compactin X (X closure_of U) \ x \ U)" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis assms closed_compactin closedin_closure_of closure_of_eq closure_of_mono locally_compact_space_compact_closedin) next assume ?rhs then show ?lhs by (meson closure_of_subset locally_compact_space_def openin_subset) qed lemma locally_compact_space_neighbourhood_base_closedin: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast with L have "neighbourhood_base_of (compactin X) X" by (simp add: locally_compact_imp_neighbourhood_base) with \regular_space X\ show ?rhs by (smt (verit, ccfv_threshold) closed_compactin neighbourhood_base_of subset_trans neighbourhood_base_of_closedin) next assume ?rhs then show ?lhs using neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_mono by blast qed lemma locally_compact_space_neighbourhood_base_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\T. compactin X (X closure_of T)) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast with L have "neighbourhood_base_of (\A. compactin X A \ closedin X A) X" using locally_compact_space_neighbourhood_base_closedin by blast then show ?rhs by (simp add: closure_of_closedin neighbourhood_base_of_mono) next assume ?rhs then show ?lhs unfolding locally_compact_space_def neighbourhood_base_of by (meson closure_of_subset openin_topspace subset_trans) qed lemma locally_compact_space_neighbourhood_base_open_closure_of: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ neighbourhood_base_of (\U. openin X U \ compactin X (X closure_of U)) X" (is "?lhs=?rhs") proof assume L: ?lhs then have "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space by blast then have "neighbourhood_base_of (\T. compactin X (X closure_of T)) X" using L locally_compact_space_neighbourhood_base_closure_of by auto with L show ?rhs unfolding neighbourhood_base_of by (meson closed_compactin closure_of_closure_of closure_of_eq closure_of_mono subset_trans) next assume ?rhs then show ?lhs unfolding locally_compact_space_def neighbourhood_base_of by (meson closure_of_subset openin_topspace subset_trans) qed lemma locally_compact_space_compact_closed_compact: assumes "Hausdorff_space X \ regular_space X" shows "locally_compact_space X \ (\K. compactin X K \ (\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L))" (is "?lhs=?rhs") proof assume L: ?lhs then obtain U L where UL: "\x \ topspace X. openin X (U x) \ compactin X (L x) \ closedin X (L x) \ x \ U x \ U x \ L x" unfolding locally_compact_space_compact_closedin [OF assms] by metis show ?rhs proof clarify fix K assume "compactin X K" then have "K \ topspace X" by (simp add: compactin_subset_topspace) then have *: "(\U\U ` K. openin X U) \ K \ \ (U ` K)" using UL by blast with \compactin X K\ obtain KF where KF: "finite KF" "KF \ K" "K \ \(U ` KF)" by (metis compactinD finite_subset_image) show "\U L. openin X U \ compactin X L \ closedin X L \ K \ U \ U \ L" proof (intro conjI exI) show "openin X (\ (U ` KF))" using "*" \KF \ K\ by fastforce show "compactin X (\ (L ` KF))" by (smt (verit) UL \K \ topspace X\ KF compactin_Union finite_imageI imageE subset_iff) show "closedin X (\ (L ` KF))" by (smt (verit) UL \K \ topspace X\ KF closedin_Union finite_imageI imageE subsetD) qed (use UL \K \ topspace X\ KF in auto) qed next assume ?rhs then show ?lhs by (metis compactin_sing insert_subset locally_compact_space_def) qed lemma locally_compact_regular_space_neighbourhood_base: "locally_compact_space X \ regular_space X \ neighbourhood_base_of (\C. compactin X C \ closedin X C) X" using locally_compact_space_neighbourhood_base_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono by blast lemma locally_compact_kc_space: "neighbourhood_base_of (compactin X) X \ kc_space X \ locally_compact_space X \ Hausdorff_space X" using Hausdorff_imp_kc_space locally_compact_imp_kc_eq_Hausdorff_space locally_compact_space_neighbourhood_base by blast lemma locally_compact_kc_space_alt: "neighbourhood_base_of (compactin X) X \ kc_space X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" using Hausdorff_regular locally_compact_kc_space by blast lemma locally_compact_kc_imp_regular_space: "\neighbourhood_base_of (compactin X) X; kc_space X\ \ regular_space X" using Hausdorff_regular locally_compact_imp_kc_eq_Hausdorff_space by blast lemma kc_locally_compact_space: "kc_space X \ neighbourhood_base_of (compactin X) X \ locally_compact_space X \ Hausdorff_space X \ regular_space X" using Hausdorff_regular locally_compact_kc_space by blast lemma locally_compact_space_closed_subset: assumes loc: "locally_compact_space X" and "closedin X S" shows "locally_compact_space (subtopology X S)" proof (clarsimp simp: locally_compact_space_def) fix x assume x: "x \ topspace X" "x \ S" then obtain U K where UK: "openin X U \ compactin X K \ x \ U \ U \ K" by (meson loc locally_compact_space_def) show "\U. openin (subtopology X S) U \ (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" proof (intro conjI exI) show "openin (subtopology X S) (S \ U)" by (simp add: UK openin_subtopology_Int2) show "compactin (subtopology X S) (S \ K)" by (simp add: UK assms(2) closed_Int_compactin compactin_subtopology) qed (use UK x in auto) qed lemma locally_compact_space_open_subset: - assumes reg: "regular_space X" and loc: "locally_compact_space X" and "openin X S" + assumes X: "Hausdorff_space X \ regular_space X" and loc: "locally_compact_space X" and "openin X S" shows "locally_compact_space (subtopology X S)" proof (clarsimp simp: locally_compact_space_def) fix x assume x: "x \ topspace X" "x \ S" then obtain U K where UK: "openin X U" "compactin X K" "x \ U" "U \ K" by (meson loc locally_compact_space_def) - have "openin X (U \ S)" + moreover have reg: "regular_space X" + using X loc locally_compact_Hausdorff_imp_regular_space by blast + moreover have "openin X (U \ S)" by (simp add: UK \openin X S\ openin_Int) - with UK reg x obtain V C + ultimately obtain V C where VC: "openin X V" "closedin X C" "x \ V" "V \ C" "C \ U" "C \ S" - by (metis IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin) + by (metis \x \ S\ IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin) show "\U. openin (subtopology X S) U \ (\K. compactin (subtopology X S) K \ x \ U \ U \ K)" proof (intro conjI exI) show "openin (subtopology X S) V" using VC by (meson \openin X S\ openin_open_subtopology order_trans) show "compactin (subtopology X S) (C \ K)" using UK VC closed_Int_compactin compactin_subtopology by fastforce qed (use UK VC x in auto) qed lemma locally_compact_space_discrete_topology: "locally_compact_space (discrete_topology U)" by (simp add: neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_discrete_topology) lemma locally_compact_space_continuous_open_map_image: "\continuous_map X X' f; open_map X X' f; f ` topspace X = topspace X'; locally_compact_space X\ \ locally_compact_space X'" unfolding locally_compact_space_def open_map_def by (smt (verit, ccfv_SIG) image_compactin image_iff image_mono) lemma locally_compact_subspace_openin_closure_of: assumes "Hausdorff_space X" and S: "S \ topspace X" and loc: "locally_compact_space (subtopology X S)" shows "openin (subtopology X (X closure_of S)) S" unfolding openin_subopen [where S=S] proof clarify fix a assume "a \ S" then obtain T K where *: "openin X T" "compactin X K" "K \ S" "a \ S" "a \ T" "S \ T \ K" using loc unfolding locally_compact_space_def by (metis IntE S compactin_subtopology inf_commute openin_subtopology topspace_subtopology_subset) have "T \ X closure_of S \ X closure_of (T \ S)" by (simp add: "*"(1) openin_Int_closure_of_subset) also have "... \ S" using * \Hausdorff_space X\ by (metis closure_of_minimal compactin_imp_closedin order.trans inf_commute) finally have "T \ X closure_of S \ T \ S" by simp then have "openin (subtopology X (X closure_of S)) (T \ S)" unfolding openin_subtopology using \openin X T\ S closure_of_subset by fastforce with * show "\T. openin (subtopology X (X closure_of S)) T \ a \ T \ T \ S" by blast qed lemma locally_compact_subspace_closed_Int_openin: "\Hausdorff_space X \ S \ topspace X \ locally_compact_space(subtopology X S)\ \ \C U. closedin X C \ openin X U \ C \ U = S" by (metis closedin_closure_of inf_commute locally_compact_subspace_openin_closure_of openin_subtopology) lemma locally_compact_subspace_open_in_closure_of_eq: assumes "Hausdorff_space X" and loc: "locally_compact_space X" shows "openin (subtopology X (X closure_of S)) S \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") proof assume L: ?lhs then obtain "S \ topspace X" "regular_space X" using assms locally_compact_Hausdorff_imp_regular_space openin_subset by fastforce then have "locally_compact_space (subtopology (subtopology X (X closure_of S)) S)" by (simp add: L loc locally_compact_space_closed_subset locally_compact_space_open_subset regular_space_subtopology) then show ?rhs by (metis L inf.orderE inf_commute le_inf_iff openin_subset subtopology_subtopology topspace_subtopology) next assume ?rhs then show ?lhs using assms locally_compact_subspace_openin_closure_of by blast qed lemma locally_compact_subspace_closed_Int_openin_eq: assumes "Hausdorff_space X" and loc: "locally_compact_space X" shows "(\C U. closedin X C \ openin X U \ C \ U = S) \ S \ topspace X \ locally_compact_space(subtopology X S)" (is "?lhs=?rhs") proof assume L: ?lhs then obtain C U where "closedin X C" "openin X U" and Seq: "S = C \ U" by blast then have "C \ topspace X" by (simp add: closedin_subset) have "locally_compact_space (subtopology (subtopology X C) (topspace (subtopology X C) \ U))" - proof (rule locally_compact_space_open_subset) - show "regular_space (subtopology X C)" - by (simp add: \Hausdorff_space X\ loc locally_compact_Hausdorff_imp_regular_space regular_space_subtopology) - show "locally_compact_space (subtopology X C)" - by (simp add: \closedin X C\ loc locally_compact_space_closed_subset) - show "openin (subtopology X C) (topspace (subtopology X C) \ U)" - by (simp add: \openin X U\ Int_left_commute inf_commute openin_Int openin_subtopology_Int2) -qed - then show ?rhs - by (metis Seq \C \ topspace X\ inf.coboundedI1 subtopology_subtopology subtopology_topspace) + proof (rule locally_compact_space_open_subset) + show "locally_compact_space (subtopology X C)" + by (simp add: \closedin X C\ loc locally_compact_space_closed_subset) + show "openin (subtopology X C) (topspace (subtopology X C) \ U)" + by (simp add: \openin X U\ Int_left_commute inf_commute openin_Int openin_subtopology_Int2) + qed (simp add: Hausdorff_space_subtopology \Hausdorff_space X\) + then show ?rhs + by (metis Seq \C \ topspace X\ inf.coboundedI1 subtopology_subtopology subtopology_topspace) next assume ?rhs then show ?lhs - using assms locally_compact_subspace_closed_Int_openin by blast + using assms locally_compact_subspace_closed_Int_openin by blast qed lemma dense_locally_compact_openin_Hausdorff_space: "\Hausdorff_space X; S \ topspace X; X closure_of S = topspace X; locally_compact_space (subtopology X S)\ \ openin X S" by (metis locally_compact_subspace_openin_closure_of subtopology_topspace) lemma locally_compact_space_prod_topology: "locally_compact_space (prod_topology X Y) \ topspace (prod_topology X Y) = {} \ locally_compact_space X \ locally_compact_space Y" (is "?lhs=?rhs") proof (cases "topspace (prod_topology X Y) = {}") case True then show ?thesis unfolding locally_compact_space_def by blast next case False then obtain w z where wz: "w \ topspace X" "z \ topspace Y" by auto show ?thesis proof assume L: ?lhs then show ?rhs by (metis wz empty_iff locally_compact_space_retraction_map_image retraction_map_fst retraction_map_snd) next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix x y assume "x \ topspace X" and "y \ topspace Y" obtain U C where "openin X U" "compactin X C" "x \ U" "U \ C" by (meson False R \x \ topspace X\ locally_compact_space_def) obtain V D where "openin Y V" "compactin Y D" "y \ V" "V \ D" by (meson False R \y \ topspace Y\ locally_compact_space_def) show "\U. openin (prod_topology X Y) U \ (\K. compactin (prod_topology X Y) K \ (x, y) \ U \ U \ K)" proof (intro exI conjI) show "openin (prod_topology X Y) (U \ V)" by (simp add: \openin X U\ \openin Y V\ openin_prod_Times_iff) show "compactin (prod_topology X Y) (C \ D)" by (simp add: \compactin X C\ \compactin Y D\ compactin_Times) show "(x, y) \ U \ V" by (simp add: \x \ U\ \y \ V\) show "U \ V \ C \ D" by (simp add: Sigma_mono \U \ C\ \V \ D\) qed qed qed qed lemma locally_compact_space_product_topology: "locally_compact_space(product_topology X I) \ topspace(product_topology X I) = {} \ finite {i \ I. \ compact_space(X i)} \ (\i \ I. locally_compact_space(X i))" (is "?lhs=?rhs") proof (cases "topspace (product_topology X I) = {}") case True then show ?thesis by (simp add: locally_compact_space_def) next case False show ?thesis proof assume L: ?lhs obtain z where z: "z \ topspace (product_topology X I)" using False by auto with L z obtain U C where "openin (product_topology X I) U" "compactin (product_topology X I) C" "z \ U" "U \ C" by (meson locally_compact_space_def) then obtain V where finV: "finite {i \ I. V i \ topspace (X i)}" and "\i \ I. openin (X i) (V i)" and "z \ PiE I V" "PiE I V \ U" by (auto simp: openin_product_topology_alt) have "compact_space (X i)" if "i \ I" "V i = topspace (X i)" for i proof - have "compactin (X i) ((\x. x i) ` C)" using \compactin (product_topology X I) C\ image_compactin by (metis continuous_map_product_projection \i \ I\) moreover have "V i \ (\x. x i) ` C" proof - have "V i \ (\x. x i) ` Pi\<^sub>E I V" by (metis \z \ Pi\<^sub>E I V\ empty_iff image_projection_PiE order_refl \i \ I\) also have "\ \ (\x. x i) ` C" using \U \ C\ \Pi\<^sub>E I V \ U\ by blast finally show ?thesis . qed ultimately show ?thesis by (metis closed_compactin closedin_topspace compact_space_def that(2)) qed with finV have "finite {i \ I. \ compact_space (X i)}" by (metis (mono_tags, lifting) mem_Collect_eq finite_subset subsetI) moreover have "locally_compact_space (X i)" if "i \ I" for i by (meson False L locally_compact_space_retraction_map_image retraction_map_product_projection that) ultimately show ?rhs by metis next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix z assume z: "z \ (\\<^sub>E i\I. topspace (X i))" have "\U C. openin (X i) U \ compactin (X i) C \ z i \ U \ U \ C \ (compact_space(X i) \ U = topspace(X i) \ C = topspace(X i))" if "i \ I" for i using that R z unfolding locally_compact_space_def compact_space_def by (metis (no_types, lifting) False PiE_mem openin_topspace set_eq_subset) then obtain U C where UC: "\i. i \ I \ openin (X i) (U i) \ compactin (X i) (C i) \ z i \ U i \ U i \ C i \ (compact_space(X i) \ U i = topspace(X i) \ C i = topspace(X i))" by metis show "\U. openin (product_topology X I) U \ (\K. compactin (product_topology X I) K \ z \ U \ U \ K)" proof (intro exI conjI) show "openin (product_topology X I) (Pi\<^sub>E I U)" by (smt (verit) Collect_cong False R UC compactin_subspace openin_PiE_gen subset_antisym subtopology_topspace) show "compactin (product_topology X I) (Pi\<^sub>E I C)" by (simp add: UC compactin_PiE) qed (use UC z in blast)+ qed qed qed lemma locally_compact_space_sum_topology: "locally_compact_space (sum_topology X I) \ (\i \ I. locally_compact_space (X i))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (metis closed_map_component_injection embedding_map_imp_homeomorphic_space embedding_map_component_injection embedding_imp_closed_map_eq homeomorphic_locally_compact_space locally_compact_space_closed_subset) next assume R: ?rhs show ?lhs unfolding locally_compact_space_def proof clarsimp fix i y assume "i \ I" and y: "y \ topspace (X i)" then obtain U K where UK: "openin (X i) U" "compactin (X i) K" "y \ U" "U \ K" using R by (fastforce simp: locally_compact_space_def) then show "\U. openin (sum_topology X I) U \ (\K. compactin (sum_topology X I) K \ (i, y) \ U \ U \ K)" by (metis \i \ I\ continuous_map_component_injection image_compactin image_mono imageI open_map_component_injection open_map_def) qed qed proposition quotient_map_prod_right: assumes loc: "locally_compact_space Z" and reg: "Hausdorff_space Z \ regular_space Z" and f: "quotient_map X Y f" shows "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x,y). (x,f y))" proof - define h where "h \ (\(x::'a,y). (x,f y))" have "continuous_map (prod_topology Z X) Y (f o snd)" by (simp add: continuous_map_of_snd f quotient_imp_continuous_map) then have cmh: "continuous_map (prod_topology Z X) (prod_topology Z Y) h" by (simp add: h_def continuous_map_paired split_def continuous_map_fst o_def) have fim: "f ` topspace X = topspace Y" by (simp add: f quotient_imp_surjective_map) moreover have "openin (prod_topology Z X) {u \ topspace Z \ topspace X. h u \ W} \ openin (prod_topology Z Y) W" (is "?lhs=?rhs") if W: "W \ topspace Z \ topspace Y" for W proof define S where "S \ {u \ topspace Z \ topspace X. h u \ W}" assume ?lhs then have L: "openin (prod_topology Z X) S" using S_def by blast have "\T. openin (prod_topology Z Y) T \ (x0, z0) \ T \ T \ W" if \
: "(x0,z0) \ W" for x0 z0 proof - have x0: "x0 \ topspace Z" using W that by blast obtain y0 where y0: "y0 \ topspace X" "f y0 = z0" by (metis W fim imageE insert_absorb insert_subset mem_Sigma_iff \
) then have "(x0, y0) \ S" by (simp add: S_def h_def that x0) have "continuous_map Z (prod_topology Z X) (\x. (x, y0))" by (simp add: continuous_map_paired y0) with openin_continuous_map_preimage [OF _ L] have ope_ZS: "openin Z {x \ topspace Z. (x,y0) \ S}" by blast obtain U U' where "openin Z U" "compactin Z U'" "closedin Z U'" "x0 \ U" "U \ U'" "U' \ {x \ topspace Z. (x,y0) \ S}" using loc ope_ZS x0 \(x0, y0) \ S\ by (force simp: locally_compact_space_neighbourhood_base_closedin [OF reg] neighbourhood_base_of) then have D: "U' \ {y0} \ S" by (auto simp: ) define V where "V \ {z \ topspace Y. U' \ {y \ topspace X. f y = z} \ S}" have "z0 \ V" using D y0 Int_Collect fim by (fastforce simp: h_def V_def S_def) have "openin X {x \ topspace X. f x \ V} \ openin Y V" using f unfolding V_def quotient_map_def subset_iff by (smt (verit, del_insts) Collect_cong mem_Collect_eq) moreover have "openin X {x \ topspace X. f x \ V}" proof - let ?Z = "subtopology Z U'" have *: "{x \ topspace X. f x \ V} = topspace X - snd ` (U' \ topspace X - S)" by (force simp: V_def S_def h_def simp flip: fim) have "compact_space ?Z" using \compactin Z U'\ compactin_subspace by auto moreover have "closedin (prod_topology ?Z X) (U' \ topspace X - S)" by (simp add: L \closedin Z U'\ closedin_closed_subtopology closedin_diff closedin_prod_Times_iff prod_topology_subtopology(1)) ultimately show ?thesis using "*" closed_map_snd closed_map_def by fastforce qed ultimately have "openin Y V" by metis show ?thesis proof (intro conjI exI) show "openin (prod_topology Z Y) (U \ V)" by (simp add: openin_prod_Times_iff \openin Z U\ \openin Y V\) show "(x0, z0) \ U \ V" by (simp add: \x0 \ U\ \z0 \ V\) show "U \ V \ W" using \U \ U'\ by (force simp: V_def S_def h_def simp flip: fim) qed qed with openin_subopen show ?rhs by force next assume ?rhs then show ?lhs using openin_continuous_map_preimage cmh by fastforce qed ultimately show ?thesis by (fastforce simp: image_iff quotient_map_def h_def) qed lemma quotient_map_prod_left: assumes loc: "locally_compact_space Z" and reg: "Hausdorff_space Z \ regular_space Z" and f: "quotient_map X Y f" shows "quotient_map (prod_topology X Z) (prod_topology Y Z) (\(x,y). (f x,y))" proof - have "(\(x,y). (f x,y)) = prod.swap \ (\(x,y). (x,f y)) \ prod.swap" by force then show ?thesis apply (rule ssubst) proof (intro quotient_map_compose) show "quotient_map (prod_topology X Z) (prod_topology Z X) prod.swap" "quotient_map (prod_topology Z Y) (prod_topology Y Z) prod.swap" using homeomorphic_map_def homeomorphic_map_swap quotient_map_eq by fastforce+ show "quotient_map (prod_topology Z X) (prod_topology Z Y) (\(x, y). (x, f y))" by (simp add: f loc quotient_map_prod_right reg) qed qed lemma locally_compact_space_perfect_map_preimage: assumes "locally_compact_space X'" and f: "perfect_map X X' f" shows "locally_compact_space X" unfolding locally_compact_space_def proof (intro strip) fix x assume x: "x \ topspace X" then obtain U K where "openin X' U" "compactin X' K" "f x \ U" "U \ K" using assms unfolding locally_compact_space_def perfect_map_def by (metis (no_types, lifting) continuous_map_closedin) show "\U K. openin X U \ compactin X K \ x \ U \ U \ K" proof (intro exI conjI) have "continuous_map X X' f" using f perfect_map_def by blast then show "openin X {x \ topspace X. f x \ U}" by (simp add: \openin X' U\ continuous_map) show "compactin X {x \ topspace X. f x \ K}" using \compactin X' K\ f perfect_imp_proper_map proper_map_alt by blast qed (use x \f x \ U\ \U \ K\ in auto) qed subsection\Special characterizations of classes of functions into and out of R\ lemma monotone_map_into_euclideanreal_alt: assumes "continuous_map X euclideanreal f" shows "(\k. is_interval k \ connectedin X {x \ topspace X. f x \ k}) \ connected_space X \ monotone_map X euclideanreal f" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof show "connected_space X" using L connected_space_subconnected by blast have "connectedin X {x \ topspace X. f x \ {y}}" for y by (metis L is_interval_1 nle_le singletonD) then show "monotone_map X euclideanreal f" by (simp add: monotone_map) qed next assume R: ?rhs then have *: False if "a < b" "closedin X U" "closedin X V" "U \ {}" "V \ {}" "disjnt U V" and UV: "{x \ topspace X. f x \ {a..b}} = U \ V" and dis: "disjnt U {x \ topspace X. f x = b}" "disjnt V {x \ topspace X. f x = a}" for a b U V proof - define E1 where "E1 \ U \ {x \ topspace X. f x \ {c. c \ a}}" define E2 where "E2 \ V \ {x \ topspace X. f x \ {c. b \ c}}" have "closedin X {x \ topspace X. f x \ a}" "closedin X {x \ topspace X. b \ f x}" using assms continuous_map_upper_lower_semicontinuous_le by blast+ then have "closedin X E1" "closedin X E2" unfolding E1_def E2_def using that by auto moreover have "E1 \ E2 = {}" unfolding E1_def E2_def using \a \disjnt U V\ dis UV by (simp add: disjnt_def set_eq_iff) (smt (verit)) have "topspace X \ E1 \ E2" unfolding E1_def E2_def using UV by fastforce have "E1 = {} \ E2 = {}" using R connected_space_closedin using \E1 \ E2 = {}\ \closedin X E1\ \closedin X E2\ \topspace X \ E1 \ E2\ by blast then show False using E1_def E2_def \U \ {}\ \V \ {}\ by fastforce qed show ?lhs proof (intro strip) fix K :: "real set" assume "is_interval K" have False if "a \ K" "b \ K" and clo: "closedin X U" "closedin X V" and UV: "{x. x \ topspace X \ f x \ K} \ U \ V" "U \ V \ {x. x \ topspace X \ f x \ K} = {}" and nondis: "\ disjnt U {x. x \ topspace X \ f x = a}" "\ disjnt V {x. x \ topspace X \ f x = b}" for a b U V proof - have "\y. connectedin X {x. x \ topspace X \ f x = y}" using R monotone_map by fastforce then have **: False if "p \ U \ q \ V \ f p = f q \ f q \ K" for p q unfolding connectedin_closedin using \a \ K\ \b \ K\ UV clo that by (smt (verit, ccfv_threshold) closedin_subset disjoint_iff mem_Collect_eq subset_iff) consider "a < b" | "a = b" | "b < a" by linarith then show ?thesis proof cases case 1 define W where "W \ {x \ topspace X. f x \ {a..b}}" have "closedin X W" unfolding W_def by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) show ?thesis proof (rule * [OF 1 , of "U \ W" "V \ W"]) show "closedin X (U \ W)" "closedin X (V \ W)" using \closedin X W\ clo by auto show "U \ W \ {}" "V \ W \ {}" using nondis 1 by (auto simp: disjnt_iff W_def) show "disjnt (U \ W) (V \ W)" using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) have "\x. \x \ topspace X; a \ f x; f x \ b\ \ x \ U \ x \ V" using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff by blast then show "{x \ topspace X. f x \ {a..b}} = U \ W \ V \ W" by (auto simp: W_def) show "disjnt (U \ W) {x \ topspace X. f x = b}" "disjnt (V \ W) {x \ topspace X. f x = a}" using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ qed next case 2 then show ?thesis using ** nondis \b \ K\ by (force simp add: disjnt_iff) next case 3 define W where "W \ {x \ topspace X. f x \ {b..a}}" have "closedin X W" unfolding W_def by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) show ?thesis proof (rule * [OF 3, of "V \ W" "U \ W"]) show "closedin X (U \ W)" "closedin X (V \ W)" using \closedin X W\ clo by auto show "U \ W \ {}" "V \ W \ {}" using nondis 3 by (auto simp: disjnt_iff W_def) show "disjnt (V \ W) (U \ W)" using \is_interval K\ unfolding is_interval_1 disjnt_iff W_def by (metis (mono_tags, lifting) \a \ K\ \b \ K\ ** Int_Collect atLeastAtMost_iff) have "\x. \x \ topspace X; b \ f x; f x \ a\ \ x \ U \ x \ V" using \a \ K\ \b \ K\ \is_interval K\ UV unfolding is_interval_1 disjnt_iff by blast then show "{x \ topspace X. f x \ {b..a}} = V \ W \ U \ W" by (auto simp: W_def) show "disjnt (V \ W) {x \ topspace X. f x = a}" "disjnt (U \ W) {x \ topspace X. f x = b}" using ** \a \ K\ \b \ K\ nondis by (force simp: disjnt_iff)+ qed qed qed then show "connectedin X {x \ topspace X. f x \ K}" unfolding connectedin_closedin disjnt_iff by blast qed qed lemma monotone_map_into_euclideanreal: "\connected_space X; continuous_map X euclideanreal f\ \ monotone_map X euclideanreal f \ (\k. is_interval k \ connectedin X {x \ topspace X. f x \ k})" by (simp add: monotone_map_into_euclideanreal_alt) lemma monotone_map_euclideanreal_alt: "(\I::real set. is_interval I \ is_interval {x::real. x \ S \ f x \ I}) \ is_interval S \ (mono_on S f \ antimono_on S f)" (is "?lhs=?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof show "is_interval S" using L is_interval_1 by auto have False if "a \ S" "b \ S" "c \ S" "a f c < f b \ f a > f b \ f c > f b" for a b c using d proof assume "f a < f b \ f c < f b" then show False using L [of "{y. y < f b}"] unfolding is_interval_1 by (smt (verit, best) mem_Collect_eq that) next assume "f b < f a \ f b < f c" then show False using L [of "{y. y > f b}"] unfolding is_interval_1 by (smt (verit, best) mem_Collect_eq that) qed then show "mono_on S f \ monotone_on S (\) (\) f" unfolding monotone_on_def by (smt (verit)) qed next assume ?rhs then show ?lhs unfolding is_interval_1 monotone_on_def by simp meson qed lemma monotone_map_euclideanreal: fixes S :: "real set" shows "\is_interval S; continuous_on S f\ \ monotone_map (top_of_set S) euclideanreal f \ (mono_on S f \ monotone_on S (\) (\) f)" using monotone_map_euclideanreal_alt by (simp add: monotone_map_into_euclideanreal connectedin_subtopology is_interval_connected_1) lemma injective_eq_monotone_map: fixes f :: "real \ real" assumes "is_interval S" "continuous_on S f" shows "inj_on f S \ strict_mono_on S f \ strict_antimono_on S f" by (metis assms injective_imp_monotone_map monotone_map_euclideanreal strict_antimono_iff_antimono strict_mono_iff_mono top_greatest topspace_euclidean topspace_euclidean_subtopology) subsection\Normal spaces including Urysohn's lemma and the Tietze extension theorem\ definition normal_space where "normal_space X \ \S T. closedin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V)" lemma normal_space_retraction_map_image: assumes r: "retraction_map X Y r" and X: "normal_space X" shows "normal_space Y" unfolding normal_space_def proof clarify fix S T assume "closedin Y S" and "closedin Y T" and "disjnt S T" obtain r' where r': "retraction_maps X Y r r'" using r retraction_map_def by blast have "closedin X {x \ topspace X. r x \ S}" "closedin X {x \ topspace X. r x \ T}" using closedin_continuous_map_preimage \closedin Y S\ \closedin Y T\ r' by (auto simp: retraction_maps_def) moreover have "disjnt {x \ topspace X. r x \ S} {x \ topspace X. r x \ T}" using \disjnt S T\ by (auto simp: disjnt_def) ultimately obtain U V where UV: "openin X U \ openin X V \ {x \ topspace X. r x \ S} \ U \ {x \ topspace X. r x \ T} \ V" "disjnt U V" by (meson X normal_space_def) show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" proof (intro exI conjI) show "openin Y {x \ topspace Y. r' x \ U}" "openin Y {x \ topspace Y. r' x \ V}" using openin_continuous_map_preimage UV r' by (auto simp: retraction_maps_def) show "S \ {x \ topspace Y. r' x \ U}" "T \ {x \ topspace Y. r' x \ V}" using openin_continuous_map_preimage UV r' \closedin Y S\ \closedin Y T\ by (auto simp add: closedin_def continuous_map_closedin retraction_maps_def subset_iff) show "disjnt {x \ topspace Y. r' x \ U} {x \ topspace Y. r' x \ V}" using \disjnt U V\ by (auto simp: disjnt_def) qed qed lemma homeomorphic_normal_space: "X homeomorphic_space Y \ normal_space X \ normal_space Y" unfolding homeomorphic_space_def by (meson homeomorphic_imp_retraction_maps homeomorphic_maps_sym normal_space_retraction_map_image retraction_map_def) lemma normal_space: "normal_space X \ (\S T. closedin X S \ closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U)))" proof - have "(\V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V) \ openin X U \ S \ U \ disjnt T (X closure_of U)" (is "?lhs=?rhs") if "closedin X S" "closedin X T" "disjnt S T" for S T U proof show "?lhs \ ?rhs" by (smt (verit, best) disjnt_iff in_closure_of subsetD) assume R: ?rhs then have "(U \ S) \ (topspace X - X closure_of U) = {}" by (metis Diff_eq_empty_iff Int_Diff Int_Un_eq(4) closure_of_subset inf.orderE openin_subset) moreover have "T \ topspace X - X closure_of U" by (meson DiffI R closedin_subset disjnt_iff subsetD subsetI that(2)) ultimately show ?lhs by (metis R closedin_closure_of closedin_def disjnt_def sup.orderE) qed then show ?thesis unfolding normal_space_def by meson qed lemma normal_space_alt: "normal_space X \ (\S U. closedin X S \ openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U))" proof - have "\V. openin X V \ S \ V \ X closure_of V \ U" if "\T. closedin X T \ disjnt S T \ (\U. openin X U \ S \ U \ disjnt T (X closure_of U))" "closedin X S" "openin X U" "S \ U" for S U using that by (smt (verit) Diff_eq_empty_iff Int_Diff closure_of_subset_topspace disjnt_def inf.orderE inf_commute openin_closedin_eq) moreover have "\U. openin X U \ S \ U \ disjnt T (X closure_of U)" if "\U. openin X U \ S \ U \ (\V. openin X V \ S \ V \ X closure_of V \ U)" and "closedin X S" "closedin X T" "disjnt S T" for S T using that by (smt (verit) Diff_Diff_Int Diff_eq_empty_iff Int_Diff closedin_def disjnt_def inf.absorb_iff2 inf.orderE) ultimately show ?thesis by (fastforce simp: normal_space) qed lemma normal_space_closures: "normal_space X \ (\S T. S \ topspace X \ T \ topspace X \ disjnt (X closure_of S) (X closure_of T) \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt U V))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (meson closedin_closure_of closure_of_subset normal_space_def order.trans) show "?rhs \ ?lhs" by (metis closedin_subset closure_of_eq normal_space_def) qed lemma normal_space_disjoint_closures: "normal_space X \ (\S T. closedin X S \ closedin X T \ disjnt S T \ (\U V. openin X U \ openin X V \ S \ U \ T \ V \ disjnt (X closure_of U) (X closure_of V)))" (is "?lhs=?rhs") proof show "?lhs \ ?rhs" by (metis closedin_closure_of normal_space) show "?rhs \ ?lhs" by (smt (verit) closure_of_subset disjnt_iff normal_space openin_subset subset_eq) qed lemma normal_space_dual: "normal_space X \ (\U V. openin X U \ openin X V \ U \ V = topspace X \ (\S T. closedin X S \ closedin X T \ S \ U \ T \ V \ S \ T = topspace X))" (is "_ = ?rhs") proof - have "normal_space X \ (\U V. closedin X U \ closedin X V \ disjnt U V \ (\S T. \ (openin X S \ openin X T \ \ (U \ S \ V \ T \ disjnt S T))))" unfolding normal_space_def by meson also have "... \ (\U V. openin X U \ openin X V \ disjnt (topspace X - U) (topspace X - V) \ (\S T. \ (openin X S \ openin X T \ \ (topspace X - U \ S \ topspace X - V \ T \ disjnt S T))))" by (auto simp: all_closedin) also have "... \ ?rhs" proof - have *: "disjnt (topspace X - U) (topspace X - V) \ U \ V = topspace X" if "U \ topspace X" "V \ topspace X" for U V using that by (auto simp: disjnt_iff) show ?thesis using ex_closedin * apply (simp add: ex_closedin * [OF openin_subset openin_subset] cong: conj_cong) apply (intro all_cong1 ex_cong1 imp_cong refl) by (smt (verit, best) "*" Diff_Diff_Int Diff_subset Diff_subset_conv inf.orderE inf_commute openin_subset sup_commute) qed finally show ?thesis . qed lemma normal_t1_imp_Hausdorff_space: assumes "normal_space X" "t1_space X" shows "Hausdorff_space X" unfolding Hausdorff_space_def proof clarify fix x y assume xy: "x \ topspace X" "y \ topspace X" "x \ y" then have "disjnt {x} {y}" by (auto simp: disjnt_iff) then show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" using assms xy closedin_t1_singleton normal_space_def by (metis singletonI subsetD) qed lemma normal_t1_eq_Hausdorff_space: "normal_space X \ t1_space X \ Hausdorff_space X" using normal_t1_imp_Hausdorff_space t1_or_Hausdorff_space by blast lemma normal_t1_imp_regular_space: "\normal_space X; t1_space X\ \ regular_space X" by (metis compactin_imp_closedin normal_space_def normal_t1_eq_Hausdorff_space regular_space_compact_closed_sets) lemma compact_Hausdorff_or_regular_imp_normal_space: "\compact_space X; Hausdorff_space X \ regular_space X\ \ normal_space X" by (metis Hausdorff_space_compact_sets closedin_compact_space normal_space_def regular_space_compact_closed_sets) lemma normal_space_discrete_topology: "normal_space(discrete_topology U)" by (metis discrete_topology_closure_of inf_le2 normal_space_alt) lemma normal_space_fsigmas: "normal_space X \ (\S T. fsigma_in X S \ fsigma_in X T \ separatedin X S T \ (\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof clarify fix S T assume "fsigma_in X S" then obtain C where C: "\n. closedin X (C n)" "\n. C n \ C (Suc n)" "\ (range C) = S" by (meson fsigma_in_ascending) assume "fsigma_in X T" then obtain D where D: "\n. closedin X (D n)" "\n. D n \ D (Suc n)" "\ (range D) = T" by (meson fsigma_in_ascending) assume "separatedin X S T" have "\n. disjnt (D n) (X closure_of S)" by (metis D(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) then have "\n. \V V'. openin X V \ openin X V' \ D n \ V \ X closure_of S \ V' \ disjnt V V'" by (metis D(1) L closedin_closure_of normal_space_def) then obtain V V' where V: "\n. openin X (V n)" and "\n. openin X (V' n)" "\n. disjnt (V n) (V' n)" and DV: "\n. D n \ V n" and subV': "\n. X closure_of S \ V' n" by metis then have VV: "V' n \ X closure_of V n = {}" for n using openin_Int_closure_of_eq_empty [of X "V' n" "V n"] by (simp add: Int_commute disjnt_def) have "\n. disjnt (C n) (X closure_of T)" by (metis C(3) \separatedin X S T\ disjnt_Union1 disjnt_def rangeI separatedin_def) then have "\n. \U U'. openin X U \ openin X U' \ C n \ U \ X closure_of T \ U' \ disjnt U U'" by (metis C(1) L closedin_closure_of normal_space_def) then obtain U U' where U: "\n. openin X (U n)" and "\n. openin X (U' n)" "\n. disjnt (U n) (U' n)" and CU: "\n. C n \ U n" and subU': "\n. X closure_of T \ U' n" by metis then have UU: "U' n \ X closure_of U n = {}" for n using openin_Int_closure_of_eq_empty [of X "U' n" "U n"] by (simp add: Int_commute disjnt_def) show "\U B. openin X U \ openin X B \ S \ U \ T \ B \ disjnt U B" proof (intro conjI exI) have "\S n. closedin X (\m\n. X closure_of V m)" by (force intro: closedin_Union) then show "openin X (\n. U n - (\m\n. X closure_of V m))" using U by blast have "\S n. closedin X (\m\n. X closure_of U m)" by (force intro: closedin_Union) then show "openin X (\n. V n - (\m\n. X closure_of U m))" using V by blast have "S \ topspace X" by (simp add: \fsigma_in X S\ fsigma_in_subset) then show "S \ (\n. U n - (\m\n. X closure_of V m))" apply (clarsimp simp: Ball_def) by (metis VV C(3) CU IntI UN_E closure_of_subset empty_iff subV' subsetD) have "T \ topspace X" by (simp add: \fsigma_in X T\ fsigma_in_subset) then show "T \ (\n. V n - (\m\n. X closure_of U m))" apply (clarsimp simp: Ball_def) by (metis UU D(3) DV IntI UN_E closure_of_subset empty_iff subU' subsetD) have "\x m n. \x \ U n; x \ V m; \k\m. x \ X closure_of U k\ \ \k\n. x \ X closure_of V k" by (meson U V closure_of_subset nat_le_linear openin_subset subsetD) then show "disjnt (\n. U n - (\m\n. X closure_of V m)) (\n. V n - (\m\n. X closure_of U m))" by (force simp: disjnt_iff) qed qed next show "?rhs \ ?lhs" by (simp add: closed_imp_fsigma_in normal_space_def separatedin_closed_sets) qed lemma normal_space_fsigma_subtopology: assumes "normal_space X" "fsigma_in X S" shows "normal_space (subtopology X S)" unfolding normal_space_fsigmas proof clarify fix T U assume "fsigma_in (subtopology X S) T" and "fsigma_in (subtopology X S) U" and TU: "separatedin (subtopology X S) T U" then obtain A B where "openin X A \ openin X B \ T \ A \ U \ B \ disjnt A B" by (metis assms fsigma_in_fsigma_subtopology normal_space_fsigmas separatedin_subtopology) then show "\A B. openin (subtopology X S) A \ openin (subtopology X S) B \ T \ A \ U \ B \ disjnt A B" using TU by (force simp add: separatedin_subtopology openin_subtopology_alt disjnt_iff) qed lemma normal_space_closed_subtopology: assumes "normal_space X" "closedin X S" shows "normal_space (subtopology X S)" by (simp add: assms closed_imp_fsigma_in normal_space_fsigma_subtopology) lemma normal_space_continuous_closed_map_image: assumes "normal_space X" and contf: "continuous_map X Y f" and clof: "closed_map X Y f" and fim: "f ` topspace X = topspace Y" shows "normal_space Y" unfolding normal_space_def proof clarify fix S T assume "closedin Y S" and "closedin Y T" and "disjnt S T" have "closedin X {x \ topspace X. f x \ S}" "closedin X {x \ topspace X. f x \ T}" using \closedin Y S\ \closedin Y T\ closedin_continuous_map_preimage contf by auto moreover have "disjnt {x \ topspace X. f x \ S} {x \ topspace X. f x \ T}" using \disjnt S T\ by (auto simp: disjnt_iff) ultimately obtain U V where "closedin X U" "closedin X V" and subXU: "{x \ topspace X. f x \ S} \ topspace X - U" and subXV: "{x \ topspace X. f x \ T} \ topspace X - V" and dis: "disjnt (topspace X - U) (topspace X -V)" using \normal_space X\ by (force simp add: normal_space_def ex_openin) have "closedin Y (f ` U)" "closedin Y (f ` V)" using \closedin X U\ \closedin X V\ clof closed_map_def by blast+ moreover have "S \ topspace Y - f ` U" using \closedin Y S\ \closedin X U\ subXU by (force dest: closedin_subset) moreover have "T \ topspace Y - f ` V" using \closedin Y T\ \closedin X V\ subXV by (force dest: closedin_subset) moreover have "disjnt (topspace Y - f ` U) (topspace Y - f ` V)" using fim dis by (force simp add: disjnt_iff) ultimately show "\U V. openin Y U \ openin Y V \ S \ U \ T \ V \ disjnt U V" by (force simp add: ex_openin) qed subsection \Hereditary topological properties\ definition hereditarily where "hereditarily P X \ \S. S \ topspace X \ P(subtopology X S)" lemma hereditarily: "hereditarily P X \ (\S. P(subtopology X S))" by (metis Int_lower1 hereditarily_def subtopology_restrict) lemma hereditarily_mono: "\hereditarily P X; \x. P x \ Q x\ \ hereditarily Q X" by (simp add: hereditarily) lemma hereditarily_inc: "hereditarily P X \ P X" by (metis hereditarily subtopology_topspace) lemma hereditarily_subtopology: "hereditarily P X \ hereditarily P (subtopology X S)" by (simp add: hereditarily subtopology_subtopology) lemma hereditarily_normal_space_continuous_closed_map_image: assumes X: "hereditarily normal_space X" and contf: "continuous_map X Y f" and clof: "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" shows "hereditarily normal_space Y" unfolding hereditarily_def proof (intro strip) fix T assume "T \ topspace Y" then have nx: "normal_space (subtopology X {x \ topspace X. f x \ T})" by (meson X hereditarily) moreover have "continuous_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" by (simp add: contf continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) moreover have "closed_map (subtopology X {x \ topspace X. f x \ T}) (subtopology Y T) f" by (simp add: clof closed_map_restriction) ultimately show "normal_space (subtopology Y T)" using fim normal_space_continuous_closed_map_image by fastforce qed lemma homeomorphic_hereditarily_normal_space: "X homeomorphic_space Y \ (hereditarily normal_space X \ hereditarily normal_space Y)" by (meson hereditarily_normal_space_continuous_closed_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) lemma hereditarily_normal_space_retraction_map_image: "\retraction_map X Y r; hereditarily normal_space X\ \ hereditarily normal_space Y" by (smt (verit) hereditarily_subtopology hereditary_imp_retractive_property homeomorphic_hereditarily_normal_space) subsection\Limits in a topological space\ lemma limitin_const_iff: assumes "t1_space X" "\ trivial_limit F" shows "limitin X (\k. a) l F \ l \ topspace X \ a = l" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs using assms unfolding limitin_def t1_space_def by (metis eventually_const openin_topspace) next assume ?rhs then show ?lhs using assms by (auto simp: limitin_def t1_space_def) qed lemma compactin_sequence_with_limit: assumes lim: "limitin X \ l sequentially" and "S \ range \" and SX: "S \ topspace X" shows "compactin X (insert l S)" unfolding compactin_def proof (intro conjI strip) show "insert l S \ topspace X" by (meson SX insert_subset lim limitin_topspace) fix \ assume \
: "Ball \ (openin X) \ insert l S \ \ \" have "\V. finite V \ V \ \ \ (\t \ V. l \ t) \ S \ \ V" if *: "\x \ S. \T \ \. x \ T" and "T \ \" "l \ T" for T proof - obtain V where V: "\x. x \ S \ V x \ \ \ x \ V x" using * by metis obtain N where N: "\n. N \ n \ \ n \ T" by (meson "\
" \T \ \\ \l \ T\ lim limitin_sequentially) show ?thesis proof (intro conjI exI) have "x \ T" if "x \ S" and "\A. (\x \ S. (\n\N. x \ \ n) \ A \ V x) \ x \ A" for x by (metis (no_types) N V that assms(2) imageE nle_le subsetD) then show "S \ \ (insert T (V ` (S \ \ ` {0..N})))" by force qed (use V that in auto) qed then show "\\. finite \ \ \ \ \ \ insert l S \ \ \" by (smt (verit, best) Union_iff \
insert_subset subsetD) qed lemma limitin_Hausdorff_unique: assumes "limitin X f l1 F" "limitin X f l2 F" "\ trivial_limit F" "Hausdorff_space X" shows "l1 = l2" proof (rule ccontr) assume "l1 \ l2" with assms obtain U V where "openin X U" "openin X V" "l1 \ U" "l2 \ V" "disjnt U V" by (metis Hausdorff_space_def limitin_topspace) then have "eventually (\x. f x \ U) F" "eventually (\x. f x \ V) F" using assms by (fastforce simp: limitin_def)+ then have "\x. f x \ U \ f x \ V" using assms eventually_elim2 filter_eq_iff by fastforce with assms \disjnt U V\ show False by (meson disjnt_iff) qed lemma limitin_kc_unique: assumes "kc_space X" and lim1: "limitin X f l1 sequentially" and lim2: "limitin X f l2 sequentially" shows "l1 = l2" proof (rule ccontr) assume "l1 \ l2" define A where "A \ insert l1 (range f - {l2})" have "l1 \ topspace X" using lim1 limitin_def by fastforce moreover have "compactin X (insert l1 (topspace X \ (range f - {l2})))" by (meson Diff_subset compactin_sequence_with_limit inf_le1 inf_le2 lim1 subset_trans) ultimately have "compactin X (topspace X \ A)" by (simp add: A_def) then have OXA: "openin X (topspace X - A)" by (metis Diff_Diff_Int Diff_subset \kc_space X\ kc_space_def openin_closedin_eq) have "l2 \ topspace X - A" using \l1 \ l2\ A_def lim2 limitin_topspace by fastforce then have "\\<^sub>F x in sequentially. f x = l2" using limitinD [OF lim2 OXA] by (auto simp: A_def eventually_conj_iff) then show False using limitin_transform_eventually [OF _ lim1] limitin_const_iff [OF kc_imp_t1_space trivial_limit_sequentially] using \l1 \ l2\ \kc_space X\ by fastforce qed lemma limitin_closedin: assumes lim: "limitin X f l F" and "closedin X S" and ev: "eventually (\x. f x \ S) F" "\ trivial_limit F" shows "l \ S" proof (rule ccontr) assume "l \ S" have "\\<^sub>F x in F. f x \ topspace X - S" by (metis Diff_iff \l \ S\ \closedin X S\ closedin_def lim limitin_def) with ev eventually_elim2 trivial_limit_def show False by force qed subsection\Quasi-components\ definition quasi_component_of :: "'a topology \ 'a \ 'a \ bool" where "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ (\T. closedin X T \ openin X T \ (x \ T \ y \ T))" abbreviation "quasi_component_of_set S x \ Collect (quasi_component_of S x)" definition quasi_components_of :: "'a topology \ ('a set) set" where "quasi_components_of X = quasi_component_of_set X ` topspace X" lemma quasi_component_in_topspace: "quasi_component_of X x y \ x \ topspace X \ y \ topspace X" by (simp add: quasi_component_of_def) lemma quasi_component_of_refl [simp]: "quasi_component_of X x x \ x \ topspace X" by (simp add: quasi_component_of_def) lemma quasi_component_of_sym: "quasi_component_of X x y \ quasi_component_of X y x" by (meson quasi_component_of_def) lemma quasi_component_of_trans: "\quasi_component_of X x y; quasi_component_of X y z\ \ quasi_component_of X x z" by (simp add: quasi_component_of_def) lemma quasi_component_of_subset_topspace: "quasi_component_of_set X x \ topspace X" using quasi_component_of_def by fastforce lemma quasi_component_of_eq_empty: "quasi_component_of_set X x = {} \ (x \ topspace X)" using quasi_component_of_def by fastforce lemma quasi_component_of: "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ (\T. x \ T \ closedin X T \ openin X T \ y \ T)" - unfolding quasi_component_of_def by blast + unfolding quasi_component_of_def by (metis Diff_iff closedin_def openin_closedin_eq) lemma quasi_component_of_alt: "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ \ (\U V. openin X U \ openin X V \ U \ V = topspace X \ disjnt U V \ x \ U \ y \ V)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding quasi_component_of_def by (metis disjnt_iff separatedin_full separatedin_open_sets) show "?rhs \ ?lhs" unfolding quasi_component_of_def by (metis Diff_disjoint Diff_iff Un_Diff_cancel closedin_def disjnt_def inf_commute sup.orderE sup_commute) qed lemma quasi_component_of_separated: "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ \ (\U V. separatedin X U V \ U \ V = topspace X \ x \ U \ y \ V)" by (meson quasi_component_of_alt separatedin_full separatedin_open_sets) lemma quasi_component_of_subtopology: "quasi_component_of (subtopology X s) x y \ quasi_component_of X x y" unfolding quasi_component_of_def by (simp add: closedin_subtopology) (metis Int_iff inf_commute openin_subtopology_Int2) lemma quasi_component_of_mono: "quasi_component_of (subtopology X S) x y \ S \ T \ quasi_component_of (subtopology X T) x y" by (metis inf.absorb_iff2 quasi_component_of_subtopology subtopology_subtopology) lemma quasi_component_of_equiv: "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ quasi_component_of X x = quasi_component_of X y" using quasi_component_of_def by fastforce lemma quasi_component_of_disjoint [simp]: "disjnt (quasi_component_of_set X x) (quasi_component_of_set X y) \ \ (quasi_component_of X x y)" by (metis disjnt_iff quasi_component_of_equiv mem_Collect_eq) lemma quasi_component_of_eq: "quasi_component_of X x = quasi_component_of X y \ (x \ topspace X \ y \ topspace X) \ x \ topspace X \ y \ topspace X \ quasi_component_of X x y" by (metis Collect_empty_eq_bot quasi_component_of_eq_empty quasi_component_of_equiv) lemma topspace_imp_quasi_components_of: assumes "x \ topspace X" obtains C where "C \ quasi_components_of X" "x \ C" by (metis assms imageI mem_Collect_eq quasi_component_of_refl quasi_components_of_def) lemma Union_quasi_components_of: "\ (quasi_components_of X) = topspace X" by (auto simp: quasi_components_of_def quasi_component_of_def) lemma pairwise_disjoint_quasi_components_of: "pairwise disjnt (quasi_components_of X)" by (auto simp: quasi_components_of_def quasi_component_of_def disjoint_def) lemma complement_quasi_components_of_Union: assumes "C \ quasi_components_of X" shows "topspace X - C = \ (quasi_components_of X - {C})" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using Union_quasi_components_of by fastforce show "?rhs \ ?lhs" using assms using quasi_component_of_equiv by (fastforce simp add: quasi_components_of_def image_iff subset_iff) qed lemma nonempty_quasi_components_of: "C \ quasi_components_of X \ C \ {}" by (metis imageE quasi_component_of_eq_empty quasi_components_of_def) lemma quasi_components_of_subset: "C \ quasi_components_of X \ C \ topspace X" using Union_quasi_components_of by force lemma quasi_component_in_quasi_components_of: "quasi_component_of_set X a \ quasi_components_of X \ a \ topspace X" by (metis (no_types, lifting) image_iff quasi_component_of_eq_empty quasi_components_of_def) lemma quasi_components_of_eq_empty [simp]: "quasi_components_of X = {} \ topspace X = {}" by (simp add: quasi_components_of_def) lemma quasi_components_of_empty_space: "topspace X = {} \ quasi_components_of X = {}" by simp lemma quasi_component_of_set: "quasi_component_of_set X x = (if x \ topspace X then \ {t. closedin X t \ openin X t \ x \ t} else {})" by (auto simp: quasi_component_of) lemma closedin_quasi_component_of: "closedin X (quasi_component_of_set X x)" by (auto simp: quasi_component_of_set) lemma closedin_quasi_components_of: "C \ quasi_components_of X \ closedin X C" by (auto simp: quasi_components_of_def closedin_quasi_component_of) lemma openin_finite_quasi_components: "\finite(quasi_components_of X); C \ quasi_components_of X\ \ openin X C" apply (simp add:openin_closedin_eq quasi_components_of_subset complement_quasi_components_of_Union) by (meson DiffD1 closedin_Union closedin_quasi_components_of finite_Diff) lemma quasi_component_of_eq_overlap: "quasi_component_of X x = quasi_component_of X y \ (x \ topspace X \ y \ topspace X) \ \ (quasi_component_of_set X x \ quasi_component_of_set X y = {})" using quasi_component_of_equiv by fastforce lemma quasi_component_of_nonoverlap: "quasi_component_of_set X x \ quasi_component_of_set X y = {} \ (x \ topspace X) \ (y \ topspace X) \ \ (quasi_component_of X x = quasi_component_of X y)" by (metis inf.idem quasi_component_of_eq_empty quasi_component_of_eq_overlap) lemma quasi_component_of_overlap: "\ (quasi_component_of_set X x \ quasi_component_of_set X y = {}) \ x \ topspace X \ y \ topspace X \ quasi_component_of X x = quasi_component_of X y" by (meson quasi_component_of_nonoverlap) lemma quasi_components_of_disjoint: "\C \ quasi_components_of X; D \ quasi_components_of X\ \ disjnt C D \ C \ D" by (metis disjnt_self_iff_empty nonempty_quasi_components_of pairwiseD pairwise_disjoint_quasi_components_of) lemma quasi_components_of_overlap: "\C \ quasi_components_of X; D \ quasi_components_of X\ \ \ (C \ D = {}) \ C = D" by (metis disjnt_def quasi_components_of_disjoint) lemma pairwise_separated_quasi_components_of: "pairwise (separatedin X) (quasi_components_of X)" by (metis closedin_quasi_components_of pairwise_def pairwise_disjoint_quasi_components_of separatedin_closed_sets) lemma finite_quasi_components_of_finite: "finite(topspace X) \ finite(quasi_components_of X)" by (simp add: Union_quasi_components_of finite_UnionD) lemma connected_imp_quasi_component_of: assumes "connected_component_of X x y" shows "quasi_component_of X x y" proof - have "x \ topspace X" "y \ topspace X" by (meson assms connected_component_of_equiv)+ with assms show ?thesis apply (clarsimp simp add: quasi_component_of connected_component_of_def) by (meson connectedin_clopen_cases disjnt_iff subsetD) qed lemma connected_component_subset_quasi_component_of: "connected_component_of_set X x \ quasi_component_of_set X x" using connected_imp_quasi_component_of by force lemma quasi_component_as_connected_component_Union: "quasi_component_of_set X x = \ (connected_component_of_set X ` quasi_component_of_set X x)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using connected_component_of_refl quasi_component_of by fastforce show "?rhs \ ?lhs" apply (rule SUP_least) by (simp add: connected_component_subset_quasi_component_of quasi_component_of_equiv) qed lemma quasi_components_as_connected_components_Union: assumes "C \ quasi_components_of X" obtains \ where "\ \ connected_components_of X" "\\ = C" proof - obtain x where "x \ topspace X" and Ceq: "C = quasi_component_of_set X x" by (metis assms imageE quasi_components_of_def) define \ where "\ \ connected_component_of_set X ` quasi_component_of_set X x" show thesis proof show "\ \ connected_components_of X" by (simp add: \_def connected_components_of_def image_mono quasi_component_of_subset_topspace) show "\\ = C" by (metis \_def Ceq quasi_component_as_connected_component_Union) qed qed lemma path_imp_quasi_component_of: "path_component_of X x y \ quasi_component_of X x y" by (simp add: connected_imp_quasi_component_of path_imp_connected_component_of) lemma path_component_subset_quasi_component_of: "path_component_of_set X x \ quasi_component_of_set X x" by (simp add: Collect_mono path_imp_quasi_component_of) lemma connected_space_iff_quasi_component: "connected_space X \ (\x \ topspace X. \y \ topspace X. quasi_component_of X x y)" unfolding connected_space_clopen_in closedin_def quasi_component_of by blast lemma connected_space_imp_quasi_component_of: " \connected_space X; a \ topspace X; b \ topspace X\ \ quasi_component_of X a b" by (simp add: connected_space_iff_quasi_component) lemma connected_space_quasi_component_set: "connected_space X \ (\x \ topspace X. quasi_component_of_set X x = topspace X)" by (metis Ball_Collect connected_space_iff_quasi_component quasi_component_of_subset_topspace subset_antisym) lemma connected_space_iff_quasi_components_eq: "connected_space X \ (\C \ quasi_components_of X. \D \ quasi_components_of X. C = D)" apply (simp add: quasi_components_of_def) by (metis connected_space_iff_quasi_component mem_Collect_eq quasi_component_of_equiv) lemma quasi_components_of_subset_sing: "quasi_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "quasi_components_of X = {}") case True then show ?thesis by (simp add: connected_space_topspace_empty subset_singleton_iff) next case False then show ?thesis apply (simp add: connected_space_iff_quasi_components_eq subset_iff Ball_def) by (metis quasi_components_of_subset subsetI subset_antisym subset_empty topspace_imp_quasi_components_of) qed lemma connected_space_iff_quasi_components_subset_sing: "connected_space X \ (\a. quasi_components_of X \ {a})" by (simp add: quasi_components_of_subset_sing) lemma quasi_components_of_eq_singleton: "quasi_components_of X = {S} \ connected_space X \ \ (topspace X = {}) \ S = topspace X" by (metis ccpo_Sup_singleton insert_not_empty quasi_components_of_subset_sing subset_singleton_iff) lemma quasi_components_of_connected_space: "connected_space X \ quasi_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: quasi_components_of_eq_singleton) lemma separated_between_singletons: "separated_between X {x} {y} \ x \ topspace X \ y \ topspace X \ \ (quasi_component_of X x y)" proof (cases "x \ topspace X \ y \ topspace X") case True then show ?thesis by (auto simp add: separated_between_def quasi_component_of_alt) qed (use separated_between_imp_subset in blast) lemma quasi_component_nonseparated: "quasi_component_of X x y \ x \ topspace X \ y \ topspace X \ \ (separated_between X {x} {y})" by (metis quasi_component_of_equiv separated_between_singletons) lemma separated_between_quasi_component_pointwise_left: assumes "C \ quasi_components_of X" shows "separated_between X C S \ (\x \ C. separated_between X {x} S)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms quasi_components_of_disjoint separated_between_mono by fastforce next assume ?rhs then obtain y where "separated_between X {y} S" and "y \ C" by metis with assms show ?lhs by (force simp add: separated_between quasi_components_of_def quasi_component_of_def) qed lemma separated_between_quasi_component_pointwise_right: "C \ quasi_components_of X \ separated_between X S C \ (\x \ C. separated_between X S {x})" by (simp add: separated_between_quasi_component_pointwise_left separated_between_sym) lemma separated_between_quasi_component_point: assumes "C \ quasi_components_of X" shows "separated_between X C {x} \ x \ topspace X - C" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (meson DiffI disjnt_insert2 insert_subset separated_between_imp_disjoint separated_between_imp_subset) next assume ?rhs with assms show ?lhs unfolding quasi_components_of_def image_iff Diff_iff separated_between_quasi_component_pointwise_left [OF assms] by (metis mem_Collect_eq quasi_component_of_refl separated_between_singletons) qed lemma separated_between_point_quasi_component: "C \ quasi_components_of X \ separated_between X {x} C \ x \ topspace X - C" by (simp add: separated_between_quasi_component_point separated_between_sym) lemma separated_between_quasi_component_compact: "\C \ quasi_components_of X; compactin X K\ \ (separated_between X C K \ disjnt C K)" unfolding disjnt_iff using compactin_subset_topspace quasi_components_of_subset separated_between_pointwise_right separated_between_quasi_component_point by fastforce lemma separated_between_compact_quasi_component: "\compactin X K; C \ quasi_components_of X\ \ separated_between X K C \ disjnt K C" using disjnt_sym separated_between_quasi_component_compact separated_between_sym by blast lemma separated_between_quasi_components: assumes C: "C \ quasi_components_of X" and D: "D \ quasi_components_of X" shows "separated_between X C D \ disjnt C D" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: separated_between_imp_disjoint) next assume ?rhs obtain x y where x: "C = quasi_component_of_set X x" and "x \ C" and y: "D = quasi_component_of_set X y" and "y \ D" using assms by (auto simp: quasi_components_of_def) then have "separated_between X {x} {y}" using \disjnt C D\ separated_between_singletons by fastforce with \x \ C\ \y \ D\ show ?lhs by (auto simp: assms separated_between_quasi_component_pointwise_left separated_between_quasi_component_pointwise_right) qed lemma quasi_eq_connected_component_of_eq: "quasi_component_of X x = connected_component_of X x \ connectedin X (quasi_component_of_set X x)" (is "?lhs = ?rhs") proof (cases "x \ topspace X") case True show ?thesis proof show "?lhs \ ?rhs" by (simp add: connectedin_connected_component_of) next assume ?rhs then have "\y. quasi_component_of X x y = connected_component_of X x y" by (metis connected_component_of_def connected_imp_quasi_component_of mem_Collect_eq quasi_component_of_equiv) then show ?lhs by force qed next case False then show ?thesis by (metis Collect_empty_eq_bot connected_component_of_eq_empty connectedin_empty quasi_component_of_eq_empty) qed lemma connected_quasi_component_of: assumes "C \ quasi_components_of X" shows "C \ connected_components_of X \ connectedin X C" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (simp add: connectedin_connected_components_of) next assume ?rhs with assms show ?lhs unfolding quasi_components_of_def connected_components_of_def image_iff by (metis quasi_eq_connected_component_of_eq) qed lemma quasi_component_of_clopen_cases: "\C \ quasi_components_of X; closedin X T; openin X T\ \ C \ T \ disjnt C T" by (smt (verit) disjnt_iff image_iff mem_Collect_eq quasi_component_of_def quasi_components_of_def subset_iff) lemma quasi_components_of_set: assumes "C \ quasi_components_of X" shows "\ {T. closedin X T \ openin X T \ C \ T} = C" (is "?lhs = ?rhs") proof have "x \ C" if "x \ \ {T. closedin X T \ openin X T \ C \ T}" for x proof (rule ccontr) assume "x \ C" have "x \ topspace X" using assms quasi_components_of_subset that by force then have "separated_between X C {x}" by (simp add: \x \ C\ assms separated_between_quasi_component_point) with that show False by (auto simp: separated_between) qed then show "?lhs \ ?rhs" by auto qed blast lemma open_quasi_eq_connected_components_of: assumes "openin X C" shows "C \ quasi_components_of X \ C \ connected_components_of X" (is "?lhs = ?rhs") proof (cases "closedin X C") case True show ?thesis proof assume L: ?lhs have "T = {} \ T = topspace X \ C" if "openin (subtopology X C) T" "closedin (subtopology X C) T" for T proof - have "C \ T \ disjnt C T" by (meson L True assms closedin_trans_full openin_trans_full quasi_component_of_clopen_cases that) with that show ?thesis by (metis Int_absorb2 True closedin_imp_subset closure_of_subset_eq disjnt_def inf_absorb2) qed with L assms show "?rhs" by (simp add: connected_quasi_component_of connected_space_clopen_in connectedin_def openin_subset) next assume ?rhs then obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" by (metis connected_components_of_def imageE) have "C = quasi_component_of_set X x" using True assms connected_component_of_refl connected_imp_quasi_component_of quasi_component_of_def x by fastforce then show ?lhs using \x \ topspace X\ quasi_components_of_def by fastforce qed next case False then show ?thesis using closedin_connected_components_of closedin_quasi_components_of by blast qed lemma quasi_component_of_continuous_image: assumes f: "continuous_map X Y f" and qc: "quasi_component_of X x y" shows "quasi_component_of Y (f x) (f y)" unfolding quasi_component_of_def proof (intro strip conjI) show "f x \ topspace Y" "f y \ topspace Y" using assms by (simp_all add: continuous_map_def quasi_component_of_def) fix T assume "closedin Y T \ openin Y T" with assms show "(f x \ T) = (f y \ T)" by (smt (verit) continuous_map_closedin continuous_map_def mem_Collect_eq quasi_component_of_def) qed lemma quasi_component_of_discrete_topology: "quasi_component_of_set (discrete_topology U) x = (if x \ U then {x} else {})" proof - have "quasi_component_of_set (discrete_topology U) y = {y}" if "y \ U" for y using that apply (simp add: set_eq_iff quasi_component_of_def) by (metis Set.set_insert insertE subset_insertI) then show ?thesis by (simp add: quasi_component_of) qed lemma quasi_components_of_discrete_topology: "quasi_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp add: quasi_components_of_def quasi_component_of_discrete_topology) lemma homeomorphic_map_quasi_component_of: assumes hmf: "homeomorphic_map X Y f" and "x \ topspace X" shows "quasi_component_of_set Y (f x) = f ` (quasi_component_of_set X x)" proof - obtain g where hmg: "homeomorphic_map Y X g" and contf: "continuous_map X Y f" and contg: "continuous_map Y X g" and fg: "(\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" by (smt (verit, best) hmf homeomorphic_map_maps homeomorphic_maps_def) show ?thesis proof show "quasi_component_of_set Y (f x) \ f ` quasi_component_of_set X x" using quasi_component_of_continuous_image [OF contg] \x \ topspace X\ fg image_iff quasi_component_of_subset_topspace by fastforce show "f ` quasi_component_of_set X x \ quasi_component_of_set Y (f x)" using quasi_component_of_continuous_image [OF contf] by blast qed qed lemma homeomorphic_map_quasi_components_of: assumes "homeomorphic_map X Y f" shows "quasi_components_of Y = image (image f) (quasi_components_of X)" using assms proof - have "\x\topspace X. quasi_component_of_set Y y = f ` quasi_component_of_set X x" if "y \ topspace Y" for y by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of image_iff) moreover have "\x\topspace Y. f ` quasi_component_of_set X u = quasi_component_of_set Y x" if "u \ topspace X" for u by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of imageI) ultimately show ?thesis by (auto simp: quasi_components_of_def image_iff) qed lemma openin_quasi_component_of_locally_connected_space: assumes "locally_connected_space X" shows "openin X (quasi_component_of_set X x)" proof - have *: "openin X (connected_component_of_set X x)" by (simp add: assms openin_connected_component_of_locally_connected_space) moreover have "connected_component_of_set X x = quasi_component_of_set X x" using * closedin_connected_component_of connected_component_of_refl connected_imp_quasi_component_of quasi_component_of_def by fastforce ultimately show ?thesis by simp qed lemma openin_quasi_components_of_locally_connected_space: "locally_connected_space X \ c \ quasi_components_of X \ openin X c" by (smt (verit, best) image_iff openin_quasi_component_of_locally_connected_space quasi_components_of_def) lemma quasi_eq_connected_components_of_alt: "quasi_components_of X = connected_components_of X \ (\C \ quasi_components_of X. connectedin X C)" (is "?lhs = ?rhs") proof assume R: ?rhs moreover have "connected_components_of X \ quasi_components_of X" using R unfolding quasi_components_of_def connected_components_of_def by (force simp flip: quasi_eq_connected_component_of_eq) ultimately show ?lhs using connected_quasi_component_of by blast qed (use connected_quasi_component_of in blast) lemma connected_subset_quasi_components_of_pointwise: "connected_components_of X \ quasi_components_of X \ (\x \ topspace X. quasi_component_of X x = connected_component_of X x)" (is "?lhs = ?rhs") proof assume L: ?lhs have "connectedin X (quasi_component_of_set X x)" if "x \ topspace X" for x proof - have "\y\topspace X. connected_component_of_set X x = quasi_component_of_set X y" using L that by (force simp: quasi_components_of_def connected_components_of_def image_subset_iff) then show ?thesis by (metis connected_component_of_equiv connectedin_connected_component_of mem_Collect_eq quasi_component_of_eq) qed then show ?rhs by (simp add: quasi_eq_connected_component_of_eq) qed (simp add: connected_components_of_def quasi_components_of_def) lemma quasi_subset_connected_components_of_pointwise: "quasi_components_of X \ connected_components_of X \ (\x \ topspace X. quasi_component_of X x = connected_component_of X x)" by (simp add: connected_quasi_component_of image_subset_iff quasi_components_of_def quasi_eq_connected_component_of_eq) lemma quasi_eq_connected_components_of_pointwise: "quasi_components_of X = connected_components_of X \ (\x \ topspace X. quasi_component_of X x = connected_component_of X x)" using connected_subset_quasi_components_of_pointwise quasi_subset_connected_components_of_pointwise by fastforce lemma quasi_eq_connected_components_of_pointwise_alt: "quasi_components_of X = connected_components_of X \ (\x. quasi_component_of X x = connected_component_of X x)" unfolding quasi_eq_connected_components_of_pointwise by (metis connectedin_empty quasi_component_of_eq_empty quasi_eq_connected_component_of_eq) lemma quasi_eq_connected_components_of_inclusion: "quasi_components_of X = connected_components_of X \ connected_components_of X \ quasi_components_of X \ quasi_components_of X \ connected_components_of X" by (simp add: connected_subset_quasi_components_of_pointwise dual_order.eq_iff quasi_subset_connected_components_of_pointwise) lemma quasi_eq_connected_components_of: "finite(connected_components_of X) \ finite(quasi_components_of X) \ locally_connected_space X \ compact_space X \ (Hausdorff_space X \ regular_space X \ normal_space X) \ quasi_components_of X = connected_components_of X" proof (elim disjE) show "quasi_components_of X = connected_components_of X" if "finite (connected_components_of X)" unfolding quasi_eq_connected_components_of_inclusion using that open_in_finite_connected_components open_quasi_eq_connected_components_of by blast show "quasi_components_of X = connected_components_of X" if "finite (quasi_components_of X)" unfolding quasi_eq_connected_components_of_inclusion using that open_quasi_eq_connected_components_of openin_finite_quasi_components by blast show "quasi_components_of X = connected_components_of X" if "locally_connected_space X" unfolding quasi_eq_connected_components_of_inclusion using that open_quasi_eq_connected_components_of openin_quasi_components_of_locally_connected_space by auto show "quasi_components_of X = connected_components_of X" if "compact_space X \ (Hausdorff_space X \ regular_space X \ normal_space X)" proof - show ?thesis unfolding quasi_eq_connected_components_of_alt proof (intro strip) fix C assume C: "C \ quasi_components_of X" then have cloC: "closedin X C" by (simp add: closedin_quasi_components_of) have "normal_space X" using that compact_Hausdorff_or_regular_imp_normal_space by blast show "connectedin X C" proof (clarsimp simp add: connectedin_def connected_space_closedin_eq closedin_closed_subtopology cloC closedin_subset [OF cloC]) fix S T assume "S \ C" and "closedin X S" and "S \ T = {}" and SUT: "S \ T = topspace X \ C" and T: "T \ C" "T \ {}" and "closedin X T" with \normal_space X\ obtain U V where UV: "openin X U" "openin X V" "S \ U" "T \ V" "disjnt U V" by (meson disjnt_def normal_space_def) moreover have "compactin X (topspace X - (U \ V))" using UV that by (intro closedin_compact_space closedin_diff openin_Un) auto ultimately have "separated_between X C (topspace X - (U \ V)) \ disjnt C (topspace X - (U \ V))" by (simp add: \C \ quasi_components_of X\ separated_between_quasi_component_compact) moreover have "disjnt C (topspace X - (U \ V))" using UV SUT disjnt_def by fastforce ultimately have "separated_between X C (topspace X - (U \ V))" by simp then obtain A B where "openin X A" "openin X B" "A \ B = topspace X" "disjnt A B" "C \ A" and subB: "topspace X - (U \ V) \ B" by (meson separated_between_def) have "B \ U = topspace X - (A \ V)" proof show "B \ U \ topspace X - A \ V" using \openin X U\ \disjnt U V\ \disjnt A B\ \openin X B\ disjnt_iff openin_closedin_eq by fastforce show "topspace X - A \ V \ B \ U" using \A \ B = topspace X\ subB by fastforce qed then have "closedin X (B \ U)" using \openin X V\ \openin X A\ by auto then have "C \ B \ U \ disjnt C (B \ U)" using quasi_component_of_clopen_cases [OF C] \openin X U\ \openin X B\ by blast with UV show "S = {}" by (metis UnE \C \ A\ \S \ C\ T \disjnt A B\ all_not_in_conv disjnt_Un2 disjnt_iff subset_eq) qed qed qed qed lemma quasi_eq_connected_component_of: "finite(connected_components_of X) \ finite(quasi_components_of X) \ locally_connected_space X \ compact_space X \ (Hausdorff_space X \ regular_space X \ normal_space X) \ quasi_component_of X x = connected_component_of X x" by (metis quasi_eq_connected_components_of quasi_eq_connected_components_of_pointwise_alt) subsection\Additional quasicomponent and continuum properties like Boundary Bumping\ lemma cut_wire_fence_theorem_gen: assumes "compact_space X" and X: "Hausdorff_space X \ regular_space X \ normal_space X" and S: "compactin X S" and T: "closedin X T" and dis: "\C. connectedin X C \ disjnt C S \ disjnt C T" shows "separated_between X S T" proof - have "x \ topspace X" if "x \ S" and "T = {}" for x using that S compactin_subset_topspace by auto moreover have "separated_between X {x} {y}" if "x \ S" and "y \ T" for x y proof (cases "x \ topspace X \ y \ topspace X") case True then have "\ connected_component_of X x y" by (meson dis connected_component_of_def disjnt_iff that) with True X \compact_space X\ show ?thesis by (metis quasi_component_nonseparated quasi_eq_connected_component_of) next case False then show ?thesis using S T compactin_subset_topspace closedin_subset that by blast qed ultimately show ?thesis using assms by (simp add: separated_between_pointwise_left separated_between_pointwise_right closedin_compact_space closedin_subset) qed lemma cut_wire_fence_theorem: "\compact_space X; Hausdorff_space X; closedin X S; closedin X T; \C. connectedin X C \ disjnt C S \ disjnt C T\ \ separated_between X S T" by (simp add: closedin_compact_space cut_wire_fence_theorem_gen) lemma separated_between_from_closed_subtopology: assumes XC: "separated_between (subtopology X C) S (X frontier_of C)" and ST: "separated_between (subtopology X C) S T" shows "separated_between X S T" proof - obtain U where clo: "closedin (subtopology X C) U" and ope: "openin (subtopology X C) U" and "S \ U" and sub: "X frontier_of C \ T \ topspace (subtopology X C) - U" by (meson assms separated_between separated_between_Un) then have "X frontier_of C \ T \ topspace X \ C - U" by auto have "closedin X (topspace X \ C)" by (metis XC frontier_of_restrict frontier_of_subset_eq inf_le1 separated_between_imp_subset topspace_subtopology) then have "closedin X U" by (metis clo closedin_closed_subtopology subtopology_restrict) moreover have "openin (subtopology X C) U \ openin X U \ U \ C" using disjnt_iff sub by (force intro!: openin_subset_topspace_eq) with ope have "openin X U" by blast moreover have "T \ topspace X - U" using ope openin_closedin_eq sub by auto ultimately show ?thesis using \S \ U\ separated_between by blast qed lemma separated_between_from_closed_subtopology_frontier: "separated_between (subtopology X T) S (X frontier_of T) \ separated_between X S (X frontier_of T)" using separated_between_from_closed_subtopology by blast lemma separated_between_from_frontier_of_closed_subtopology: assumes "separated_between (subtopology X T) S (X frontier_of T)" shows "separated_between X S (topspace X - T)" proof - have "disjnt S (topspace X - T)" using assms disjnt_iff separated_between_imp_subset by fastforce then show ?thesis by (metis Diff_subset assms frontier_of_complement separated_between_from_closed_subtopology separated_between_frontier_of_eq') qed lemma separated_between_compact_connected_component: assumes "locally_compact_space X" "Hausdorff_space X" and C: "C \ connected_components_of X" and "compactin X C" "closedin X T" "disjnt C T" shows "separated_between X C T" proof - have Csub: "C \ topspace X" by (simp add: assms(4) compactin_subset_topspace) have "Hausdorff_space (subtopology X (topspace X - T))" using Hausdorff_space_subtopology assms(2) by blast moreover have "compactin (subtopology X (topspace X - T)) C" using assms Csub by (metis Diff_Int_distrib Diff_empty compact_imp_compactin_subtopology disjnt_def le_iff_inf) moreover have "locally_compact_space (subtopology X (topspace X - T))" by (meson assms closedin_def locally_compact_Hausdorff_imp_regular_space locally_compact_space_open_subset) ultimately obtain N L where "openin X N" "compactin X L" "closedin X L" "C \ N" "N \ L" and Lsub: "L \ topspace X - T" using \Hausdorff_space X\ \closedin X T\ apply (simp add: locally_compact_space_compact_closed_compact compactin_subtopology) by (meson closedin_def compactin_imp_closedin openin_trans_full) then have disC: "disjnt C (topspace X - L)" by (meson DiffD2 disjnt_iff subset_iff) have "separated_between (subtopology X L) C (X frontier_of L)" proof (rule cut_wire_fence_theorem) show "compact_space (subtopology X L)" by (simp add: \compactin X L\ compact_space_subtopology) show "Hausdorff_space (subtopology X L)" by (simp add: Hausdorff_space_subtopology \Hausdorff_space X\) show "closedin (subtopology X L) C" by (meson \C \ N\ \N \ L\ \Hausdorff_space X\ \compactin X C\ closedin_subset_topspace compactin_imp_closedin subset_trans) show "closedin (subtopology X L) (X frontier_of L)" by (simp add: \closedin X L\ closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin) show "disjnt D C \ disjnt D (X frontier_of L)" if "connectedin (subtopology X L) D" for D proof (rule ccontr) assume "\ (disjnt D C \ disjnt D (X frontier_of L))" moreover have "connectedin X D" using connectedin_subtopology that by blast ultimately show False using that connected_components_of_maximal [of C X D] C apply (simp add: disjnt_iff) by (metis Diff_eq_empty_iff \C \ N\ \N \ L\ \openin X N\ disjoint_iff frontier_of_openin_straddle_Int(2) subsetD) qed qed then have "separated_between X (X frontier_of C) (topspace X - L)" using separated_between_from_frontier_of_closed_subtopology separated_between_frontier_of_eq by blast with \closedin X T\ separated_between_frontier_of [OF Csub disC] show ?thesis unfolding separated_between by (smt (verit) Diff_iff Lsub closedin_subset subset_iff) qed lemma wilder_locally_compact_component_thm: assumes "locally_compact_space X" "Hausdorff_space X" and "C \ connected_components_of X" "compactin X C" "openin X W" "C \ W" obtains U V where "openin X U" "openin X V" "disjnt U V" "U \ V = topspace X" "C \ U" "U \ W" proof - have "closedin X (topspace X - W)" using \openin X W\ by blast moreover have "disjnt C (topspace X - W)" using \C \ W\ disjnt_def by fastforce ultimately have "separated_between X C (topspace X - W)" using separated_between_compact_connected_component assms by blast then show thesis by (smt (verit, del_insts) DiffI disjnt_iff openin_subset separated_between_def subset_iff that) qed lemma compact_quasi_eq_connected_components_of: assumes "locally_compact_space X" "Hausdorff_space X" "compactin X C" shows "C \ quasi_components_of X \ C \ connected_components_of X" proof - have "compactin X (connected_component_of_set X x)" if "x \ topspace X" "compactin X (quasi_component_of_set X x)" for x proof (rule closed_compactin) show "compactin X (quasi_component_of_set X x)" by (simp add: that) show "connected_component_of_set X x \ quasi_component_of_set X x" by (simp add: connected_component_subset_quasi_component_of) show "closedin X (connected_component_of_set X x)" by (simp add: closedin_connected_component_of) qed moreover have "connected_component_of X x = quasi_component_of X x" if \
: "x \ topspace X" "compactin X (connected_component_of_set X x)" for x proof - have "\y. connected_component_of X x y \ quasi_component_of X x y" by (simp add: connected_imp_quasi_component_of) moreover have False if non: "\ connected_component_of X x y" and quasi: "quasi_component_of X x y" for y proof - have "y \ topspace X" by (meson quasi_component_of_equiv that) then have "closedin X {y}" by (simp add: \Hausdorff_space X\ compactin_imp_closedin) moreover have "disjnt (connected_component_of_set X x) {y}" by (simp add: non) moreover have "\ separated_between X (connected_component_of_set X x) {y}" using \
quasi separated_between_pointwise_left by (fastforce simp: quasi_component_nonseparated connected_component_of_refl) ultimately show False using assms by (metis \
connected_component_in_connected_components_of separated_between_compact_connected_component) qed ultimately show ?thesis by blast qed ultimately show ?thesis using \compactin X C\ unfolding connected_components_of_def image_iff quasi_components_of_def by metis qed lemma boundary_bumping_theorem_closed_gen: assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" "closedin X S" "S \ topspace X" and C: "compactin X C" "C \ connected_components_of (subtopology X S)" shows "C \ X frontier_of S \ {}" proof assume \
: "C \ X frontier_of S = {}" consider "C \ {}" "X frontier_of S \ topspace X" | "C \ topspace X" "S = {}" using C by (metis frontier_of_subset_topspace nonempty_connected_components_of) then show False proof cases case 1 have "separated_between (subtopology X S) C (X frontier_of S)" proof (rule separated_between_compact_connected_component) show "compactin (subtopology X S) C" using C compact_imp_compactin_subtopology connected_components_of_subset by fastforce show "closedin (subtopology X S) (X frontier_of S)" by (simp add: \closedin X S\ closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin) show "disjnt C (X frontier_of S)" using \
by (simp add: disjnt_def) qed (use assms Hausdorff_space_subtopology locally_compact_space_closed_subset in auto) then have "separated_between X C (X frontier_of S)" using separated_between_from_closed_subtopology by auto then have "X frontier_of S = {}" using \C \ {}\ \connected_space X\ connected_space_separated_between by blast moreover have "C \ S" using C connected_components_of_subset by fastforce ultimately show False using 1 assms by (metis closedin_subset connected_space_eq_frontier_eq_empty subset_empty) next case 2 then show False using C connected_components_of_eq_empty by fastforce qed qed lemma boundary_bumping_theorem_closed: assumes "connected_space X" "compact_space X" "Hausdorff_space X" "closedin X S" "S \ topspace X" "C \ connected_components_of(subtopology X S)" shows "C \ X frontier_of S \ {}" by (meson assms boundary_bumping_theorem_closed_gen closedin_compact_space closedin_connected_components_of closedin_trans_full compact_imp_locally_compact_space) lemma intermediate_continuum_exists: assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" and C: "compactin X C" "connectedin X C" "C \ {}" "C \ topspace X" and U: "openin X U" "C \ U" obtains D where "compactin X D" "connectedin X D" "C \ D" "D \ U" proof - have "C \ topspace X" by (simp add: C compactin_subset_topspace) with C obtain a where a: "a \ topspace X" "a \ C" by blast moreover have "compactin (subtopology X (U - {a})) C" by (simp add: C U a compact_imp_compactin_subtopology subset_Diff_insert) moreover have "Hausdorff_space (subtopology X (U - {a}))" using Hausdorff_space_subtopology assms(3) by blast moreover have "locally_compact_space (subtopology X (U - {a}))" by (rule locally_compact_space_open_subset) (auto simp: locally_compact_Hausdorff_imp_regular_space open_in_Hausdorff_delete assms) ultimately obtain V K where V: "openin X V" "a \ V" "V \ U" and K: "compactin X K" "a \ K" "K \ U" and cloK: "closedin (subtopology X (U - {a})) K" and "C \ V" "V \ K" using locally_compact_space_compact_closed_compact [of "subtopology X (U - {a})"] assms by (smt (verit, del_insts) Diff_empty compactin_subtopology open_in_Hausdorff_delete openin_open_subtopology subset_Diff_insert) then obtain D where D: "D \ connected_components_of (subtopology X K)" and "C \ D" using C by (metis bot.extremum_unique connectedin_subtopology order.trans exists_connected_component_of_superset subtopology_topspace) show thesis proof have cloD: "closedin (subtopology X K) D" by (simp add: D closedin_connected_components_of) then have XKD: "compactin (subtopology X K) D" by (simp add: K closedin_compact_space compact_space_subtopology) then show "compactin X D" using compactin_subtopology_imp_compact by blast show "connectedin X D" using D connectedin_connected_components_of connectedin_subtopology by blast have "K \ topspace X" using K a by blast moreover have "V \ X interior_of K" by (simp add: \openin X V\ \V \ K\ interior_of_maximal) ultimately have "C \ D" using boundary_bumping_theorem_closed_gen [of X K C] D \C \ V\ by (auto simp add: assms K compactin_imp_closedin frontier_of_def) then show "C \ D" using \C \ D\ by blast have "D \ U" using K(3) \closedin (subtopology X K) D\ closedin_imp_subset by blast moreover have "D \ U" using K XKD \C \ D\ assms by (metis \K \ topspace X\ cloD closedin_imp_subset compactin_imp_closedin connected_space_clopen_in inf_bot_left inf_le2 subset_antisym) ultimately show "D \ U" by blast qed qed lemma boundary_bumping_theorem_gen: assumes X: "connected_space X" "locally_compact_space X" "Hausdorff_space X" and "S \ topspace X" and C: "C \ connected_components_of(subtopology X S)" and compC: "compactin X (X closure_of C)" shows "X frontier_of C \ X frontier_of S \ {}" proof - have Csub: "C \ topspace X" "C \ S" and "connectedin X C" using C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology by fastforce+ have "C \ {}" using C nonempty_connected_components_of by blast obtain "X interior_of C \ X interior_of S" "X closure_of C \ X closure_of S" by (simp add: Csub closure_of_mono interior_of_mono) moreover have False if "X closure_of C \ X interior_of S" proof - have "X closure_of C = C" by (meson C closedin_connected_component_of_subtopology closure_of_eq interior_of_subset order_trans that) with that have "C \ X interior_of S" by simp then obtain D where "compactin X D" and "connectedin X D" and "C \ D" and "D \ X interior_of S" using intermediate_continuum_exists assms \X closure_of C = C\ compC Csub by (metis \C \ {}\ \connectedin X C\ openin_interior_of psubsetE) then have "D \ C" by (metis C \C \ {}\ connected_components_of_maximal connectedin_subtopology disjnt_def inf.orderE interior_of_subset order_trans psubsetE) then show False using \C \ D\ by blast qed ultimately show ?thesis by (smt (verit, ccfv_SIG) DiffI disjoint_iff_not_equal frontier_of_def subset_eq) qed lemma boundary_bumping_theorem: "\connected_space X; compact_space X; Hausdorff_space X; S \ topspace X; C \ connected_components_of(subtopology X S)\ \ X frontier_of C \ X frontier_of S \ {}" by (simp add: boundary_bumping_theorem_gen closedin_compact_space compact_imp_locally_compact_space) +subsection \Compactly generated spaces (k-spaces)\ + +text \These don't have to be Hausdorff\ + +definition k_space where + "k_space X \ + \S. S \ topspace X \ + (closedin X S \ (\K. compactin X K \ closedin (subtopology X K) (K \ S)))" + +lemma k_space: + "k_space X \ + (\S. S \ topspace X \ + (\K. compactin X K \ closedin (subtopology X K) (K \ S)) \ closedin X S)" + by (metis closedin_subtopology inf_commute k_space_def) + +lemma k_space_open: + "k_space X \ + (\S. S \ topspace X \ + (\K. compactin X K \ openin (subtopology X K) (K \ S)) \ openin X S)" +proof - + have "openin X S" + if "k_space X" "S \ topspace X" + and "\K. compactin X K \ openin (subtopology X K) (K \ S)" for S + using that unfolding k_space openin_closedin_eq + by (metis Diff_Int_distrib2 Diff_subset inf_commute topspace_subtopology) + moreover have "k_space X" + if "\S. S \ topspace X \ (\K. compactin X K \ openin (subtopology X K) (K \ S)) \ openin X S" + unfolding k_space openin_closedin_eq + by (simp add: Diff_Int_distrib closedin_def inf_commute that) + ultimately show ?thesis + by blast +qed + +lemma k_space_alt: + "k_space X \ + (\S. S \ topspace X + \ (openin X S \ (\K. compactin X K \ openin (subtopology X K) (K \ S))))" + by (meson k_space_open openin_subtopology_Int2) + +lemma k_space_quotient_map_image: + assumes q: "quotient_map X Y q" and X: "k_space X" + shows "k_space Y" + unfolding k_space +proof clarify + fix S + assume "S \ topspace Y" and S: "\K. compactin Y K \ closedin (subtopology Y K) (K \ S)" + then have iff: "closedin X {x \ topspace X. q x \ S} \ closedin Y S" + using q quotient_map_closedin by fastforce + have "closedin (subtopology X K) (K \ {x \ topspace X. q x \ S})" if "compactin X K" for K + proof - + have "{x \ topspace X. q x \ q ` K} \ K = K" + using compactin_subset_topspace that by blast + then have *: "subtopology X K = subtopology (subtopology X {x \ topspace X. q x \ q ` K}) K" + by (simp add: subtopology_subtopology) + have **: "K \ {x \ topspace X. q x \ S} = + K \ {x \ topspace (subtopology X {x \ topspace X. q x \ q ` K}). q x \ q ` K \ S}" + by auto + have "K \ topspace X" + by (simp add: compactin_subset_topspace that) + show ?thesis + unfolding * ** + proof (intro closedin_continuous_map_preimage closedin_subtopology_Int_closed) + show "continuous_map (subtopology X {x \ topspace X. q x \ q ` K}) (subtopology Y (q ` K)) q" + by (auto simp add: continuous_map_in_subtopology continuous_map_from_subtopology q quotient_imp_continuous_map) + show "closedin (subtopology Y (q ` K)) (q ` K \ S)" + by (meson S image_compactin q quotient_imp_continuous_map that) + qed + qed + then have "closedin X {x \ topspace X. q x \ S}" + by (metis (no_types, lifting) X k_space mem_Collect_eq subsetI) + with iff show "closedin Y S" by simp +qed + +lemma k_space_retraction_map_image: + "\retraction_map X Y r; k_space X\ \ k_space Y" + using k_space_quotient_map_image retraction_imp_quotient_map by blast + +lemma homeomorphic_k_space: + "X homeomorphic_space Y \ k_space X \ k_space Y" + by (meson homeomorphic_map_def homeomorphic_space homeomorphic_space_sym k_space_quotient_map_image) + +lemma k_space_perfect_map_image: + "\k_space X; perfect_map X Y f\ \ k_space Y" + using k_space_quotient_map_image perfect_imp_quotient_map by blast + +lemma locally_compact_imp_k_space: + assumes "locally_compact_space X" + shows "k_space X" + unfolding k_space +proof clarify + fix S + assume "S \ topspace X" and S: "\K. compactin X K \ closedin (subtopology X K) (K \ S)" + have False if non: "\ (X closure_of S \ S)" + proof - + obtain x where "x \ X closure_of S" "x \ S" + using non by blast + then have "x \ topspace X" + by (simp add: in_closure_of) + then obtain K U where "openin X U" "compactin X K" "x \ U" "U \ K" + by (meson assms locally_compact_space_def) + then show False + using \x \ X closure_of S\ openin_Int_closure_of_eq [OF \openin X U\] + by (smt (verit, ccfv_threshold) Int_iff S \x \ S\ closedin_Int_closure_of inf.orderE inf_assoc) + qed + then show "closedin X S" + using S \S \ topspace X\ closure_of_subset_eq by blast +qed + +lemma compact_imp_k_space: + "compact_space X \ k_space X" + by (simp add: compact_imp_locally_compact_space locally_compact_imp_k_space) + +lemma k_space_discrete_topology: "k_space(discrete_topology U)" + by (simp add: k_space_open) + +lemma k_space_closed_subtopology: + assumes "k_space X" "closedin X C" + shows "k_space (subtopology X C)" +unfolding k_space compactin_subtopology +proof clarsimp + fix S + assume Ssub: "S \ topspace X" "S \ C" + and S: "\K. compactin X K \ K \ C \ closedin (subtopology (subtopology X C) K) (K \ S)" + have "closedin (subtopology X K) (K \ S)" if "compactin X K" for K + proof - + have "closedin (subtopology (subtopology X C) (K \ C)) ((K \ C) \ S)" + by (simp add: S \closedin X C\ compact_Int_closedin that) + then show ?thesis + using \closedin X C\ Ssub by (auto simp add: closedin_subtopology) + qed + then show "closedin (subtopology X C) S" + by (metis Ssub \k_space X\ closedin_subset_topspace k_space_def) +qed + +lemma k_space_subtopology: + assumes 1: "\T. \T \ topspace X; T \ S; + \K. compactin X K \ closedin (subtopology X (K \ S)) (K \ T)\ \ closedin (subtopology X S) T" + assumes 2: "\K. compactin X K \ k_space(subtopology X (K \ S))" + shows "k_space (subtopology X S)" + unfolding k_space +proof (intro conjI strip) + fix U + assume \
: "U \ topspace (subtopology X S) \ (\K. compactin (subtopology X S) K \ closedin (subtopology (subtopology X S) K) (K \ U))" + have "closedin (subtopology X (K \ S)) (K \ U)" if "compactin X K" for K + proof - + have "K \ U \ topspace (subtopology X (K \ S))" + using "\
" by auto + moreover + have "\K'. compactin (subtopology X (K \ S)) K' \ closedin (subtopology (subtopology X (K \ S)) K') (K' \ K \ U)" + by (metis "\
" compactin_subtopology inf.orderE inf_commute subtopology_subtopology) + ultimately show ?thesis + by (metis (no_types, opaque_lifting) "2" inf.assoc k_space_def that) + qed + then show "closedin (subtopology X S) U" + using "1" \
by auto +qed + +lemma k_space_subtopology_open: + assumes 1: "\T. \T \ topspace X; T \ S; + \K. compactin X K \ openin (subtopology X (K \ S)) (K \ T)\ \ openin (subtopology X S) T" + assumes 2: "\K. compactin X K \ k_space(subtopology X (K \ S))" + shows "k_space (subtopology X S)" + unfolding k_space_open +proof (intro conjI strip) + fix U + assume \
: "U \ topspace (subtopology X S) \ (\K. compactin (subtopology X S) K \ openin (subtopology (subtopology X S) K) (K \ U))" + have "openin (subtopology X (K \ S)) (K \ U)" if "compactin X K" for K + proof - + have "K \ U \ topspace (subtopology X (K \ S))" + using "\
" by auto + moreover + have "\K'. compactin (subtopology X (K \ S)) K' \ openin (subtopology (subtopology X (K \ S)) K') (K' \ K \ U)" + by (metis "\
" compactin_subtopology inf.orderE inf_commute subtopology_subtopology) + ultimately show ?thesis + by (metis (no_types, opaque_lifting) "2" inf.assoc k_space_open that) + qed + then show "openin (subtopology X S) U" + using "1" \
by auto +qed + + +lemma k_space_open_subtopology_aux: + assumes "kc_space X" "compact_space X" "openin X V" + shows "k_space (subtopology X V)" +proof (clarsimp simp: k_space subtopology_subtopology compactin_subtopology Int_absorb1) + fix S + assume "S \ topspace X" + and "S \ V" + and S: "\K. compactin X K \ K \ V \ closedin (subtopology X K) (K \ S)" + then have "V \ topspace X" + using assms openin_subset by blast + have "S = V \ ((topspace X - V) \ S)" + using \S \ V\ by auto + moreover have "closedin (subtopology X V) (V \ ((topspace X - V) \ S))" + proof (intro closedin_subtopology_Int_closed compactin_imp_closedin_gen \kc_space X\) + show "compactin X (topspace X - V \ S)" + unfolding compactin_def + proof (intro conjI strip) + show "topspace X - V \ S \ topspace X" + by (simp add: \S \ topspace X\) + fix \ + assume \: "Ball \ (openin X) \ topspace X - V \ S \ \\" + moreover + have "compactin X (topspace X - V)" + using assms closedin_compact_space by blast + ultimately obtain \ where "finite \" "\ \ \" and \: "topspace X - V \ \\" + unfolding compactin_def using \V \ topspace X\ by (metis le_sup_iff) + then have "topspace X - \\ \ V" + by blast + then have "closedin (subtopology X (topspace X - \\)) ((topspace X - \\) \ S)" + by (meson S \ \\ \ \\ \compact_space X\ closedin_compact_space openin_Union openin_closedin_eq subset_iff) + then have "compactin X ((topspace X - \\) \ S)" + by (meson \ \\ \ \\\compact_space X\ closedin_compact_space closedin_trans_full openin_Union openin_closedin_eq subset_iff) + then obtain \ where "finite \" "\ \ \" "(topspace X - \\) \ S \ \\" + unfolding compactin_def by (smt (verit, best) \ inf_le2 subset_trans sup.boundedE) + with \ have "topspace X - V \ S \ \(\ \ \)" + using \S \ topspace X\ by auto + then show "\\. finite \ \ \ \ \ \ topspace X - V \ S \ \\" + by (metis \\ \ \\ \\ \ \\ \finite \\ \finite \\ finite_Un le_sup_iff) + qed + qed + ultimately show "closedin (subtopology X V) S" + by metis +qed + + +lemma k_space_open_subtopology: + assumes X: "kc_space X \ Hausdorff_space X \ regular_space X" and "k_space X" "openin X S" + shows "k_space(subtopology X S)" +proof (rule k_space_subtopology_open) + fix T + assume "T \ topspace X" + and "T \ S" + and T: "\K. compactin X K \ openin (subtopology X (K \ S)) (K \ T)" + have "openin (subtopology X K) (K \ T)" if "compactin X K" for K + by (smt (verit, ccfv_threshold) T assms(3) inf_assoc inf_commute openin_Int openin_subtopology that) + then show "openin (subtopology X S) T" + by (metis \T \ S\ \T \ topspace X\ assms(2) k_space_alt subset_openin_subtopology) +next + fix K + assume "compactin X K" + then have KS: "openin (subtopology X K) (K \ S)" + by (simp add: \openin X S\ openin_subtopology_Int2) + have XK: "compact_space (subtopology X K)" + by (simp add: \compactin X K\ compact_space_subtopology) + show "k_space (subtopology X (K \ S))" + using X + proof (rule disjE) + assume "kc_space X" + then show "k_space (subtopology X (K \ S))" + using k_space_open_subtopology_aux [of "subtopology X K" "K \ S"] + by (simp add: KS XK kc_space_subtopology subtopology_subtopology) + next + assume "Hausdorff_space X \ regular_space X" + then have "locally_compact_space (subtopology (subtopology X K) (K \ S))" + using locally_compact_space_open_subset Hausdorff_space_subtopology KS XK + compact_imp_locally_compact_space regular_space_subtopology by blast + then show "k_space (subtopology X (K \ S))" + by (simp add: locally_compact_imp_k_space subtopology_subtopology) + qed +qed + +lemma k_kc_space_subtopology: + "\k_space X; kc_space X; openin X S \ closedin X S\ \ k_space(subtopology X S) \ kc_space(subtopology X S)" + by (metis k_space_closed_subtopology k_space_open_subtopology kc_space_subtopology) + + +lemma k_space_as_quotient_explicit: + "k_space X \ quotient_map (sum_topology (subtopology X) {K. compactin X K}) X snd" +proof - + have [simp]: "{x \ topspace X. x \ K \ x \ U} = K \ U" if "U \ topspace X" for K U + using that by blast + show "?thesis" + apply (simp add: quotient_map_def openin_sum_topology snd_image_Sigma k_space_alt) + by (smt (verit, del_insts) Union_iff compactin_sing inf.orderE mem_Collect_eq singletonI subsetI) +qed + +lemma k_space_as_quotient: + fixes X :: "'a topology" + shows "k_space X \ (\q. \Y:: ('a set * 'a) topology. locally_compact_space Y \ quotient_map Y X q)" + (is "?lhs=?rhs") +proof + show "k_space X" if ?rhs + using that k_space_quotient_map_image locally_compact_imp_k_space by blast +next + assume "k_space X" + show ?rhs + proof (intro exI conjI) + show "locally_compact_space (sum_topology (subtopology X) {K. compactin X K})" + by (simp add: compact_imp_locally_compact_space compact_space_subtopology locally_compact_space_sum_topology) + show "quotient_map (sum_topology (subtopology X) {K. compactin X K}) X snd" + using \k_space X\ k_space_as_quotient_explicit by blast + qed +qed + +lemma k_space_prod_topology_left: + assumes X: "locally_compact_space X" "Hausdorff_space X \ regular_space X" and "k_space Y" + shows "k_space (prod_topology X Y)" +proof - + obtain q and Z :: "('b set * 'b) topology" where "locally_compact_space Z" and q: "quotient_map Z Y q" + using \k_space Y\ k_space_as_quotient by blast + then show ?thesis + using quotient_map_prod_right [OF X q] X k_space_quotient_map_image locally_compact_imp_k_space + locally_compact_space_prod_topology by blast +qed + +text \Essentially the same proof\ +lemma k_space_prod_topology_right: + assumes "k_space X" and Y: "locally_compact_space Y" "Hausdorff_space Y \ regular_space Y" + shows "k_space (prod_topology X Y)" +proof - + obtain q and Z :: "('a set * 'a) topology" where "locally_compact_space Z" and q: "quotient_map Z X q" + using \k_space X\ k_space_as_quotient by blast + then show ?thesis + using quotient_map_prod_left [OF Y q] using Y k_space_quotient_map_image locally_compact_imp_k_space + locally_compact_space_prod_topology by blast +qed + + +lemma continuous_map_from_k_space: + assumes "k_space X" and f: "\K. compactin X K \ continuous_map(subtopology X K) Y f" + shows "continuous_map X Y f" +proof - + have "\x. x \ topspace X \ f x \ topspace Y" + by (metis compactin_absolute compactin_sing f image_compactin image_empty image_insert) + moreover have "closedin X {x \ topspace X. f x \ C}" if "closedin Y C" for C + proof - + have "{x \ topspace X. f x \ C} \ topspace X" + by fastforce + moreover + have eq: "K \ {x \ topspace X. f x \ C} = {x. x \ topspace(subtopology X K) \ f x \ (f ` K \ C)}" for K + by auto + have "closedin (subtopology X K) (K \ {x \ topspace X. f x \ C})" if "compactin X K" for K + unfolding eq + proof (rule closedin_continuous_map_preimage) + show "continuous_map (subtopology X K) (subtopology Y (f`K)) f" + by (simp add: continuous_map_in_subtopology f image_mono that) + show "closedin (subtopology Y (f`K)) (f ` K \ C)" + using \closedin Y C\ closedin_subtopology by blast + qed + ultimately show ?thesis + using \k_space X\ k_space by blast + qed + ultimately show ?thesis + by (simp add: continuous_map_closedin) +qed + +lemma closed_map_into_k_space: + assumes "k_space Y" and fim: "f ` (topspace X) \ topspace Y" + and f: "\K. compactin Y K + \ closed_map(subtopology X {x \ topspace X. f x \ K}) (subtopology Y K) f" + shows "closed_map X Y f" + unfolding closed_map_def +proof (intro strip) + fix C + assume "closedin X C" + have "closedin (subtopology Y K) (K \ f ` C)" + if "compactin Y K" for K + proof - + have eq: "K \ f ` C = f ` ({x \ topspace X. f x \ K} \ C)" + using \closedin X C\ closedin_subset by auto + show ?thesis + unfolding eq + by (metis (no_types, lifting) \closedin X C\ closed_map_def closedin_subtopology f inf_commute that) + qed + then show "closedin Y (f ` C)" + using \k_space Y\ unfolding k_space + by (meson \closedin X C\ closedin_subset dual_order.trans fim image_mono) +qed + + +text \Essentially the same proof\ +lemma open_map_into_k_space: + assumes "k_space Y" and fim: "f ` (topspace X) \ topspace Y" + and f: "\K. compactin Y K + \ open_map (subtopology X {x \ topspace X. f x \ K}) (subtopology Y K) f" + shows "open_map X Y f" + unfolding open_map_def +proof (intro strip) + fix C + assume "openin X C" + have "openin (subtopology Y K) (K \ f ` C)" + if "compactin Y K" for K + proof - + have eq: "K \ f ` C = f ` ({x \ topspace X. f x \ K} \ C)" + using \openin X C\ openin_subset by auto + show ?thesis + unfolding eq + by (metis (no_types, lifting) \openin X C\ open_map_def openin_subtopology f inf_commute that) + qed + then show "openin Y (f ` C)" + using \k_space Y\ unfolding k_space_open + by (meson \openin X C\ openin_subset dual_order.trans fim image_mono) +qed + +lemma quotient_map_into_k_space: + fixes f :: "'a \ 'b" + assumes "k_space Y" and cmf: "continuous_map X Y f" + and fim: "f ` (topspace X) = topspace Y" + and f: "\k. compactin Y k + \ quotient_map (subtopology X {x \ topspace X. f x \ k}) + (subtopology Y k) f" + shows "quotient_map X Y f" +proof - + have "closedin Y C" + if "C \ topspace Y" and K: "closedin X {x \ topspace X. f x \ C}" for C + proof - + have "closedin (subtopology Y K) (K \ C)" if "compactin Y K" for K + proof - + define Kf where "Kf \ {x \ topspace X. f x \ K}" + have *: "K \ C \ topspace Y \ K \ C \ K" + using \C \ topspace Y\ by blast + then have eq: "closedin (subtopology X Kf) (Kf \ {x \ topspace X. f x \ C}) = + closedin (subtopology Y K) (K \ C)" + using f [OF that] * unfolding quotient_map_closedin Kf_def + by (smt (verit, ccfv_SIG) Collect_cong Int_def compactin_subset_topspace mem_Collect_eq that topspace_subtopology topspace_subtopology_subset) + have dd: "{x \ topspace X \ Kf. f x \ K \ C} = Kf \ {x \ topspace X. f x \ C}" + by (auto simp add: Kf_def) + have "closedin (subtopology X Kf) {x \ topspace X. x \ Kf \ f x \ K \ f x \ C}" + using K closedin_subtopology by (fastforce simp add: Kf_def) + with K closedin_subtopology_Int_closed eq show ?thesis + by blast + qed + then show ?thesis + using \k_space Y\ that unfolding k_space by blast + qed + moreover have "closedin X {x \ topspace X. f x \ K}" + if "K \ topspace Y" "closedin Y K" for K + using that cmf continuous_map_closedin by fastforce + ultimately show ?thesis + unfolding quotient_map_closedin using fim by blast +qed + +lemma quotient_map_into_k_space_eq: + assumes "k_space Y" "kc_space Y" + shows "quotient_map X Y f \ + continuous_map X Y f \ f ` (topspace X) = topspace Y \ + (\K. compactin Y K + \ quotient_map (subtopology X {x \ topspace X. f x \ K}) (subtopology Y K) f)" + using assms + by (auto simp: kc_space_def intro: quotient_map_into_k_space quotient_map_restriction + dest: quotient_imp_continuous_map quotient_imp_surjective_map) + +lemma open_map_into_k_space_eq: + assumes "k_space Y" + shows "open_map X Y f \ + f ` (topspace X) \ topspace Y \ + (\k. compactin Y k + \ open_map (subtopology X {x \ topspace X. f x \ k}) (subtopology Y k) f)" + (is "?lhs=?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: open_map_imp_subset_topspace open_map_restriction) + show "?rhs \ ?lhs" + by (simp add: assms open_map_into_k_space) +qed + +lemma closed_map_into_k_space_eq: + assumes "k_space Y" + shows "closed_map X Y f \ + f ` (topspace X) \ topspace Y \ + (\k. compactin Y k + \ closed_map (subtopology X {x \ topspace X. f x \ k}) (subtopology Y k) f)" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: closed_map_imp_subset_topspace closed_map_restriction) + show "?rhs \ ?lhs" + by (simp add: assms closed_map_into_k_space) +qed + +lemma proper_map_into_k_space: + assumes "k_space Y" and fim: "f ` (topspace X) \ topspace Y" + and f: "\K. compactin Y K + \ proper_map (subtopology X {x \ topspace X. f x \ K}) + (subtopology Y K) f" + shows "proper_map X Y f" +proof - + have "closed_map X Y f" + by (meson assms closed_map_into_k_space fim proper_map_def) + with f topspace_subtopology_subset show ?thesis + apply (simp add: proper_map_alt) + by (smt (verit, best) Collect_cong compactin_absolute) +qed + +lemma proper_map_into_k_space_eq: + assumes "k_space Y" + shows "proper_map X Y f \ + f ` (topspace X) \ topspace Y \ + (\K. compactin Y K + \ proper_map (subtopology X {x \ topspace X. f x \ K}) (subtopology Y K) f)" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: proper_map_imp_subset_topspace proper_map_restriction) + show "?rhs \ ?lhs" + by (simp add: assms proper_map_into_k_space) +qed + +lemma compact_imp_proper_map: + assumes "k_space Y" "kc_space Y" and fim: "f ` (topspace X) \ topspace Y" + and f: "continuous_map X Y f \ kc_space X" + and comp: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" + shows "proper_map X Y f" +proof (rule compact_imp_proper_map_gen) + fix S + assume "S \ topspace Y" + and "\K. compactin Y K \ compactin Y (S \ K)" + with assms show "closedin Y S" + by (simp add: closedin_subset_topspace inf_commute k_space kc_space_def) +qed (use assms in auto) + +lemma proper_eq_compact_map: + assumes "k_space Y" "kc_space Y" + and f: "continuous_map X Y f \ kc_space X" + shows "proper_map X Y f \ + f ` (topspace X) \ topspace Y \ + (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" + (is "?lhs \ ?rhs") +proof + show "?lhs \ ?rhs" + by (simp add: proper_map_alt proper_map_imp_subset_topspace) +qed (use assms compact_imp_proper_map in auto) + +lemma compact_imp_perfect_map: + assumes "k_space Y" "kc_space Y" and "f ` (topspace X) = topspace Y" + and "continuous_map X Y f" + and "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" + shows "perfect_map X Y f" + by (simp add: assms compact_imp_proper_map perfect_map_def) + end diff --git a/src/HOL/Analysis/Abstract_Topology.thy b/src/HOL/Analysis/Abstract_Topology.thy --- a/src/HOL/Analysis/Abstract_Topology.thy +++ b/src/HOL/Analysis/Abstract_Topology.thy @@ -1,5024 +1,5090 @@ (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) section \Operators involving abstract topology\ theory Abstract_Topology imports Complex_Main "HOL-Library.Set_Idioms" "HOL-Library.FuncSet" begin subsection \General notion of a topology as a value\ definition\<^marker>\tag important\ istopology :: "('a set \ bool) \ bool" where "istopology L \ (\S T. L S \ L T \ L (S \ T)) \ (\\. (\K\\. L K) \ L (\\))" typedef\<^marker>\tag important\ 'a topology = "{L::('a set) \ bool. istopology L}" morphisms "openin" "topology" unfolding istopology_def by blast lemma istopology_openin[iff]: "istopology(openin U)" using openin[of U] by blast lemma istopology_open[iff]: "istopology open" by (auto simp: istopology_def) lemma topology_inverse' [simp]: "istopology U \ openin (topology U) = U" using topology_inverse[unfolded mem_Collect_eq] . lemma topology_inverse_iff: "istopology U \ openin (topology U) = U" by (metis istopology_openin topology_inverse') lemma topology_eq: "T1 = T2 \ (\S. openin T1 S \ openin T2 S)" proof assume "T1 = T2" then show "\S. openin T1 S \ openin T2 S" by simp next assume H: "\S. openin T1 S \ openin T2 S" then have "openin T1 = openin T2" by (simp add: fun_eq_iff) then have "topology (openin T1) = topology (openin T2)" by simp then show "T1 = T2" unfolding openin_inverse . qed text\The "universe": the union of all sets in the topology.\ definition "topspace T = \{S. openin T S}" subsubsection \Main properties of open sets\ proposition openin_clauses: fixes U :: "'a topology" shows "openin U {}" "\S T. openin U S \ openin U T \ openin U (S\T)" "\K. (\S \ K. openin U S) \ openin U (\K)" using openin[of U] unfolding istopology_def by auto lemma openin_subset: "openin U S \ S \ topspace U" unfolding topspace_def by blast lemma openin_empty[simp]: "openin U {}" by (rule openin_clauses) lemma openin_Int[intro]: "openin U S \ openin U T \ openin U (S \ T)" by (rule openin_clauses) lemma openin_Union[intro]: "(\S. S \ K \ openin U S) \ openin U (\K)" using openin_clauses by blast lemma openin_Un[intro]: "openin U S \ openin U T \ openin U (S \ T)" using openin_Union[of "{S,T}" U] by auto lemma openin_topspace[intro, simp]: "openin U (topspace U)" by (force simp: openin_Union topspace_def) lemma openin_subopen: "openin U S \ (\x \ S. \T. openin U T \ x \ T \ T \ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs by auto next assume H: ?rhs let ?t = "\{T. openin U T \ T \ S}" have "openin U ?t" by (force simp: openin_Union) also have "?t = S" using H by auto finally show "openin U S" . qed lemma openin_INT [intro]: assumes "finite I" "\i. i \ I \ openin T (U i)" shows "openin T ((\i \ I. U i) \ topspace T)" using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) lemma openin_INT2 [intro]: assumes "finite I" "I \ {}" "\i. i \ I \ openin T (U i)" shows "openin T (\i \ I. U i)" proof - have "(\i \ I. U i) \ topspace T" using \I \ {}\ openin_subset[OF assms(3)] by auto then show ?thesis using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) qed lemma openin_Inter [intro]: assumes "finite \" "\ \ {}" "\X. X \ \ \ openin T X" shows "openin T (\\)" by (metis (full_types) assms openin_INT2 image_ident) lemma openin_Int_Inter: assumes "finite \" "openin T U" "\X. X \ \ \ openin T X" shows "openin T (U \ \\)" using openin_Inter [of "insert U \"] assms by auto subsubsection \Closed sets\ definition\<^marker>\tag important\ closedin :: "'a topology \ 'a set \ bool" where "closedin U S \ S \ topspace U \ openin U (topspace U - S)" lemma closedin_subset: "closedin U S \ S \ topspace U" by (metis closedin_def) lemma closedin_empty[simp]: "closedin U {}" by (simp add: closedin_def) lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" by (simp add: closedin_def) lemma closedin_Un[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" by (auto simp: Diff_Un closedin_def) lemma Diff_Inter[intro]: "A - \S = \{A - s|s. s\S}" by auto lemma closedin_Union: assumes "finite S" "\T. T \ S \ closedin U T" shows "closedin U (\S)" using assms by induction auto lemma closedin_Inter[intro]: assumes Ke: "K \ {}" and Kc: "\S. S \K \ closedin U S" shows "closedin U (\K)" using Ke Kc unfolding closedin_def Diff_Inter by auto lemma closedin_INT[intro]: assumes "A \ {}" "\x. x \ A \ closedin U (B x)" shows "closedin U (\x\A. B x)" using assms by blast lemma closedin_Int[intro]: "closedin U S \ closedin U T \ closedin U (S \ T)" using closedin_Inter[of "{S,T}" U] by auto lemma openin_closedin_eq: "openin U S \ S \ topspace U \ closedin U (topspace U - S)" by (metis Diff_subset closedin_def double_diff equalityD1 openin_subset) lemma topology_finer_closedin: "topspace X = topspace Y \ (\S. openin Y S \ openin X S) \ (\S. closedin Y S \ closedin X S)" by (metis closedin_def openin_closedin_eq) lemma openin_closedin: "S \ topspace U \ (openin U S \ closedin U (topspace U - S))" by (simp add: openin_closedin_eq) lemma openin_diff[intro]: assumes oS: "openin U S" and cT: "closedin U T" shows "openin U (S - T)" by (metis Int_Diff cT closedin_def inf.orderE oS openin_Int openin_subset) lemma closedin_diff[intro]: assumes oS: "closedin U S" and cT: "openin U T" shows "closedin U (S - T)" by (metis Int_Diff cT closedin_Int closedin_subset inf.orderE oS openin_closedin_eq) lemma all_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) lemma all_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) lemma ex_openin: "(\U. openin X U \ P U) \ (\U. closedin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) lemma ex_closedin: "(\U. closedin X U \ P U) \ (\U. openin X U \ P (topspace X - U))" by (metis Diff_Diff_Int closedin_subset inf.absorb_iff2 openin_closedin_eq) subsection\The discrete topology\ definition discrete_topology where "discrete_topology U \ topology (\S. S \ U)" lemma openin_discrete_topology [simp]: "openin (discrete_topology U) S \ S \ U" proof - have "istopology (\S. S \ U)" by (auto simp: istopology_def) then show ?thesis by (simp add: discrete_topology_def topology_inverse') qed lemma topspace_discrete_topology [simp]: "topspace(discrete_topology U) = U" by (meson openin_discrete_topology openin_subset openin_topspace order_refl subset_antisym) lemma closedin_discrete_topology [simp]: "closedin (discrete_topology U) S \ S \ U" by (simp add: closedin_def) lemma discrete_topology_unique: "discrete_topology U = X \ topspace X = U \ (\x \ U. openin X {x})" (is "?lhs = ?rhs") proof assume R: ?rhs then have "openin X S" if "S \ U" for S using openin_subopen subsetD that by fastforce then show ?lhs by (metis R openin_discrete_topology openin_subset topology_eq) qed auto lemma discrete_topology_unique_alt: "discrete_topology U = X \ topspace X \ U \ (\x \ U. openin X {x})" using openin_subset by (auto simp: discrete_topology_unique) lemma subtopology_eq_discrete_topology_empty: "X = discrete_topology {} \ topspace X = {}" using discrete_topology_unique [of "{}" X] by auto lemma subtopology_eq_discrete_topology_sing: "X = discrete_topology {a} \ topspace X = {a}" by (metis discrete_topology_unique openin_topspace singletonD) subsection \Subspace topology\ definition\<^marker>\tag important\ subtopology :: "'a topology \ 'a set \ 'a topology" where "subtopology U V = topology (\T. \S. T = S \ V \ openin U S)" lemma istopology_subtopology: "istopology (\T. \S. T = S \ V \ openin U S)" (is "istopology ?L") proof - have "?L {}" by blast { fix A B assume A: "?L A" and B: "?L B" from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \ V" and Sb: "openin U Sb" "B = Sb \ V" by blast have "A \ B = (Sa \ Sb) \ V" "openin U (Sa \ Sb)" using Sa Sb by blast+ then have "?L (A \ B)" by blast } moreover { fix K assume K: "K \ Collect ?L" have th0: "Collect ?L = (\S. S \ V) ` Collect (openin U)" by blast from K[unfolded th0 subset_image_iff] obtain Sk where Sk: "Sk \ Collect (openin U)" "K = (\S. S \ V) ` Sk" by blast have "\K = (\Sk) \ V" using Sk by auto moreover have "openin U (\Sk)" using Sk by (auto simp: subset_eq) ultimately have "?L (\K)" by blast } ultimately show ?thesis unfolding subset_eq mem_Collect_eq istopology_def by auto qed lemma openin_subtopology: "openin (subtopology U V) S \ (\T. openin U T \ S = T \ V)" unfolding subtopology_def topology_inverse'[OF istopology_subtopology] by auto +lemma subset_openin_subtopology: + "\openin X S; S \ T\ \ openin (subtopology X T) S" + by (metis inf.orderE openin_subtopology) + lemma openin_subtopology_Int: "openin X S \ openin (subtopology X T) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_Int2: "openin X T \ openin (subtopology X S) (S \ T)" using openin_subtopology by auto lemma openin_subtopology_diff_closed: "\S \ topspace X; closedin X T\ \ openin (subtopology X S) (S - T)" unfolding closedin_def openin_subtopology by (rule_tac x="topspace X - T" in exI) auto lemma openin_relative_to: "(openin X relative_to S) = openin (subtopology X S)" by (force simp: relative_to_def openin_subtopology) lemma topspace_subtopology [simp]: "topspace (subtopology U V) = topspace U \ V" by (auto simp: topspace_def openin_subtopology) lemma topspace_subtopology_subset: "S \ topspace X \ topspace(subtopology X S) = S" by (simp add: inf.absorb_iff2) lemma closedin_subtopology: "closedin (subtopology U V) S \ (\T. closedin U T \ S = T \ V)" unfolding closedin_def topspace_subtopology by (auto simp: openin_subtopology) +lemma closedin_subtopology_Int_closed: + "closedin X T \ closedin (subtopology X S) (S \ T)" + using closedin_subtopology inf_commute by blast + lemma closedin_subset_topspace: "\closedin X S; S \ T\ \ closedin (subtopology X T) S" using closedin_subtopology by fastforce lemma closedin_relative_to: "(closedin X relative_to S) = closedin (subtopology X S)" by (force simp: relative_to_def closedin_subtopology) lemma openin_subtopology_refl: "openin (subtopology U V) V \ V \ topspace U" unfolding openin_subtopology by auto (metis IntD1 in_mono openin_subset) lemma subtopology_subtopology: "subtopology (subtopology X S) T = subtopology X (S \ T)" proof - have eq: "\T'. (\S'. T' = S' \ T \ (\T. openin X T \ S' = T \ S)) = (\Sa. T' = Sa \ (S \ T) \ openin X Sa)" by (metis inf_assoc) have "subtopology (subtopology X S) T = topology (\Ta. \Sa. Ta = Sa \ T \ openin (subtopology X S) Sa)" by (simp add: subtopology_def) also have "\ = subtopology X (S \ T)" by (simp add: openin_subtopology eq) (simp add: subtopology_def) finally show ?thesis . qed lemma openin_subtopology_alt: "openin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (openin X)" by (simp add: image_iff inf_commute openin_subtopology) lemma closedin_subtopology_alt: "closedin (subtopology X U) S \ S \ (\T. U \ T) ` Collect (closedin X)" by (simp add: image_iff inf_commute closedin_subtopology) lemma subtopology_superset: assumes UV: "topspace U \ V" shows "subtopology U V = U" proof - { fix S have "openin U S" if "openin U T" "S = T \ V" for T by (metis Int_subset_iff assms inf.orderE openin_subset that) then have "(\T. openin U T \ S = T \ V) \ openin U S" by (metis assms inf.orderE inf_assoc openin_subset) } then show ?thesis unfolding topology_eq openin_subtopology by blast qed lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" by (simp add: subtopology_superset) lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" by (simp add: subtopology_superset) lemma subtopology_restrict: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) lemma openin_subtopology_empty: "openin (subtopology U {}) S \ S = {}" by (metis Int_empty_right openin_empty openin_subtopology) lemma closedin_subtopology_empty: "closedin (subtopology U {}) S \ S = {}" by (metis Int_empty_right closedin_empty closedin_subtopology) lemma closedin_subtopology_refl [simp]: "closedin (subtopology U X) X \ X \ topspace U" by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) lemma closedin_topspace_empty: "topspace T = {} \ (closedin T S \ S = {})" by (simp add: closedin_def) lemma open_in_topspace_empty: "topspace X = {} \ openin X S \ S = {}" by (simp add: openin_closedin_eq) lemma openin_imp_subset: "openin (subtopology U S) T \ T \ S" by (metis Int_iff openin_subtopology subsetI) lemma closedin_imp_subset: "closedin (subtopology U S) T \ T \ S" by (simp add: closedin_def) lemma openin_open_subtopology: "openin X S \ openin (subtopology X S) T \ openin X T \ T \ S" by (metis inf.orderE openin_Int openin_imp_subset openin_subtopology) lemma closedin_closed_subtopology: "closedin X S \ (closedin (subtopology X S) T \ closedin X T \ T \ S)" by (metis closedin_Int closedin_imp_subset closedin_subtopology inf.orderE) lemma closedin_trans_full: "\closedin (subtopology X U) S; closedin X U\ \ closedin X S" using closedin_closed_subtopology by blast lemma openin_subtopology_Un: "\openin (subtopology X T) S; openin (subtopology X U) S\ \ openin (subtopology X (T \ U)) S" by (simp add: openin_subtopology) blast lemma closedin_subtopology_Un: "\closedin (subtopology X T) S; closedin (subtopology X U) S\ \ closedin (subtopology X (T \ U)) S" by (simp add: closedin_subtopology) blast lemma openin_trans_full: "\openin (subtopology X U) S; openin X U\ \ openin X S" by (simp add: openin_open_subtopology) subsection \The canonical topology from the underlying type class\ abbreviation\<^marker>\tag important\ euclidean :: "'a::topological_space topology" where "euclidean \ topology open" abbreviation top_of_set :: "'a::topological_space set \ 'a topology" where "top_of_set \ subtopology (topology open)" lemma open_openin: "open S \ openin euclidean S" by simp declare open_openin [symmetric, simp] lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" by (force simp: topspace_def) lemma topspace_euclidean_subtopology[simp]: "topspace (top_of_set S) = S" by (simp) lemma closed_closedin: "closed S \ closedin euclidean S" by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) declare closed_closedin [symmetric, simp] lemma openin_subtopology_self [simp]: "openin (top_of_set S) S" by (metis openin_topspace topspace_euclidean_subtopology) subsubsection\The most basic facts about the usual topology and metric on R\ abbreviation euclideanreal :: "real topology" where "euclideanreal \ topology open" subsection \Basic "localization" results are handy for connectedness.\ lemma openin_open: "openin (top_of_set U) S \ (\T. open T \ (S = U \ T))" by (auto simp: openin_subtopology) lemma openin_Int_open: "\openin (top_of_set U) S; open T\ \ openin (top_of_set U) (S \ T)" by (metis open_Int Int_assoc openin_open) lemma openin_open_Int[intro]: "open S \ openin (top_of_set U) (U \ S)" by (auto simp: openin_open) lemma open_openin_trans[trans]: "open S \ open T \ T \ S \ openin (top_of_set S) T" by (metis Int_absorb1 openin_open_Int) lemma open_subset: "S \ T \ open S \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_closed: "closedin (top_of_set U) S \ (\T. closed T \ S = U \ T)" by (simp add: closedin_subtopology Int_ac) lemma closedin_closed_Int: "closed S \ closedin (top_of_set U) (U \ S)" by (metis closedin_closed) lemma closed_subset: "S \ T \ closed S \ closedin (top_of_set T) S" by (auto simp: closedin_closed) lemma closedin_closed_subset: "\closedin (top_of_set U) V; T \ U; S = V \ T\ \ closedin (top_of_set T) S" by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) lemma finite_imp_closedin: fixes S :: "'a::t1_space set" shows "\finite S; S \ T\ \ closedin (top_of_set T) S" by (simp add: finite_imp_closed closed_subset) lemma closedin_singleton [simp]: fixes a :: "'a::t1_space" shows "closedin (top_of_set U) {a} \ a \ U" using closedin_subset by (force intro: closed_subset) lemma openin_euclidean_subtopology_iff: fixes S U :: "'a::metric_space set" shows "openin (top_of_set U) S \ S \ U \ (\x\S. \e>0. \x'\U. dist x' x < e \ x'\ S)" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs unfolding openin_open open_dist by blast next define T where "T = {x. \a\S. \d>0. (\y\U. dist y a < d \ y \ S) \ dist x a < d}" have 1: "\x\T. \e>0. \y. dist y x < e \ y \ T" unfolding T_def apply clarsimp apply (rule_tac x="d - dist x a" in exI) by (metis add_0_left dist_commute dist_triangle_lt less_diff_eq) assume ?rhs then have 2: "S = U \ T" unfolding T_def by auto (metis dist_self) from 1 2 show ?lhs unfolding openin_open open_dist by fast qed lemma connected_openin: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_def openin_open disjoint_iff_not_equal by blast lemma connected_openin_eq: "connected S \ \(\E1 E2. openin (top_of_set S) E1 \ openin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_openin by (metis (no_types, lifting) Un_subset_iff openin_imp_subset subset_antisym) lemma connected_closedin: "connected S \ (\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ S \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp add: connected_closed closedin_closed) next assume R: ?rhs then show ?lhs proof (clarsimp simp add: connected_closed closedin_closed) fix A B assume s_sub: "S \ A \ B" "B \ S \ {}" and disj: "A \ B \ S = {}" and cl: "closed A" "closed B" have "S - A = B \ S" using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto then show "A \ S = {}" by (metis Int_Diff_Un Int_Diff_disjoint R cl closedin_closed_Int dual_order.refl inf_commute s_sub(2)) qed qed lemma connected_closedin_eq: "connected S \ \(\E1 E2. closedin (top_of_set S) E1 \ closedin (top_of_set S) E2 \ E1 \ E2 = S \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_closedin by (metis Un_subset_iff closedin_imp_subset subset_antisym) text \These "transitivity" results are handy too\ lemma openin_trans[trans]: "openin (top_of_set T) S \ openin (top_of_set U) T \ openin (top_of_set U) S" by (metis openin_Int_open openin_open) lemma openin_open_trans: "openin (top_of_set T) S \ open T \ open S" by (auto simp: openin_open intro: openin_trans) lemma closedin_trans[trans]: "closedin (top_of_set T) S \ closedin (top_of_set U) T \ closedin (top_of_set U) S" by (auto simp: closedin_closed closed_Inter Int_assoc) lemma closedin_closed_trans: "closedin (top_of_set T) S \ closed T \ closed S" by (auto simp: closedin_closed intro: closedin_trans) lemma openin_subtopology_Int_subset: "\openin (top_of_set u) (u \ S); v \ u\ \ openin (top_of_set v) (v \ S)" by (auto simp: openin_subtopology) lemma openin_open_eq: "open s \ (openin (top_of_set s) t \ open t \ t \ s)" using open_subset openin_open_trans openin_subset by fastforce subsection\Derived set (set of limit points)\ definition derived_set_of :: "'a topology \ 'a set \ 'a set" (infixl "derived'_set'_of" 80) where "X derived_set_of S \ {x \ topspace X. (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))}" lemma derived_set_of_restrict [simp]: "X derived_set_of (topspace X \ S) = X derived_set_of S" by (simp add: derived_set_of_def) (metis openin_subset subset_iff) lemma in_derived_set_of: "x \ X derived_set_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y\x. y \ S \ y \ T))" by (simp add: derived_set_of_def) lemma derived_set_of_subset_topspace: "X derived_set_of S \ topspace X" by (auto simp add: derived_set_of_def) lemma derived_set_of_subtopology: "(subtopology X U) derived_set_of S = U \ (X derived_set_of (U \ S))" by (simp add: derived_set_of_def openin_subtopology) blast lemma derived_set_of_subset_subtopology: "(subtopology X S) derived_set_of T \ S" by (simp add: derived_set_of_subtopology) lemma derived_set_of_empty [simp]: "X derived_set_of {} = {}" by (auto simp: derived_set_of_def) lemma derived_set_of_mono: "S \ T \ X derived_set_of S \ X derived_set_of T" unfolding derived_set_of_def by blast lemma derived_set_of_Un: "X derived_set_of (S \ T) = X derived_set_of S \ X derived_set_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (clarsimp simp: in_derived_set_of) (metis IntE IntI openin_Int) show "?rhs \ ?lhs" by (simp add: derived_set_of_mono) qed lemma derived_set_of_Union: "finite \ \ X derived_set_of (\\) = (\S \ \. X derived_set_of S)" proof (induction \ rule: finite_induct) case (insert S \) then show ?case by (simp add: derived_set_of_Un) qed auto lemma derived_set_of_topspace: "X derived_set_of (topspace X) = {x \ topspace X. \ openin X {x}}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: in_derived_set_of) show "?rhs \ ?lhs" by (clarsimp simp: in_derived_set_of) (metis openin_closedin_eq openin_subopen singletonD subset_eq) qed lemma discrete_topology_unique_derived_set: "discrete_topology U = X \ topspace X = U \ X derived_set_of U = {}" by (auto simp: discrete_topology_unique derived_set_of_topspace) lemma subtopology_eq_discrete_topology_eq: "subtopology X U = discrete_topology U \ U \ topspace X \ U \ X derived_set_of U = {}" using discrete_topology_unique_derived_set [of U "subtopology X U"] by (auto simp: eq_commute derived_set_of_subtopology) lemma subtopology_eq_discrete_topology: "S \ topspace X \ S \ X derived_set_of S = {} \ subtopology X S = discrete_topology S" by (simp add: subtopology_eq_discrete_topology_eq) lemma subtopology_eq_discrete_topology_gen: assumes "S \ X derived_set_of S = {}" shows "subtopology X S = discrete_topology(topspace X \ S)" proof - have "subtopology X S = subtopology X (topspace X \ S)" by (simp add: subtopology_restrict) then show ?thesis using assms by (simp add: inf.assoc subtopology_eq_discrete_topology_eq) qed lemma subtopology_discrete_topology [simp]: "subtopology (discrete_topology U) S = discrete_topology(U \ S)" proof - have "(\T. \Sa. T = Sa \ S \ Sa \ U) = (\Sa. Sa \ U \ Sa \ S)" by force then show ?thesis by (simp add: subtopology_def) (simp add: discrete_topology_def) qed lemma openin_Int_derived_set_of_subset: "openin X S \ S \ X derived_set_of T \ X derived_set_of (S \ T)" by (auto simp: derived_set_of_def) lemma openin_Int_derived_set_of_eq: assumes "openin X S" shows "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_derived_set_of_subset) show "?rhs \ ?lhs" by (metis derived_set_of_mono inf_commute inf_le1 inf_mono order_refl) qed subsection\ Closure with respect to a topological space\ definition closure_of :: "'a topology \ 'a set \ 'a set" (infixr "closure'_of" 80) where "X closure_of S \ {x \ topspace X. \T. x \ T \ openin X T \ (\y \ S. y \ T)}" lemma closure_of_restrict: "X closure_of S = X closure_of (topspace X \ S)" unfolding closure_of_def using openin_subset by blast lemma in_closure_of: "x \ X closure_of S \ x \ topspace X \ (\T. x \ T \ openin X T \ (\y. y \ S \ y \ T))" by (auto simp: closure_of_def) lemma closure_of: "X closure_of S = topspace X \ (S \ X derived_set_of S)" by (fastforce simp: in_closure_of in_derived_set_of) lemma closure_of_alt: "X closure_of S = topspace X \ S \ X derived_set_of S" using derived_set_of_subset_topspace [of X S] unfolding closure_of_def in_derived_set_of by safe (auto simp: in_derived_set_of) lemma derived_set_of_subset_closure_of: "X derived_set_of S \ X closure_of S" by (fastforce simp: closure_of_def in_derived_set_of) lemma closure_of_subtopology: "(subtopology X U) closure_of S = U \ (X closure_of (U \ S))" unfolding closure_of_def topspace_subtopology openin_subtopology by safe (metis (full_types) IntI Int_iff inf.commute)+ lemma closure_of_empty [simp]: "X closure_of {} = {}" by (simp add: closure_of_alt) lemma closure_of_topspace [simp]: "X closure_of topspace X = topspace X" by (simp add: closure_of) lemma closure_of_UNIV [simp]: "X closure_of UNIV = topspace X" by (simp add: closure_of) lemma closure_of_subset_topspace: "X closure_of S \ topspace X" by (simp add: closure_of) lemma closure_of_subset_subtopology: "(subtopology X S) closure_of T \ S" by (simp add: closure_of_subtopology) lemma closure_of_mono: "S \ T \ X closure_of S \ X closure_of T" by (fastforce simp add: closure_of_def) lemma closure_of_subtopology_subset: "(subtopology X U) closure_of S \ (X closure_of S)" unfolding closure_of_subtopology by clarsimp (meson closure_of_mono contra_subsetD inf.cobounded2) lemma closure_of_subtopology_mono: "T \ U \ (subtopology X T) closure_of S \ (subtopology X U) closure_of S" unfolding closure_of_subtopology by auto (meson closure_of_mono inf_mono subset_iff) lemma closure_of_Un [simp]: "X closure_of (S \ T) = X closure_of S \ X closure_of T" by (simp add: Un_assoc Un_left_commute closure_of_alt derived_set_of_Un inf_sup_distrib1) lemma closure_of_Union: "finite \ \ X closure_of (\\) = (\S \ \. X closure_of S)" by (induction \ rule: finite_induct) auto lemma closure_of_subset: "S \ topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_Int: "topspace X \ S \ X closure_of S" by (auto simp: closure_of_def) lemma closure_of_subset_eq: "S \ topspace X \ X closure_of S \ S \ closedin X S" proof - have "openin X (topspace X - S)" if "\x. \x \ topspace X; \T. x \ T \ openin X T \ S \ T \ {}\ \ x \ S" apply (subst openin_subopen) by (metis Diff_iff Diff_mono Diff_triv inf.commute openin_subset order_refl that) then show ?thesis by (auto simp: closedin_def closure_of_def disjoint_iff_not_equal) qed lemma closure_of_eq: "X closure_of S = S \ closedin X S" by (metis closure_of_subset closure_of_subset_eq closure_of_subset_topspace subset_antisym) lemma closedin_contains_derived_set: "closedin X S \ X derived_set_of S \ S \ S \ topspace X" proof (intro iffI conjI) show "closedin X S \ X derived_set_of S \ S" using closure_of_eq derived_set_of_subset_closure_of by fastforce show "closedin X S \ S \ topspace X" using closedin_subset by blast show "X derived_set_of S \ S \ S \ topspace X \ closedin X S" by (metis closure_of closure_of_eq inf.absorb_iff2 sup.orderE) qed lemma derived_set_subset_gen: "X derived_set_of S \ S \ closedin X (topspace X \ S)" by (simp add: closedin_contains_derived_set derived_set_of_subset_topspace) lemma derived_set_subset: "S \ topspace X \ (X derived_set_of S \ S \ closedin X S)" by (simp add: closedin_contains_derived_set) lemma closedin_derived_set: "closedin (subtopology X T) S \ S \ topspace X \ S \ T \ (\x. x \ X derived_set_of S \ x \ T \ x \ S)" by (auto simp: closedin_contains_derived_set derived_set_of_subtopology Int_absorb1) lemma closedin_Int_closure_of: "closedin (subtopology X S) T \ S \ X closure_of T = T" by (metis Int_left_absorb closure_of_eq closure_of_subtopology) lemma closure_of_closedin: "closedin X S \ X closure_of S = S" by (simp add: closure_of_eq) lemma closure_of_eq_diff: "X closure_of S = topspace X - \{T. openin X T \ disjnt S T}" by (auto simp: closure_of_def disjnt_iff) lemma closedin_closure_of [simp]: "closedin X (X closure_of S)" unfolding closure_of_eq_diff by blast lemma closure_of_closure_of [simp]: "X closure_of (X closure_of S) = X closure_of S" by (simp add: closure_of_eq) lemma closure_of_hull: assumes "S \ topspace X" shows "X closure_of S = (closedin X) hull S" by (metis assms closedin_closure_of closure_of_eq closure_of_mono closure_of_subset hull_unique) lemma closure_of_minimal: "\S \ T; closedin X T\ \ (X closure_of S) \ T" by (metis closure_of_eq closure_of_mono) lemma closure_of_minimal_eq: "\S \ topspace X; closedin X T\ \ (X closure_of S) \ T \ S \ T" by (meson closure_of_minimal closure_of_subset subset_trans) lemma closure_of_unique: "\S \ T; closedin X T; \T'. \S \ T'; closedin X T'\ \ T \ T'\ \ X closure_of S = T" by (meson closedin_closure_of closedin_subset closure_of_minimal closure_of_subset eq_iff order.trans) lemma closure_of_eq_empty_gen: "X closure_of S = {} \ disjnt (topspace X) S" unfolding disjnt_def closure_of_restrict [where S=S] using closure_of by fastforce lemma closure_of_eq_empty: "S \ topspace X \ X closure_of S = {} \ S = {}" using closure_of_subset by fastforce lemma openin_Int_closure_of_subset: assumes "openin X S" shows "S \ X closure_of T \ X closure_of (S \ T)" proof - have "S \ X derived_set_of T = S \ X derived_set_of (S \ T)" by (meson assms openin_Int_derived_set_of_eq) moreover have "S \ (S \ T) = S \ T" by fastforce ultimately show ?thesis by (metis closure_of_alt inf.cobounded2 inf_left_commute inf_sup_distrib1) qed lemma closure_of_openin_Int_closure_of: assumes "openin X S" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" proof show "X closure_of (S \ X closure_of T) \ X closure_of (S \ T)" by (simp add: assms closure_of_minimal openin_Int_closure_of_subset) next show "X closure_of (S \ T) \ X closure_of (S \ X closure_of T)" by (metis Int_subset_iff assms closure_of_alt closure_of_mono inf_mono openin_subset subset_refl sup.coboundedI1) qed lemma openin_Int_closure_of_eq: assumes "openin X S" shows "S \ X closure_of T = S \ X closure_of (S \ T)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: assms openin_Int_closure_of_subset) show "?rhs \ ?lhs" by (metis closure_of_mono inf_commute inf_le1 inf_mono order_refl) qed lemma openin_Int_closure_of_eq_empty: assumes "openin X S" shows "S \ X closure_of T = {} \ S \ T = {}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding disjoint_iff by (meson assms in_closure_of in_mono openin_subset) show "?rhs \ ?lhs" by (simp add: assms openin_Int_closure_of_eq) qed lemma closure_of_openin_Int_superset: "openin X S \ S \ X closure_of T \ X closure_of (S \ T) = X closure_of S" by (metis closure_of_openin_Int_closure_of inf.orderE) lemma closure_of_openin_subtopology_Int_closure_of: assumes S: "openin (subtopology X U) S" and "T \ U" shows "X closure_of (S \ X closure_of T) = X closure_of (S \ T)" (is "?lhs = ?rhs") proof obtain S0 where S0: "openin X S0" "S = S0 \ U" using assms by (auto simp: openin_subtopology) then show "?lhs \ ?rhs" proof - have "S0 \ X closure_of T = S0 \ X closure_of (S0 \ T)" by (meson S0(1) openin_Int_closure_of_eq) moreover have "S0 \ T = S0 \ U \ T" using \T \ U\ by fastforce ultimately have "S \ X closure_of T \ X closure_of (S \ T)" using S0(2) by auto then show ?thesis by (meson closedin_closure_of closure_of_minimal) qed next show "?rhs \ ?lhs" proof - have "T \ S \ T \ X derived_set_of T" by force then show ?thesis by (smt (verit, del_insts) Int_iff in_closure_of inf.orderE openin_subset subsetI) qed qed lemma closure_of_subtopology_open: "openin X U \ S \ U \ (subtopology X U) closure_of S = U \ X closure_of S" by (metis closure_of_subtopology inf_absorb2 openin_Int_closure_of_eq) lemma discrete_topology_closure_of: "(discrete_topology U) closure_of S = U \ S" by (metis closedin_discrete_topology closure_of_restrict closure_of_unique discrete_topology_unique inf_sup_ord(1) order_refl) text\ Interior with respect to a topological space. \ definition interior_of :: "'a topology \ 'a set \ 'a set" (infixr "interior'_of" 80) where "X interior_of S \ {x. \T. openin X T \ x \ T \ T \ S}" lemma interior_of_restrict: "X interior_of S = X interior_of (topspace X \ S)" using openin_subset by (auto simp: interior_of_def) lemma interior_of_eq: "(X interior_of S = S) \ openin X S" unfolding interior_of_def using openin_subopen by blast lemma interior_of_openin: "openin X S \ X interior_of S = S" by (simp add: interior_of_eq) lemma interior_of_empty [simp]: "X interior_of {} = {}" by (simp add: interior_of_eq) lemma interior_of_topspace [simp]: "X interior_of (topspace X) = topspace X" by (simp add: interior_of_eq) lemma openin_interior_of [simp]: "openin X (X interior_of S)" unfolding interior_of_def using openin_subopen by fastforce lemma interior_of_interior_of [simp]: "X interior_of X interior_of S = X interior_of S" by (simp add: interior_of_eq) lemma interior_of_subset: "X interior_of S \ S" by (auto simp: interior_of_def) lemma interior_of_subset_closure_of: "X interior_of S \ X closure_of S" by (metis closure_of_subset_Int dual_order.trans interior_of_restrict interior_of_subset) lemma subset_interior_of_eq: "S \ X interior_of S \ openin X S" by (metis interior_of_eq interior_of_subset subset_antisym) lemma interior_of_mono: "S \ T \ X interior_of S \ X interior_of T" by (auto simp: interior_of_def) lemma interior_of_maximal: "\T \ S; openin X T\ \ T \ X interior_of S" by (auto simp: interior_of_def) lemma interior_of_maximal_eq: "openin X T \ T \ X interior_of S \ T \ S" by (meson interior_of_maximal interior_of_subset order_trans) lemma interior_of_unique: "\T \ S; openin X T; \T'. \T' \ S; openin X T'\ \ T' \ T\ \ X interior_of S = T" by (simp add: interior_of_maximal_eq interior_of_subset subset_antisym) lemma interior_of_subset_topspace: "X interior_of S \ topspace X" by (simp add: openin_subset) lemma interior_of_subset_subtopology: "(subtopology X S) interior_of T \ S" by (meson openin_imp_subset openin_interior_of) lemma interior_of_Int: "X interior_of (S \ T) = X interior_of S \ X interior_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: interior_of_mono) show "?rhs \ ?lhs" by (meson inf_mono interior_of_maximal interior_of_subset openin_Int openin_interior_of) qed lemma interior_of_Inter_subset: "X interior_of (\\) \ (\S \ \. X interior_of S)" by (simp add: INT_greatest Inf_lower interior_of_mono) lemma union_interior_of_subset: "X interior_of S \ X interior_of T \ X interior_of (S \ T)" by (simp add: interior_of_mono) lemma interior_of_eq_empty: "X interior_of S = {} \ (\T. openin X T \ T \ S \ T = {})" by (metis bot.extremum_uniqueI interior_of_maximal interior_of_subset openin_interior_of) lemma interior_of_eq_empty_alt: "X interior_of S = {} \ (\T. openin X T \ T \ {} \ T - S \ {})" by (auto simp: interior_of_eq_empty) lemma interior_of_Union_openin_subsets: "\{T. openin X T \ T \ S} = X interior_of S" by (rule interior_of_unique [symmetric]) auto lemma interior_of_complement: "X interior_of (topspace X - S) = topspace X - X closure_of S" by (auto simp: interior_of_def closure_of_def) lemma interior_of_closure_of: "X interior_of S = topspace X - X closure_of (topspace X - S)" unfolding interior_of_complement [symmetric] by (metis Diff_Diff_Int interior_of_restrict) lemma closure_of_interior_of: "X closure_of S = topspace X - X interior_of (topspace X - S)" by (simp add: interior_of_complement Diff_Diff_Int closure_of) lemma closure_of_complement: "X closure_of (topspace X - S) = topspace X - X interior_of S" unfolding interior_of_def closure_of_def by (blast dest: openin_subset) lemma interior_of_eq_empty_complement: "X interior_of S = {} \ X closure_of (topspace X - S) = topspace X" using interior_of_subset_topspace [of X S] closure_of_complement by fastforce lemma closure_of_eq_topspace: "X closure_of S = topspace X \ X interior_of (topspace X - S) = {}" using closure_of_subset_topspace [of X S] interior_of_complement by fastforce lemma interior_of_subtopology_subset: "U \ X interior_of S \ (subtopology X U) interior_of S" by (auto simp: interior_of_def openin_subtopology) lemma interior_of_subtopology_subsets: "T \ U \ T \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis inf.absorb_iff2 interior_of_subtopology_subset subtopology_subtopology) lemma interior_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X U) interior_of S \ (subtopology X T) interior_of S" by (metis dual_order.trans inf.orderE inf_commute interior_of_subset interior_of_subtopology_subsets) lemma interior_of_subtopology_open: assumes "openin X U" shows "(subtopology X U) interior_of S = U \ X interior_of S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (meson assms interior_of_maximal interior_of_subset le_infI openin_interior_of openin_open_subtopology) show "?rhs \ ?lhs" by (simp add: interior_of_subtopology_subset) qed lemma dense_intersects_open: "X closure_of S = topspace X \ (\T. openin X T \ T \ {} \ S \ T \ {})" proof - have "X closure_of S = topspace X \ (topspace X - X interior_of (topspace X - S) = topspace X)" by (simp add: closure_of_interior_of) also have "\ \ X interior_of (topspace X - S) = {}" by (simp add: closure_of_complement interior_of_eq_empty_complement) also have "\ \ (\T. openin X T \ T \ {} \ S \ T \ {})" unfolding interior_of_eq_empty_alt using openin_subset by fastforce finally show ?thesis . qed lemma interior_of_closedin_union_empty_interior_of: assumes "closedin X S" and disj: "X interior_of T = {}" shows "X interior_of (S \ T) = X interior_of S" proof - have "X closure_of (topspace X - T) = topspace X" by (metis Diff_Diff_Int disj closure_of_eq_topspace closure_of_restrict interior_of_closure_of) then show ?thesis unfolding interior_of_closure_of by (metis Diff_Un Diff_subset assms(1) closedin_def closure_of_openin_Int_superset) qed lemma interior_of_union_eq_empty: "closedin X S \ (X interior_of (S \ T) = {} \ X interior_of S = {} \ X interior_of T = {})" by (metis interior_of_closedin_union_empty_interior_of le_sup_iff subset_empty union_interior_of_subset) lemma discrete_topology_interior_of [simp]: "(discrete_topology U) interior_of S = U \ S" by (simp add: interior_of_restrict [of _ S] interior_of_eq) subsection \Frontier with respect to topological space \ definition frontier_of :: "'a topology \ 'a set \ 'a set" (infixr "frontier'_of" 80) where "X frontier_of S \ X closure_of S - X interior_of S" lemma frontier_of_closures: "X frontier_of S = X closure_of S \ X closure_of (topspace X - S)" by (metis Diff_Diff_Int closure_of_complement closure_of_subset_topspace double_diff frontier_of_def interior_of_subset_closure_of) lemma interior_of_union_frontier_of [simp]: "X interior_of S \ X frontier_of S = X closure_of S" by (simp add: frontier_of_def interior_of_subset_closure_of subset_antisym) lemma frontier_of_restrict: "X frontier_of S = X frontier_of (topspace X \ S)" by (metis closure_of_restrict frontier_of_def interior_of_restrict) lemma closedin_frontier_of: "closedin X (X frontier_of S)" by (simp add: closedin_Int frontier_of_closures) lemma frontier_of_subset_topspace: "X frontier_of S \ topspace X" by (simp add: closedin_frontier_of closedin_subset) lemma frontier_of_subset_subtopology: "(subtopology X S) frontier_of T \ S" by (metis (no_types) closedin_derived_set closedin_frontier_of) lemma frontier_of_subtopology_subset: "U \ (subtopology X U) frontier_of S \ (X frontier_of S)" proof - have "U \ X interior_of S - subtopology X U interior_of S = {}" by (simp add: interior_of_subtopology_subset) moreover have "X closure_of S \ subtopology X U closure_of S = subtopology X U closure_of S" by (meson closure_of_subtopology_subset inf.absorb_iff2) ultimately show ?thesis unfolding frontier_of_def by blast qed lemma frontier_of_subtopology_mono: "\S \ T; T \ U\ \ (subtopology X T) frontier_of S \ (subtopology X U) frontier_of S" by (simp add: frontier_of_def Diff_mono closure_of_subtopology_mono interior_of_subtopology_mono) lemma clopenin_eq_frontier_of: "closedin X S \ openin X S \ S \ topspace X \ X frontier_of S = {}" proof (cases "S \ topspace X") case True then show ?thesis by (metis Diff_eq_empty_iff closure_of_eq closure_of_subset_eq frontier_of_def interior_of_eq interior_of_subset interior_of_union_frontier_of sup_bot_right) next case False then show ?thesis by (simp add: frontier_of_closures openin_closedin_eq) qed lemma frontier_of_eq_empty: "S \ topspace X \ (X frontier_of S = {} \ closedin X S \ openin X S)" by (simp add: clopenin_eq_frontier_of) lemma frontier_of_openin: "openin X S \ X frontier_of S = X closure_of S - S" by (metis (no_types) frontier_of_def interior_of_eq) lemma frontier_of_openin_straddle_Int: assumes "openin X U" "U \ X frontier_of S \ {}" shows "U \ S \ {}" "U - S \ {}" proof - have "U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}" using assms by (simp add: frontier_of_closures) then show "U \ S \ {}" using assms openin_Int_closure_of_eq_empty by fastforce show "U - S \ {}" proof - have "\A. X closure_of (A - S) \ U \ {}" using \U \ (X closure_of S \ X closure_of (topspace X - S)) \ {}\ by blast then have "\ U \ S" by (metis Diff_disjoint Diff_eq_empty_iff Int_Diff assms(1) inf_commute openin_Int_closure_of_eq_empty) then show ?thesis by blast qed qed lemma frontier_of_subset_closedin: "closedin X S \ (X frontier_of S) \ S" using closure_of_eq frontier_of_def by fastforce lemma frontier_of_empty [simp]: "X frontier_of {} = {}" by (simp add: frontier_of_def) lemma frontier_of_topspace [simp]: "X frontier_of topspace X = {}" by (simp add: frontier_of_def) lemma frontier_of_subset_eq: assumes "S \ topspace X" shows "(X frontier_of S) \ S \ closedin X S" proof show "X frontier_of S \ S \ closedin X S" by (metis assms closure_of_subset_eq interior_of_subset interior_of_union_frontier_of le_sup_iff) show "closedin X S \ X frontier_of S \ S" by (simp add: frontier_of_subset_closedin) qed lemma frontier_of_complement: "X frontier_of (topspace X - S) = X frontier_of S" by (metis Diff_Diff_Int closure_of_restrict frontier_of_closures inf_commute) lemma frontier_of_disjoint_eq: assumes "S \ topspace X" shows "((X frontier_of S) \ S = {} \ openin X S)" proof assume "X frontier_of S \ S = {}" then have "closedin X (topspace X - S)" using assms closure_of_subset frontier_of_def interior_of_eq interior_of_subset by fastforce then show "openin X S" using assms by (simp add: openin_closedin) next show "openin X S \ X frontier_of S \ S = {}" by (simp add: Diff_Diff_Int closedin_def frontier_of_openin inf.absorb_iff2 inf_commute) qed lemma frontier_of_disjoint_eq_alt: "S \ (topspace X - X frontier_of S) \ openin X S" proof (cases "S \ topspace X") case True show ?thesis using True frontier_of_disjoint_eq by auto next case False then show ?thesis by (meson Diff_subset openin_subset subset_trans) qed lemma frontier_of_Int: "X frontier_of (S \ T) = X closure_of (S \ T) \ (X frontier_of S \ X frontier_of T)" proof - have *: "U \ S \ U \ T \ U \ (S \ A \ T \ B) = U \ (A \ B)" for U S T A B :: "'a set" by blast show ?thesis by (simp add: frontier_of_closures closure_of_mono Diff_Int * flip: closure_of_Un) qed lemma frontier_of_Int_subset: "X frontier_of (S \ T) \ X frontier_of S \ X frontier_of T" by (simp add: frontier_of_Int) lemma frontier_of_Int_closedin: assumes "closedin X S" "closedin X T" shows "X frontier_of(S \ T) = X frontier_of S \ T \ S \ X frontier_of T" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms by (force simp add: frontier_of_Int closedin_Int closure_of_closedin) show "?rhs \ ?lhs" using assms frontier_of_subset_closedin by (auto simp add: frontier_of_Int closedin_Int closure_of_closedin) qed lemma frontier_of_Un_subset: "X frontier_of(S \ T) \ X frontier_of S \ X frontier_of T" by (metis Diff_Un frontier_of_Int_subset frontier_of_complement) lemma frontier_of_Union_subset: "finite \ \ X frontier_of (\\) \ (\T \ \. X frontier_of T)" proof (induction \ rule: finite_induct) case (insert A \) then show ?case using frontier_of_Un_subset by fastforce qed simp lemma frontier_of_frontier_of_subset: "X frontier_of (X frontier_of S) \ X frontier_of S" by (simp add: closedin_frontier_of frontier_of_subset_closedin) lemma frontier_of_subtopology_open: "openin X U \ (subtopology X U) frontier_of S = U \ X frontier_of S" by (simp add: Diff_Int_distrib closure_of_subtopology_open frontier_of_def interior_of_subtopology_open) lemma discrete_topology_frontier_of [simp]: "(discrete_topology U) frontier_of S = {}" by (simp add: Diff_eq discrete_topology_closure_of frontier_of_closures) lemma openin_subset_topspace_eq: assumes "disjnt S (X frontier_of U)" shows "openin (subtopology X U) S \ openin X S \ S \ U" proof assume S: "openin (subtopology X U) S" show "openin X S \ S \ U" proof show "S \ U" using S openin_imp_subset by blast have "disjnt S (X frontier_of (topspace X \ U))" by (metis assms frontier_of_restrict) moreover have "openin (subtopology X (topspace X \ U)) S" by (simp add: S subtopology_restrict) moreover have "openin X S" if dis: "disjnt S (X frontier_of U)" and ope: "openin (subtopology X U) S" and "U \ topspace X" for S U proof - obtain T where T: "openin X T" "S = T \ U" using ope by (auto simp: openin_subtopology) have "T \ U = T \ X interior_of U" using that T interior_of_subset in_closure_of by (fastforce simp: disjnt_iff frontier_of_def) then show ?thesis using \S = T \ U\ \openin X T\ by auto qed ultimately show "openin X S" by blast qed qed (metis inf.absorb_iff1 openin_subtopology_Int) subsection\Locally finite collections\ definition locally_finite_in where "locally_finite_in X \ \ (\\ \ topspace X) \ (\x \ topspace X. \V. openin X V \ x \ V \ finite {U \ \. U \ V \ {}})" lemma finite_imp_locally_finite_in: "\finite \; \\ \ topspace X\ \ locally_finite_in X \" by (auto simp: locally_finite_in_def) lemma locally_finite_in_subset: assumes "locally_finite_in X \" "\ \ \" shows "locally_finite_in X \" proof - have "finite (\ \ {U. U \ V \ {}}) \ finite (\ \ {U. U \ V \ {}})" for V by (meson \\ \ \\ finite_subset inf_le1 inf_le2 le_inf_iff subset_trans) then show ?thesis using assms unfolding locally_finite_in_def Int_def by fastforce qed lemma locally_finite_in_refinement: assumes \: "locally_finite_in X \" and f: "\S. S \ \ \ f S \ S" shows "locally_finite_in X (f ` \)" proof - show ?thesis unfolding locally_finite_in_def proof safe fix x assume "x \ topspace X" then obtain V where "openin X V" "x \ V" "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast moreover have "{U \ \. f U \ V \ {}} \ {U \ \. U \ V \ {}}" for V using f by blast ultimately have "finite {U \ \. f U \ V \ {}}" using finite_subset by blast moreover have "f ` {U \ \. f U \ V \ {}} = {U \ f ` \. U \ V \ {}}" by blast ultimately have "finite {U \ f ` \. U \ V \ {}}" by (metis (no_types, lifting) finite_imageI) then show "\V. openin X V \ x \ V \ finite {U \ f ` \. U \ V \ {}}" using \openin X V\ \x \ V\ by blast next show "\x xa. \xa \ \; x \ f xa\ \ x \ topspace X" by (meson Sup_upper \ f locally_finite_in_def subset_iff) qed qed lemma locally_finite_in_subtopology: assumes \: "locally_finite_in X \" "\\ \ S" shows "locally_finite_in (subtopology X S) \" unfolding locally_finite_in_def proof safe fix x assume x: "x \ topspace (subtopology X S)" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def topspace_subtopology by blast show "\V. openin (subtopology X S) V \ x \ V \ finite {U \ \. U \ V \ {}}" proof (intro exI conjI) show "openin (subtopology X S) (S \ V)" by (simp add: \openin X V\ openin_subtopology_Int2) have "{U \ \. U \ (S \ V) \ {}} \ {U \ \. U \ V \ {}}" by auto with fin show "finite {U \ \. U \ (S \ V) \ {}}" using finite_subset by auto show "x \ S \ V" using x \x \ V\ by (simp) qed next show "\x A. \x \ A; A \ \\ \ x \ topspace (subtopology X S)" using assms unfolding locally_finite_in_def topspace_subtopology by blast qed lemma closedin_locally_finite_Union: assumes clo: "\S. S \ \ \ closedin X S" and \: "locally_finite_in X \" shows "closedin X (\\)" using \ unfolding locally_finite_in_def closedin_def proof clarify show "openin X (topspace X - \\)" proof (subst openin_subopen, clarify) fix x assume "x \ topspace X" and "x \ \\" then obtain V where "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast let ?T = "V - \{S \ \. S \ V \ {}}" show "\T. openin X T \ x \ T \ T \ topspace X - \\" proof (intro exI conjI) show "openin X ?T" by (metis (no_types, lifting) fin \openin X V\ clo closedin_Union mem_Collect_eq openin_diff) show "x \ ?T" using \x \ \\\ \x \ V\ by auto show "?T \ topspace X - \\" using \openin X V\ openin_subset by auto qed qed qed lemma locally_finite_in_closure: assumes \: "locally_finite_in X \" shows "locally_finite_in X ((\S. X closure_of S) ` \)" using \ unfolding locally_finite_in_def proof (intro conjI; clarsimp) fix x A assume "x \ X closure_of A" then show "x \ topspace X" by (meson in_closure_of) next fix x assume "x \ topspace X" and "\\ \ topspace X" then obtain V where V: "openin X V" "x \ V" and fin: "finite {U \ \. U \ V \ {}}" using \ unfolding locally_finite_in_def by blast have eq: "{y \ f ` \. Q y} = f ` {x. x \ \ \ Q(f x)}" for f and Q :: "'a set \ bool" by blast have eq2: "{A \ \. X closure_of A \ V \ {}} = {A \ \. A \ V \ {}}" using openin_Int_closure_of_eq_empty V by blast have "finite {U \ (closure_of) X ` \. U \ V \ {}}" by (simp add: eq eq2 fin) with V show "\V. openin X V \ x \ V \ finite {U \ (closure_of) X ` \. U \ V \ {}}" by blast qed lemma closedin_Union_locally_finite_closure: "locally_finite_in X \ \ closedin X (\((\S. X closure_of S) ` \))" by (metis (mono_tags) closedin_closure_of closedin_locally_finite_Union imageE locally_finite_in_closure) lemma closure_of_Union_subset: "\((\S. X closure_of S) ` \) \ X closure_of (\\)" by (simp add: SUP_le_iff Sup_upper closure_of_mono) lemma closure_of_locally_finite_Union: assumes "locally_finite_in X \" shows "X closure_of (\\) = \((\S. X closure_of S) ` \)" proof (rule closure_of_unique) show "\ \ \ \ ((closure_of) X ` \)" using assms by (simp add: SUP_upper2 Sup_le_iff closure_of_subset locally_finite_in_def) show "closedin X (\ ((closure_of) X ` \))" using assms by (simp add: closedin_Union_locally_finite_closure) show "\T'. \\ \ \ T'; closedin X T'\ \ \ ((closure_of) X ` \) \ T'" by (simp add: Sup_le_iff closure_of_minimal) qed subsection\<^marker>\tag important\ \Continuous maps\ text \We will need to deal with continuous maps in terms of topologies and not in terms of type classes, as defined below.\ definition continuous_map where "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" lemma continuous_map: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X {x \ topspace X. f x \ U})" by (auto simp: continuous_map_def) lemma continuous_map_image_subset_topspace: "continuous_map X Y f \ f ` (topspace X) \ topspace Y" by (auto simp: continuous_map_def) lemma continuous_map_on_empty: "topspace X = {} \ continuous_map X Y f" by (auto simp: continuous_map_def) lemma continuous_map_closedin: "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" proof - have "(\U. openin Y U \ openin X {x \ topspace X. f x \ U}) = (\C. closedin Y C \ closedin X {x \ topspace X. f x \ C})" if "\x. x \ topspace X \ f x \ topspace Y" proof - have eq: "{x \ topspace X. f x \ topspace Y \ f x \ C} = (topspace X - {x \ topspace X. f x \ C})" for C using that by blast show ?thesis proof (intro iffI allI impI) fix C assume "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" by (auto simp add: closedin_def eq) next fix U assume "\C. closedin Y C \ closedin X {x \ topspace X. f x \ C}" and "openin Y U" then show "openin X {x \ topspace X. f x \ U}" by (auto simp add: openin_closedin_eq eq) qed qed then show ?thesis by (auto simp: continuous_map_def) qed lemma openin_continuous_map_preimage: "\continuous_map X Y f; openin Y U\ \ openin X {x \ topspace X. f x \ U}" by (simp add: continuous_map_def) lemma closedin_continuous_map_preimage: "\continuous_map X Y f; closedin Y C\ \ closedin X {x \ topspace X. f x \ C}" by (simp add: continuous_map_closedin) lemma openin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "openin X U" "openin Y V" shows "openin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) openin_closedin_eq by fastforce show ?thesis unfolding eq using assms openin_continuous_map_preimage by fastforce qed lemma closedin_continuous_map_preimage_gen: assumes "continuous_map X Y f" "closedin X U" "closedin Y V" shows "closedin X {x \ U. f x \ V}" proof - have eq: "{x \ U. f x \ V} = U \ {x \ topspace X. f x \ V}" using assms(2) closedin_def by fastforce show ?thesis unfolding eq using assms closedin_continuous_map_preimage by fastforce qed lemma continuous_map_image_closure_subset: assumes "continuous_map X Y f" shows "f ` (X closure_of S) \ Y closure_of f ` S" proof - have *: "f ` (topspace X) \ topspace Y" by (meson assms continuous_map) have "X closure_of T \ {x \ X closure_of T. f x \ Y closure_of (f ` T)}" if "T \ topspace X" for T proof (rule closure_of_minimal) show "T \ {x \ X closure_of T. f x \ Y closure_of f ` T}" using closure_of_subset * that by (fastforce simp: in_closure_of) next show "closedin X {x \ X closure_of T. f x \ Y closure_of f ` T}" using assms closedin_continuous_map_preimage_gen by fastforce qed then show ?thesis by (smt (verit, ccfv_threshold) assms continuous_map image_eqI image_subset_iff in_closure_of mem_Collect_eq) qed lemma continuous_map_subset_aux1: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_image_closure_subset by blast lemma continuous_map_subset_aux2: assumes "\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S" shows "continuous_map X Y f" unfolding continuous_map_closedin proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "f x \ topspace Y" using assms closure_of_subset_topspace by fastforce next fix C assume "closedin Y C" then show "closedin X {x \ topspace X. f x \ C}" proof (clarsimp simp flip: closure_of_subset_eq, intro conjI) fix x assume x: "x \ X closure_of {x \ topspace X. f x \ C}" and "C \ topspace Y" and "Y closure_of C \ C" show "x \ topspace X" by (meson x in_closure_of) have "{a \ topspace X. f a \ C} \ topspace X" by simp moreover have "Y closure_of f ` {a \ topspace X. f a \ C} \ C" by (simp add: \closedin Y C\ closure_of_minimal image_subset_iff) ultimately show "f x \ C" using x assms by blast qed qed lemma continuous_map_eq_image_closure_subset: "continuous_map X Y f \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_alt: "continuous_map X Y f \ (\S. S \ topspace X \ f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 by metis lemma continuous_map_eq_image_closure_subset_gen: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\S. f ` (X closure_of S) \ Y closure_of f ` S)" using continuous_map_subset_aux1 continuous_map_subset_aux2 continuous_map_image_subset_topspace by metis lemma continuous_map_closure_preimage_subset: "continuous_map X Y f \ X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" unfolding continuous_map_closedin by (rule closure_of_minimal) (use in_closure_of in \fastforce+\) lemma continuous_map_frontier_frontier_preimage_subset: assumes "continuous_map X Y f" shows "X frontier_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y frontier_of T}" proof - have eq: "topspace X - {x \ topspace X. f x \ T} = {x \ topspace X. f x \ topspace Y - T}" using assms unfolding continuous_map_def by blast have "X closure_of {x \ topspace X. f x \ T} \ {x \ topspace X. f x \ Y closure_of T}" by (simp add: assms continuous_map_closure_preimage_subset) moreover have "X closure_of (topspace X - {x \ topspace X. f x \ T}) \ {x \ topspace X. f x \ Y closure_of (topspace Y - T)}" using continuous_map_closure_preimage_subset [OF assms] eq by presburger ultimately show ?thesis by (auto simp: frontier_of_closures) qed lemma topology_finer_continuous_id: assumes "topspace X = topspace Y" shows "(\S. openin X S \ openin Y S) \ continuous_map Y X id" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding continuous_map_def using assms openin_subopen openin_subset by fastforce show "?rhs \ ?lhs" unfolding continuous_map_def using assms openin_subopen topspace_def by fastforce qed lemma continuous_map_const [simp]: "continuous_map X Y (\x. C) \ topspace X = {} \ C \ topspace Y" proof (cases "topspace X = {}") case False show ?thesis proof (cases "C \ topspace Y") case True with openin_subopen show ?thesis by (auto simp: continuous_map_def) next case False then show ?thesis unfolding continuous_map_def by fastforce qed qed (auto simp: continuous_map_on_empty) declare continuous_map_const [THEN iffD2, continuous_intros] lemma continuous_map_compose [continuous_intros]: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" shows "continuous_map X X'' (g \ f)" unfolding continuous_map_def proof (intro conjI ballI allI impI) fix x assume "x \ topspace X" then show "(g \ f) x \ topspace X''" using assms unfolding continuous_map_def by force next fix U assume "openin X'' U" have eq: "{x \ topspace X. (g \ f) x \ U} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U}}" by auto (meson f continuous_map_def) show "openin X {x \ topspace X. (g \ f) x \ U}" unfolding eq using assms unfolding continuous_map_def using \openin X'' U\ by blast qed lemma continuous_map_eq: assumes "continuous_map X X' f" and "\x. x \ topspace X \ f x = g x" shows "continuous_map X X' g" proof - have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. g x \ U}" for U using assms by auto show ?thesis using assms by (simp add: continuous_map_def eq) qed lemma restrict_continuous_map [simp]: "topspace X \ S \ continuous_map X X' (restrict f S) \ continuous_map X X' f" by (auto simp: elim!: continuous_map_eq) lemma continuous_map_in_subtopology: "continuous_map X (subtopology X' S) f \ continuous_map X X' f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof - have "\A. f ` (X closure_of A) \ subtopology X' S closure_of f ` A" by (meson L continuous_map_image_closure_subset) then show ?thesis by (metis (no_types) closure_of_subset_subtopology closure_of_subtopology_subset closure_of_topspace continuous_map_eq_image_closure_subset order.trans) qed next assume R: ?rhs then have eq: "{x \ topspace X. f x \ U} = {x \ topspace X. f x \ U \ f x \ S}" for U by auto show ?lhs using R unfolding continuous_map by (auto simp: openin_subtopology eq) qed lemma continuous_map_from_subtopology: "continuous_map X X' f \ continuous_map (subtopology X S) X' f" by (auto simp: continuous_map openin_subtopology) lemma continuous_map_into_fulltopology: "continuous_map X (subtopology X' T) f \ continuous_map X X' f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_into_subtopology: "\continuous_map X X' f; f ` topspace X \ T\ \ continuous_map X (subtopology X' T) f" by (auto simp: continuous_map_in_subtopology) lemma continuous_map_from_subtopology_mono: "\continuous_map (subtopology X T) X' f; S \ T\ \ continuous_map (subtopology X S) X' f" by (metis inf.absorb_iff2 continuous_map_from_subtopology subtopology_subtopology) lemma continuous_map_from_discrete_topology [simp]: "continuous_map (discrete_topology U) X f \ f ` U \ topspace X" by (auto simp: continuous_map_def) lemma continuous_map_iff_continuous [simp]: "continuous_map (top_of_set S) euclidean g = continuous_on S g" by (fastforce simp add: continuous_map openin_subtopology continuous_on_open_invariant) lemma continuous_map_iff_continuous2 [simp]: "continuous_map euclidean euclidean g = continuous_on UNIV g" by (metis continuous_map_iff_continuous subtopology_UNIV) lemma continuous_map_openin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. openin Y U \ openin X (topspace X \ f -` U))" by (auto simp: continuous_map_def vimage_def Int_def) lemma continuous_map_closedin_preimage_eq: "continuous_map X Y f \ f ` (topspace X) \ topspace Y \ (\U. closedin Y U \ closedin X (topspace X \ f -` U))" by (auto simp: continuous_map_closedin vimage_def Int_def) lemma continuous_map_square_root: "continuous_map euclideanreal euclideanreal sqrt" by (simp add: continuous_at_imp_continuous_on isCont_real_sqrt) lemma continuous_map_sqrt [continuous_intros]: "continuous_map X euclideanreal f \ continuous_map X euclideanreal (\x. sqrt(f x))" by (meson continuous_map_compose continuous_map_eq continuous_map_square_root o_apply) lemma continuous_map_id [simp, continuous_intros]: "continuous_map X X id" unfolding continuous_map_def using openin_subopen topspace_def by fastforce declare continuous_map_id [unfolded id_def, simp, continuous_intros] lemma continuous_map_id_subt [simp]: "continuous_map (subtopology X S) X id" by (simp add: continuous_map_from_subtopology) declare continuous_map_id_subt [unfolded id_def, simp] lemma\<^marker>\tag important\ continuous_map_alt: "continuous_map T1 T2 f = ((\U. openin T2 U \ openin T1 (f -` U \ topspace T1)) \ f ` topspace T1 \ topspace T2)" by (auto simp: continuous_map_def vimage_def image_def Collect_conj_eq inf_commute) lemma continuous_map_open [intro]: "continuous_map T1 T2 f \ openin T2 U \ openin T1 (f-`U \ topspace(T1))" unfolding continuous_map_alt by auto lemma continuous_map_preimage_topspace [intro]: assumes "continuous_map T1 T2 f" shows "f-`(topspace T2) \ topspace T1 = topspace T1" using assms unfolding continuous_map_def by auto subsection\Open and closed maps (not a priori assumed continuous)\ definition open_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "open_map X1 X2 f \ \U. openin X1 U \ openin X2 (f ` U)" definition closed_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "closed_map X1 X2 f \ \U. closedin X1 U \ closedin X2 (f ` U)" lemma open_map_imp_subset_topspace: "open_map X1 X2 f \ f ` (topspace X1) \ topspace X2" unfolding open_map_def by (simp add: openin_subset) lemma open_map_on_empty: "topspace X = {} \ open_map X Y f" by (metis empty_iff imageE in_mono open_map_def openin_subopen openin_subset) lemma closed_map_on_empty: "topspace X = {} \ closed_map X Y f" by (simp add: closed_map_def closedin_topspace_empty) lemma closed_map_const: "closed_map X Y (\x. c) \ topspace X = {} \ closedin Y {c}" by (metis closed_map_def closed_map_on_empty closedin_empty closedin_topspace image_constant_conv) lemma open_map_imp_subset: "\open_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" by (meson order_trans open_map_imp_subset_topspace subset_image_iff) lemma topology_finer_open_id: "(\S. openin X S \ openin X' S) \ open_map X X' id" unfolding open_map_def by auto lemma open_map_id: "open_map X X id" unfolding open_map_def by auto lemma open_map_eq: "\open_map X X' f; \x. x \ topspace X \ f x = g x\ \ open_map X X' g" unfolding open_map_def by (metis image_cong openin_subset subset_iff) lemma open_map_inclusion_eq: "open_map (subtopology X S) X id \ openin X (topspace X \ S)" by (metis openin_topspace openin_trans_full subtopology_restrict topology_finer_open_id topspace_subtopology) lemma open_map_inclusion: "openin X S \ open_map (subtopology X S) X id" by (simp add: open_map_inclusion_eq openin_Int) lemma open_map_compose: "\open_map X X' f; open_map X' X'' g\ \ open_map X X'' (g \ f)" by (metis (no_types, lifting) image_comp open_map_def) lemma closed_map_imp_subset_topspace: "closed_map X1 X2 f \ f ` (topspace X1) \ topspace X2" by (simp add: closed_map_def closedin_subset) lemma closed_map_imp_subset: "\closed_map X1 X2 f; S \ topspace X1\ \ f ` S \ topspace X2" using closed_map_imp_subset_topspace by blast lemma topology_finer_closed_id: "(\S. closedin X S \ closedin X' S) \ closed_map X X' id" by (simp add: closed_map_def) lemma closed_map_id: "closed_map X X id" by (simp add: closed_map_def) lemma closed_map_eq: "\closed_map X X' f; \x. x \ topspace X \ f x = g x\ \ closed_map X X' g" unfolding closed_map_def by (metis image_cong closedin_subset subset_iff) lemma closed_map_compose: "\closed_map X X' f; closed_map X' X'' g\ \ closed_map X X'' (g \ f)" by (metis (no_types, lifting) closed_map_def image_comp) lemma closed_map_inclusion_eq: "closed_map (subtopology X S) X id \ closedin X (topspace X \ S)" proof - have *: "closedin X (T \ S)" if "closedin X (S \ topspace X)" "closedin X T" for T by (smt (verit, best) closedin_Int closure_of_subset_eq inf_sup_aci le_iff_inf that) then show ?thesis by (fastforce simp add: closed_map_def Int_commute closedin_subtopology_alt intro: *) qed lemma closed_map_inclusion: "closedin X S \ closed_map (subtopology X S) X id" by (simp add: closed_map_inclusion_eq closedin_Int) lemma open_map_into_subtopology: "\open_map X X' f; f ` topspace X \ S\ \ open_map X (subtopology X' S) f" unfolding open_map_def openin_subtopology using openin_subset by fastforce lemma closed_map_into_subtopology: "\closed_map X X' f; f ` topspace X \ S\ \ closed_map X (subtopology X' S) f" unfolding closed_map_def closedin_subtopology using closedin_subset by fastforce lemma open_map_into_discrete_topology: "open_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding open_map_def openin_discrete_topology using openin_subset by blast lemma closed_map_into_discrete_topology: "closed_map X (discrete_topology U) f \ f ` (topspace X) \ U" unfolding closed_map_def closedin_discrete_topology using closedin_subset by blast lemma bijective_open_imp_closed_map: "\open_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ closed_map X X' f" unfolding open_map_def closed_map_def closedin_def by auto (metis Diff_subset inj_on_image_set_diff) lemma bijective_closed_imp_open_map: "\closed_map X X' f; f ` (topspace X) = topspace X'; inj_on f (topspace X)\ \ open_map X X' f" unfolding closed_map_def open_map_def openin_closedin_eq by auto (metis Diff_subset inj_on_image_set_diff) lemma open_map_from_subtopology: "\open_map X X' f; openin X U\ \ open_map (subtopology X U) X' f" unfolding open_map_def openin_subtopology_alt by blast lemma closed_map_from_subtopology: "\closed_map X X' f; closedin X U\ \ closed_map (subtopology X U) X' f" unfolding closed_map_def closedin_subtopology_alt by blast lemma open_map_restriction: assumes f: "open_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "open_map (subtopology X U) (subtopology X' V) f" unfolding open_map_def proof clarsimp fix W assume "openin (subtopology X U) W" then obtain T where "openin X T" "W = T \ U" by (meson openin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding open_map_def openin_closedin_eq by auto then show "openin (subtopology X' V) (f ` W)" by (metis \openin X T\ f open_map_def openin_subtopology_Int) qed lemma closed_map_restriction: assumes f: "closed_map X X' f" and U: "{x \ topspace X. f x \ V} = U" shows "closed_map (subtopology X U) (subtopology X' V) f" unfolding closed_map_def proof clarsimp fix W assume "closedin (subtopology X U) W" then obtain T where "closedin X T" "W = T \ U" by (meson closedin_subtopology) with f U have "f ` W = (f ` T) \ V" unfolding closed_map_def closedin_def by auto then show "closedin (subtopology X' V) (f ` W)" by (metis \closedin X T\ closed_map_def closedin_subtopology f) qed lemma closed_map_closure_of_image: "closed_map X Y f \ f ` topspace X \ topspace Y \ (\S. S \ topspace X \ Y closure_of (f ` S) \ image f (X closure_of S))" (is "?lhs=?rhs") proof assume ?lhs then show ?rhs by (simp add: closed_map_def closed_map_imp_subset_topspace closure_of_minimal closure_of_subset image_mono) next assume ?rhs then show ?lhs by (metis closed_map_def closed_map_into_discrete_topology closure_of_eq closure_of_subset_eq topspace_discrete_topology) qed lemma open_map_interior_of_image_subset: "open_map X Y f \ (\S. image f (X interior_of S) \ Y interior_of (f ` S))" by (metis image_mono interior_of_eq interior_of_maximal interior_of_subset open_map_def openin_interior_of subset_antisym) lemma open_map_interior_of_image_subset_alt: "open_map X Y f \ (\S\topspace X. f ` (X interior_of S) \ Y interior_of f ` S)" by (metis interior_of_eq open_map_def open_map_interior_of_image_subset openin_subset subset_interior_of_eq) lemma open_map_interior_of_image_subset_gen: "open_map X Y f \ (f ` topspace X \ topspace Y \ (\S. f ` (X interior_of S) \ Y interior_of f ` S))" by (meson open_map_imp_subset_topspace open_map_interior_of_image_subset) lemma open_map_preimage_neighbourhood: "open_map X Y f \ (f ` topspace X \ topspace Y \ (\U T. closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` topspace X \ topspace Y" by (simp add: \open_map X Y f\ open_map_imp_subset_topspace) next fix U T assume UT: "closedin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" have "closedin Y (topspace Y - f ` (topspace X - U))" by (meson UT L open_map_def openin_closedin_eq openin_diff openin_topspace) with UT show "\V. closedin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" using image_iff by auto qed next assume R: ?rhs show ?lhs unfolding open_map_def proof (intro strip) fix U assume "openin X U" have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" by blast then obtain V where V: "closedin Y V" and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" using R \openin X U\ by (meson Diff_subset openin_closedin_eq) then have "f ` U \ topspace Y - V" using R \openin X U\ openin_subset by fastforce with sub have "f ` U = topspace Y - V" by auto then show "openin Y (f ` U)" using V(1) by auto qed qed lemma closed_map_preimage_neighbourhood: "closed_map X Y f \ image f (topspace X) \ topspace Y \ (\U T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U))" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "f ` topspace X \ topspace Y" by (simp add: L closed_map_imp_subset_topspace) next fix U T assume UT: "openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U" then have "openin Y (topspace Y - f ` (topspace X - U))" by (meson L closed_map_def closedin_def closedin_diff closedin_topspace) then show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" using UT image_iff by auto qed next assume R: ?rhs show ?lhs unfolding closed_map_def proof (intro strip) fix U assume "closedin X U" have "{x \ topspace X. f x \ topspace Y - f ` U} \ topspace X - U" by blast then obtain V where V: "openin Y V" and sub: "topspace Y - f ` U \ V" "{x \ topspace X. f x \ V} \ topspace X - U" using R Diff_subset \closedin X U\ unfolding closedin_def by (smt (verit, ccfv_threshold) Collect_mem_eq Collect_mono_iff) then have "f ` U \ topspace Y - V" using R \closedin X U\ closedin_subset by fastforce with sub have "f ` U = topspace Y - V" by auto with V show "closedin Y (f ` U)" by auto qed qed lemma closed_map_fibre_neighbourhood: "closed_map X Y f \ f ` (topspace X) \ topspace Y \ (\U y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" unfolding closed_map_preimage_neighbourhood proof (intro conj_cong refl all_cong1) fix U assume "f ` topspace X \ topspace Y" show "(\T. openin X U \ T \ topspace Y \ {x \ topspace X. f x \ T} \ U \ (\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U)) = (\y. openin X U \ y \ topspace Y \ {x \ topspace X. f x = y} \ U \ (\V. openin Y V \ y \ V \ {x \ topspace X. f x \ V} \ U))" (is "(\T. ?P T) \ (\y. ?Q y)") proof assume L [rule_format]: "\T. ?P T" show "\y. ?Q y" proof fix y show "?Q y" using L [of "{y}"] by blast qed next assume R: "\y. ?Q y" show "\T. ?P T" proof (cases "openin X U") case True note [[unify_search_bound=3]] obtain V where V: "\y. \y \ topspace Y; {x \ topspace X. f x = y} \ U\ \ openin Y (V y) \ y \ V y \ {x \ topspace X. f x \ V y} \ U" using R by (simp add: True) meson show ?thesis proof clarify fix T assume "openin X U" and "T \ topspace Y" and "{x \ topspace X. f x \ T} \ U" with V show "\V. openin Y V \ T \ V \ {x \ topspace X. f x \ V} \ U" by (rule_tac x="\y\T. V y" in exI) fastforce qed qed auto qed qed lemma open_map_in_subtopology: "openin Y S \ (open_map X (subtopology Y S) f \ open_map X Y f \ f ` (topspace X) \ S)" by (metis le_inf_iff open_map_def open_map_imp_subset_topspace open_map_into_subtopology openin_trans_full topspace_subtopology) lemma open_map_from_open_subtopology: "\openin Y S; open_map X (subtopology Y S) f\ \ open_map X Y f" using open_map_in_subtopology by blast lemma closed_map_in_subtopology: "closedin Y S \ closed_map X (subtopology Y S) f \ (closed_map X Y f \ f ` topspace X \ S)" by (metis closed_map_def closed_map_imp_subset_topspace closed_map_into_subtopology closedin_closed_subtopology closedin_subset topspace_subtopology_subset) lemma closed_map_from_closed_subtopology: "\closedin Y S; closed_map X (subtopology Y S) f\ \ closed_map X Y f" using closed_map_in_subtopology by blast lemma closed_map_from_composition_left: assumes cmf: "closed_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "closed_map Y Z g" unfolding closed_map_def proof (intro strip) fix U assume "closedin Y U" then have "closedin X {x \ topspace X. f x \ U}" using contf closedin_continuous_map_preimage by blast then have "closedin Z ((g \ f) ` {x \ topspace X. f x \ U})" using cmf closed_map_def by blast moreover have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" by (smt (verit, ccfv_SIG) \closedin Y U\ closedin_subset fim image_iff subsetD) then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto ultimately show "closedin Z (g ` U)" by metis qed text \identical proof as the above\ lemma open_map_from_composition_left: assumes cmf: "open_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" shows "open_map Y Z g" unfolding open_map_def proof (intro strip) fix U assume "openin Y U" then have "openin X {x \ topspace X. f x \ U}" using contf openin_continuous_map_preimage by blast then have "openin Z ((g \ f) ` {x \ topspace X. f x \ U})" using cmf open_map_def by blast moreover have "\y. y \ U \ \x \ topspace X. f x \ U \ g y = g (f x)" by (smt (verit, ccfv_SIG) \openin Y U\ openin_subset fim image_iff subsetD) then have "(g \ f) ` {x \ topspace X. f x \ U} = g`U" by auto ultimately show "openin Z (g ` U)" by metis qed lemma closed_map_from_composition_right: assumes cmf: "closed_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" shows "closed_map X Y f" unfolding closed_map_def proof (intro strip) fix C assume "closedin X C" have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" using \closedin X C\ assms closedin_subset inj_onD by fastforce then have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" using \closedin X C\ assms(2) closedin_subset by fastforce moreover have "closedin Z ((g \ f) ` C)" using \closedin X C\ cmf closed_map_def by blast ultimately show "closedin Y (f ` C)" using assms(3) closedin_continuous_map_preimage by fastforce qed text \identical proof as the above\ lemma open_map_from_composition_right: assumes "open_map X Z (g \ f)" "f ` topspace X \ topspace Y" "continuous_map Y Z g" "inj_on g (topspace Y)" shows "open_map X Y f" unfolding open_map_def proof (intro strip) fix C assume "openin X C" have "\y c. \y \ topspace Y; g y = g (f c); c \ C\ \ y \ f ` C" using \openin X C\ assms openin_subset inj_onD by fastforce then have "f ` C = {x \ topspace Y. g x \ (g \ f) ` C}" using \openin X C\ assms(2) openin_subset by fastforce moreover have "openin Z ((g \ f) ` C)" using \openin X C\ assms(1) open_map_def by blast ultimately show "openin Y (f ` C)" using assms(3) openin_continuous_map_preimage by fastforce qed subsection\Quotient maps\ definition quotient_map where "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (openin X {x. x \ topspace X \ f x \ U} \ openin X' U))" lemma quotient_map_eq: assumes "quotient_map X X' f" "\x. x \ topspace X \ f x = g x" shows "quotient_map X X' g" by (smt (verit) Collect_cong assms image_cong quotient_map_def) lemma quotient_map_compose: assumes f: "quotient_map X X' f" and g: "quotient_map X' X'' g" shows "quotient_map X X'' (g \ f)" unfolding quotient_map_def proof (intro conjI allI impI) show "(g \ f) ` topspace X = topspace X''" using assms by (simp only: image_comp [symmetric]) (simp add: quotient_map_def) next fix U'' assume U'': "U'' \ topspace X''" define U' where "U' \ {y \ topspace X'. g y \ U''}" have "U' \ topspace X'" by (auto simp add: U'_def) then have U': "openin X {x \ topspace X. f x \ U'} = openin X' U'" using assms unfolding quotient_map_def by simp have "{x \ topspace X. f x \ topspace X' \ g (f x) \ U''} = {x \ topspace X. (g \ f) x \ U''}" using f quotient_map_def by fastforce then show "openin X {x \ topspace X. (g \ f) x \ U''} = openin X'' U''" by (smt (verit, best) Collect_cong U' U'_def U'' g mem_Collect_eq quotient_map_def) qed lemma quotient_map_from_composition: assumes f: "continuous_map X X' f" and g: "continuous_map X' X'' g" and gf: "quotient_map X X'' (g \ f)" shows "quotient_map X' X'' g" unfolding quotient_map_def proof (intro conjI allI impI) show "g ` topspace X' = topspace X''" using assms unfolding continuous_map_def quotient_map_def by fastforce next fix U'' :: "'c set" assume U'': "U'' \ topspace X''" have eq: "{x \ topspace X. g (f x) \ U''} = {x \ topspace X. f x \ {y. y \ topspace X' \ g y \ U''}}" using continuous_map_def f by fastforce show "openin X' {x \ topspace X'. g x \ U''} = openin X'' U''" using assms unfolding continuous_map_def quotient_map_def by (metis (mono_tags, lifting) Collect_cong U'' comp_apply eq) qed lemma quotient_imp_continuous_map: "quotient_map X X' f \ continuous_map X X' f" by (simp add: continuous_map openin_subset quotient_map_def) lemma quotient_imp_surjective_map: "quotient_map X X' f \ f ` (topspace X) = topspace X'" by (simp add: quotient_map_def) lemma quotient_map_closedin: "quotient_map X X' f \ f ` (topspace X) = topspace X' \ (\U. U \ topspace X' \ (closedin X {x. x \ topspace X \ f x \ U} \ closedin X' U))" proof - have eq: "(topspace X - {x \ topspace X. f x \ U'}) = {x \ topspace X. f x \ topspace X' \ f x \ U'}" if "f ` topspace X = topspace X'" "U' \ topspace X'" for U' using that by auto have "(\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U) = (\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U)" if "f ` topspace X = topspace X'" proof (rule iffI; intro allI impI subsetI) fix U' assume *[rule_format]: "\U\topspace X'. openin X {x \ topspace X. f x \ U} = openin X' U" and U': "U' \ topspace X'" show "closedin X {x \ topspace X. f x \ U'} = closedin X' U'" using U' by (auto simp add: closedin_def simp flip: * [of "topspace X' - U'"] eq [OF that]) next fix U' :: "'b set" assume *[rule_format]: "\U\topspace X'. closedin X {x \ topspace X. f x \ U} = closedin X' U" and U': "U' \ topspace X'" show "openin X {x \ topspace X. f x \ U'} = openin X' U'" using U' by (auto simp add: openin_closedin_eq simp flip: * [of "topspace X' - U'"] eq [OF that]) qed then show ?thesis unfolding quotient_map_def by force qed lemma continuous_open_imp_quotient_map: assumes "continuous_map X X' f" and om: "open_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - { fix U assume U: "U \ topspace X'" and "openin X {x \ topspace X. f x \ U}" then have ope: "openin X' (f ` {x \ topspace X. f x \ U})" using om unfolding open_map_def by blast then have "openin X' U" using U feq by (subst openin_subopen) force } moreover have "openin X {x \ topspace X. f x \ U}" if "U \ topspace X'" and "openin X' U" for U using that assms unfolding continuous_map_def by blast ultimately show ?thesis unfolding quotient_map_def using assms by blast qed lemma continuous_closed_imp_quotient_map: assumes "continuous_map X X' f" and om: "closed_map X X' f" and feq: "f ` (topspace X) = topspace X'" shows "quotient_map X X' f" proof - have "f ` {x \ topspace X. f x \ U} = U" if "U \ topspace X'" for U using that feq by auto with assms show ?thesis unfolding quotient_map_closedin closed_map_def continuous_map_closedin by auto qed lemma continuous_open_quotient_map: "\continuous_map X X' f; open_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_open_imp_quotient_map quotient_map_def) lemma continuous_closed_quotient_map: "\continuous_map X X' f; closed_map X X' f\ \ quotient_map X X' f \ f ` (topspace X) = topspace X'" by (meson continuous_closed_imp_quotient_map quotient_map_def) lemma injective_quotient_map: assumes "inj_on f (topspace X)" shows "quotient_map X X' f \ continuous_map X X' f \ open_map X X' f \ closed_map X X' f \ f ` (topspace X) = topspace X'" (is "?lhs = ?rhs") proof assume L: ?lhs have om: "open_map X X' f" proof (clarsimp simp add: open_map_def) fix U assume "openin X U" then have "U \ topspace X" by (simp add: openin_subset) moreover have "{x \ topspace X. f x \ f ` U} = U" using \U \ topspace X\ assms inj_onD by fastforce ultimately show "openin X' (f ` U)" using L unfolding quotient_map_def by (metis (no_types, lifting) Collect_cong \openin X U\ image_mono) qed then have "closed_map X X' f" by (simp add: L assms bijective_open_imp_closed_map quotient_imp_surjective_map) then show ?rhs using L om by (simp add: quotient_imp_continuous_map quotient_imp_surjective_map) next assume ?rhs then show ?lhs by (simp add: continuous_closed_imp_quotient_map) qed lemma continuous_compose_quotient_map: assumes f: "quotient_map X X' f" and g: "continuous_map X X'' (g \ f)" shows "continuous_map X' X'' g" unfolding quotient_map_def continuous_map_def proof (intro conjI ballI allI impI) show "\x'. x' \ topspace X' \ g x' \ topspace X''" using assms unfolding quotient_map_def by (metis (no_types, opaque_lifting) continuous_map_image_subset_topspace image_comp image_subset_iff) next fix U'' :: "'c set" assume U'': "openin X'' U''" have "f ` topspace X = topspace X'" by (simp add: f quotient_imp_surjective_map) then have eq: "{x \ topspace X. f x \ topspace X' \ g (f x) \ U} = {x \ topspace X. g (f x) \ U}" for U by auto have "openin X {x \ topspace X. f x \ topspace X' \ g (f x) \ U''}" unfolding eq using U'' g openin_continuous_map_preimage by fastforce then have *: "openin X {x \ topspace X. f x \ {x \ topspace X'. g x \ U''}}" by auto show "openin X' {x \ topspace X'. g x \ U''}" using f unfolding quotient_map_def by (metis (no_types) Collect_subset *) qed lemma continuous_compose_quotient_map_eq: "quotient_map X X' f \ continuous_map X X'' (g \ f) \ continuous_map X' X'' g" using continuous_compose_quotient_map continuous_map_compose quotient_imp_continuous_map by blast lemma quotient_map_compose_eq: "quotient_map X X' f \ quotient_map X X'' (g \ f) \ quotient_map X' X'' g" by (meson continuous_compose_quotient_map_eq quotient_imp_continuous_map quotient_map_compose quotient_map_from_composition) lemma quotient_map_restriction: assumes quo: "quotient_map X Y f" and U: "{x \ topspace X. f x \ V} = U" and disj: "openin Y V \ closedin Y V" shows "quotient_map (subtopology X U) (subtopology Y V) f" using disj proof assume V: "openin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: openin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" using quo unfolding quotient_map_def by auto have "openin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_def proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "openin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = openin (subtopology Y V) Y'" using U V Y \openin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: openin_open_subtopology eq) (auto simp: openin_closedin_eq) qed next assume V: "closedin Y V" with U have sub: "U \ topspace X" "V \ topspace Y" by (auto simp: closedin_subset) have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin X {x \ topspace X. f x \ U} = closedin Y U" using quo unfolding quotient_map_closedin by auto have "closedin X U" using U V Y sub(2) by blast show ?thesis unfolding quotient_map_closedin proof (intro conjI allI impI) show "f ` topspace (subtopology X U) = topspace (subtopology Y V)" using sub U fim by (auto) next fix Y' :: "'b set" assume "Y' \ topspace (subtopology Y V)" then have "Y' \ topspace Y" "Y' \ V" by (simp_all) then have eq: "{x \ topspace X. x \ U \ f x \ Y'} = {x \ topspace X. f x \ Y'}" using U by blast then show "closedin (subtopology X U) {x \ topspace (subtopology X U). f x \ Y'} = closedin (subtopology Y V) Y'" using U V Y \closedin X U\ \Y' \ topspace Y\ \Y' \ V\ by (simp add: closedin_closed_subtopology eq) (auto simp: closedin_def) qed qed lemma quotient_map_saturated_open: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. openin X U \ {x \ topspace X. f x \ f ` U} \ U \ openin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then have fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ openin Y U = openin X {x \ topspace X. f x \ U}" unfolding quotient_map_def by auto show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "openin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim openin_subset by fastforce+ show "openin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then have YX: "\U. openin Y U \ openin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \openin X U; {x \ topspace X. f x \ f ` U} \ U\ \ openin Y (f ` U)" by (auto simp: quotient_map_def continuous_map_def) show ?lhs proof (simp add: quotient_map_def fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "openin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "openin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "openin Y U" show "openin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed lemma quotient_map_saturated_closed: "quotient_map X Y f \ continuous_map X Y f \ f ` (topspace X) = topspace Y \ (\U. closedin X U \ {x \ topspace X. f x \ f ` U} \ U \ closedin Y (f ` U))" (is "?lhs = ?rhs") proof assume L: ?lhs then obtain fim: "f ` topspace X = topspace Y" and Y: "\U. U \ topspace Y \ closedin Y U = closedin X {x \ topspace X. f x \ U}" by (simp add: quotient_map_closedin) show ?rhs proof (intro conjI allI impI) show "continuous_map X Y f" by (simp add: L quotient_imp_continuous_map) show "f ` topspace X = topspace Y" by (simp add: fim) next fix U :: "'a set" assume U: "closedin X U \ {x \ topspace X. f x \ f ` U} \ U" then have sub: "f ` U \ topspace Y" and eq: "{x \ topspace X. f x \ f ` U} = U" using fim closedin_subset by fastforce+ show "closedin Y (f ` U)" by (simp add: sub Y eq U) qed next assume ?rhs then obtain YX: "\U. closedin Y U \ closedin X {x \ topspace X. f x \ U}" and fim: "f ` topspace X = topspace Y" and XY: "\U. \closedin X U; {x \ topspace X. f x \ f ` U} \ U\ \ closedin Y (f ` U)" by (simp add: continuous_map_closedin) show ?lhs proof (simp add: quotient_map_closedin fim, intro allI impI iffI) fix U :: "'b set" assume "U \ topspace Y" and X: "closedin X {x \ topspace X. f x \ U}" have feq: "f ` {x \ topspace X. f x \ U} = U" using \U \ topspace Y\ fim by auto show "closedin Y U" using XY [OF X] by (simp add: feq) next fix U :: "'b set" assume "U \ topspace Y" and Y: "closedin Y U" show "closedin X {x \ topspace X. f x \ U}" by (metis YX [OF Y]) qed qed lemma quotient_map_onto_image: assumes "f ` topspace X \ topspace Y" and U: "\U. U \ topspace Y \ openin X {x \ topspace X. f x \ U} = openin Y U" shows "quotient_map X (subtopology Y (f ` topspace X)) f" unfolding quotient_map_def topspace_subtopology proof (intro conjI strip) fix U assume "U \ topspace Y \ f ` topspace X" with U have "openin X {x \ topspace X. f x \ U} \ \x. openin Y x \ U = f ` topspace X \ x" by fastforce moreover have "\x. openin Y x \ U = f ` topspace X \ x \ openin X {x \ topspace X. f x \ U}" by (metis (mono_tags, lifting) Collect_cong IntE IntI U image_eqI openin_subset) ultimately show "openin X {x \ topspace X. f x \ U} = openin (subtopology Y (f ` topspace X)) U" by (force simp: openin_subtopology_alt image_iff) qed (use assms in auto) lemma quotient_map_lift_exists: assumes f: "quotient_map X Y f" and h: "continuous_map X Z h" and "\x y. \x \ topspace X; y \ topspace X; f x = f y\ \ h x = h y" obtains g where "continuous_map Y Z g" "g ` topspace Y = h ` topspace X" "\x. x \ topspace X \ g(f x) = h x" proof - obtain g where g: "\x. x \ topspace X \ h x = g(f x)" using function_factors_left_gen[of "\x. x \ topspace X" f h] assms by blast show ?thesis proof show "g ` topspace Y = h ` topspace X" using f g by (force dest!: quotient_imp_surjective_map) show "continuous_map Y Z g" by (smt (verit) f g h continuous_compose_quotient_map_eq continuous_map_eq o_def) qed (simp add: g) qed subsection\ Separated Sets\ definition separatedin :: "'a topology \ 'a set \ 'a set \ bool" where "separatedin X S T \ S \ topspace X \ T \ topspace X \ S \ X closure_of T = {} \ T \ X closure_of S = {}" lemma separatedin_empty [simp]: "separatedin X S {} \ S \ topspace X" "separatedin X {} S \ S \ topspace X" by (simp_all add: separatedin_def) lemma separatedin_refl [simp]: "separatedin X S S \ S = {}" by (metis closure_of_subset empty_subsetI inf.orderE separatedin_def) lemma separatedin_sym: "separatedin X S T \ separatedin X T S" by (auto simp: separatedin_def) lemma separatedin_imp_disjoint: "separatedin X S T \ disjnt S T" by (meson closure_of_subset disjnt_def disjnt_subset2 separatedin_def) lemma separatedin_mono: "\separatedin X S T; S' \ S; T' \ T\ \ separatedin X S' T'" unfolding separatedin_def using closure_of_mono by blast lemma separatedin_open_sets: "\openin X S; openin X T\ \ separatedin X S T \ disjnt S T" unfolding disjnt_def separatedin_def by (auto simp: openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closed_sets: "\closedin X S; closedin X T\ \ separatedin X S T \ disjnt S T" unfolding closure_of_eq disjnt_def separatedin_def by (metis closedin_def closure_of_eq inf_commute) lemma separatedin_subtopology: "separatedin (subtopology X U) S T \ S \ U \ T \ U \ separatedin X S T" by (auto simp: separatedin_def closure_of_subtopology Int_ac disjoint_iff elim!: inf.orderE) lemma separatedin_discrete_topology: "separatedin (discrete_topology U) S T \ S \ U \ T \ U \ disjnt S T" by (metis openin_discrete_topology separatedin_def separatedin_open_sets topspace_discrete_topology) lemma separated_eq_distinguishable: "separatedin X {x} {y} \ x \ topspace X \ y \ topspace X \ (\U. openin X U \ x \ U \ (y \ U)) \ (\v. openin X v \ y \ v \ (x \ v))" by (force simp: separatedin_def closure_of_def) lemma separatedin_Un [simp]: "separatedin X S (T \ U) \ separatedin X S T \ separatedin X S U" "separatedin X (S \ T) U \ separatedin X S U \ separatedin X T U" by (auto simp: separatedin_def) lemma separatedin_Union: "finite \ \ separatedin X S (\\) \ S \ topspace X \ (\T \ \. separatedin X S T)" "finite \ \ separatedin X (\\) S \ (\T \ \. separatedin X S T) \ S \ topspace X" by (auto simp: separatedin_def closure_of_Union) lemma separatedin_openin_diff: "\openin X S; openin X T\ \ separatedin X (S - T) (T - S)" unfolding separatedin_def by (metis Diff_Int_distrib2 Diff_disjoint Diff_empty Diff_mono empty_Diff empty_subsetI openin_Int_closure_of_eq_empty openin_subset) lemma separatedin_closedin_diff: assumes "closedin X S" "closedin X T" shows "separatedin X (S - T) (T - S)" proof - have "S - T \ topspace X" "T - S \ topspace X" using assms closedin_subset by auto with assms show ?thesis by (simp add: separatedin_def Diff_Int_distrib2 closure_of_minimal inf_absorb2) qed lemma separation_closedin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ closedin (subtopology X (S \ T)) S \ closedin (subtopology X (S \ T)) T" by (auto simp add: separatedin_def closedin_Int_closure_of disjnt_iff dest: closure_of_subset) lemma separation_openin_Un_gen: "separatedin X S T \ S \ topspace X \ T \ topspace X \ disjnt S T \ openin (subtopology X (S \ T)) S \ openin (subtopology X (S \ T)) T" unfolding openin_closedin_eq topspace_subtopology separation_closedin_Un_gen disjnt_def by (auto simp: Diff_triv Int_commute Un_Diff inf_absorb1 topspace_def) lemma separatedin_full: "S \ T = topspace X \ separatedin X S T \ disjnt S T \ closedin X S \ openin X S \ closedin X T \ openin X T" by (metis separatedin_open_sets separation_closedin_Un_gen separation_openin_Un_gen subtopology_topspace) subsection\Homeomorphisms\ text\(1-way and 2-way versions may be useful in places)\ definition homeomorphic_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "homeomorphic_map X Y f \ quotient_map X Y f \ inj_on f (topspace X)" definition homeomorphic_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "homeomorphic_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" lemma homeomorphic_map_eq: "\homeomorphic_map X Y f; \x. x \ topspace X \ f x = g x\ \ homeomorphic_map X Y g" by (meson homeomorphic_map_def inj_on_cong quotient_map_eq) lemma homeomorphic_maps_eq: "\homeomorphic_maps X Y f g; \x. x \ topspace X \ f x = f' x; \y. y \ topspace Y \ g y = g' y\ \ homeomorphic_maps X Y f' g'" unfolding homeomorphic_maps_def by (metis continuous_map_eq continuous_map_eq_image_closure_subset_gen image_subset_iff) lemma homeomorphic_maps_sym: "homeomorphic_maps X Y f g \ homeomorphic_maps Y X g f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_maps_id: "homeomorphic_maps X Y id id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have "topspace X = topspace Y" by (auto simp: homeomorphic_maps_def continuous_map_def) with L show ?rhs unfolding homeomorphic_maps_def by (metis topology_finer_continuous_id topology_eq) next assume ?rhs then show ?lhs unfolding homeomorphic_maps_def by auto qed lemma homeomorphic_map_id [simp]: "homeomorphic_map X Y id \ Y = X" (is "?lhs = ?rhs") proof assume L: ?lhs then have eq: "topspace X = topspace Y" by (auto simp: homeomorphic_map_def continuous_map_def quotient_map_def) then have "\S. openin X S \ openin Y S" by (meson L homeomorphic_map_def injective_quotient_map topology_finer_open_id) then show ?rhs using L unfolding homeomorphic_map_def by (metis eq quotient_imp_continuous_map topology_eq topology_finer_continuous_id) next assume ?rhs then show ?lhs unfolding homeomorphic_map_def by (simp add: closed_map_id continuous_closed_imp_quotient_map) qed lemma homeomorphic_map_compose: assumes "homeomorphic_map X Y f" "homeomorphic_map Y X'' g" shows "homeomorphic_map X X'' (g \ f)" proof - have "inj_on g (f ` topspace X)" by (metis (no_types) assms homeomorphic_map_def quotient_imp_surjective_map) then show ?thesis using assms by (meson comp_inj_on homeomorphic_map_def quotient_map_compose_eq) qed lemma homeomorphic_maps_compose: "homeomorphic_maps X Y f h \ homeomorphic_maps Y X'' g k \ homeomorphic_maps X X'' (g \ f) (h \ k)" unfolding homeomorphic_maps_def by (auto simp: continuous_map_compose; simp add: continuous_map_def) lemma homeomorphic_eq_everything_map: "homeomorphic_map X Y f \ continuous_map X Y f \ open_map X Y f \ closed_map X Y f \ f ` (topspace X) = topspace Y \ inj_on f (topspace X)" unfolding homeomorphic_map_def by (force simp: injective_quotient_map intro: injective_quotient_map) lemma homeomorphic_imp_continuous_map: "homeomorphic_map X Y f \ continuous_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_open_map: "homeomorphic_map X Y f \ open_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_closed_map: "homeomorphic_map X Y f \ closed_map X Y f" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_surjective_map: "homeomorphic_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: homeomorphic_eq_everything_map) lemma homeomorphic_imp_injective_map: "homeomorphic_map X Y f \ inj_on f (topspace X)" by (simp add: homeomorphic_eq_everything_map) lemma bijective_open_imp_homeomorphic_map: "\continuous_map X Y f; open_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: homeomorphic_map_def continuous_open_imp_quotient_map) lemma bijective_closed_imp_homeomorphic_map: "\continuous_map X Y f; closed_map X Y f; f ` (topspace X) = topspace Y; inj_on f (topspace X)\ \ homeomorphic_map X Y f" by (simp add: continuous_closed_quotient_map homeomorphic_map_def) lemma open_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "open_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y open_map_def continuous_map_def eq) qed lemma closed_eq_continuous_inverse_map: assumes X: "\x. x \ topspace X \ f x \ topspace Y \ g(f x) = x" and Y: "\y. y \ topspace Y \ g y \ topspace X \ f(g y) = y" shows "closed_map X Y f \ continuous_map Y X g" proof - have eq: "{x \ topspace Y. g x \ U} = f ` U" if "closedin X U" for U using closedin_subset [OF that] by (force simp: X Y image_iff) show ?thesis by (auto simp: Y closed_map_def continuous_map_closedin eq) qed lemma homeomorphic_maps_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f \ homeomorphic_map Y X g \ (\x \ topspace X. g(f x) = x) \ (\y \ topspace Y. f(g y) = y)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "continuous_map Y X g" "\x\topspace X. g (f x) = x" "\x'\topspace Y. f (g x') = x'" by (auto simp: homeomorphic_maps_def) show ?rhs proof (intro conjI bijective_open_imp_homeomorphic_map L) show "open_map X Y f" using L using open_eq_continuous_inverse_map [of concl: X Y f g] by (simp add: continuous_map_def) show "open_map Y X g" using L using open_eq_continuous_inverse_map [of concl: Y X g f] by (simp add: continuous_map_def) show "f ` topspace X = topspace Y" "g ` topspace Y = topspace X" using L by (force simp: continuous_map_closedin)+ show "inj_on f (topspace X)" "inj_on g (topspace Y)" using L unfolding inj_on_def by metis+ qed next assume ?rhs then show ?lhs by (auto simp: homeomorphic_maps_def homeomorphic_imp_continuous_map) qed lemma homeomorphic_maps_imp_map: "homeomorphic_maps X Y f g \ homeomorphic_map X Y f" using homeomorphic_maps_map by blast lemma homeomorphic_map_maps: "homeomorphic_map X Y f \ (\g. homeomorphic_maps X Y f g)" (is "?lhs = ?rhs") proof assume ?lhs then have L: "continuous_map X Y f" "open_map X Y f" "closed_map X Y f" "f ` (topspace X) = topspace Y" "inj_on f (topspace X)" by (auto simp: homeomorphic_eq_everything_map) have X: "\x. x \ topspace X \ f x \ topspace Y \ inv_into (topspace X) f (f x) = x" using L by auto have Y: "\y. y \ topspace Y \ inv_into (topspace X) f y \ topspace X \ f (inv_into (topspace X) f y) = y" by (simp add: L f_inv_into_f inv_into_into) have "homeomorphic_maps X Y f (inv_into (topspace X) f)" unfolding homeomorphic_maps_def proof (intro conjI L) show "continuous_map Y X (inv_into (topspace X) f)" by (simp add: L X Y flip: open_eq_continuous_inverse_map [where f=f]) next show "\x\topspace X. inv_into (topspace X) f (f x) = x" "\y\topspace Y. f (inv_into (topspace X) f y) = y" using X Y by auto qed then show ?rhs by metis next assume ?rhs then show ?lhs using homeomorphic_maps_map by blast qed lemma homeomorphic_maps_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_maps X X f f" by (auto simp: homeomorphic_maps_def) lemma homeomorphic_map_involution: "\continuous_map X X f; \x. x \ topspace X \ f(f x) = x\ \ homeomorphic_map X X f" using homeomorphic_maps_involution homeomorphic_maps_map by blast lemma homeomorphic_map_openness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "openin Y (f ` U) \ openin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "openin X U \ openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast show "openin Y (f ` U) = openin X U" proof assume L: "openin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "openin X U" by (metis L homeomorphic_imp_open_map open_map_def g) next assume "openin X U" then show "openin Y (f ` U)" using hom homeomorphic_imp_open_map open_map_def by blast qed qed lemma homeomorphic_map_closedness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "closedin Y (f ` U) \ closedin X U" proof - obtain g where "homeomorphic_maps X Y f g" using assms by (auto simp: homeomorphic_map_maps) then have g: "homeomorphic_map Y X g" and gf: "\x. x \ topspace X \ g(f x) = x" by (auto simp: homeomorphic_maps_map) then have "closedin X U \ closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast show "closedin Y (f ` U) = closedin X U" proof assume L: "closedin Y (f ` U)" have "U = g ` (f ` U)" using U gf by force then show "closedin X U" by (metis L homeomorphic_imp_closed_map closed_map_def g) next assume "closedin X U" then show "closedin Y (f ` U)" using hom homeomorphic_imp_closed_map closed_map_def by blast qed qed lemma homeomorphic_map_openness_eq: "homeomorphic_map X Y f \ openin X U \ U \ topspace X \ openin Y (f ` U)" by (meson homeomorphic_map_openness openin_closedin_eq) lemma homeomorphic_map_closedness_eq: "homeomorphic_map X Y f \ closedin X U \ U \ topspace X \ closedin Y (f ` U)" by (meson closedin_subset homeomorphic_map_closedness) lemma all_openin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. openin Y V \ P V) \ (\U. openin X U \ P(f ` U))" by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_openness openin_subset subset_image_iff) lemma all_closedin_homeomorphic_image: assumes "homeomorphic_map X Y f" shows "(\V. closedin Y V \ P V) \ (\U. closedin X U \ P(f ` U))" (is "?lhs = ?rhs") by (metis (no_types, lifting) assms homeomorphic_imp_surjective_map homeomorphic_map_closedness closedin_subset subset_image_iff) lemma homeomorphic_map_derived_set_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y derived_set_of (f ` S) = f ` (X derived_set_of S)" proof - have fim: "f ` (topspace X) = topspace Y" and inj: "inj_on f (topspace X)" using hom by (auto simp: homeomorphic_eq_everything_map) have iff: "(\T. x \ T \ openin X T \ (\y. y \ x \ y \ S \ y \ T)) = (\T. T \ topspace Y \ f x \ T \ openin Y T \ (\y. y \ f x \ y \ f ` S \ y \ T))" if "x \ topspace X" for x proof - have \
: "(x \ T \ openin X T) = (T \ topspace X \ f x \ f ` T \ openin Y (f ` T))" for T by (meson hom homeomorphic_map_openness_eq inj inj_on_image_mem_iff that) moreover have "(\y. y \ x \ y \ S \ y \ T) = (\y. y \ f x \ y \ f ` S \ y \ f ` T)" (is "?lhs = ?rhs") if "T \ topspace X \ f x \ f ` T \ openin Y (f ` T)" for T by (smt (verit, del_insts) S \x \ topspace X\ image_iff inj inj_on_def subsetD that) ultimately show ?thesis by (auto simp flip: fim simp: all_subset_image) qed have *: "\T = f ` S; \x. x \ S \ P x \ Q(f x)\ \ {y. y \ T \ Q y} = f ` {x \ S. P x}" for T S P Q by auto show ?thesis unfolding derived_set_of_def by (rule *) (use fim iff openin_subset in force)+ qed lemma homeomorphic_map_closure_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y closure_of (f ` S) = f ` (X closure_of S)" unfolding closure_of using homeomorphic_imp_surjective_map [OF hom] S by (auto simp: in_derived_set_of homeomorphic_map_derived_set_of [OF assms]) lemma homeomorphic_map_interior_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y interior_of (f ` S) = f ` (X interior_of S)" proof - { fix y assume "y \ topspace Y" and "y \ Y closure_of (topspace Y - f ` S)" then have "y \ f ` (topspace X - X closure_of (topspace X - S))" using homeomorphic_eq_everything_map [THEN iffD1, OF hom] homeomorphic_map_closure_of [OF hom] by (metis DiffI Diff_subset S closure_of_subset_topspace inj_on_image_set_diff) } moreover { fix x assume "x \ topspace X" then have "f x \ topspace Y" using hom homeomorphic_imp_surjective_map by blast } moreover { fix x assume "x \ topspace X" and "x \ X closure_of (topspace X - S)" and "f x \ Y closure_of (topspace Y - f ` S)" then have "False" using homeomorphic_map_closure_of [OF hom] hom unfolding homeomorphic_eq_everything_map by (metis Diff_subset S closure_of_subset_topspace inj_on_image_mem_iff inj_on_image_set_diff) } ultimately show ?thesis by (auto simp: interior_of_closure_of) qed lemma homeomorphic_map_frontier_of: assumes hom: "homeomorphic_map X Y f" and S: "S \ topspace X" shows "Y frontier_of (f ` S) = f ` (X frontier_of S)" unfolding frontier_of_def proof (intro equalityI subsetI DiffI) fix y assume "y \ Y closure_of f ` S - Y interior_of f ` S" then show "y \ f ` (X closure_of S - X interior_of S)" using S hom homeomorphic_map_closure_of homeomorphic_map_interior_of by fastforce next fix y assume "y \ f ` (X closure_of S - X interior_of S)" then show "y \ Y closure_of f ` S" using S hom homeomorphic_map_closure_of by fastforce next fix x assume "x \ f ` (X closure_of S - X interior_of S)" then obtain y where y: "x = f y" "y \ X closure_of S" "y \ X interior_of S" by blast then show "x \ Y interior_of f ` S" using S hom homeomorphic_map_interior_of y(1) unfolding homeomorphic_map_def by (smt (verit, ccfv_SIG) in_closure_of inj_on_image_mem_iff interior_of_subset_topspace) qed lemma homeomorphic_maps_subtopologies: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_maps_subtopologies_alt: "\homeomorphic_maps X Y f g; f ` (topspace X \ S) \ T; g ` (topspace Y \ T) \ S\ \ homeomorphic_maps (subtopology X S) (subtopology Y T) f g" unfolding homeomorphic_maps_def by (force simp: continuous_map_from_subtopology continuous_map_in_subtopology) lemma homeomorphic_map_subtopologies: "\homeomorphic_map X Y f; f ` (topspace X \ S) = topspace Y \ T\ \ homeomorphic_map (subtopology X S) (subtopology Y T) f" by (meson homeomorphic_map_maps homeomorphic_maps_subtopologies) lemma homeomorphic_map_subtopologies_alt: assumes hom: "homeomorphic_map X Y f" and S: "\x. \x \ topspace X; f x \ topspace Y\ \ f x \ T \ x \ S" shows "homeomorphic_map (subtopology X S) (subtopology Y T) f" proof - have "homeomorphic_maps (subtopology X S) (subtopology Y T) f g" if "homeomorphic_maps X Y f g" for g proof (rule homeomorphic_maps_subtopologies [OF that]) have "f ` (topspace X \ S) \ topspace Y \ T" using S hom homeomorphic_imp_surjective_map by fastforce then show "f ` (topspace X \ S) = topspace Y \ T" using that unfolding homeomorphic_maps_def continuous_map_def by (smt (verit, del_insts) Int_iff S image_iff subsetI subset_antisym) qed then show ?thesis using hom by (meson homeomorphic_map_maps) qed subsection\Relation of homeomorphism between topological spaces\ definition homeomorphic_space (infixr "homeomorphic'_space" 50) where "X homeomorphic_space Y \ \f g. homeomorphic_maps X Y f g" lemma homeomorphic_space_refl: "X homeomorphic_space X" by (meson homeomorphic_maps_id homeomorphic_space_def) lemma homeomorphic_space_sym: "X homeomorphic_space Y \ Y homeomorphic_space X" unfolding homeomorphic_space_def by (metis homeomorphic_maps_sym) lemma homeomorphic_space_trans [trans]: "\X1 homeomorphic_space X2; X2 homeomorphic_space X3\ \ X1 homeomorphic_space X3" unfolding homeomorphic_space_def by (metis homeomorphic_maps_compose) lemma homeomorphic_space: "X homeomorphic_space Y \ (\f. homeomorphic_map X Y f)" by (simp add: homeomorphic_map_maps homeomorphic_space_def) lemma homeomorphic_maps_imp_homeomorphic_space: "homeomorphic_maps X Y f g \ X homeomorphic_space Y" unfolding homeomorphic_space_def by metis lemma homeomorphic_map_imp_homeomorphic_space: "homeomorphic_map X Y f \ X homeomorphic_space Y" unfolding homeomorphic_map_maps using homeomorphic_space_def by blast lemma homeomorphic_empty_space: "X homeomorphic_space Y \ topspace X = {} \ topspace Y = {}" by (metis homeomorphic_imp_surjective_map homeomorphic_space image_is_empty) lemma homeomorphic_empty_space_eq: assumes "topspace X = {}" shows "X homeomorphic_space Y \ topspace Y = {}" unfolding homeomorphic_maps_def homeomorphic_space_def by (metis assms continuous_map_on_empty continuous_map_closedin ex_in_conv) subsection\Connected topological spaces\ definition connected_space :: "'a topology \ bool" where "connected_space X \ \(\E1 E2. openin X E1 \ openin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" definition connectedin :: "'a topology \ 'a set \ bool" where "connectedin X S \ S \ topspace X \ connected_space (subtopology X S)" lemma connected_spaceD: "\connected_space X; openin X U; openin X V; topspace X \ U \ V; U \ V = {}; U \ {}; V \ {}\ \ False" by (auto simp: connected_space_def) lemma connectedin_subset_topspace: "connectedin X S \ S \ topspace X" by (simp add: connectedin_def) lemma connectedin_topspace: "connectedin X (topspace X) \ connected_space X" by (simp add: connectedin_def) lemma connected_space_subtopology: "connectedin X S \ connected_space (subtopology X S)" by (simp add: connectedin_def) lemma connectedin_subtopology: "connectedin (subtopology X S) T \ connectedin X T \ T \ S" by (force simp: connectedin_def subtopology_subtopology inf_absorb2) lemma connected_space_eq: "connected_space X \ (\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" unfolding connected_space_def by (metis openin_Un openin_subset subset_antisym) lemma connected_space_closedin: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ topspace X \ E1 \ E2 \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then have "\E1 E2. \openin X E1; E1 \ E2 = {}; topspace X \ E1 \ E2; openin X E2\ \ E1 = {} \ E2 = {}" by (simp add: connected_space_def) then show ?rhs unfolding connected_space_def by (metis disjnt_def separatedin_closed_sets separation_openin_Un_gen subtopology_superset) next assume R: ?rhs then show ?lhs unfolding connected_space_def by (metis Diff_triv Int_commute separatedin_openin_diff separation_closedin_Un_gen subtopology_superset) qed lemma connected_space_closedin_eq: "connected_space X \ (\E1 E2. closedin X E1 \ closedin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ E1 \ {} \ E2 \ {})" by (metis closedin_Un closedin_def connected_space_closedin subset_antisym) lemma connected_space_clopen_in: "connected_space X \ (\T. openin X T \ closedin X T \ T = {} \ T = topspace X)" proof - have eq: "openin X E1 \ openin X E2 \ E1 \ E2 = topspace X \ E1 \ E2 = {} \ P \ E2 = topspace X - E1 \ openin X E1 \ openin X E2 \ P" for E1 E2 P using openin_subset by blast show ?thesis unfolding connected_space_eq eq closedin_def by (auto simp: openin_closedin_eq) qed lemma connectedin: "connectedin X S \ S \ topspace X \ (\E1 E2. openin X E1 \ openin X E2 \ S \ E1 \ E2 \ E1 \ E2 \ S = {} \ E1 \ S \ {} \ E2 \ S \ {})" (is "?lhs = ?rhs") proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_def openin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedinD: "\connectedin X S; openin X E1; openin X E2; S \ E1 \ E2; E1 \ E2 \ S = {}; E1 \ S \ {}; E2 \ S \ {}\ \ False" by (meson connectedin) lemma connectedin_iff_connected [simp]: "connectedin euclidean S \ connected S" by (simp add: connected_def connectedin) lemma connectedin_closedin: "connectedin X S \ S \ topspace X \ \(\E1 E2. closedin X E1 \ closedin X E2 \ S \ (E1 \ E2) \ (E1 \ E2 \ S = {}) \ \(E1 \ S = {}) \ \(E2 \ S = {}))" proof - have *: "(\E1:: 'a set. \E2:: 'a set. (\T1:: 'a set. P1 T1 \ E1 = f1 T1) \ (\T2:: 'a set. P2 T2 \ E2 = f2 T2) \ R E1 E2) \ (\T1 T2. P1 T1 \ P2 T2 \ R(f1 T1) (f2 T2))" for P1 f1 P2 f2 R by auto show ?thesis unfolding connectedin_def connected_space_closedin closedin_subtopology topspace_subtopology * by (intro conj_cong arg_cong [where f=Not] ex_cong1; blast dest!: openin_subset) qed lemma connectedin_empty [simp]: "connectedin X {}" by (simp add: connectedin) lemma connected_space_topspace_empty: "topspace X = {} \ connected_space X" using connectedin_topspace by fastforce lemma connectedin_sing [simp]: "connectedin X {a} \ a \ topspace X" by (simp add: connectedin) lemma connectedin_absolute [simp]: "connectedin (subtopology X S) S \ connectedin X S" by (simp add: connectedin_subtopology) lemma connectedin_Union: assumes \: "\S. S \ \ \ connectedin X S" and ne: "\\ \ {}" shows "connectedin X (\\)" proof - have "\\ \ topspace X" using \ by (simp add: Union_least connectedin_def) moreover have False if "openin X E1" "openin X E2" and cover: "\\ \ E1 \ E2" and disj: "E1 \ E2 \ \\ = {}" and overlap1: "E1 \ \\ \ {}" and overlap2: "E2 \ \\ \ {}" for E1 E2 proof - have disjS: "E1 \ E2 \ S = {}" if "S \ \" for S using Diff_triv that disj by auto have coverS: "S \ E1 \ E2" if "S \ \" for S using that cover by blast have "\ \ {}" using overlap1 by blast obtain a where a: "\U. U \ \ \ a \ U" using ne by force with \\ \ {}\ have "a \ \\" by blast then consider "a \ E1" | "a \ E2" using \\\ \ E1 \ E2\ by auto then show False proof cases case 1 then obtain b S where "b \ E2" "b \ S" "S \ \" using overlap2 by blast then show ?thesis using "1" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) next case 2 then obtain b S where "b \ E1" "b \ S" "S \ \" using overlap1 by blast then show ?thesis using "2" \openin X E1\ \openin X E2\ disjS coverS a [OF \S \ \\] \[OF \S \ \\] unfolding connectedin by (meson disjoint_iff_not_equal) qed qed ultimately show ?thesis unfolding connectedin by blast qed lemma connectedin_Un: "\connectedin X S; connectedin X T; S \ T \ {}\ \ connectedin X (S \ T)" using connectedin_Union [of "{S,T}"] by auto lemma connected_space_subconnected: "connected_space X \ (\x \ topspace X. \y \ topspace X. \S. connectedin X S \ x \ S \ y \ S)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connectedin_topspace by blast next assume R [rule_format]: ?rhs have False if "openin X U" "openin X V" and disj: "U \ V = {}" and cover: "topspace X \ U \ V" and "U \ {}" "V \ {}" for U V proof - obtain u v where "u \ U" "v \ V" using \U \ {}\ \V \ {}\ by auto then obtain T where "u \ T" "v \ T" and T: "connectedin X T" using R [of u v] that by (meson \openin X U\ \openin X V\ subsetD openin_subset) then show False using that unfolding connectedin by (metis IntI \u \ U\ \v \ V\ empty_iff inf_bot_left subset_trans) qed then show ?lhs by (auto simp: connected_space_def) qed lemma connectedin_intermediate_closure_of: assumes "connectedin X S" "S \ T" "T \ X closure_of S" shows "connectedin X T" proof - have S: "S \ topspace X" and T: "T \ topspace X" using assms by (meson closure_of_subset_topspace dual_order.trans)+ have \
: "\E1 E2. \openin X E1; openin X E2; E1 \ S = {} \ E2 \ S = {}\ \ E1 \ T = {} \ E2 \ T = {}" using assms unfolding disjoint_iff by (meson in_closure_of subsetD) then show ?thesis using assms unfolding connectedin closure_of_subset_topspace S T by (metis Int_empty_right T dual_order.trans inf.orderE inf_left_commute) qed lemma connectedin_closure_of: "connectedin X S \ connectedin X (X closure_of S)" by (meson closure_of_subset connectedin_def connectedin_intermediate_closure_of subset_refl) lemma connectedin_separation: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ C1 \ X closure_of C2 = {} \ C2 \ X closure_of C1 = {})" unfolding connectedin_def connected_space_closedin_eq closedin_Int_closure_of topspace_subtopology apply (intro conj_cong refl arg_cong [where f=Not]) apply (intro ex_cong1 iffI, blast) using closure_of_subset_Int by force lemma connectedin_eq_not_separated: "connectedin X S \ S \ topspace X \ (\C1 C2. C1 \ C2 = S \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" unfolding separatedin_def by (metis connectedin_separation sup.boundedE) lemma connectedin_eq_not_separated_subset: "connectedin X S \ S \ topspace X \ (\C1 C2. S \ C1 \ C2 \ S \ C1 \ {} \ S \ C2 \ {} \ separatedin X C1 C2)" proof - have "\C1 C2. S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" if "\C1 C2. C1 \ C2 = S \ C1 = {} \ C2 = {} \ \ separatedin X C1 C2" proof (intro allI) fix C1 C2 show "S \ C1 \ C2 \ S \ C1 = {} \ S \ C2 = {} \ \ separatedin X C1 C2" using that [of "S \ C1" "S \ C2"] by (auto simp: separatedin_mono) qed then show ?thesis by (metis Un_Int_eq(1) Un_Int_eq(2) connectedin_eq_not_separated order_refl) qed lemma connected_space_eq_not_separated: "connected_space X \ (\C1 C2. C1 \ C2 = topspace X \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (simp add: connectedin_eq_not_separated flip: connectedin_topspace) lemma connected_space_eq_not_separated_subset: "connected_space X \ (\C1 C2. topspace X \ C1 \ C2 \ C1 \ {} \ C2 \ {} \ separatedin X C1 C2)" by (metis connected_space_eq_not_separated le_sup_iff separatedin_def subset_antisym) lemma connectedin_subset_separated_union: "\connectedin X C; separatedin X S T; C \ S \ T\ \ C \ S \ C \ T" unfolding connectedin_eq_not_separated_subset by blast lemma connectedin_nonseparated_union: assumes "connectedin X S" "connectedin X T" "\separatedin X S T" shows "connectedin X (S \ T)" proof - have "\C1 C2. \T \ C1 \ C2; S \ C1 \ C2\ \ S \ C1 = {} \ T \ C1 = {} \ S \ C2 = {} \ T \ C2 = {} \ \ separatedin X C1 C2" using assms unfolding connectedin_eq_not_separated_subset by (metis (no_types, lifting) assms connectedin_subset_separated_union inf.orderE separatedin_empty(1) separatedin_mono separatedin_sym) then show ?thesis unfolding connectedin_eq_not_separated_subset by (simp add: assms connectedin_subset_topspace Int_Un_distrib2) qed lemma connected_space_closures: "connected_space X \ (\e1 e2. e1 \ e2 = topspace X \ X closure_of e1 \ X closure_of e2 = {} \ e1 \ {} \ e2 \ {})" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding connected_space_closedin_eq by (metis Un_upper1 Un_upper2 closedin_closure_of closure_of_Un closure_of_eq_empty closure_of_topspace) next assume ?rhs then show ?lhs unfolding connected_space_closedin_eq by (metis closure_of_eq) qed lemma connectedin_Int_frontier_of: assumes "connectedin X S" "S \ T \ {}" "S - T \ {}" shows "S \ X frontier_of T \ {}" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) moreover have "S - (topspace X \ T) \ {}" using assms(3) by blast moreover have "S \ topspace X \ T \ {}" using assms connectedin by fastforce moreover have False if "S \ T \ {}" "S - T \ {}" "T \ topspace X" "S \ X frontier_of T = {}" for T proof - have null: "S \ (X closure_of T - X interior_of T) = {}" using that unfolding frontier_of_def by blast have "X interior_of T \ (topspace X - X closure_of T) \ S = {}" by (metis Diff_disjoint inf_bot_left interior_of_Int interior_of_complement interior_of_empty) moreover have "S \ X interior_of T \ (topspace X - X closure_of T)" using that \S \ topspace X\ null by auto moreover have "S \ X interior_of T \ {}" using closure_of_subset that(1) that(3) null by fastforce ultimately have "S \ X interior_of (topspace X - T) = {}" by (metis "*" inf_commute interior_of_complement openin_interior_of) then have "topspace (subtopology X S) \ X interior_of T = S" using \S \ topspace X\ interior_of_complement null by fastforce then show ?thesis using that by (metis Diff_eq_empty_iff inf_le2 interior_of_subset subset_trans) qed ultimately show ?thesis by (metis Int_lower1 frontier_of_restrict inf_assoc) qed lemma connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and "connectedin X S" shows "connectedin Y (f ` S)" proof - have "S \ topspace X" and *: "\E1 E2. openin X E1 \ openin X E2 \ E1 \ E2 \ S = {} \ S \ E1 \ E2 \ E1 \ S = {} \ E2 \ S = {}" using \connectedin X S\ by (auto simp: connectedin) show ?thesis unfolding connectedin connected_space_def proof (intro conjI notI; clarify) show "f x \ topspace Y" if "x \ S" for x using \S \ topspace X\ continuous_map_image_subset_topspace f that by blast next fix U V let ?U = "{x \ topspace X. f x \ U}" let ?V = "{x \ topspace X. f x \ V}" assume UV: "openin Y U" "openin Y V" "f ` S \ U \ V" "U \ V \ f ` S = {}" "U \ f ` S \ {}" "V \ f ` S \ {}" then have 1: "?U \ ?V \ S = {}" by auto have 2: "openin X ?U" "openin X ?V" using \openin Y U\ \openin Y V\ continuous_map f by fastforce+ show "False" using * [of ?U ?V] UV \S \ topspace X\ by (auto simp: 1 2) qed qed lemma connected_space_quotient_map_image: "\quotient_map X X' q; connected_space X\ \ connected_space X'" by (metis connectedin_continuous_map_image connectedin_topspace quotient_imp_continuous_map quotient_imp_surjective_map) lemma homeomorphic_connected_space: "X homeomorphic_space Y \ connected_space X \ connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis connected_space_subconnected connectedin_continuous_map_image connectedin_topspace continuous_map_image_subset_topspace image_eqI image_subset_iff) lemma homeomorphic_map_connectedness: assumes f: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "connectedin Y (f ` U) \ connectedin X U" proof - have 1: "f ` U \ topspace Y \ U \ topspace X" using U f homeomorphic_imp_surjective_map by blast moreover have "connected_space (subtopology Y (f ` U)) \ connected_space (subtopology X U)" proof (rule homeomorphic_connected_space) have "f ` U \ topspace Y" by (simp add: U 1) then have "topspace Y \ f ` U = f ` U" by (simp add: subset_antisym) then show "subtopology Y (f ` U) homeomorphic_space subtopology X U" by (metis U f homeomorphic_map_imp_homeomorphic_space homeomorphic_map_subtopologies homeomorphic_space_sym inf.absorb_iff2) qed ultimately show ?thesis by (auto simp: connectedin_def) qed lemma homeomorphic_map_connectedness_eq: "homeomorphic_map X Y f \ connectedin X U \ U \ topspace X \ connectedin Y (f ` U)" using homeomorphic_map_connectedness connectedin_subset_topspace by metis lemma connectedin_discrete_topology: "connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" proof (cases "S \ U") case True show ?thesis proof (cases "S = {}") case False moreover have "connectedin (discrete_topology U) S \ (\a. S = {a})" proof show "connectedin (discrete_topology U) S \ \a. S = {a}" using False connectedin_Int_frontier_of insert_Diff by fastforce qed (use True in auto) ultimately show ?thesis by auto qed simp next case False then show ?thesis by (simp add: connectedin_def) qed lemma connected_space_discrete_topology: "connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis connectedin_discrete_topology connectedin_topspace order_refl topspace_discrete_topology) subsection\Compact sets\ definition compactin where "compactin X S \ S \ topspace X \ (\\. (\U \ \. openin X U) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\))" definition compact_space where "compact_space X \ compactin X (topspace X)" lemma compact_space_alt: "compact_space X \ (\\. (\U \ \. openin X U) \ topspace X \ \\ \ (\\. finite \ \ \ \ \ \ topspace X \ \\))" by (simp add: compact_space_def compactin_def) lemma compact_space: "compact_space X \ (\\. (\U \ \. openin X U) \ \\ = topspace X \ (\\. finite \ \ \ \ \ \ \\ = topspace X))" unfolding compact_space_alt using openin_subset by fastforce lemma compactinD: "\compactin X S; \U. U \ \ \ openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" by (auto simp: compactin_def) lemma compactin_euclidean_iff [simp]: "compactin euclidean S \ compact S" by (simp add: compact_eq_Heine_Borel compactin_def) meson lemma compactin_absolute [simp]: "compactin (subtopology X S) S \ compactin X S" proof - have eq: "(\U \ \. \Y. openin X Y \ U = Y \ S) \ \ \ (\Y. Y \ S) ` {y. openin X y}" for \ by auto show ?thesis by (auto simp: compactin_def openin_subtopology eq imp_conjL all_subset_image ex_finite_subset_image) qed lemma compactin_subspace: "compactin X S \ S \ topspace X \ compact_space (subtopology X S)" unfolding compact_space_def topspace_subtopology by (metis compactin_absolute compactin_def inf.absorb2) lemma compact_space_subtopology: "compactin X S \ compact_space (subtopology X S)" by (simp add: compactin_subspace) lemma compactin_subtopology: "compactin (subtopology X S) T \ compactin X T \ T \ S" by (metis compactin_subspace inf.absorb_iff2 le_inf_iff subtopology_subtopology topspace_subtopology) lemma compactin_subset_topspace: "compactin X S \ S \ topspace X" by (simp add: compactin_subspace) lemma compactin_contractive: "\compactin X' S; topspace X' = topspace X; \U. openin X U \ openin X' U\ \ compactin X S" by (simp add: compactin_def) lemma finite_imp_compactin: "\S \ topspace X; finite S\ \ compactin X S" by (metis compactin_subspace compact_space finite_UnionD inf.absorb_iff2 order_refl topspace_subtopology) lemma compactin_empty [iff]: "compactin X {}" by (simp add: finite_imp_compactin) lemma compact_space_topspace_empty: "topspace X = {} \ compact_space X" by (simp add: compact_space_def) lemma finite_imp_compactin_eq: "finite S \ (compactin X S \ S \ topspace X)" using compactin_subset_topspace finite_imp_compactin by blast lemma compactin_sing [simp]: "compactin X {a} \ a \ topspace X" by (simp add: finite_imp_compactin_eq) lemma closed_compactin: assumes XK: "compactin X K" and "C \ K" and XC: "closedin X C" shows "compactin X C" unfolding compactin_def proof (intro conjI allI impI) show "C \ topspace X" by (simp add: XC closedin_subset) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ C \ \\" have "(\U\insert (topspace X - C) \. openin X U)" using XC \ by blast moreover have "K \ \(insert (topspace X - C) \)" using \ XK compactin_subset_topspace by fastforce ultimately obtain \ where "finite \" "\ \ insert (topspace X - C) \" "K \ \\" using assms unfolding compactin_def by metis moreover have "openin X (topspace X - C)" using XC by auto ultimately show "\\. finite \ \ \ \ \ \ C \ \\" using \C \ K\ by (rule_tac x="\ - {topspace X - C}" in exI) auto qed lemma closedin_compact_space: "\compact_space X; closedin X S\ \ compactin X S" by (simp add: closed_compactin closedin_subset compact_space_def) lemma compact_Int_closedin: assumes "compactin X S" "closedin X T" shows "compactin X (S \ T)" proof - have "compactin (subtopology X S) (S \ T)" by (metis assms closedin_compact_space closedin_subtopology compactin_subspace inf_commute) then show ?thesis by (simp add: compactin_subtopology) qed lemma closed_Int_compactin: "\closedin X S; compactin X T\ \ compactin X (S \ T)" by (metis compact_Int_closedin inf_commute) lemma compactin_Un: assumes S: "compactin X S" and T: "compactin X T" shows "compactin X (S \ T)" unfolding compactin_def proof (intro conjI allI impI) show "S \ T \ topspace X" using assms by (auto simp: compactin_def) next fix \ :: "'a set set" assume \: "Ball \ (openin X) \ S \ T \ \\" with S obtain \ where \: "finite \" "\ \ \" "S \ \\" unfolding compactin_def by (meson sup.bounded_iff) obtain \ where "finite \" "\ \ \" "T \ \\" using \ T unfolding compactin_def by (meson sup.bounded_iff) with \ show "\\. finite \ \ \ \ \ \ S \ T \ \\" by (rule_tac x="\ \ \" in exI) auto qed lemma compactin_Union: "\finite \; \S. S \ \ \ compactin X S\ \ compactin X (\\)" by (induction rule: finite_induct) (simp_all add: compactin_Un) lemma compactin_subtopology_imp_compact: assumes "compactin (subtopology X S) K" shows "compactin X K" using assms proof (clarsimp simp add: compactin_def) fix \ define \ where "\ \ (\U. U \ S) ` \" assume "K \ topspace X" and "K \ S" and "\x\\. openin X x" and "K \ \\" then have "\V \ \. openin (subtopology X S) V" "K \ \\" unfolding \_def by (auto simp: openin_subtopology) moreover assume "\\. (\x\\. openin (subtopology X S) x) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson then have \: "\U. U \ \ \ V = U \ S" if "V \ \" for V unfolding \_def using that by blast let ?\ = "(\F. @U. U \ \ \ F = U \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using someI_ex [OF \] by blast show "K \ \?\" proof clarsimp fix x assume "x \ K" then show "\V \ \. x \ (SOME U. U \ \ \ V = U \ S)" using \K \ \\\ someI_ex [OF \] by (metis (no_types, lifting) IntD1 Union_iff subsetCE) qed qed qed lemma compact_imp_compactin_subtopology: assumes "compactin X K" "K \ S" shows "compactin (subtopology X S) K" using assms proof (clarsimp simp add: compactin_def) fix \ :: "'a set set" define \ where "\ \ {V. openin X V \ (\U \ \. U = V \ S)}" assume "K \ S" and "K \ topspace X" and "\U\\. openin (subtopology X S) U" and "K \ \\" then have "\V \ \. openin X V" "K \ \\" unfolding \_def by (fastforce simp: subset_eq openin_subtopology)+ moreover assume "\\. (\U\\. openin X U) \ K \ \\ \ (\\. finite \ \ \ \ \ \ K \ \\)" ultimately obtain \ where "finite \" "\ \ \" "K \ \\" by meson let ?\ = "(\F. F \ S) ` \" show "\\. finite \ \ \ \ \ \ K \ \\" proof (intro exI conjI) show "finite ?\" using \finite \\ by blast show "?\ \ \" using \_def \\ \ \\ by blast show "K \ \?\" using \K \ \\\ assms(2) by auto qed qed proposition compact_space_fip: "compact_space X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" (is "_ = ?rhs") proof (cases "topspace X = {}") case True then show ?thesis unfolding compact_space_def by (metis Sup_bot_conv(1) closedin_topspace_empty compactin_empty finite.emptyI finite_UnionD order_refl) next case False show ?thesis proof safe fix \ :: "'a set set" assume * [rule_format]: "\\. finite \ \ \ \ \ \ \\ \ {}" define \ where "\ \ (\S. topspace X - S) ` \" assume clo: "\C\\. closedin X C" and [simp]: "\\ = {}" then have "\V \ \. openin X V" "topspace X \ \\" by (auto simp: \_def) moreover assume [unfolded compact_space_alt, rule_format, of \]: "compact_space X" ultimately obtain \ where \: "finite \" "\ \ \" "topspace X \ topspace X - \\" by (auto simp: ex_finite_subset_image \_def) moreover have "\ \ {}" using \ \topspace X \ {}\ by blast ultimately show "False" using * [of \] by auto (metis Diff_iff Inter_iff clo closedin_def subsetD) next assume R [rule_format]: ?rhs show "compact_space X" unfolding compact_space_alt proof clarify fix \ :: "'a set set" define \ where "\ \ (\S. topspace X - S) ` \" assume "\C\\. openin X C" and "topspace X \ \\" with \topspace X \ {}\ have *: "\V \ \. closedin X V" "\ \ {}" by (auto simp: \_def) show "\\. finite \ \ \ \ \ \ topspace X \ \\" proof (rule ccontr; simp) assume "\\\\. finite \ \ \ topspace X \ \\" then have "\\. finite \ \ \ \ \ \ \\ \ {}" by (simp add: \_def all_finite_subset_image) with \topspace X \ \\\ show False using R [of \] * by (simp add: \_def) qed qed qed qed corollary compactin_fip: "compactin X S \ S \ topspace X \ (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof (cases "S = {}") case False show ?thesis proof (cases "S \ topspace X") case True then have "compactin X S \ (\\. \ \ (\T. S \ T) ` {T. closedin X T} \ (\\. finite \ \ \ \ \ \ \\ \ {}) \ \\ \ {})" by (simp add: compact_space_fip compactin_subspace closedin_subtopology image_def subset_eq Int_commute imp_conjL) also have "\ = (\\\Collect (closedin X). (\\. finite \ \ \ \ (\) S ` \ \ \\ \ {}) \ \ ((\) S ` \) \ {})" by (simp add: all_subset_image) also have "\ = (\\. (\C\\. closedin X C) \ (\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" proof - have eq: "((\\. finite \ \ \ \ \ \ \ ((\) S ` \) \ {}) \ \ ((\) S ` \) \ {}) \ ((\\. finite \ \ \ \ \ \ S \ \\ \ {}) \ S \ \\ \ {})" for \ by simp (use \S \ {}\ in blast) show ?thesis unfolding imp_conjL [symmetric] all_finite_subset_image eq by blast qed finally show ?thesis using True by simp qed (simp add: compactin_subspace) qed force corollary compact_space_imp_nest: fixes C :: "nat \ 'a set" assumes "compact_space X" and clo: "\n. closedin X (C n)" and ne: "\n. C n \ {}" and dec: "decseq C" shows "(\n. C n) \ {}" proof - let ?\ = "range (\n. \m \ n. C m)" have "closedin X A" if "A \ ?\" for A using that clo by auto moreover have "(\n\K. \m \ n. C m) \ {}" if "finite K" for K proof - obtain n where "\k. k \ K \ k \ n" using Max.coboundedI \finite K\ by blast with dec have "C n \ (\n\K. \m \ n. C m)" unfolding decseq_def by blast with ne [of n] show ?thesis by blast qed ultimately show ?thesis using \compact_space X\ [unfolded compact_space_fip, rule_format, of ?\] by (simp add: all_finite_subset_image INT_extend_simps UN_atMost_UNIV del: INT_simps) qed lemma compactin_discrete_topology: "compactin (discrete_topology X) S \ S \ X \ finite S" (is "?lhs = ?rhs") proof (intro iffI conjI) assume L: ?lhs then show "S \ X" by (auto simp: compactin_def) have *: "\\. Ball \ (openin (discrete_topology X)) \ S \ \\ \ (\\. finite \ \ \ \ \ \ S \ \\)" using L by (auto simp: compactin_def) show "finite S" using * [of "(\x. {x}) ` X"] \S \ X\ by clarsimp (metis UN_singleton finite_subset_image infinite_super) next assume ?rhs then show ?lhs by (simp add: finite_imp_compactin) qed lemma compact_space_discrete_topology: "compact_space(discrete_topology X) \ finite X" by (simp add: compactin_discrete_topology compact_space_def) lemma compact_space_imp_Bolzano_Weierstrass: assumes "compact_space X" "infinite S" "S \ topspace X" shows "X derived_set_of S \ {}" proof assume X: "X derived_set_of S = {}" then have "closedin X S" by (simp add: closedin_contains_derived_set assms) then have "compactin X S" by (rule closedin_compact_space [OF \compact_space X\]) with X show False by (metis \infinite S\ compactin_subspace compact_space_discrete_topology inf_bot_right subtopology_eq_discrete_topology_eq) qed lemma compactin_imp_Bolzano_Weierstrass: "\compactin X S; infinite T \ T \ S\ \ S \ X derived_set_of T \ {}" using compact_space_imp_Bolzano_Weierstrass [of "subtopology X S"] by (simp add: compactin_subspace derived_set_of_subtopology inf_absorb2) lemma compact_closure_of_imp_Bolzano_Weierstrass: "\compactin X (X closure_of S); infinite T; T \ S; T \ topspace X\ \ X derived_set_of T \ {}" using closure_of_mono closure_of_subset compactin_imp_Bolzano_Weierstrass by fastforce lemma discrete_compactin_eq_finite: "S \ X derived_set_of S = {} \ compactin X S \ S \ topspace X \ finite S" by (meson compactin_imp_Bolzano_Weierstrass finite_imp_compactin_eq order_refl) lemma discrete_compact_space_eq_finite: "X derived_set_of (topspace X) = {} \ (compact_space X \ finite(topspace X))" by (metis compact_space_discrete_topology discrete_topology_unique_derived_set) lemma image_compactin: assumes cpt: "compactin X S" and cont: "continuous_map X Y f" shows "compactin Y (f ` S)" unfolding compactin_def proof (intro conjI allI impI) show "f ` S \ topspace Y" using compactin_subset_topspace cont continuous_map_image_subset_topspace cpt by blast next fix \ :: "'b set set" assume \: "Ball \ (openin Y) \ f ` S \ \\" define \ where "\ \ (\U. {x \ topspace X. f x \ U}) ` \" have "S \ topspace X" and *: "\\. \\U\\. openin X U; S \ \\\ \ \\. finite \ \ \ \ \ \ S \ \\" using cpt by (auto simp: compactin_def) obtain \ where \: "finite \" "\ \ \" "S \ \\" proof - have 1: "\U\\. openin X U" unfolding \_def using \ cont[unfolded continuous_map] by blast have 2: "S \ \\" unfolding \_def using compactin_subset_topspace cpt \ by fastforce show thesis using * [OF 1 2] that by metis qed have "\v \ \. \U. U \ \ \ v = {x \ topspace X. f x \ U}" using \_def by blast then obtain U where U: "\v \ \. U v \ \ \ v = {x \ topspace X. f x \ U v}" by metis show "\\. finite \ \ \ \ \ \ f ` S \ \\" proof (intro conjI exI) show "finite (U ` \)" by (simp add: \finite \\) next show "U ` \ \ \" using \\ \ \\ U by auto next show "f ` S \ \ (U ` \)" using \(2-3) U UnionE subset_eq U by fastforce qed qed lemma homeomorphic_compact_space: assumes "X homeomorphic_space Y" shows "compact_space X \ compact_space Y" using homeomorphic_space_sym by (metis assms compact_space_def homeomorphic_eq_everything_map homeomorphic_space image_compactin) lemma homeomorphic_map_compactness: assumes hom: "homeomorphic_map X Y f" and U: "U \ topspace X" shows "compactin Y (f ` U) \ compactin X U" proof - have "f ` U \ topspace Y" using hom U homeomorphic_imp_surjective_map by blast moreover have "homeomorphic_map (subtopology X U) (subtopology Y (f ` U)) f" using U hom homeomorphic_imp_surjective_map by (blast intro: homeomorphic_map_subtopologies) then have "compact_space (subtopology Y (f ` U)) = compact_space (subtopology X U)" using homeomorphic_compact_space homeomorphic_map_imp_homeomorphic_space by blast ultimately show ?thesis by (simp add: compactin_subspace U) qed lemma homeomorphic_map_compactness_eq: "homeomorphic_map X Y f \ compactin X U \ U \ topspace X \ compactin Y (f ` U)" by (meson compactin_subset_topspace homeomorphic_map_compactness) subsection\Embedding maps\ definition embedding_map where "embedding_map X Y f \ homeomorphic_map X (subtopology Y (f ` (topspace X))) f" lemma embedding_map_eq: "\embedding_map X Y f; \x. x \ topspace X \ f x = g x\ \ embedding_map X Y g" unfolding embedding_map_def by (metis homeomorphic_map_eq image_cong) lemma embedding_map_compose: assumes "embedding_map X X' f" "embedding_map X' X'' g" shows "embedding_map X X'' (g \ f)" proof - have hm: "homeomorphic_map X (subtopology X' (f ` topspace X)) f" "homeomorphic_map X' (subtopology X'' (g ` topspace X')) g" using assms by (auto simp: embedding_map_def) then obtain C where "g ` topspace X' \ C = (g \ f) ` topspace X" by (metis homeomorphic_imp_surjective_map image_comp image_mono inf.absorb_iff2 topspace_subtopology) then have "homeomorphic_map (subtopology X' (f ` topspace X)) (subtopology X'' ((g \ f) ` topspace X)) g" by (metis hm homeomorphic_imp_surjective_map homeomorphic_map_subtopologies image_comp subtopology_subtopology topspace_subtopology) then show ?thesis unfolding embedding_map_def using hm(1) homeomorphic_map_compose by blast qed lemma surjective_embedding_map: "embedding_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (force simp: embedding_map_def homeomorphic_eq_everything_map) lemma embedding_map_in_subtopology: "embedding_map X (subtopology Y S) f \ embedding_map X Y f \ f ` (topspace X) \ S" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" unfolding embedding_map_def by (metis continuous_map_in_subtopology homeomorphic_imp_continuous_map inf_absorb2 subtopology_subtopology) qed (simp add: embedding_map_def inf.absorb_iff2 subtopology_subtopology) lemma injective_open_imp_embedding_map: "\continuous_map X Y f; open_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: continuous_map_in_subtopology continuous_open_quotient_map eq_iff homeomorphic_map_def open_map_imp_subset open_map_into_subtopology) lemma injective_closed_imp_embedding_map: "\continuous_map X Y f; closed_map X Y f; inj_on f (topspace X)\ \ embedding_map X Y f" unfolding embedding_map_def by (simp add: closed_map_imp_subset closed_map_into_subtopology continuous_closed_quotient_map continuous_map_in_subtopology dual_order.eq_iff homeomorphic_map_def) lemma embedding_map_imp_homeomorphic_space: "embedding_map X Y f \ X homeomorphic_space (subtopology Y (f ` (topspace X)))" unfolding embedding_map_def using homeomorphic_space by blast lemma embedding_imp_closed_map: "\embedding_map X Y f; closedin Y (f ` topspace X)\ \ closed_map X Y f" unfolding closed_map_def by (auto simp: closedin_closed_subtopology embedding_map_def homeomorphic_map_closedness_eq) lemma embedding_imp_closed_map_eq: "embedding_map X Y f \ (closed_map X Y f \ closedin Y (f ` topspace X))" using closed_map_def embedding_imp_closed_map by blast subsection\Retraction and section maps\ definition retraction_maps :: "'a topology \ 'b topology \ ('a \ 'b) \ ('b \ 'a) \ bool" where "retraction_maps X Y f g \ continuous_map X Y f \ continuous_map Y X g \ (\x \ topspace Y. f(g x) = x)" definition section_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "section_map X Y f \ \g. retraction_maps Y X g f" definition retraction_map :: "'a topology \ 'b topology \ ('a \ 'b) \ bool" where "retraction_map X Y f \ \g. retraction_maps X Y f g" lemma retraction_maps_eq: "\retraction_maps X Y f g; \x. x \ topspace X \ f x = f' x; \x. x \ topspace Y \ g x = g' x\ \ retraction_maps X Y f' g'" unfolding retraction_maps_def by (metis (no_types, lifting) continuous_map_def continuous_map_eq) lemma section_map_eq: "\section_map X Y f; \x. x \ topspace X \ f x = g x\ \ section_map X Y g" unfolding section_map_def using retraction_maps_eq by blast lemma retraction_map_eq: "\retraction_map X Y f; \x. x \ topspace X \ f x = g x\ \ retraction_map X Y g" unfolding retraction_map_def using retraction_maps_eq by blast lemma homeomorphic_imp_retraction_maps: "homeomorphic_maps X Y f g \ retraction_maps X Y f g" by (simp add: homeomorphic_maps_def retraction_maps_def) lemma section_and_retraction_eq_homeomorphic_map: "section_map X Y f \ retraction_map X Y f \ homeomorphic_map X Y f" (is "?lhs = ?rhs") proof assume ?lhs then obtain g where "homeomorphic_maps X Y f g" unfolding homeomorphic_maps_def retraction_map_def section_map_def by (smt (verit, best) continuous_map_def retraction_maps_def) then show ?rhs using homeomorphic_map_maps by blast next assume ?rhs then show ?lhs unfolding retraction_map_def section_map_def by (meson homeomorphic_imp_retraction_maps homeomorphic_map_maps homeomorphic_maps_sym) qed lemma section_imp_embedding_map: "section_map X Y f \ embedding_map X Y f" unfolding section_map_def embedding_map_def homeomorphic_map_maps retraction_maps_def homeomorphic_maps_def by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology) lemma retraction_imp_quotient_map: assumes "retraction_map X Y f" shows "quotient_map X Y f" unfolding quotient_map_def proof (intro conjI subsetI allI impI) show "f ` topspace X = topspace Y" using assms by (force simp: retraction_map_def retraction_maps_def continuous_map_def) next fix U assume U: "U \ topspace Y" have "openin Y U" if "\x\topspace Y. g x \ topspace X" "\x\topspace Y. f (g x) = x" "openin Y {x \ topspace Y. g x \ {x \ topspace X. f x \ U}}" for g using openin_subopen U that by fastforce then show "openin X {x \ topspace X. f x \ U} = openin Y U" using assms by (auto simp: retraction_map_def retraction_maps_def continuous_map_def) qed lemma retraction_maps_compose: "\retraction_maps X Y f f'; retraction_maps Y Z g g'\ \ retraction_maps X Z (g \ f) (f' \ g')" by (clarsimp simp: retraction_maps_def continuous_map_compose) (simp add: continuous_map_def) lemma retraction_map_compose: "\retraction_map X Y f; retraction_map Y Z g\ \ retraction_map X Z (g \ f)" by (meson retraction_map_def retraction_maps_compose) lemma section_map_compose: "\section_map X Y f; section_map Y Z g\ \ section_map X Z (g \ f)" by (meson retraction_maps_compose section_map_def) lemma surjective_section_eq_homeomorphic_map: "section_map X Y f \ f ` (topspace X) = topspace Y \ homeomorphic_map X Y f" by (meson section_and_retraction_eq_homeomorphic_map section_imp_embedding_map surjective_embedding_map) lemma surjective_retraction_or_section_map: "f ` (topspace X) = topspace Y \ retraction_map X Y f \ section_map X Y f \ retraction_map X Y f" using section_and_retraction_eq_homeomorphic_map surjective_section_eq_homeomorphic_map by fastforce lemma retraction_imp_surjective_map: "retraction_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: retraction_imp_quotient_map quotient_imp_surjective_map) lemma section_imp_injective_map: "\section_map X Y f; x \ topspace X; y \ topspace X\ \ f x = f y \ x = y" by (metis (mono_tags, opaque_lifting) retraction_maps_def section_map_def) lemma retraction_maps_to_retract_maps: "retraction_maps X Y r s \ retraction_maps X (subtopology X (s ` (topspace Y))) (s \ r) id" unfolding retraction_maps_def by (auto simp: continuous_map_compose continuous_map_into_subtopology continuous_map_from_subtopology) subsection \Continuity\ lemma continuous_on_open: "continuous_on S f \ (\T. openin (top_of_set (f ` S)) T \ openin (top_of_set S) (S \ f -` T))" unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_closed: "continuous_on S f \ (\T. closedin (top_of_set (f ` S)) T \ closedin (top_of_set S) (S \ f -` T))" unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) lemma continuous_on_imp_closedin: assumes "continuous_on S f" "closedin (top_of_set (f ` S)) T" shows "closedin (top_of_set S) (S \ f -` T)" using assms continuous_on_closed by blast lemma continuous_map_subtopology_eu [simp]: "continuous_map (top_of_set S) (subtopology euclidean T) h \ continuous_on S h \ h ` S \ T" by (simp add: continuous_map_in_subtopology) lemma continuous_map_euclidean_top_of_set: assumes eq: "f -` S = UNIV" and cont: "continuous_on UNIV f" shows "continuous_map euclidean (top_of_set S) f" by (simp add: cont continuous_map_into_subtopology eq image_subset_iff_subset_vimage) subsection\<^marker>\tag unimportant\ \Half-global and completely global cases\ lemma continuous_openin_preimage_gen: assumes "continuous_on S f" "open T" shows "openin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "openin (top_of_set (f ` S)) (T \ f ` S)" using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto then show ?thesis using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_closedin_preimage: assumes "continuous_on S f" and "closed T" shows "closedin (top_of_set S) (S \ f -` T)" proof - have *: "(S \ f -` T) = (S \ f -` (T \ f ` S))" by auto have "closedin (top_of_set (f ` S)) (T \ f ` S)" using closedin_closed_Int[of T "f ` S", OF assms(2)] by (simp add: Int_commute) then show ?thesis using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \ f ` S"]] using * by auto qed lemma continuous_openin_preimage_eq: "continuous_on S f \ (\T. open T \ openin (top_of_set S) (S \ f -` T))" by (metis Int_commute continuous_on_open_invariant open_openin openin_subtopology) lemma continuous_closedin_preimage_eq: "continuous_on S f \ (\T. closed T \ closedin (top_of_set S) (S \ f -` T))" by (metis Int_commute closedin_closed continuous_on_closed_invariant) lemma continuous_open_preimage: assumes contf: "continuous_on S f" and "open S" "open T" shows "open (S \ f -` T)" proof- obtain U where "open U" "(S \ f -` T) = S \ U" using continuous_openin_preimage_gen[OF contf \open T\] unfolding openin_open by auto then show ?thesis using open_Int[of S U, OF \open S\] by auto qed lemma continuous_closed_preimage: assumes contf: "continuous_on S f" and "closed S" "closed T" shows "closed (S \ f -` T)" proof- obtain U where "closed U" "(S \ f -` T) = S \ U" using continuous_closedin_preimage[OF contf \closed T\] unfolding closedin_closed by auto then show ?thesis using closed_Int[of S U, OF \closed S\] by auto qed lemma continuous_open_vimage: "open S \ (\x. continuous (at x) f) \ open (f -` S)" by (metis continuous_on_eq_continuous_within open_vimage) lemma continuous_closed_vimage: "closed S \ (\x. continuous (at x) f) \ closed (f -` S)" by (simp add: closed_vimage continuous_on_eq_continuous_within) lemma Times_in_interior_subtopology: assumes "(x, y) \ U" "openin (top_of_set (S \ T)) U" obtains V W where "openin (top_of_set S) V" "x \ V" "openin (top_of_set T) W" "y \ W" "(V \ W) \ U" proof - from assms obtain E where "open E" "U = S \ T \ E" "(x, y) \ E" "x \ S" "y \ T" by (auto simp: openin_open) from open_prod_elim[OF \open E\ \(x, y) \ E\] obtain E1 E2 where "open E1" "open E2" "(x, y) \ E1 \ E2" "E1 \ E2 \ E" by blast show ?thesis proof show "openin (top_of_set S) (E1 \ S)" "openin (top_of_set T) (E2 \ T)" using \open E1\ \open E2\ by (auto simp: openin_open) show "x \ E1 \ S" "y \ E2 \ T" using \(x, y) \ E1 \ E2\ \x \ S\ \y \ T\ by auto show "(E1 \ S) \ (E2 \ T) \ U" using \E1 \ E2 \ E\ \U = _\ by auto qed qed lemma closedin_Times: "closedin (top_of_set S) S' \ closedin (top_of_set T) T' \ closedin (top_of_set (S \ T)) (S' \ T')" unfolding closedin_closed using closed_Times by blast lemma openin_Times: "openin (top_of_set S) S' \ openin (top_of_set T) T' \ openin (top_of_set (S \ T)) (S' \ T')" unfolding openin_open using open_Times by blast lemma openin_Times_eq: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" shows "openin (top_of_set (S \ T)) (S' \ T') \ S' = {} \ T' = {} \ openin (top_of_set S) S' \ openin (top_of_set T) T'" (is "?lhs = ?rhs") proof (cases "S' = {} \ T' = {}") case True then show ?thesis by auto next case False then obtain x y where "x \ S'" "y \ T'" by blast show ?thesis proof assume ?lhs have "openin (top_of_set S) S'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set S) U \ x \ U \ U \ S'" if "x \ S'" for x using that \y \ T'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed moreover have "openin (top_of_set T) T'" proof (subst openin_subopen, clarify) show "\U. openin (top_of_set T) U \ y \ U \ U \ T'" if "y \ T'" for y using that \x \ S'\ Times_in_interior_subtopology [OF _ \?lhs\, of x y] by simp (metis mem_Sigma_iff subsetD subsetI) qed ultimately show ?rhs by simp next assume ?rhs with False show ?lhs by (simp add: openin_Times) qed qed lemma Lim_transform_within_openin: assumes f: "(f \ l) (at a within T)" and "openin (top_of_set T) S" "a \ S" and eq: "\x. \x \ S; x \ a\ \ f x = g x" shows "(g \ l) (at a within T)" proof - have "\\<^sub>F x in at a within T. x \ T \ x \ a" by (simp add: eventually_at_filter) moreover from \openin _ _\ obtain U where "open U" "S = T \ U" by (auto simp: openin_open) then have "a \ U" using \a \ S\ by auto from topological_tendstoD[OF tendsto_ident_at \open U\ \a \ U\] have "\\<^sub>F x in at a within T. x \ U" by auto ultimately have "\\<^sub>F x in at a within T. f x = g x" by eventually_elim (auto simp: \S = _\ eq) with f show ?thesis by (rule Lim_transform_eventually) qed lemma continuous_on_open_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. openin (top_of_set T) U \ openin (top_of_set S) (S \ f -` U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (clarsimp simp add: continuous_openin_preimage_eq openin_open) (metis Int_assoc assms image_subset_iff_subset_vimage inf.absorb_iff1) next assume R [rule_format]: ?rhs show ?lhs proof (clarsimp simp add: continuous_openin_preimage_eq) fix U::"'a set" assume "open U" then have "openin (top_of_set S) (S \ f -` (U \ T))" by (metis R inf_commute openin_open) then show "openin (top_of_set S) (S \ f -` U)" by (metis Int_assoc Int_commute assms image_subset_iff_subset_vimage inf.absorb_iff2 vimage_Int) qed qed lemma continuous_openin_preimage: "\continuous_on S f; f ` S \ T; openin (top_of_set T) U\ \ openin (top_of_set S) (S \ f -` U)" by (simp add: continuous_on_open_gen) lemma continuous_on_closed_gen: assumes "f ` S \ T" shows "continuous_on S f \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) (S \ f -` U))" proof - have *: "U \ T \ S \ f -` (T - U) = S - (S \ f -` U)" for U using assms by blast then show ?thesis unfolding continuous_on_open_gen [OF assms] by (metis closedin_def inf.cobounded1 openin_closedin_eq topspace_euclidean_subtopology) qed lemma continuous_closedin_preimage_gen: assumes "continuous_on S f" "f ` S \ T" "closedin (top_of_set T) U" shows "closedin (top_of_set S) (S \ f -` U)" using assms continuous_on_closed_gen by blast lemma continuous_transform_within_openin: assumes "continuous (at a within T) f" and "openin (top_of_set T) S" "a \ S" and eq: "\x. x \ S \ f x = g x" shows "continuous (at a within T) g" using assms by (simp add: Lim_transform_within_openin continuous_within) subsection\<^marker>\tag important\ \The topology generated by some (open) subsets\ text \In the definition below of a generated topology, the \Empty\ case is not necessary, as it follows from \UN\ taking for \K\ the empty set. However, it is convenient to have, and is never a problem in proofs, so I prefer to write it down explicitly. We do not require \UNIV\ to be an open set, as this will not be the case in applications. (We are thinking of a topology on a subset of \UNIV\, the remaining part of \UNIV\ being irrelevant.)\ inductive generate_topology_on for S where Empty: "generate_topology_on S {}" | Int: "generate_topology_on S a \ generate_topology_on S b \ generate_topology_on S (a \ b)" | UN: "(\k. k \ K \ generate_topology_on S k) \ generate_topology_on S (\K)" | Basis: "s \ S \ generate_topology_on S s" lemma istopology_generate_topology_on: "istopology (generate_topology_on S)" unfolding istopology_def by (auto intro: generate_topology_on.intros) text \The basic property of the topology generated by a set \S\ is that it is the smallest topology containing all the elements of \S\:\ lemma generate_topology_on_coarsest: assumes T: "istopology T" "\s. s \ S \ T s" and gen: "generate_topology_on S s0" shows "T s0" using gen by (induct rule: generate_topology_on.induct) (use T in \auto simp: istopology_def\) abbreviation\<^marker>\tag unimportant\ topology_generated_by::"('a set set) \ ('a topology)" where "topology_generated_by S \ topology (generate_topology_on S)" lemma openin_topology_generated_by_iff: "openin (topology_generated_by S) s \ generate_topology_on S s" using topology_inverse'[OF istopology_generate_topology_on[of S]] by simp lemma openin_topology_generated_by: "openin (topology_generated_by S) s \ generate_topology_on S s" using openin_topology_generated_by_iff by auto lemma topology_generated_by_topspace [simp]: "topspace (topology_generated_by S) = (\S)" proof { fix s assume "openin (topology_generated_by S) s" then have "generate_topology_on S s" by (rule openin_topology_generated_by) then have "s \ (\S)" by (induct, auto) } then show "topspace (topology_generated_by S) \ (\S)" unfolding topspace_def by auto next have "generate_topology_on S (\S)" using generate_topology_on.UN[OF generate_topology_on.Basis, of S S] by simp then show "(\S) \ topspace (topology_generated_by S)" unfolding topspace_def using openin_topology_generated_by_iff by auto qed lemma topology_generated_by_Basis: "s \ S \ openin (topology_generated_by S) s" by (simp add: Basis openin_topology_generated_by_iff) lemma generate_topology_on_Inter: "\finite \; \K. K \ \ \ generate_topology_on \ K; \ \ {}\ \ generate_topology_on \ (\\)" by (induction \ rule: finite_induct; force intro: generate_topology_on.intros) subsection\Topology bases and sub-bases\ lemma istopology_base_alt: "istopology (arbitrary union_of P) \ (\S T. (arbitrary union_of P) S \ (arbitrary union_of P) T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_def) (blast intro: arbitrary_union_of_Union) lemma istopology_base_eq: "istopology (arbitrary union_of P) \ (\S T. P S \ P T \ (arbitrary union_of P) (S \ T))" by (simp add: istopology_base_alt arbitrary_union_of_Int_eq) lemma istopology_base: "(\S T. \P S; P T\ \ P(S \ T)) \ istopology (arbitrary union_of P)" by (simp add: arbitrary_def istopology_base_eq union_of_inc) lemma openin_topology_base_unique: "openin X = arbitrary union_of P \ (\V. P V \ openin X V) \ (\U x. openin X U \ x \ U \ (\V. P V \ x \ V \ V \ U))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: union_of_def arbitrary_def) next assume R: ?rhs then have *: "\\\Collect P. \\ = S" if "openin X S" for S using that by (rule_tac x="{V. P V \ V \ S}" in exI) fastforce from R show ?lhs by (fastforce simp add: union_of_def arbitrary_def intro: *) qed lemma topology_base_unique: assumes "\S. P S \ openin X S" "\U x. \openin X U; x \ U\ \ \B. P B \ x \ B \ B \ U" shows "topology (arbitrary union_of P) = X" proof - have "X = topology (openin X)" by (simp add: openin_inverse) also from assms have "openin X = arbitrary union_of P" by (subst openin_topology_base_unique) auto finally show ?thesis .. qed lemma topology_bases_eq_aux: "\(arbitrary union_of P) S; \U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U\ \ (arbitrary union_of Q) S" by (metis arbitrary_union_of_alt arbitrary_union_of_idempot) lemma topology_bases_eq: "\\U x. \P U; x \ U\ \ \V. Q V \ x \ V \ V \ U; \V x. \Q V; x \ V\ \ \U. P U \ x \ U \ U \ V\ \ topology (arbitrary union_of P) = topology (arbitrary union_of Q)" by (fastforce intro: arg_cong [where f=topology] elim: topology_bases_eq_aux) lemma istopology_subbase: "istopology (arbitrary union_of (finite intersection_of P relative_to S))" by (simp add: finite_intersection_of_Int istopology_base relative_to_Int) lemma openin_subbase: "openin (topology (arbitrary union_of (finite intersection_of B relative_to U))) S \ (arbitrary union_of (finite intersection_of B relative_to U)) S" by (simp add: istopology_subbase topology_inverse') lemma topspace_subbase [simp]: "topspace(topology (arbitrary union_of (finite intersection_of B relative_to U))) = U" (is "?lhs = _") proof show "?lhs \ U" by (metis arbitrary_union_of_relative_to openin_subbase openin_topspace relative_to_imp_subset) show "U \ ?lhs" by (metis arbitrary_union_of_inc finite_intersection_of_empty inf.orderE istopology_subbase openin_subset relative_to_inc subset_UNIV topology_inverse') qed lemma minimal_topology_subbase: assumes X: "\S. P S \ openin X S" and "openin X U" and S: "openin(topology(arbitrary union_of (finite intersection_of P relative_to U))) S" shows "openin X S" proof - have "(arbitrary union_of (finite intersection_of P relative_to U)) S" using S openin_subbase by blast with X \openin X U\ show ?thesis by (force simp add: union_of_def intersection_of_def relative_to_def intro: openin_Int_Inter) qed lemma istopology_subbase_UNIV: "istopology (arbitrary union_of (finite intersection_of P))" by (simp add: istopology_base finite_intersection_of_Int) lemma generate_topology_on_eq: "generate_topology_on S = arbitrary union_of finite' intersection_of (\x. x \ S)" (is "?lhs = ?rhs") proof (intro ext iffI) fix A assume "?lhs A" then show "?rhs A" proof induction case (Int a b) then show ?case by (metis (mono_tags, lifting) istopology_base_alt finite'_intersection_of_Int istopology_base) next case (UN K) then show ?case by (simp add: arbitrary_union_of_Union) next case (Basis s) then show ?case by (simp add: Sup_upper arbitrary_union_of_inc finite'_intersection_of_inc relative_to_subset) qed auto next fix A assume "?rhs A" then obtain \ where \: "\T. T \ \ \ \\. finite' \ \ \ \ S \ \\ = T" and eq: "A = \\" unfolding union_of_def intersection_of_def by auto show "?lhs A" unfolding eq proof (rule generate_topology_on.UN) fix T assume "T \ \" with \ obtain \ where "finite' \" "\ \ S" "\\ = T" by blast have "generate_topology_on S (\\)" proof (rule generate_topology_on_Inter) show "finite \" "\ \ {}" by (auto simp: \finite' \\) show "\K. K \ \ \ generate_topology_on S K" by (metis \\ \ S\ generate_topology_on.simps subset_iff) qed then show "generate_topology_on S T" using \\\ = T\ by blast qed qed lemma continuous_on_generated_topo_iff: "continuous_map T1 (topology_generated_by S) f \ ((\U. U \ S \ openin T1 (f-`U \ topspace(T1))) \ (f`(topspace T1) \ (\ S)))" unfolding continuous_map_alt topology_generated_by_topspace proof (auto simp add: topology_generated_by_Basis) assume H: "\U. U \ S \ openin T1 (f -` U \ topspace T1)" fix U assume "openin (topology_generated_by S) U" then have "generate_topology_on S U" by (rule openin_topology_generated_by) then show "openin T1 (f -` U \ topspace T1)" proof (induct) fix a b assume H: "openin T1 (f -` a \ topspace T1)" "openin T1 (f -` b \ topspace T1)" have "f -` (a \ b) \ topspace T1 = (f-`a \ topspace T1) \ (f-`b \ topspace T1)" by auto then show "openin T1 (f -` (a \ b) \ topspace T1)" using H by auto next fix K assume H: "openin T1 (f -` k \ topspace T1)" if "k\ K" for k define L where "L = {f -` k \ topspace T1|k. k \ K}" have *: "openin T1 l" if "l \L" for l using that H unfolding L_def by auto have "openin T1 (\L)" using openin_Union[OF *] by simp moreover have "(\L) = (f -` \K \ topspace T1)" unfolding L_def by auto ultimately show "openin T1 (f -` \K \ topspace T1)" by simp qed (auto simp add: H) qed lemma continuous_on_generated_topo: assumes "\U. U \S \ openin T1 (f-`U \ topspace(T1))" "f`(topspace T1) \ (\ S)" shows "continuous_map T1 (topology_generated_by S) f" using assms continuous_on_generated_topo_iff by blast subsection\Continuity via bases/subbases, hence upper and lower semicontinuity\ lemma continuous_map_into_topology_base: assumes P: "openin Y = arbitrary union_of P" and f: "\x. x \ topspace X \ f x \ topspace Y" and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" shows "continuous_map X Y f" proof - have *: "\\. (\t. t \ \ \ P t) \ openin X {x \ topspace X. \U\\. f x \ U}" by (smt (verit) Ball_Collect ope mem_Collect_eq openin_subopen) show ?thesis using P by (auto simp: continuous_map_def arbitrary_def union_of_def intro!: f *) qed lemma continuous_map_into_topology_base_eq: assumes P: "openin Y = arbitrary union_of P" shows "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" (is "?lhs=?rhs") proof assume L: ?lhs then have "\x. x \ topspace X \ f x \ topspace Y" by (meson continuous_map_def) moreover have "\U. P U \ openin X {x \ topspace X. f x \ U}" using L assms continuous_map openin_topology_base_unique by fastforce ultimately show ?rhs by auto qed (simp add: assms continuous_map_into_topology_base) lemma continuous_map_into_topology_subbase: fixes U P defines "Y \ topology(arbitrary union_of (finite intersection_of P relative_to U))" assumes f: "\x. x \ topspace X \ f x \ topspace Y" and ope: "\U. P U \ openin X {x \ topspace X. f x \ U}" shows "continuous_map X Y f" proof (intro continuous_map_into_topology_base) show "openin Y = arbitrary union_of (finite intersection_of P relative_to U)" unfolding Y_def using istopology_subbase topology_inverse' by blast show "openin X {x \ topspace X. f x \ V}" if \
: "(finite intersection_of P relative_to U) V" for V proof - define finv where "finv \ \V. {x \ topspace X. f x \ V}" obtain \ where \: "finite \" "\V. V \ \ \ P V" "{x \ topspace X. f x \ V} = (\V \ insert U \. finv V)" using \
by (fastforce simp: finv_def intersection_of_def relative_to_def) show ?thesis unfolding \ proof (intro openin_Inter ope) have U: "U = topspace Y" unfolding Y_def using topspace_subbase by fastforce fix V assume V: "V \ finv ` insert U \" with U f have "openin X {x \ topspace X. f x \ U}" by (auto simp: openin_subopen [of X "Collect _"]) then show "openin X V" using V \(2) ope by (fastforce simp: finv_def) qed (use \finite \\ in auto) qed qed (use f in auto) lemma continuous_map_into_topology_subbase_eq: assumes "Y = topology(arbitrary union_of (finite intersection_of P relative_to U))" shows "continuous_map X Y f \ (\x \ topspace X. f x \ topspace Y) \ (\U. P U \ openin X {x \ topspace X. f x \ U})" (is "?lhs=?rhs") proof assume L: ?lhs show ?rhs proof (intro conjI strip) show "\x. x \ topspace X \ f x \ topspace Y" using L continuous_map_def by fastforce fix V assume "P V" have "U = topspace Y" unfolding assms using topspace_subbase by fastforce then have eq: "{x \ topspace X. f x \ V} = {x \ topspace X. f x \ U \ V}" using L by (auto simp: continuous_map) have "openin Y (U \ V)" unfolding assms openin_subbase by (meson \P V\ arbitrary_union_of_inc finite_intersection_of_inc relative_to_inc) show "openin X {x \ topspace X. f x \ V}" using L \openin Y (U \ V)\ continuous_map eq by fastforce qed next show "?rhs \ ?lhs" unfolding assms by (intro continuous_map_into_topology_subbase) auto qed lemma subbase_subtopology_euclidean: fixes U :: "'a::order_topology set" shows "topology (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ range lessThan) relative_to U)) = subtopology euclidean U" proof - have "\V. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ x \ V \ V \ W" if "open W" "x \ W" for W and x::'a using \open W\ [unfolded open_generated_order] \x \ W\ proof (induct rule: generate_topology.induct) case UNIV then show ?case using finite_intersection_of_empty by blast next case (Int a b) then show ?case by (meson Int_iff finite_intersection_of_Int inf_mono) next case (UN K) then show ?case by (meson Union_iff subset_iff) next case (Basis s) then show ?case by (metis (no_types, lifting) Un_iff finite_intersection_of_inc order_refl) qed moreover have "\V::'a set. (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)) V \ open V" by (force simp: intersection_of_def subset_iff) ultimately have *: "openin (euclidean::'a topology) = (arbitrary union_of (finite intersection_of (\x. x \ range greaterThan \ x \ range lessThan)))" by (smt (verit, best) openin_topology_base_unique open_openin) show ?thesis unfolding subtopology_def arbitrary_union_of_relative_to [symmetric] apply (simp add: relative_to_def flip: *) by (metis Int_commute) qed lemma continuous_map_upper_lower_semicontinuous_lt_gen: fixes U :: "'a::order_topology set" shows "continuous_map X (subtopology euclidean U) f \ (\x \ topspace X. f x \ U) \ (\a. openin X {x \ topspace X. f x > a}) \ (\a. openin X {x \ topspace X. f x < a})" by (auto simp: continuous_map_into_topology_subbase_eq [OF subbase_subtopology_euclidean [symmetric, of U]] greaterThan_def lessThan_def image_iff simp flip: all_simps) lemma continuous_map_upper_lower_semicontinuous_lt: fixes f :: "'a \ 'b::order_topology" shows "continuous_map X euclidean f \ (\a. openin X {x \ topspace X. f x > a}) \ (\a. openin X {x \ topspace X. f x < a})" using continuous_map_upper_lower_semicontinuous_lt_gen [where U=UNIV] by auto lemma Int_Collect_imp_eq: "A \ {x. x\A \ P x} = {x\A. P x}" by blast lemma continuous_map_upper_lower_semicontinuous_le_gen: shows "continuous_map X (subtopology euclideanreal U) f \ (\x \ topspace X. f x \ U) \ (\a. closedin X {x \ topspace X. f x \ a}) \ (\a. closedin X {x \ topspace X. f x \ a})" unfolding continuous_map_upper_lower_semicontinuous_lt_gen by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) lemma continuous_map_upper_lower_semicontinuous_le: "continuous_map X euclideanreal f \ (\a. closedin X {x \ topspace X. f x \ a}) \ (\a. closedin X {x \ topspace X. f x \ a})" using continuous_map_upper_lower_semicontinuous_le_gen [where U=UNIV] by auto lemma continuous_map_upper_lower_semicontinuous_lte_gen: "continuous_map X (subtopology euclideanreal U) f \ (\x \ topspace X. f x \ U) \ (\a. openin X {x \ topspace X. f x < a}) \ (\a. closedin X {x \ topspace X. f x \ a})" unfolding continuous_map_upper_lower_semicontinuous_lt_gen by (auto simp: closedin_def Diff_eq Compl_eq not_le Int_Collect_imp_eq) lemma continuous_map_upper_lower_semicontinuous_lte: "continuous_map X euclideanreal f \ (\a. openin X {x \ topspace X. f x < a}) \ (\a. closedin X {x \ topspace X. f x \ a})" using continuous_map_upper_lower_semicontinuous_lte_gen [where U=UNIV] by auto subsection\<^marker>\tag important\ \Pullback topology\ text \Pulling back a topology by map gives again a topology. \subtopology\ is a special case of this notion, pulling back by the identity. We introduce the general notion as we will need it to define the strong operator topology on the space of continuous linear operators, by pulling back the product topology on the space of all functions.\ text \\pullback_topology A f T\ is the pullback of the topology \T\ by the map \f\ on the set \A\.\ definition\<^marker>\tag important\ pullback_topology::"('a set) \ ('a \ 'b) \ ('b topology) \ ('a topology)" where "pullback_topology A f T = topology (\S. \U. openin T U \ S = f-`U \ A)" lemma istopology_pullback_topology: "istopology (\S. \U. openin T U \ S = f-`U \ A)" unfolding istopology_def proof (auto) fix K assume "\S\K. \U. openin T U \ S = f -` U \ A" then have "\U. \S\K. openin T (U S) \ S = f-`(U S) \ A" by (rule bchoice) then obtain U where U: "\S\K. openin T (U S) \ S = f-`(U S) \ A" by blast define V where "V = (\S\K. U S)" have "openin T V" "\K = f -` V \ A" unfolding V_def using U by auto then show "\V. openin T V \ \K = f -` V \ A" by auto qed lemma openin_pullback_topology: "openin (pullback_topology A f T) S \ (\U. openin T U \ S = f-`U \ A)" unfolding pullback_topology_def topology_inverse'[OF istopology_pullback_topology] by auto lemma topspace_pullback_topology: "topspace (pullback_topology A f T) = f-`(topspace T) \ A" by (auto simp add: topspace_def openin_pullback_topology) proposition continuous_map_pullback [intro]: assumes "continuous_map T1 T2 g" shows "continuous_map (pullback_topology A f T1) T2 (g o f)" unfolding continuous_map_alt proof (auto) fix U::"'b set" assume "openin T2 U" then have "openin T1 (g-`U \ topspace T1)" using assms unfolding continuous_map_alt by auto have "(g o f)-`U \ topspace (pullback_topology A f T1) = (g o f)-`U \ A \ f-`(topspace T1)" unfolding topspace_pullback_topology by auto also have "... = f-`(g-`U \ topspace T1) \ A " by auto also have "openin (pullback_topology A f T1) (...)" unfolding openin_pullback_topology using \openin T1 (g-`U \ topspace T1)\ by auto finally show "openin (pullback_topology A f T1) ((g \ f) -` U \ topspace (pullback_topology A f T1))" by auto next fix x assume "x \ topspace (pullback_topology A f T1)" then have "f x \ topspace T1" unfolding topspace_pullback_topology by auto then show "g (f x) \ topspace T2" using assms unfolding continuous_map_def by auto qed proposition continuous_map_pullback' [intro]: assumes "continuous_map T1 T2 (f o g)" "topspace T1 \ g-`A" shows "continuous_map T1 (pullback_topology A f T2) g" unfolding continuous_map_alt proof (auto) fix U assume "openin (pullback_topology A f T2) U" then have "\V. openin T2 V \ U = f-`V \ A" unfolding openin_pullback_topology by auto then obtain V where "openin T2 V" "U = f-`V \ A" by blast then have "g -` U \ topspace T1 = g-`(f-`V \ A) \ topspace T1" by blast also have "... = (f o g)-`V \ (g-`A \ topspace T1)" by auto also have "... = (f o g)-`V \ topspace T1" using assms(2) by auto also have "openin T1 (...)" using assms(1) \openin T2 V\ by auto finally show "openin T1 (g -` U \ topspace T1)" by simp next fix x assume "x \ topspace T1" have "(f o g) x \ topspace T2" using assms(1) \x \ topspace T1\ unfolding continuous_map_def by auto then have "g x \ f-`(topspace T2)" unfolding comp_def by blast moreover have "g x \ A" using assms(2) \x \ topspace T1\ by blast ultimately show "g x \ topspace (pullback_topology A f T2)" unfolding topspace_pullback_topology by blast qed subsection\Proper maps (not a priori assumed continuous) \ definition proper_map where "proper_map X Y f \ closed_map X Y f \ (\y \ topspace Y. compactin X {x \ topspace X. f x = y})" lemma proper_imp_closed_map: "proper_map X Y f \ closed_map X Y f" by (simp add: proper_map_def) lemma proper_map_imp_subset_topspace: "proper_map X Y f \ f ` (topspace X) \ topspace Y" by (simp add: closed_map_imp_subset_topspace proper_map_def) +lemma proper_map_restriction: + assumes "proper_map X Y f" "{x \ topspace X. f x \ V} = U" + shows "proper_map (subtopology X U) (subtopology Y V) f" +proof - + have [simp]: "{x \ topspace X. f x \ V \ f x = y} = {x \ topspace X. f x = y}" + if "y \ V" for y + using that by auto + show ?thesis + using assms unfolding proper_map_def using closed_map_restriction + by (force simp: compactin_subtopology) +qed + lemma closed_injective_imp_proper_map: assumes f: "closed_map X Y f" and inj: "inj_on f (topspace X)" shows "proper_map X Y f" unfolding proper_map_def proof (clarsimp simp: f) show "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y using inj_on_eq_iff [OF inj] that proof - have "{x \ topspace X. f x = y} = {} \ (\a \ topspace X. {x \ topspace X. f x = y} = {a})" using inj_on_eq_iff [OF inj] by auto then show ?thesis using that by (metis (no_types, lifting) compactin_empty compactin_sing) qed qed lemma injective_imp_proper_eq_closed_map: "inj_on f (topspace X) \ (proper_map X Y f \ closed_map X Y f)" using closed_injective_imp_proper_map proper_imp_closed_map by blast lemma homeomorphic_imp_proper_map: "homeomorphic_map X Y f \ proper_map X Y f" by (simp add: closed_injective_imp_proper_map homeomorphic_eq_everything_map) lemma compactin_proper_map_preimage: assumes f: "proper_map X Y f" and "compactin Y K" shows "compactin X {x. x \ topspace X \ f x \ K}" proof - have "f ` (topspace X) \ topspace Y" by (simp add: f proper_map_imp_subset_topspace) have *: "\y. y \ topspace Y \ compactin X {x \ topspace X. f x = y}" using f by (auto simp: proper_map_def) show ?thesis unfolding compactin_def proof clarsimp show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x \ K} \ \\" if \: "\U\\. openin X U" and sub: "{x \ topspace X. f x \ K} \ \\" for \ proof - have "\y \ K. \\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" proof fix y assume "y \ K" then have "compactin X {x \ topspace X. f x = y}" by (metis "*" \compactin Y K\ compactin_subspace subsetD) with \y \ K\ show "\\. finite \ \ \ \ \ \ {x \ topspace X. f x = y} \ \\" unfolding compactin_def using \ sub by fastforce qed then obtain \ where \: "\y. y \ K \ finite (\ y) \ \ y \ \ \ {x \ topspace X. f x = y} \ \(\ y)" by (metis (full_types)) define F where "F \ \y. topspace Y - f ` (topspace X - \(\ y))" have "\\. finite \ \ \ \ F ` K \ K \ \\" proof (rule compactinD [OF \compactin Y K\]) have "\x. x \ K \ closedin Y (f ` (topspace X - \(\ x)))" using f unfolding proper_map_def closed_map_def by (meson \ \ openin_Union openin_closedin_eq subsetD) then show "openin Y U" if "U \ F ` K" for U using that by (auto simp: F_def) show "K \ \(F ` K)" using \ \compactin Y K\ unfolding F_def compactin_def by fastforce qed then obtain J where "finite J" "J \ K" and J: "K \ \(F ` J)" by (auto simp: ex_finite_subset_image) show ?thesis unfolding F_def proof (intro exI conjI) show "finite (\(\ ` J))" using \ \J \ K\ \finite J\ by blast show "\(\ ` J) \ \" using \ \J \ K\ by blast show "{x \ topspace X. f x \ K} \ \(\(\ ` J))" using J \J \ K\ unfolding F_def by auto qed qed qed qed lemma compact_space_proper_map_preimage: assumes f: "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" and "compact_space Y" shows "compact_space X" proof - have eq: "topspace X = {x \ topspace X. f x \ topspace Y}" using fim by blast moreover have "compactin Y (topspace Y)" using \compact_space Y\ compact_space_def by auto ultimately show ?thesis unfolding compact_space_def using eq f compactin_proper_map_preimage by fastforce qed lemma proper_map_alt: "proper_map X Y f \ closed_map X Y f \ (\K. compactin Y K \ compactin X {x. x \ topspace X \ f x \ K})" proof (intro iffI conjI allI impI) show "compactin X {x \ topspace X. f x \ K}" if "proper_map X Y f" and "compactin Y K" for K using that by (simp add: compactin_proper_map_preimage) show "proper_map X Y f" if f: "closed_map X Y f \ (\K. compactin Y K \ compactin X {x \ topspace X. f x \ K})" proof - have "compactin X {x \ topspace X. f x = y}" if "y \ topspace Y" for y proof - have "compactin X {x \ topspace X. f x \ {y}}" using f compactin_sing that by fastforce then show ?thesis by auto qed with f show ?thesis by (auto simp: proper_map_def) qed qed (simp add: proper_imp_closed_map) lemma proper_map_on_empty: "topspace X = {} \ proper_map X Y f" by (auto simp: proper_map_def closed_map_on_empty) lemma proper_map_id [simp]: "proper_map X X id" proof (clarsimp simp: proper_map_alt closed_map_id) fix K assume K: "compactin X K" then have "{a \ topspace X. a \ K} = K" by (simp add: compactin_subspace subset_antisym subset_iff) then show "compactin X {a \ topspace X. a \ K}" using K by auto qed lemma proper_map_compose: assumes "proper_map X Y f" "proper_map Y Z g" shows "proper_map X Z (g \ f)" proof - have "closed_map X Y f" and f: "\K. compactin Y K \ compactin X {x \ topspace X. f x \ K}" and "closed_map Y Z g" and g: "\K. compactin Z K \ compactin Y {x \ topspace Y. g x \ K}" using assms by (auto simp: proper_map_alt) show ?thesis unfolding proper_map_alt proof (intro conjI allI impI) show "closed_map X Z (g \ f)" using \closed_map X Y f\ \closed_map Y Z g\ closed_map_compose by blast have "{x \ topspace X. g (f x) \ K} = {x \ topspace X. f x \ {b \ topspace Y. g b \ K}}" for K using \closed_map X Y f\ closed_map_imp_subset_topspace by blast then show "compactin X {x \ topspace X. (g \ f) x \ K}" if "compactin Z K" for K using f [OF g [OF that]] by auto qed qed lemma proper_map_const: "proper_map X Y (\x. c) \ compact_space X \ (topspace X = {} \ closedin Y {c})" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: compact_space_topspace_empty proper_map_on_empty) next case False have *: "compactin X {x \ topspace X. c = y}" if "compact_space X" for y using that unfolding compact_space_def by (metis (mono_tags, lifting) compactin_empty empty_subsetI mem_Collect_eq subsetI subset_antisym) then show ?thesis using closed_compactin closedin_subset by (force simp: False proper_map_def closed_map_const compact_space_def) qed lemma proper_map_inclusion: "S \ topspace X \ proper_map (subtopology X S) X id \ closedin X S \ (\k. compactin X k \ compactin X (S \ k))" by (metis closed_Int_compactin closed_map_inclusion_eq inf.absorb_iff2 inj_on_id injective_imp_proper_eq_closed_map) +lemma proper_map_into_subtopology: + "\proper_map X Y f; f ` topspace X \ C\ \ proper_map X (subtopology Y C) f" + by (simp add: closed_map_into_subtopology compactin_subtopology proper_map_alt) + +lemma proper_map_from_composition_left: + assumes gf: "proper_map X Z (g \ f)" and contf: "continuous_map X Y f" and fim: "f ` topspace X = topspace Y" + shows "proper_map Y Z g" + unfolding proper_map_def +proof (intro strip conjI) + show "closed_map Y Z g" + by (meson closed_map_from_composition_left gf contf fim proper_imp_closed_map) + fix z assume "z \ topspace Z" + have eq: "{y \ topspace Y. g y = z} = f ` {x \ topspace X. (g \ f) x = z}" + using fim by force + show "compactin Y {x \ topspace Y. g x = z}" + unfolding eq + proof (rule image_compactin [OF _ contf]) + show "compactin X {x \ topspace X. (g \ f) x = z}" + using \z \ topspace Z\ gf proper_map_def by fastforce + qed +qed + +lemma proper_map_from_composition_right_inj: + assumes gf: "proper_map X Z (g \ f)" and fim: "f ` topspace X \ topspace Y" + and contf: "continuous_map Y Z g" and inj: "inj_on g (topspace Y)" + shows "proper_map X Y f" + unfolding proper_map_def +proof (intro strip conjI) + show "closed_map X Y f" + by (meson closed_map_from_composition_right assms proper_imp_closed_map) + fix y + assume "y \ topspace Y" + with fim inj have eq: "{x \ topspace X. f x = y} = {x \ topspace X. (g \ f) x = g y}" + by (auto simp: image_subset_iff inj_onD) + show "compactin X {x \ topspace X. f x = y}" + unfolding eq + by (smt (verit) Collect_cong \y \ topspace Y\ contf continuous_map_closedin gf proper_map_def) +qed + subsection\Perfect maps (proper, continuous and surjective)\ definition perfect_map where "perfect_map X Y f \ continuous_map X Y f \ proper_map X Y f \ f ` (topspace X) = topspace Y" lemma homeomorphic_imp_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f" by (simp add: homeomorphic_eq_everything_map homeomorphic_imp_proper_map perfect_map_def) lemma perfect_imp_quotient_map: "perfect_map X Y f \ quotient_map X Y f" by (simp add: continuous_closed_imp_quotient_map perfect_map_def proper_map_def) lemma homeomorphic_eq_injective_perfect_map: "homeomorphic_map X Y f \ perfect_map X Y f \ inj_on f (topspace X)" using homeomorphic_imp_perfect_map homeomorphic_map_def perfect_imp_quotient_map by blast lemma perfect_injective_eq_homeomorphic_map: "perfect_map X Y f \ inj_on f (topspace X) \ homeomorphic_map X Y f" by (simp add: homeomorphic_eq_injective_perfect_map) lemma perfect_map_id [simp]: "perfect_map X X id" by (simp add: homeomorphic_imp_perfect_map) lemma perfect_map_compose: "\perfect_map X Y f; perfect_map Y Z g\ \ perfect_map X Z (g \ f)" by (meson continuous_map_compose perfect_imp_quotient_map perfect_map_def proper_map_compose quotient_map_compose_eq quotient_map_def) lemma perfect_imp_continuous_map: "perfect_map X Y f \ continuous_map X Y f" using perfect_map_def by blast lemma perfect_imp_closed_map: "perfect_map X Y f \ closed_map X Y f" by (simp add: perfect_map_def proper_map_def) lemma perfect_imp_proper_map: "perfect_map X Y f \ proper_map X Y f" by (simp add: perfect_map_def) lemma perfect_imp_surjective_map: "perfect_map X Y f \ f ` (topspace X) = topspace Y" by (simp add: perfect_map_def) +lemma perfect_map_from_composition_left: + assumes "perfect_map X Z (g \ f)" and "continuous_map X Y f" + and "continuous_map Y Z g" and "f ` topspace X = topspace Y" + shows "perfect_map Y Z g" + using assms unfolding perfect_map_def + by (metis image_comp proper_map_from_composition_left) + 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,57 +1,58 @@ theory Analysis imports (* Linear Algebra *) Convex Determinants (* Topology *) FSigma Sum_Topology Abstract_Topological_Spaces + Abstract_Metric_Spaces Connected Abstract_Limits Isolated (* 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_Sum 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/Measure_Space.thy b/src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy +++ b/src/HOL/Analysis/Measure_Space.thy @@ -1,3767 +1,3763 @@ (* Title: HOL/Analysis/Measure_Space.thy Author: Lawrence C Paulson Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Measure Spaces\ theory Measure_Space imports Measurable "HOL-Library.Extended_Nonnegative_Real" begin subsection\<^marker>\tag unimportant\ "Relate extended reals and the indicator function" lemma suminf_cmult_indicator: fixes f :: "nat \ ennreal" assumes "disjoint_family A" "x \ A i" shows "(\n. f n * indicator (A n) x) = f i" proof - have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)" using \x \ A i\ assms unfolding disjoint_family_on_def indicator_def by auto then have "\n. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" from this[of "Suc i"] show "f i \ y" by auto qed (use assms in simp) ultimately show ?thesis using assms by (simp add: suminf_eq_SUP) qed lemma suminf_indicator: assumes "disjoint_family A" shows "(\n. indicator (A n) x :: ennreal) = indicator (\i. A i) x" proof cases assume *: "x \ (\i. A i)" then obtain i where "x \ A i" by auto from suminf_cmult_indicator[OF assms(1), OF \x \ A i\, of "\k. 1"] show ?thesis using * by simp qed simp lemma sum_indicator_disjoint_family: fixes f :: "'d \ 'e::semiring_1" assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" shows "(\i\P. f i * indicator (A i) x) = f j" proof - have "P \ {i. x \ A i} = {j}" using d \x \ A j\ \j \ P\ unfolding disjoint_family_on_def by auto with \finite P\ show ?thesis by (simp add: indicator_def) qed text \ The type for emeasure spaces is already defined in \<^theory>\HOL-Analysis.Sigma_Algebra\, as it is also used to represent sigma algebras (with an arbitrary emeasure). \ subsection\<^marker>\tag unimportant\ "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" shows "(\n. \i f A + f B" proof - have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" proof fix n show "(\i < Suc (Suc n). f (binaryset A B i)) = f A + f B" by (induct n) (auto simp add: binaryset_def f) qed moreover have "\ \ f A + f B" by (rule tendsto_const) ultimately have "(\n. \i< n+2. f (binaryset A B i)) \ f A + f B" by simp thus ?thesis by (rule LIMSEQ_offset [where k=2]) qed lemma binaryset_sums: assumes f: "f {} = 0" shows "(\n. f (binaryset A B n)) sums (f A + f B)" using LIMSEQ_binaryset f sums_def by blast lemma suminf_binaryset_eq: fixes f :: "'a set \ 'b::{comm_monoid_add, t2_space}" shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) subsection\<^marker>\tag unimportant\ \Properties of a premeasure \<^term>\\\\ text \ The definitions for \<^const>\positive\ and \<^const>\countably_additive\ should be here, by they are necessary to define \<^typ>\'a measure\ in \<^theory>\HOL-Analysis.Sigma_Algebra\. \ definition subadditive where "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" lemma subadditiveD: "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) definition countably_subadditive where "countably_subadditive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" lemma (in ring_of_sets) countably_subadditive_subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" and cs: "countably_subadditive M f" shows "subadditive M f" proof (auto simp add: subadditive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) \ (\ n. f (binaryset x y n))" using cs by (auto simp add: countably_subadditive_def) hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) \ (\ n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) \ f x + f y" using f x y by (auto simp add: Un o_def suminf_binaryset_eq positive_def) qed definition additive where "additive M \ \ (\x\M. \y\M. x \ y = {} \ \ (x \ y) = \ x + \ y)" definition increasing where "increasing M \ \ (\x\M. \y\M. x \ y \ \ x \ \ y)" lemma positiveD1: "positive M f \ f {} = 0" by (auto simp: positive_def) lemma positiveD_empty: "positive M f \ f {} = 0" by (auto simp add: positive_def) lemma additiveD: "additive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) = f x + f y" by (auto simp add: additive_def) lemma increasingD: "increasing M f \ x \ y \ x\M \ y\M \ f x \ f y" by (auto simp add: increasing_def) lemma countably_additiveI[case_names countably]: "(\A. \range A \ M; disjoint_family A; (\i. A i) \ M\ \ (\i. f(A i)) = f(\i. A i)) \ countably_additive M f" by (simp add: countably_additive_def) lemma (in ring_of_sets) disjointed_additive: assumes f: "positive M f" "additive M f" and A: "range A \ M" "incseq A" shows "(\i\n. f (disjointed A i)) = f (A n)" proof (induct n) case (Suc n) then have "(\i\Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" by simp also have "\ = f (A n \ disjointed A (Suc n))" using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono) also have "A n \ disjointed A (Suc n) = A (Suc n)" using \incseq A\ by (auto dest: incseq_SucD simp: disjointed_mono) finally show ?case . qed simp lemma (in ring_of_sets) additive_sum: fixes A:: "'i \ 'a set" assumes f: "positive M f" and ad: "additive M f" and "finite S" and A: "A`S \ M" and disj: "disjoint_family_on A S" shows "(\i\S. f (A i)) = f (\i\S. A i)" using \finite S\ disj A proof induct case empty show ?case using f by (simp add: positive_def) next case (insert s S) then have "A s \ (\i\S. A i) = {}" by (auto simp add: disjoint_family_on_def neq_iff) moreover have "A s \ M" using insert by blast moreover have "(\i\S. A i) \ M" using insert \finite S\ by auto ultimately have "f (A s \ (\i\S. A i)) = f (A s) + f(\i\S. A i)" using ad UNION_in_sets A by (auto simp add: additive_def) with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] by (auto simp add: additive_def subset_insertI) qed lemma (in ring_of_sets) additive_increasing: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and addf: "additive M f" shows "increasing M f" proof (auto simp add: increasing_def) fix x y assume xy: "x \ M" "y \ M" "x \ y" then have "y - x \ M" by auto then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono zero_le) also have "\ = f (x \ (y-x))" by (metis addf Diff_disjoint \y - x \ M\ additiveD xy(1)) also have "\ = f y" by (metis Un_Diff_cancel Un_absorb1 xy(3)) finally show "f x \ f y" by simp qed lemma (in ring_of_sets) subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and A: "A`S \ M" and S: "finite S" shows "f (\i\S. A i) \ (\i\S. f (A i))" using S A proof (induct S) case empty thus ?case using f by (auto simp: positive_def) next case (insert x F) hence in_M: "A x \ M" "(\i\F. A i) \ M" "(\i\F. A i) - A x \ M" using A by force+ have subs: "(\i\F. A i) - A x \ (\i\F. A i)" by auto have "(\i\(insert x F). A i) = A x \ ((\i\F. A i) - A x)" by auto hence "f (\i\(insert x F). A i) = f (A x \ ((\i\F. A i) - A x))" by simp also have "\ = f (A x) + f ((\i\F. A i) - A x)" using f(2) by (rule additiveD) (insert in_M, auto) also have "\ \ f (A x) + f (\i\F. A i)" using additive_increasing[OF f] in_M subs by (simp add: increasingD) also have "\ \ f (A x) + (\i\F. f (A i))" using insert by (auto intro: add_left_mono) finally show "f (\i\(insert x F). A i) \ (\i\(insert x F). f (A i))" by (simp add: insert) qed lemma (in ring_of_sets) countably_additive_additive: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and ca: "countably_additive M f" shows "additive M f" proof (auto simp add: additive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" using ca by (simp add: countably_additive_def) hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) = (\n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) = f x + f y" using posf x y by (auto simp add: Un suminf_binaryset_eq positive_def) qed lemma (in algebra) increasing_additive_bound: fixes A:: "nat \ 'a set" and f :: "'a set \ ennreal" assumes f: "positive M f" and ad: "additive M f" and inc: "increasing M f" and A: "range A \ M" and disj: "disjoint_family A" shows "(\i. f (A i)) \ f \" proof (safe intro!: suminf_le_const) fix N note disj_N = disjoint_family_on_mono[OF _ disj, of "{..ii\{.. \ f \" using space_closed A by (intro increasingD[OF inc] finite_UN) auto finally show "(\i f \" by simp qed (insert f A, auto simp: positive_def) lemma (in ring_of_sets) countably_additiveI_finite: fixes \ :: "'a set \ ennreal" assumes "finite \" "positive M \" "additive M \" shows "countably_additive M \" proof (rule countably_additiveI) fix F :: "nat \ 'a set" assume F: "range F \ M" "(\i. F i) \ M" and disj: "disjoint_family F" have "\i. F i \ {} \ (\x. x \ F i)" by auto then obtain f where f: "\i. F i \ {} \ f i \ F i" by metis have finU: "finite (\i. F i)" by (metis F(2) assms(1) infinite_super sets_into_space) have F_subset: "{i. \ (F i) \ 0} \ {i. F i \ {}}" by (auto simp: positiveD_empty[OF \positive M \\]) moreover have fin_not_empty: "finite {i. F i \ {}}" proof (rule finite_imageD) from f have "f`{i. F i \ {}} \ (\i. F i)" by auto then show "finite (f`{i. F i \ {}})" by (simp add: finU finite_subset) show inj_f: "inj_on f {i. F i \ {}}" using f disj by (simp add: inj_on_def disjoint_family_on_def disjoint_iff) metis qed ultimately have fin_not_0: "finite {i. \ (F i) \ 0}" by (rule finite_subset) have disj_not_empty: "disjoint_family_on F {i. F i \ {}}" using disj by (auto simp: disjoint_family_on_def) from fin_not_0 have "(\i. \ (F i)) = (\i | \ (F i) \ 0. \ (F i))" by (rule suminf_finite) auto also have "\ = (\i | F i \ {}. \ (F i))" using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto also have "\ = \ (\i\{i. F i \ {}}. F i)" using \positive M \\ \additive M \\ fin_not_empty disj_not_empty F by (intro additive_sum) auto also have "\ = \ (\i. F i)" by (rule arg_cong[where f=\]) auto finally show "(\i. \ (F i)) = \ (\i. F i)" . qed lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "countably_additive M f \ (\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i))" unfolding countably_additive_def proof safe assume count_sum: "\A. range A \ M \ disjoint_family A \ \(A ` UNIV) \ M \ (\i. f (A i)) = f (\(A ` UNIV))" fix A :: "nat \ 'a set" assume A: "range A \ M" "incseq A" "(\i. A i) \ M" then have dA: "range (disjointed A) \ M" by (auto simp: range_disjointed_sets) with count_sum[THEN spec, of "disjointed A"] A(3) have f_UN: "(\i. f (disjointed A i)) = f (\i. A i)" by (auto simp: UN_disjointed_eq disjoint_family_disjointed) moreover have "(\n. (\i (\i. f (disjointed A i))" by (simp add: summable_LIMSEQ) from LIMSEQ_Suc[OF this] have "(\n. (\i\n. f (disjointed A i))) \ (\i. f (disjointed A i))" unfolding lessThan_Suc_atMost . moreover have "\n. (\i\n. f (disjointed A i)) = f (A n)" using disjointed_additive[OF f A(1,2)] . ultimately show "(\i. f (A i)) \ f (\i. A i)" by simp next assume cont[rule_format]: "\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i)" fix A :: "nat \ 'a set" assume A: "range A \ M" "disjoint_family A" "(\i. A i) \ M" have *: "(\n. (\ii. A i)" by auto have "range (\i. \i M" "(\i. \i M" using A * by auto then have "(\n. f (\i f (\i. A i)" unfolding *[symmetric] by (force intro!: cont incseq_SucI)+ moreover have "\n. f (\iii. f (A i)) sums f (\i. A i)" unfolding sums_def by simp then show "(\i. f (A i)) = f (\i. A i)" by (metis sums_unique) qed lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i)) \ (\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0)" proof safe assume cont[rule_format]: "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i))" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" with cont[of A] show "(\i. f (A i)) \ 0" using \positive M f\[unfolded positive_def] by auto next assume cont[rule_format]: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) \ M" "\i. f (A i) \ \" have f_mono: "\a b. a \ M \ b \ M \ a \ b \ f a \ f b" using additive_increasing[OF f] unfolding increasing_def by simp have decseq_fA: "decseq (\i. f (A i))" using A by (auto simp: decseq_def intro!: f_mono) have decseq: "decseq (\i. A i - (\i. A i))" using A by (auto simp: decseq_def) then have decseq_f: "decseq (\i. f (A i - (\i. A i)))" using A unfolding decseq_def by (auto intro!: f_mono Diff) have "f (\x. A x) \ f (A 0)" using A by (auto intro!: f_mono) then have f_Int_fin: "f (\x. A x) \ \" using A by (auto simp: top_unique) have f_fin: "f (A i - (\i. A i)) \ \" for i using A by (metis Diff Diff_subset f_mono infinity_ennreal_def range_subsetD top_unique) have "(\i. f (A i - (\i. A i))) \ 0" proof (intro cont[ OF _ decseq _ f_fin]) show "range (\i. A i - (\i. A i)) \ M" "(\i. A i - (\i. A i)) = {}" using A by auto qed with INF_Lim decseq_f have "(INF n. f (A n - (\i. A i))) = 0" by metis moreover have "(INF n. f (\i. A i)) = f (\i. A i)" by auto ultimately have "(INF n. f (A n - (\i. A i)) + f (\i. A i)) = 0 + f (\i. A i)" using A(4) f_fin f_Int_fin using INF_ennreal_add_const by presburger moreover { fix n have "f (A n - (\i. A i)) + f (\i. A i) = f ((A n - (\i. A i)) \ (\i. A i))" using A by (subst f(2)[THEN additiveD]) auto also have "(A n - (\i. A i)) \ (\i. A i) = A n" by auto finally have "f (A n - (\i. A i)) + f (\i. A i) = f (A n)" . } ultimately have "(INF n. f (A n)) = f (\i. A i)" by simp with LIMSEQ_INF[OF decseq_fA] show "(\i. f (A i)) \ f (\i. A i)" by simp qed lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" assumes A: "range A \ M" "incseq A" "(\i. A i) \ M" shows "(\i. f (A i)) \ f (\i. A i)" proof - from A have "(\i. f ((\i. A i) - A i)) \ 0" by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) moreover { fix i have "f ((\i. A i) - A i \ A i) = f ((\i. A i) - A i) + f (A i)" using A by (intro f(2)[THEN additiveD]) auto also have "((\i. A i) - A i) \ A i = (\i. A i)" by auto finally have "f ((\i. A i) - A i) = f (\i. A i) - f (A i)" using assms f by fastforce } moreover have "\\<^sub>F i in sequentially. f (A i) \ f (\i. A i)" using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\i. A i"] A by (auto intro!: always_eventually simp: subset_eq) ultimately show "(\i. f (A i)) \ f (\i. A i)" by (auto intro: ennreal_tendsto_const_minus) qed lemma (in ring_of_sets) empty_continuous_imp_countably_additive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and fin: "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" shows "countably_additive M f" using countably_additive_iff_continuous_from_below[OF f] using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast subsection\<^marker>\tag unimportant\ \Properties of \<^const>\emeasure\\ lemma emeasure_positive: "positive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma emeasure_empty[simp, intro]: "emeasure M {} = 0" using emeasure_positive[of M] by (simp add: positive_def) lemma emeasure_single_in_space: "emeasure M {x} \ 0 \ x \ space M" using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2]) lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma suminf_emeasure: "range A \ sets M \ disjoint_family A \ (\i. emeasure M (A i)) = emeasure M (\i. A i)" using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M] by (simp add: countably_additive_def) lemma sums_emeasure: "disjoint_family F \ (\i. F i \ sets M) \ (\i. emeasure M (F i)) sums emeasure M (\i. F i)" unfolding sums_iff by (intro conjI suminf_emeasure) auto lemma emeasure_additive: "additive (sets M) (emeasure M)" by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive) lemma plus_emeasure: "a \ sets M \ b \ sets M \ a \ b = {} \ emeasure M a + emeasure M b = emeasure M (a \ b)" using additiveD[OF emeasure_additive] .. lemma emeasure_Un: "A \ sets M \ B \ sets M \ emeasure M (A \ B) = emeasure M A + emeasure M (B - A)" using plus_emeasure[of A M "B - A"] by auto lemma emeasure_Un_Int: assumes "A \ sets M" "B \ sets M" shows "emeasure M A + emeasure M B = emeasure M (A \ B) + emeasure M (A \ B)" proof - have "A = (A-B) \ (A \ B)" by auto then have "emeasure M A = emeasure M (A-B) + emeasure M (A \ B)" by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff) moreover have "A \ B = (A-B) \ B" by auto then have "emeasure M (A \ B) = emeasure M (A-B) + emeasure M B" by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff) ultimately show ?thesis by (metis add.assoc add.commute) qed lemma sum_emeasure: "F`I \ sets M \ disjoint_family_on F I \ finite I \ (\i\I. emeasure M (F i)) = emeasure M (\i\I. F i)" by (metis sets.additive_sum emeasure_positive emeasure_additive) lemma emeasure_mono: "a \ b \ b \ sets M \ emeasure M a \ emeasure M b" by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD) lemma emeasure_space: "emeasure M A \ emeasure M (space M)" by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le) lemma emeasure_Diff: assumes "emeasure M B \ \" and "A \ sets M" "B \ sets M" and "B \ A" shows "emeasure M (A - B) = emeasure M A - emeasure M B" by (smt (verit, best) add_diff_self_ennreal assms emeasure_Un emeasure_mono ennreal_add_left_cancel le_iff_sup) lemma emeasure_compl: "s \ sets M \ emeasure M s \ \ \ emeasure M (space M - s) = emeasure M (space M) - emeasure M s" by (rule emeasure_Diff) (auto dest: sets.sets_into_space) lemma Lim_emeasure_incseq: "range A \ sets M \ incseq A \ (\i. (emeasure M (A i))) \ emeasure M (\i. A i)" using emeasure_countably_additive by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive emeasure_additive) lemma incseq_emeasure: assumes "range B \ sets M" "incseq B" shows "incseq (\i. emeasure M (B i))" using assms by (auto simp: incseq_def intro!: emeasure_mono) lemma SUP_emeasure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A] by (simp add: LIMSEQ_unique) lemma decseq_emeasure: assumes "range B \ sets M" "decseq B" shows "decseq (\i. emeasure M (B i))" using assms by (auto simp: decseq_def intro!: emeasure_mono) lemma INF_emeasure_decseq: assumes A: "range A \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - have le_MI: "emeasure M (\i. A i) \ emeasure M (A 0)" using A by (auto intro!: emeasure_mono) hence *: "emeasure M (\i. A i) \ \" using finite[of 0] by (auto simp: top_unique) have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))" by (simp add: ennreal_INF_const_minus) also have "\ = (SUP n. emeasure M (A 0 - A n))" using A finite \decseq A\[unfolded decseq_def] by (subst emeasure_Diff) auto also have "\ = emeasure M (\i. A 0 - A i)" proof (rule SUP_emeasure_incseq) show "range (\n. A 0 - A n) \ sets M" using A by auto show "incseq (\n. A 0 - A n)" using \decseq A\ by (auto simp add: incseq_def decseq_def) qed also have "\ = emeasure M (A 0) - emeasure M (\i. A i)" using A finite * by (simp, subst emeasure_Diff) auto finally show ?thesis by (smt (verit, best) Inf_lower diff_diff_ennreal le_MI finite range_eqI) qed lemma INF_emeasure_decseq': assumes A: "\i. A i \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - from finite obtain i where i: "emeasure M (A i) < \" by (auto simp: less_top) have fin: "i \ j \ emeasure M (A j) < \" for j by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \decseq A\] A) have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))" proof (rule INF_eq) show "\j\UNIV. emeasure M (A (j + i)) \ emeasure M (A i')" for i' by (meson A \decseq A\ decseq_def emeasure_mono iso_tuple_UNIV_I nat_le_iff_add) qed auto also have "\ = emeasure M (INF n. (A (n + i)))" using A \decseq A\ fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top) also have "(INF n. (A (n + i))) = (INF n. A n)" by (meson INF_eq UNIV_I assms(2) decseqD le_add1) finally show ?thesis . qed lemma emeasure_INT_decseq_subset: fixes F :: "nat \ 'a set" assumes I: "I \ {}" and F: "\i j. i \ I \ j \ I \ i \ j \ F j \ F i" assumes F_sets[measurable]: "\i. i \ I \ F i \ sets M" and fin: "\i. i \ I \ emeasure M (F i) \ \" shows "emeasure M (\i\I. F i) = (INF i\I. emeasure M (F i))" proof cases assume "finite I" have "(\i\I. F i) = F (Max I)" using I \finite I\ by (intro antisym INF_lower INF_greatest F) auto moreover have "(INF i\I. emeasure M (F i)) = emeasure M (F (Max I))" using I \finite I\ by (intro antisym INF_lower INF_greatest F emeasure_mono) auto ultimately show ?thesis by simp next assume "infinite I" define L where "L n = (LEAST i. i \ I \ i \ n)" for n have L: "L n \ I \ n \ L n" for n unfolding L_def proof (rule LeastI_ex) show "\x. x \ I \ n \ x" using \infinite I\ finite_subset[of I "{..< n}"] by (rule_tac ccontr) (auto simp: not_le) qed have L_eq[simp]: "i \ I \ L i = i" for i unfolding L_def by (intro Least_equality) auto have L_mono: "i \ j \ L i \ L j" for i j using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def) have "emeasure M (\i. F (L i)) = (INF i. emeasure M (F (L i)))" proof (intro INF_emeasure_decseq[symmetric]) show "decseq (\i. F (L i))" using L by (intro antimonoI F L_mono) auto qed (insert L fin, auto) also have "\ = (INF i\I. emeasure M (F i))" proof (intro antisym INF_greatest) show "i \ I \ (INF i. emeasure M (F (L i))) \ emeasure M (F i)" for i by (intro INF_lower2[of i]) auto qed (insert L, auto intro: INF_lower) also have "(\i. F (L i)) = (\i\I. F i)" proof (intro antisym INF_greatest) show "i \ I \ (\i. F (L i)) \ F i" for i by (intro INF_lower2[of i]) auto qed (insert L, auto) finally show ?thesis . qed lemma Lim_emeasure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\i. emeasure M (A i)) \ emeasure M (\i. A i)" using LIMSEQ_INF[OF decseq_emeasure, OF A] using INF_emeasure_decseq[OF A fin] by simp lemma emeasure_lfp'[consumes 1, case_names cont measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M {x\space M. lfp F x} = (SUP i. emeasure M {x\space M. (F ^^ i) (\x. False) x})" proof - have "emeasure M {x\space M. lfp F x} = emeasure M (\i. {x\space M. (F ^^ i) (\x. False) x})" using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure]) moreover { fix i from \P M\ have "{x\space M. (F ^^ i) (\x. False) x} \ sets M" by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) } moreover have "incseq (\i. {x\space M. (F ^^ i) (\x. False) x})" proof (rule incseq_SucI) fix i have "(F ^^ i) (\x. False) \ (F ^^ (Suc i)) (\x. False)" proof (induct i) case 0 show ?case by (simp add: le_fun_def) next case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto qed then show "{x \ space M. (F ^^ i) (\x. False) x} \ {x \ space M. (F ^^ Suc i) (\x. False) x}" by auto qed ultimately show ?thesis by (subst SUP_emeasure_incseq) auto qed lemma emeasure_lfp: assumes [simp]: "\s. sets (M s) = sets N" assumes cont: "sup_continuous F" "sup_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ P \ lfp F \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" shows "emeasure (M s) {x\space N. lfp F x} = lfp f s" proof (subst lfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and f=F , symmetric]) fix C assume "incseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (SUP i. C i) x}) = (SUP i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding SUP_apply by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) qed (auto simp add: iter le_fun_def SUP_apply intro!: meas cont) lemma emeasure_subadditive_finite: "finite I \ A ` I \ sets M \ emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto lemma emeasure_subadditive: "A \ sets M \ B \ sets M \ emeasure M (A \ B) \ emeasure M A + emeasure M B" using emeasure_subadditive_finite[of "{True, False}" "\True \ A | False \ B" M] by simp lemma emeasure_subadditive_countably: assumes "range f \ sets M" shows "emeasure M (\i. f i) \ (\i. emeasure M (f i))" proof - have "emeasure M (\i. f i) = emeasure M (\i. disjointed f i)" unfolding UN_disjointed_eq .. also have "\ = (\i. emeasure M (disjointed f i))" using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"] by (simp add: disjoint_family_disjointed comp_def) also have "\ \ (\i. emeasure M (f i))" using sets.range_disjointed_sets[OF assms] assms by (auto intro!: suminf_le emeasure_mono disjointed_subset) finally show ?thesis . qed lemma emeasure_insert: assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A" proof - have "{x} \ A = {}" using \x \ A\ by auto from plus_emeasure[OF sets this] show ?thesis by simp qed lemma emeasure_insert_ne: "A \ {} \ {x} \ sets M \ A \ sets M \ x \ A \ emeasure M (insert x A) = emeasure M {x} + emeasure M A" by (rule emeasure_insert) lemma emeasure_eq_sum_singleton: assumes "finite S" "\x. x \ S \ {x} \ sets M" shows "emeasure M S = (\x\S. emeasure M {x})" using sum_emeasure[of "\x. {x}" S M] assms by (auto simp: disjoint_family_on_def subset_eq) lemma sum_emeasure_cover: assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" assumes A: "A \ (\i\S. B i)" assumes disj: "disjoint_family_on B S" shows "emeasure M A = (\i\S. emeasure M (A \ (B i)))" proof - have "(\i\S. emeasure M (A \ (B i))) = emeasure M (\i\S. A \ (B i))" proof (rule sum_emeasure) show "disjoint_family_on (\i. A \ B i) S" using \disjoint_family_on B S\ unfolding disjoint_family_on_def by auto qed (insert assms, auto) also have "(\i\S. A \ (B i)) = A" using A by auto finally show ?thesis by simp qed lemma emeasure_eq_0: "N \ sets M \ emeasure M N = 0 \ K \ N \ emeasure M K = 0" by (metis emeasure_mono order_eq_iff zero_le) lemma emeasure_UN_eq_0: assumes "\i::nat. emeasure M (N i) = 0" and "range N \ sets M" shows "emeasure M (\i. N i) = 0" proof - have "emeasure M (\i. N i) \ 0" using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp then show ?thesis by (auto intro: antisym zero_le) qed lemma measure_eqI_finite: assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A" assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" shows "M = N" proof (rule measure_eqI) fix X assume "X \ sets M" then have X: "X \ A" by auto then have "emeasure M X = (\a\X. emeasure M {a})" using \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) also have "\ = (\a\X. emeasure N {a})" using X eq by (auto intro!: sum.cong) also have "\ = emeasure N X" using X \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) finally show "emeasure M X = emeasure N X" . qed simp lemma measure_eqI_generator_eq: fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \ 'a set" assumes "Int_stable E" "E \ Pow \" and eq: "\X. X \ E \ emeasure M X = emeasure N X" and M: "sets M = sigma_sets \ E" and N: "sets N = sigma_sets \ E" and A: "range A \ E" "(\i. A i) = \" "\i. emeasure M (A i) \ \" shows "M = N" proof - let ?\ = "emeasure M" and ?\ = "emeasure N" interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact have "space M = \" using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \sets M = sigma_sets \ E\ by blast { fix F D assume "F \ E" and "?\ F \ \" then have [intro]: "F \ sigma_sets \ E" by auto have "?\ F \ \" using \?\ F \ \\ \F \ E\ eq by simp assume "D \ sets M" with \Int_stable E\ \E \ Pow \\ have "emeasure M (F \ D) = emeasure N (F \ D)" unfolding M proof (induct rule: sigma_sets_induct_disjoint) case (basic A) then have "F \ A \ E" using \Int_stable E\ \F \ E\ by (auto simp: Int_stable_def) then show ?case using eq by auto next case empty then show ?case by simp next case (compl A) then have **: "F \ (\ - A) = F - (F \ A)" and [intro]: "F \ A \ sigma_sets \ E" using \F \ E\ S.sets_into_space by (auto simp: M) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) then have "?\ (F \ (\ - A)) = ?\ F - ?\ (F \ A)" unfolding ** using \F \ A \ sigma_sets \ E\ by (auto intro!: emeasure_Diff simp: M N) also have "\ = ?\ F - ?\ (F \ A)" using eq \F \ E\ compl by simp also have "\ = ?\ (F \ (\ - A))" unfolding ** using \F \ A \ sigma_sets \ E\ \?\ (F \ A) \ \\ by (auto intro!: emeasure_Diff[symmetric] simp: M N) finally show ?case using \space M = \\ by auto next case (union A) then have "?\ (\x. F \ A x) = ?\ (\x. F \ A x)" by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N) with A show ?case by auto qed } note * = this show "M = N" proof (rule measure_eqI) show "sets M = sets N" using M N by simp have [simp, intro]: "\i. A i \ sets M" using A(1) by (auto simp: subset_eq M) fix F assume "F \ sets M" let ?D = "disjointed (\i. F \ A i)" from \space M = \\ have F_eq: "F = (\i. ?D i)" using \F \ sets M\[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) have [simp, intro]: "\i. ?D i \ sets M" using sets.range_disjointed_sets[of "\i. F \ A i" M] \F \ sets M\ by (auto simp: subset_eq) have "disjoint_family ?D" by (auto simp: disjoint_family_disjointed) moreover have "(\i. emeasure M (?D i)) = (\i. emeasure N (?D i))" proof (intro arg_cong[where f=suminf] ext) fix i have "A i \ ?D i = ?D i" by (auto simp: disjointed_def) then show "emeasure M (?D i) = emeasure N (?D i)" using *[of "A i" "?D i", OF _ A(3)] A(1) by auto qed ultimately show "emeasure M F = emeasure N F" by (simp add: image_subset_iff \sets M = sets N\[symmetric] F_eq[symmetric] suminf_emeasure) qed qed lemma space_empty: "space M = {} \ M = count_space {}" by (rule measure_eqI) (simp_all add: space_empty_iff) lemma measure_eqI_generator_eq_countable: fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set" assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" and sets: "sets M = sigma_sets \ E" "sets N = sigma_sets \ E" and A: "A \ E" "(\A) = \" "countable A" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof cases assume "\ = {}" have *: "sigma_sets \ E = sets (sigma \ E)" using E(2) by simp have "space M = \" "space N = \" using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of) then show "M = N" unfolding \\ = {}\ by (auto dest: space_empty) next assume "\ \ {}" with \\A = \\ have "A \ {}" by auto from this \countable A\ have rng: "range (from_nat_into A) = A" by (rule range_from_nat_into) show "M = N" proof (rule measure_eqI_generator_eq[OF E sets]) show "range (from_nat_into A) \ E" unfolding rng using \A \ E\ . show "(\i. from_nat_into A i) = \" unfolding rng using \\A = \\ . show "emeasure M (from_nat_into A i) \ \" for i using rng by (intro A) auto qed qed lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M" proof (intro measure_eqI emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) (emeasure M)" by (simp add: positive_def) show "countably_additive (sets M) (emeasure M)" by (simp add: emeasure_countably_additive) qed simp_all subsection \\\\-null sets\ definition\<^marker>\tag important\ null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" lemma null_setsD1[dest]: "A \ null_sets M \ emeasure M A = 0" by (simp add: null_sets_def) lemma null_setsD2[dest]: "A \ null_sets M \ A \ sets M" unfolding null_sets_def by simp lemma null_setsI[intro]: "emeasure M A = 0 \ A \ sets M \ A \ null_sets M" unfolding null_sets_def by simp interpretation null_sets: ring_of_sets "space M" "null_sets M" for M proof (rule ring_of_setsI) show "null_sets M \ Pow (space M)" using sets.sets_into_space by auto show "{} \ null_sets M" by auto fix A B assume null_sets: "A \ null_sets M" "B \ null_sets M" then have sets: "A \ sets M" "B \ sets M" by auto then have *: "emeasure M (A \ B) \ emeasure M A + emeasure M B" "emeasure M (A - B) \ emeasure M A" by (auto intro!: emeasure_subadditive emeasure_mono) then have "emeasure M B = 0" "emeasure M A = 0" using null_sets by auto with sets * show "A - B \ null_sets M" "A \ B \ null_sets M" by (auto intro!: antisym zero_le) qed lemma UN_from_nat_into: assumes I: "countable I" "I \ {}" shows "(\i\I. N i) = (\i. N (from_nat_into I i))" proof - have "(\i\I. N i) = \(N ` range (from_nat_into I))" using I by simp also have "\ = (\i. (N \ from_nat_into I) i)" by simp finally show ?thesis by simp qed lemma null_sets_UN': assumes "countable I" assumes "\i. i \ I \ N i \ null_sets M" shows "(\i\I. N i) \ null_sets M" proof cases assume "I = {}" then show ?thesis by simp next assume "I \ {}" show ?thesis proof (intro conjI CollectI null_setsI) show "(\i\I. N i) \ sets M" using assms by (intro sets.countable_UN') auto have "emeasure M (\i\I. N i) \ (\n. emeasure M (N (from_nat_into I n)))" unfolding UN_from_nat_into[OF \countable I\ \I \ {}\] using assms \I \ {}\ by (intro emeasure_subadditive_countably) (auto intro: from_nat_into) also have "(\n. emeasure M (N (from_nat_into I n))) = (\_. 0)" using assms \I \ {}\ by (auto intro: from_nat_into) finally show "emeasure M (\i\I. N i) = 0" by (intro antisym zero_le) simp qed qed lemma null_sets_UN[intro]: "(\i::'i::countable. N i \ null_sets M) \ (\i. N i) \ null_sets M" by (rule null_sets_UN') auto lemma null_set_Int1: assumes "B \ null_sets M" "A \ sets M" shows "A \ B \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (A \ B) = 0" using assms by (intro emeasure_eq_0[of B _ "A \ B"]) auto qed (insert assms, auto) lemma null_set_Int2: assumes "B \ null_sets M" "A \ sets M" shows "B \ A \ null_sets M" using assms by (subst Int_commute) (rule null_set_Int1) lemma emeasure_Diff_null_set: assumes "B \ null_sets M" "A \ sets M" shows "emeasure M (A - B) = emeasure M A" proof - have *: "A - B = (A - (A \ B))" by auto have "A \ B \ null_sets M" using assms by (rule null_set_Int1) then show ?thesis unfolding * using assms by (subst emeasure_Diff) auto qed lemma null_set_Diff: assumes "B \ null_sets M" "A \ sets M" shows "B - A \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto qed (insert assms, auto) lemma emeasure_Un_null_set: assumes "A \ sets M" "B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A" proof - have *: "A \ B = A \ (B - A)" by auto have "B - A \ null_sets M" using assms(2,1) by (rule null_set_Diff) then show ?thesis unfolding * using assms by (subst plus_emeasure[symmetric]) auto qed lemma emeasure_Un': assumes "A \ sets M" "B \ sets M" "A \ B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A + emeasure M B" proof - have "A \ B = A \ (B - A \ B)" by blast also have "emeasure M \ = emeasure M A + emeasure M (B - A \ B)" using assms by (subst plus_emeasure) auto also have "emeasure M (B - A \ B) = emeasure M B" using assms by (intro emeasure_Diff_null_set) auto finally show ?thesis . qed subsection \The almost everywhere filter (i.e.\ quantifier)\ definition\<^marker>\tag important\ ae_filter :: "'a measure \ 'a filter" where "ae_filter M = (INF N\null_sets M. principal (space M - N))" abbreviation almost_everywhere :: "'a measure \ ('a \ bool) \ bool" where "almost_everywhere M P \ eventually P (ae_filter M)" syntax "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) translations "AE x in M. P" \ "CONST almost_everywhere M (\x. P)" abbreviation "set_almost_everywhere A M P \ AE x in M. x \ A \ P x" syntax "_set_almost_everywhere" :: "pttrn \ 'a set \ 'a \ bool \ bool" ("AE _\_ in _./ _" [0,0,0,10] 10) translations "AE x\A in M. P" \ "CONST set_almost_everywhere A M (\x. P)" lemma eventually_ae_filter: "eventually P (ae_filter M) \ (\N\null_sets M. {x \ space M. \ P x} \ N)" unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq) lemma AE_I': "N \ null_sets M \ {x\space M. \ P x} \ N \ (AE x in M. P x)" unfolding eventually_ae_filter by auto lemma AE_iff_null: assumes "{x\space M. \ P x} \ sets M" (is "?P \ sets M") shows "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" proof assume "AE x in M. P x" then obtain N where N: "N \ sets M" "?P \ N" "emeasure M N = 0" unfolding eventually_ae_filter by auto have "emeasure M ?P \ emeasure M N" using assms N(1,2) by (auto intro: emeasure_mono) then have "emeasure M ?P = 0" unfolding \emeasure M N = 0\ by auto then show "?P \ null_sets M" using assms by auto next assume "?P \ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') qed lemma AE_iff_null_sets: "N \ sets M \ N \ null_sets M \ (AE x in M. x \ N)" using Int_absorb1[OF sets.sets_into_space, of N M] by (subst AE_iff_null) (auto simp: Int_def[symmetric]) lemma ae_filter_eq_bot_iff: "ae_filter M = bot \ emeasure M (space M) = 0" proof - have "ae_filter M = bot \ (AE x in M. False)" using trivial_limit_def by blast also have "\ \ space M \ null_sets M" by (simp add: AE_iff_null_sets eventually_ae_filter) also have "\ \ emeasure M (space M) = 0" by auto finally show ?thesis . qed lemma AE_not_in: "N \ null_sets M \ AE x in M. x \ N" by (metis AE_iff_null_sets null_setsD2) lemma AE_iff_measurable: "N \ sets M \ {x\space M. \ P x} = N \ (AE x in M. P x) \ emeasure M N = 0" using AE_iff_null[of _ P] by auto lemma AE_E[consumes 1]: assumes "AE x in M. P x" obtains N where "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" using assms unfolding eventually_ae_filter by auto lemma AE_E2: assumes "AE x in M. P x" shows "emeasure M {x\space M. \ P x} = 0" by (metis (mono_tags, lifting) AE_iff_null assms emeasure_notin_sets null_setsD1) lemma AE_E3: assumes "AE x in M. P x" obtains N where "\x. x \ space M - N \ P x" "N \ null_sets M" using assms unfolding eventually_ae_filter by auto lemma AE_I: assumes "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" shows "AE x in M. P x" using assms unfolding eventually_ae_filter by auto lemma AE_mp[elim!]: assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \ Q x" shows "AE x in M. Q x" using assms by (fact eventually_rev_mp) text \The next lemma is convenient to combine with a lemma whose conclusion is of the form \AE x in M. P x = Q x\: for such a lemma, there is no \[symmetric]\ variant, but using \AE_symmetric[OF\]\ will replace it.\ (* depricated replace by laws about eventually *) lemma shows AE_iffI: "AE x in M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" and AE_disjI1: "AE x in M. P x \ AE x in M. P x \ Q x" and AE_disjI2: "AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conjI: "AE x in M. P x \ AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conj_iff[simp]: "(AE x in M. P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_symmetric: assumes "AE x in M. P x = Q x" shows "AE x in M. Q x = P x" using assms by auto lemma AE_impI: "(P \ AE x in M. Q x) \ AE x in M. P \ Q x" by fastforce lemma AE_measure: assumes AE: "AE x in M. P x" and sets: "{x\space M. P x} \ sets M" (is "?P \ sets M") shows "emeasure M {x\space M. P x} = emeasure M (space M)" proof - from AE_E[OF AE] obtain N where N: "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" by auto with sets have "emeasure M (space M) \ emeasure M (?P \ N)" by (intro emeasure_mono) auto also have "\ \ emeasure M ?P + emeasure M N" using sets N by (intro emeasure_subadditive) auto also have "\ = emeasure M ?P" using N by simp finally show "emeasure M ?P = emeasure M (space M)" using emeasure_space[of M "?P"] by auto qed lemma AE_space: "AE x in M. x \ space M" by (rule AE_I[where N="{}"]) auto lemma AE_I2[simp, intro]: "(\x. x \ space M \ P x) \ AE x in M. P x" using AE_space by force lemma AE_Ball_mp: "\x\space M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" by auto lemma AE_cong[cong]: "(\x. x \ space M \ P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_cong_simp: "M = N \ (\x. x \ space N =simp=> P x = Q x) \ (AE x in M. P x) \ (AE x in N. Q x)" by (auto simp: simp_implies_def) lemma AE_all_countable: "(AE x in M. \i. P i x) \ (\i::'i::countable. AE x in M. P i x)" proof assume "\i. AE x in M. P i x" from this[unfolded eventually_ae_filter Bex_def, THEN choice] obtain N where N: "\i. N i \ null_sets M" "\i. {x\space M. \ P i x} \ N i" by auto have "{x\space M. \ (\i. P i x)} \ (\i. {x\space M. \ P i x})" by auto also have "\ \ (\i. N i)" using N by auto finally have "{x\space M. \ (\i. P i x)} \ (\i. N i)" . moreover from N have "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto ultimately show "AE x in M. \i. P i x" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable: assumes [intro]: "countable X" shows "(AE x in M. \y\X. P x y) \ (\y\X. AE x in M. P x y)" proof assume "\y\X. AE x in M. P x y" from this[unfolded eventually_ae_filter Bex_def, THEN bchoice] obtain N where N: "\y. y \ X \ N y \ null_sets M" "\y. y \ X \ {x\space M. \ P x y} \ N y" by auto have "{x\space M. \ (\y\X. P x y)} \ (\y\X. {x\space M. \ P x y})" by auto also have "\ \ (\y\X. N y)" using N by auto finally have "{x\space M. \ (\y\X. P x y)} \ (\y\X. N y)" . moreover from N have "(\y\X. N y) \ null_sets M" by (intro null_sets_UN') auto ultimately show "AE x in M. \y\X. P x y" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable': "(\N. N \ I \ AE x in M. P N x) \ countable I \ AE x in M. \N \ I. P N x" unfolding AE_ball_countable by simp lemma AE_pairwise: "countable F \ pairwise (\A B. AE x in M. R x A B) F \ (AE x in M. pairwise (R x) F)" unfolding pairwise_alt by (simp add: AE_ball_countable) lemma AE_discrete_difference: assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" shows "AE x in M. x \ X" proof - have "(\x\X. {x}) \ null_sets M" using assms by (intro null_sets_UN') auto from AE_not_in[OF this] show "AE x in M. x \ X" by auto qed lemma AE_finite_all: assumes f: "finite S" shows "(AE x in M. \i\S. P i x) \ (\i\S. AE x in M. P i x)" using f by induct auto lemma AE_finite_allI: assumes "finite S" shows "(\s. s \ S \ AE x in M. Q s x) \ AE x in M. \s\S. Q s x" using AE_finite_all[OF \finite S\] by auto lemma emeasure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "emeasure M A \ emeasure M B" proof cases assume A: "A \ sets M" from imp obtain N where N: "{x\space M. \ (x \ A \ x \ B)} \ N" "N \ null_sets M" by (auto simp: eventually_ae_filter) have "emeasure M A = emeasure M (A - N)" using N A by (subst emeasure_Diff_null_set) auto also have "emeasure M (A - N) \ emeasure M (B - N)" using N A B sets.sets_into_space by (auto intro!: emeasure_mono) also have "emeasure M (B - N) = emeasure M B" using N B by (subst emeasure_Diff_null_set) auto finally show ?thesis . qed (simp add: emeasure_notin_sets) lemma emeasure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "emeasure M A = emeasure M B" using assms by (safe intro!: antisym emeasure_mono_AE) auto lemma emeasure_Collect_eq_AE: "AE x in M. P x \ Q x \ Measurable.pred M Q \ Measurable.pred M P \ emeasure M {x\space M. P x} = emeasure M {x\space M. Q x}" by (intro emeasure_eq_AE) auto lemma emeasure_eq_0_AE: "AE x in M. \ P x \ emeasure M {x\space M. P x} = 0" using AE_iff_measurable[OF _ refl, of M "\x. \ P x"] by (cases "{x\space M. P x} \ sets M") (simp_all add: emeasure_notin_sets) lemma emeasure_0_AE: assumes "emeasure M (space M) = 0" shows "AE x in M. P x" using eventually_ae_filter assms by blast lemma emeasure_add_AE: assumes [measurable]: "A \ sets M" "B \ sets M" "C \ sets M" assumes 1: "AE x in M. x \ C \ x \ A \ x \ B" assumes 2: "AE x in M. \ (x \ A \ x \ B)" shows "emeasure M C = emeasure M A + emeasure M B" proof - have "emeasure M C = emeasure M (A \ B)" by (rule emeasure_eq_AE) (insert 1, auto) also have "\ = emeasure M A + emeasure M (B - A)" by (subst plus_emeasure) auto also have "emeasure M (B - A) = emeasure M B" by (rule emeasure_eq_AE) (insert 2, auto) finally show ?thesis . qed subsection \\\\-finite Measures\ locale\<^marker>\tag important\ sigma_finite_measure = fixes M :: "'a measure" assumes sigma_finite_countable: "\A::'a set set. countable A \ A \ sets M \ (\A) = space M \ (\a\A. emeasure M a \ \)" lemma (in sigma_finite_measure) sigma_finite: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" proof - obtain A :: "'a set set" where [simp]: "countable A" and A: "A \ sets M" "(\A) = space M" "\a. a \ A \ emeasure M a \ \" using sigma_finite_countable by metis show thesis proof cases assume "A = {}" with \(\A) = space M\ show thesis by (intro that[of "\_. {}"]) auto next assume "A \ {}" show thesis proof show "range (from_nat_into A) \ sets M" using \A \ {}\ A by auto have "(\i. from_nat_into A i) = \A" using range_from_nat_into[OF \A \ {}\ \countable A\] by auto with A show "(\i. from_nat_into A i) = space M" by auto qed (intro A from_nat_into \A \ {}\) qed qed lemma (in sigma_finite_measure) sigma_finite_disjoint: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "disjoint_family A" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and space: "(\i. A i) = space M" and measure: "\i. emeasure M (A i) \ \" using sigma_finite by blast show thesis proof (rule that[of "disjointed A"]) show "range (disjointed A) \ sets M" by (rule sets.range_disjointed_sets[OF range]) show "(\i. disjointed A i) = space M" and "disjoint_family (disjointed A)" using disjoint_family_disjointed UN_disjointed_eq[of A] space range by auto show "emeasure M (disjointed A i) \ \" for i proof - have "emeasure M (disjointed A i) \ emeasure M (A i)" using range disjointed_subset[of A i] by (auto intro!: emeasure_mono) then show ?thesis using measure[of i] by (auto simp: top_unique) qed qed qed lemma (in sigma_finite_measure) sigma_finite_incseq: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" proof - obtain F :: "nat \ 'a set" where F: "range F \ sets M" "(\i. F i) = space M" "\i. emeasure M (F i) \ \" using sigma_finite by blast show thesis proof (rule that[of "\n. \i\n. F i"]) show "range (\n. \i\n. F i) \ sets M" using F by (force simp: incseq_def) show "(\n. \i\n. F i) = space M" proof - from F have "\x. x \ space M \ \i. x \ F i" by auto with F show ?thesis by fastforce qed show "emeasure M (\i\n. F i) \ \" for n proof - have "emeasure M (\i\n. F i) \ (\i\n. emeasure M (F i))" using F by (auto intro!: emeasure_subadditive_finite) also have "\ < \" using F by (auto simp: sum_Pinfty less_top) finally show ?thesis by simp qed show "incseq (\n. \i\n. F i)" by (force simp: incseq_def) qed qed lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite: fixes C::real assumes W_meas: "W \ sets M" and W_inf: "emeasure M W = \" obtains Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" proof - obtain A :: "nat \ 'a set" where A: "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" using sigma_finite_incseq by blast define B where "B = (\i. W \ A i)" have B_meas: "\i. B i \ sets M" using W_meas \range A \ sets M\ B_def by blast have b: "\i. B i \ W" using B_def by blast { fix i have "emeasure M (B i) \ emeasure M (A i)" using A by (intro emeasure_mono) (auto simp: B_def) also have "emeasure M (A i) < \" using \\i. emeasure M (A i) \ \\ by (simp add: less_top) finally have "emeasure M (B i) < \" . } note c = this have "W = (\i. B i)" using B_def \(\i. A i) = space M\ W_meas by auto moreover have "incseq B" using B_def \incseq A\ by (simp add: incseq_def subset_eq) ultimately have "(\i. emeasure M (B i)) \ emeasure M W" using W_meas B_meas by (simp add: B_meas Lim_emeasure_incseq image_subset_iff) then have "(\i. emeasure M (B i)) \ \" using W_inf by simp from order_tendstoD(1)[OF this, of C] obtain i where d: "emeasure M (B i) > C" by (auto simp: eventually_sequentially) have "B i \ sets M" "B i \ W" "emeasure M (B i) < \" "emeasure M (B i) > C" using B_meas b c d by auto then show ?thesis using that by blast qed subsection \Measure space induced by distribution of \<^const>\measurable\-functions\ definition\<^marker>\tag important\ distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" lemma shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N" and space_distr[simp]: "space (distr M N f) = space N" by (auto simp: distr_def) lemma shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'" and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng" by (auto simp: measurable_def) lemma distr_cong: "M = K \ sets N = sets L \ (\x. x \ space M \ f x = g x) \ distr M N f = distr K L g" using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong) lemma emeasure_distr: fixes f :: "'a \ 'b" assumes f: "f \ measurable M N" and A: "A \ sets N" shows "emeasure (distr M N f) A = emeasure M (f -` A \ space M)" (is "_ = ?\ A") unfolding distr_def proof (rule emeasure_measure_of_sigma) show "positive (sets N) ?\" by (auto simp: positive_def) show "countably_additive (sets N) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'b set" assume "range A \ sets N" "disjoint_family A" then have A: "\i. A i \ sets N" "(\i. A i) \ sets N" by auto then have *: "range (\i. f -` (A i) \ space M) \ sets M" using f by (auto simp: measurable_def) moreover have "(\i. f -` A i \ space M) \ sets M" using * by blast moreover have **: "disjoint_family (\i. f -` A i \ space M)" using \disjoint_family A\ by (auto simp: disjoint_family_on_def) ultimately show "(\i. ?\ (A i)) = ?\ (\i. A i)" using suminf_emeasure[OF _ **] A f by (auto simp: comp_def vimage_UN) qed show "sigma_algebra (space N) (sets N)" .. qed fact lemma emeasure_Collect_distr: assumes X[measurable]: "X \ measurable M N" "Measurable.pred N P" shows "emeasure (distr M N X) {x\space N. P x} = emeasure M {x\space M. P (X x)}" by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space]) lemma emeasure_lfp2[consumes 1, case_names cont f measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes f: "\M. P M \ f \ measurable M' M" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M' {x\space M'. lfp F (f x)} = (SUP i. emeasure M' {x\space M'. (F ^^ i) (\x. False) (f x)})" proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f]) show "f \ measurable M' M" "f \ measurable M' M" using f[OF \P M\] by auto { fix i show "Measurable.pred M ((F ^^ i) (\x. False))" using \P M\ by (induction i arbitrary: M) (auto intro!: *) } show "Measurable.pred M (lfp F)" using \P M\ cont * by (rule measurable_lfp_coinduct[of P]) have "emeasure (distr M' M f) {x \ space (distr M' M f). lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space (distr M' M f). (F ^^ i) (\x. False) x})" using \P M\ proof (coinduction arbitrary: M rule: emeasure_lfp') case (measurable A N) then have "\N. P N \ Measurable.pred (distr M' N f) A" by metis then have "\N. P N \ Measurable.pred N A" by simp with \P N\[THEN *] show ?case by auto qed fact then show "emeasure (distr M' M f) {x \ space M. lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space M. (F ^^ i) (\x. False) x})" by simp qed lemma distr_id[simp]: "distr N N (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma distr_id2: "sets M = sets N \ distr N M (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma measure_distr: "f \ measurable M N \ S \ sets N \ measure (distr M N f) S = measure M (f -` S \ space M)" by (simp add: emeasure_distr measure_def) lemma distr_cong_AE: assumes 1: "M = K" "sets N = sets L" and 2: "(AE x in M. f x = g x)" and "f \ measurable M N" and "g \ measurable K L" shows "distr M N f = distr K L g" proof (rule measure_eqI) fix A assume "A \ sets (distr M N f)" with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A" by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets) qed (insert 1, simp) lemma AE_distrD: assumes f: "f \ measurable M M'" and AE: "AE x in distr M M' f. P x" shows "AE x in M. P (f x)" proof - from AE[THEN AE_E] obtain N where "{x \ space (distr M M' f). \ P x} \ N" "emeasure (distr M M' f) N = 0" "N \ sets (distr M M' f)" by auto with f show ?thesis by (simp add: eventually_ae_filter, intro bexI[of _ "f -` N \ space M"]) (auto simp: emeasure_distr measurable_def) qed lemma AE_distr_iff: assumes f[measurable]: "f \ measurable M N" and P[measurable]: "{x \ space N. P x} \ sets N" shows "(AE x in distr M N f. P x) \ (AE x in M. P (f x))" proof (subst (1 2) AE_iff_measurable[OF _ refl]) have "f -` {x\space N. \ P x} \ space M = {x \ space M. \ P (f x)}" using f[THEN measurable_space] by auto then show "(emeasure (distr M N f) {x \ space (distr M N f). \ P x} = 0) = (emeasure M {x \ space M. \ P (f x)} = 0)" by (simp add: emeasure_distr) qed auto lemma null_sets_distr_iff: "f \ measurable M N \ A \ null_sets (distr M N f) \ f -` A \ space M \ null_sets M \ A \ sets N" by (auto simp add: null_sets_def emeasure_distr) proposition distr_distr: "g \ measurable N L \ f \ measurable M N \ distr (distr M N f) L g = distr M L (g \ f)" by (auto simp add: emeasure_distr measurable_space intro!: arg_cong[where f="emeasure M"] measure_eqI) subsection\<^marker>\tag unimportant\ \Real measure values\ lemma ring_of_finite_sets: "ring_of_sets (space M) {A\sets M. emeasure M A \ top}" proof (rule ring_of_setsI) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a \ b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_subadditive[of a M b] by (auto simp: top_unique) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a - b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_mono[of "a - b" a M] by (auto simp: top_unique) qed (auto dest: sets.sets_into_space) lemma measure_nonneg[simp]: "0 \ measure M A" unfolding measure_def by auto lemma measure_nonneg' [simp]: "\ measure M A < 0" using measure_nonneg not_le by blast lemma zero_less_measure_iff: "0 < measure M A \ measure M A \ 0" using measure_nonneg[of M A] by (auto simp add: le_less) lemma measure_le_0_iff: "measure M X \ 0 \ measure M X = 0" using measure_nonneg[of M X] by linarith lemma measure_empty[simp]: "measure M {} = 0" unfolding measure_def by (simp add: zero_ennreal.rep_eq) lemma emeasure_eq_ennreal_measure: "emeasure M A \ top \ emeasure M A = ennreal (measure M A)" by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def) lemma measure_zero_top: "emeasure M A = top \ measure M A = 0" by (simp add: measure_def) lemma measure_eq_emeasure_eq_ennreal: "0 \ x \ emeasure M A = ennreal x \ measure M A = x" using emeasure_eq_ennreal_measure[of M A] by (cases "A \ M") (auto simp: measure_notin_sets emeasure_notin_sets) lemma enn2real_plus:"a < top \ b < top \ enn2real (a + b) = enn2real a + enn2real b" by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top del: real_of_ereal_enn2ereal) lemma enn2real_sum:"(\i. i \ I \ f i < top) \ enn2real (sum f I) = sum (enn2real \ f) I" by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus) lemma measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: measure_def) lemma measure_Union: "emeasure M A \ \ \ emeasure M B \ \ \ A \ sets M \ B \ sets M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top) lemma measure_finite_Union: "finite S \ A`S \ sets M \ disjoint_family_on A S \ (\i. i \ S \ emeasure M (A i) \ \) \ measure M (\i\S. A i) = (\i\S. measure M (A i))" by (induction S rule: finite_induct) (auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite]) lemma measure_Diff: assumes finite: "emeasure M A \ \" and measurable: "A \ sets M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" proof - have "emeasure M (A - B) \ emeasure M A" "emeasure M B \ emeasure M A" using measurable by (auto intro!: emeasure_mono) hence "measure M ((A - B) \ B) = measure M (A - B) + measure M B" using measurable finite by (rule_tac measure_Union) (auto simp: top_unique) thus ?thesis using \B \ A\ by (auto simp: Un_absorb2) qed lemma measure_UNION: assumes measurable: "range A \ sets M" "disjoint_family A" assumes finite: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" proof - have "(\i. emeasure M (A i)) sums (emeasure M (\i. A i))" unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums) moreover { fix i have "emeasure M (A i) \ emeasure M (\i. A i)" using measurable by (auto intro!: emeasure_mono) then have "emeasure M (A i) = ennreal ((measure M (A i)))" using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) } ultimately show ?thesis using finite by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all qed lemma measure_subadditive: assumes measurable: "A \ sets M" "B \ sets M" and fin: "emeasure M A \ \" "emeasure M B \ \" shows "measure M (A \ B) \ measure M A + measure M B" proof - have "emeasure M (A \ B) \ \" using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique) then show "(measure M (A \ B)) \ (measure M A) + (measure M B)" unfolding measure_def by (metis emeasure_subadditive[OF measurable] fin enn2real_mono enn2real_plus ennreal_add_less_top infinity_ennreal_def less_top) qed lemma measure_subadditive_finite: assumes A: "finite I" "A`I \ sets M" and fin: "\i. i \ I \ emeasure M (A i) \ \" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" proof - { have "emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" using emeasure_subadditive_finite[OF A] . also have "\ < \" using fin by (simp add: less_top A) finally have "emeasure M (\i\I. A i) \ top" by simp } note * = this show ?thesis using emeasure_subadditive_finite[OF A] fin unfolding emeasure_eq_ennreal_measure[OF *] by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure) qed lemma measure_subadditive_countably: assumes A: "range A \ sets M" and fin: "(\i. emeasure M (A i)) \ \" shows "measure M (\i. A i) \ (\i. measure M (A i))" proof - have **: "\i. emeasure M (A i) \ top" using fin ennreal_suminf_lessD[of "\i. emeasure M (A i)"] by (simp add: less_top) have ge0: "(\i. Sigma_Algebra.measure M (A i)) \ 0" using fin emeasure_eq_ennreal_measure[OF **] by (metis infinity_ennreal_def measure_nonneg suminf_cong suminf_nonneg summable_suminf_not_top) have "emeasure M (\i. A i) \ top" by (metis A emeasure_subadditive_countably fin infinity_ennreal_def neq_top_trans) then have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" by (rule emeasure_eq_ennreal_measure[symmetric]) also have "\ \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ = ennreal (\i. measure M (A i))" using fin unfolding emeasure_eq_ennreal_measure[OF **] by (subst suminf_ennreal) (auto simp: **) finally show ?thesis using ge0 ennreal_le_iff by blast qed lemma measure_Un_null_set: "A \ sets M \ B \ null_sets M \ measure M (A \ B) = measure M A" by (simp add: measure_def emeasure_Un_null_set) lemma measure_Diff_null_set: "A \ sets M \ B \ null_sets M \ measure M (A - B) = measure M A" by (simp add: measure_def emeasure_Diff_null_set) lemma measure_eq_sum_singleton: "finite S \ (\x. x \ S \ {x} \ sets M) \ (\x. x \ S \ emeasure M {x} \ \) \ measure M S = (\x\S. measure M {x})" using emeasure_eq_sum_singleton[of S M] by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure) lemma Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" and fin: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin by (auto simp: emeasure_eq_ennreal_measure) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using assms emeasure_mono[of "A _" "\i. A i" M] by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans) ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using A by (auto intro!: Lim_emeasure_incseq) qed auto lemma Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\n. measure M (A n)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin[of 0] A emeasure_mono[of "\i. A i" "A 0" M] by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using fin A by (auto intro!: Lim_emeasure_decseq) qed auto subsection \Set of measurable sets with finite measure\ definition\<^marker>\tag important\ fmeasurable :: "'a measure \ 'a set set" where "fmeasurable M = {A\sets M. emeasure M A < \}" lemma fmeasurableD[dest, measurable_dest]: "A \ fmeasurable M \ A \ sets M" by (auto simp: fmeasurable_def) lemma fmeasurableD2: "A \ fmeasurable M \ emeasure M A \ top" by (auto simp: fmeasurable_def) lemma fmeasurableI: "A \ sets M \ emeasure M A < \ \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI_null_sets: "A \ null_sets M \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI2: "A \ fmeasurable M \ B \ A \ B \ sets M \ B \ fmeasurable M" using emeasure_mono[of B A M] by (auto simp: fmeasurable_def) lemma measure_mono_fmeasurable: "A \ B \ A \ sets M \ B \ fmeasurable M \ measure M A \ measure M B" by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono) lemma emeasure_eq_measure2: "A \ fmeasurable M \ emeasure M A = measure M A" by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top) interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M" proof (rule ring_of_setsI) show "fmeasurable M \ Pow (space M)" "{} \ fmeasurable M" by (auto simp: fmeasurable_def dest: sets.sets_into_space) fix a b assume *: "a \ fmeasurable M" "b \ fmeasurable M" then have "emeasure M (a \ b) \ emeasure M a + emeasure M b" by (intro emeasure_subadditive) auto also have "\ < top" using * by (auto simp: fmeasurable_def) finally show "a \ b \ fmeasurable M" using * by (auto intro: fmeasurableI) show "a - b \ fmeasurable M" using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def) qed subsection\<^marker>\tag unimportant\\Measurable sets formed by unions and intersections\ lemma fmeasurable_Diff: "A \ fmeasurable M \ B \ sets M \ A - B \ fmeasurable M" using fmeasurableI2[of A M "A - B"] by auto lemma fmeasurable_Int_fmeasurable: "\S \ fmeasurable M; T \ sets M\ \ (S \ T) \ fmeasurable M" by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int) lemma fmeasurable_UN: assumes "countable I" "\i. i \ I \ F i \ A" "\i. i \ I \ F i \ sets M" "A \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "A \ fmeasurable M" "(\i\I. F i) \ A" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_UN') auto qed lemma fmeasurable_INT: assumes "countable I" "i \ I" "\i. i \ I \ F i \ sets M" "F i \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "F i \ fmeasurable M" "(\i\I. F i) \ F i" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_INT') auto qed lemma measurable_measure_Diff: assumes "A \ fmeasurable M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff) lemma measurable_Un_null_set: assumes "B \ null_sets M" shows "(A \ B \ fmeasurable M \ A \ sets M) \ A \ fmeasurable M" using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2) lemma measurable_Diff_null_set: assumes "B \ null_sets M" shows "(A - B) \ fmeasurable M \ A \ sets M \ A \ fmeasurable M" using assms by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set) lemma fmeasurable_Diff_D: assumes m: "T - S \ fmeasurable M" "S \ fmeasurable M" and sub: "S \ T" shows "T \ fmeasurable M" proof - have "T = S \ (T - S)" using assms by blast then show ?thesis by (metis m fmeasurable.Un) qed lemma measure_Un2: "A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M (B - A)" using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff) lemma measure_Un3: assumes "A \ fmeasurable M" "B \ fmeasurable M" shows "measure M (A \ B) = measure M A + measure M B - measure M (A \ B)" proof - have "measure M (A \ B) = measure M A + measure M (B - A)" using assms by (rule measure_Un2) also have "B - A = B - (A \ B)" by auto also have "measure M (B - (A \ B)) = measure M B - measure M (A \ B)" using assms by (intro measure_Diff) (auto simp: fmeasurable_def) finally show ?thesis by simp qed lemma measure_Un_AE: "AE x in M. x \ A \ x \ B \ A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M B" by (subst measure_Un2) (auto intro!: measure_eq_AE) lemma measure_UNION_AE: assumes I: "finite I" shows "(\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. AE x in M. x \ F i \ x \ F j) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" unfolding AE_pairwise[OF countable_finite, OF I] using I proof (induction I rule: finite_induct) case (insert x I) have "measure M (F x \ \(F ` I)) = measure M (F x) + measure M (\(F ` I))" by (rule measure_Un_AE) (use insert in \auto simp: pairwise_insert\) with insert show ?case by (simp add: pairwise_insert ) qed simp lemma measure_UNION': "finite I \ (\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. disjnt (F i) (F j)) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually) lemma measure_Union_AE: "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise (\S T. AE x in M. x \ S \ x \ T) F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION_AE[of F "\x. x" M] by simp lemma measure_Union': "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise disjnt F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION'[of F "\x. x" M] by simp lemma measure_Un_le: assumes "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" proof cases assume "A \ fmeasurable M \ B \ fmeasurable M" with measure_subadditive[of A M B] assms show ?thesis by (auto simp: fmeasurableD2) next assume "\ (A \ fmeasurable M \ B \ fmeasurable M)" then have "A \ B \ fmeasurable M" using fmeasurableI2[of "A \ B" M A] fmeasurableI2[of "A \ B" M B] assms by auto with assms show ?thesis by (auto simp: fmeasurable_def measure_def less_top[symmetric]) qed lemma measure_UNION_le: "finite I \ (\i. i \ I \ F i \ sets M) \ measure M (\i\I. F i) \ (\i\I. measure M (F i))" proof (induction I rule: finite_induct) case (insert i I) then have "measure M (\i\insert i I. F i) = measure M (F i \ \ (F ` I))" by simp also from insert have "measure M (F i \ \ (F ` I)) \ measure M (F i) + measure M (\ (F ` I))" by (intro measure_Un_le sets.finite_Union) auto also have "measure M (\i\I. F i) \ (\i\I. measure M (F i))" using insert by auto finally show ?case using insert by simp qed simp lemma measure_Union_le: "finite F \ (\S. S \ F \ S \ sets M) \ measure M (\F) \ (\S\F. measure M S)" using measure_UNION_le[of F "\x. x" M] by simp text\Version for indexed union over a countable set\ lemma assumes "countable I" and I: "\i. i \ I \ A i \ fmeasurable M" and bound: "\I'. I' \ I \ finite I' \ measure M (\i\I'. A i) \ B" shows fmeasurable_UN_bound: "(\i\I. A i) \ fmeasurable M" (is ?fm) and measure_UN_bound: "measure M (\i\I. A i) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof cases assume "I = {}" with \B \ 0\ show ?thesis by simp next assume "I \ {}" have "(\i\I. A i) = (\i. (\n\i. A (from_nat_into I n)))" by (subst range_from_nat_into[symmetric, OF \I \ {}\ \countable I\]) auto then have "emeasure M (\i\I. A i) = emeasure M (\i. (\n\i. A (from_nat_into I n)))" by simp also have "\ = (SUP i. emeasure M (\n\i. A (from_nat_into I n)))" using I \I \ {}\[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+ also have "\ \ B" proof (intro SUP_least) fix i :: nat have "emeasure M (\n\i. A (from_nat_into I n)) = measure M (\n\i. A (from_nat_into I n))" using I \I \ {}\[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto also have "\ = measure M (\n\from_nat_into I ` {..i}. A n)" by simp also have "\ \ B" by (intro ennreal_leI bound) (auto intro: from_nat_into[OF \I \ {}\]) finally show "emeasure M (\n\i. A (from_nat_into I n)) \ ennreal B" . qed finally have *: "emeasure M (\i\I. A i) \ B" . then have ?fm using I \countable I\ by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique) with * \0\B\ show ?thesis by (simp add: emeasure_eq_measure2) qed then show ?fm ?m by auto qed text\Version for big union of a countable set\ lemma assumes "countable \" and meas: "\D. D \ \ \ D \ fmeasurable M" and bound: "\\. \\ \ \; finite \\ \ measure M (\\) \ B" shows fmeasurable_Union_bound: "\\ \ fmeasurable M" (is ?fm) and measure_Union_bound: "measure M (\\) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof (cases "\ = {}") case True with \B \ 0\ show ?thesis by auto next case False then obtain D :: "nat \ 'a set" where D: "\ = range D" using \countable \\ uncountable_def by force have 1: "\i. D i \ fmeasurable M" by (simp add: D meas) have 2: "\I'. finite I' \ measure M (\x\I'. D x) \ B" by (simp add: D bound image_subset_iff) show ?thesis unfolding D by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto qed then show ?fm ?m by auto qed text\Version for indexed union over the type of naturals\ lemma fixes S :: "nat \ 'a set" assumes S: "\i. S i \ fmeasurable M" and B: "\n. measure M (\i\n. S i) \ B" shows fmeasurable_countable_Union: "(\i. S i) \ fmeasurable M" and measure_countable_Union_le: "measure M (\i. S i) \ B" proof - have mB: "measure M (\i\I. S i) \ B" if "finite I" for I proof - have "(\i\I. S i) \ (\i\Max I. S i)" using Max_ge that by force then have "measure M (\i\I. S i) \ measure M (\i \ Max I. S i)" by (rule measure_mono_fmeasurable) (use S in \blast+\) then show ?thesis using B order_trans by blast qed show "(\i. S i) \ fmeasurable M" by (auto intro: fmeasurable_UN_bound [OF _ S mB]) show "measure M (\n. S n) \ B" by (auto intro: measure_UN_bound [OF _ S mB]) qed lemma measure_diff_le_measure_setdiff: assumes "S \ fmeasurable M" "T \ fmeasurable M" shows "measure M S - measure M T \ measure M (S - T)" proof - have "measure M S \ measure M ((S - T) \ T)" by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable) also have "\ \ measure M (S - T) + measure M T" using assms by (blast intro: measure_Un_le) finally show ?thesis by (simp add: algebra_simps) qed lemma suminf_exist_split2: fixes f :: "nat \ 'a::real_normed_vector" assumes "summable f" shows "(\n. (\k. f(k+n))) \ 0" by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms]) lemma emeasure_union_summable: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "emeasure M (\n. A n) < \" "emeasure M (\n. A n) \ (\n. measure M (A n))" proof - define B where "B = (\N. (\n\{.. sets M" for N unfolding B_def by auto have "(\N. emeasure M (B N)) \ emeasure M (\N. B N)" apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def) moreover have "emeasure M (B N) \ ennreal (\n. measure M (A n))" for N proof - have *: "(\n (\n. measure M (A n))" using assms(3) measure_nonneg sum_le_suminf by blast have "emeasure M (B N) \ (\n = (\n = ennreal (\n \ ennreal (\n. measure M (A n))" using * by (auto simp: ennreal_leI) finally show ?thesis by simp qed ultimately have "emeasure M (\N. B N) \ ennreal (\n. measure M (A n))" by (simp add: Lim_bounded) then show "emeasure M (\n. A n) \ (\n. measure M (A n))" unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV) then show "emeasure M (\n. A n) < \" by (auto simp: less_top[symmetric] top_unique) qed lemma borel_cantelli_limsup1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "limsup A \ null_sets M" proof - have "emeasure M (limsup A) \ 0" proof (rule LIMSEQ_le_const) have "(\n. (\k. measure M (A (k+n)))) \ 0" by (rule suminf_exist_split2[OF assms(3)]) then show "(\n. ennreal (\k. measure M (A (k+n)))) \ 0" unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI) have "emeasure M (limsup A) \ (\k. measure M (A (k+n)))" for n proof - have I: "(\k\{n..}. A k) = (\k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce) have "emeasure M (limsup A) \ emeasure M (\k\{n..}. A k)" by (rule emeasure_mono, auto simp add: limsup_INF_SUP) also have "\ = emeasure M (\k. A (k+n))" using I by auto also have "\ \ (\k. measure M (A (k+n)))" apply (rule emeasure_union_summable) using assms summable_ignore_initial_segment[OF assms(3), of n] by auto finally show ?thesis by simp qed then show "\N. \n\N. emeasure M (limsup A) \ (\k. measure M (A (k+n)))" by auto qed then show ?thesis using assms(1) measurable_limsup by auto qed lemma borel_cantelli_AE1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "AE x in M. eventually (\n. x \ space M - A n) sequentially" proof - have "AE x in M. x \ limsup A" using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto moreover have "\\<^sub>F n in sequentially. x \ A n" if "x \ limsup A" for x using that by (auto simp: limsup_INF_SUP eventually_sequentially) ultimately show ?thesis by auto qed subsection \Measure spaces with \<^term>\emeasure M (space M) < \\\ locale\<^marker>\tag important\ finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ top" lemma finite_measureI[Pure.intro!]: "emeasure M (space M) \ \ \ finite_measure M" proof qed (auto intro!: exI[of _ "{space M}"]) lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \ top" using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique) lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M" by (auto simp: fmeasurable_def less_top[symmetric]) lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)" by (intro emeasure_eq_ennreal_measure) simp lemma (in finite_measure) emeasure_real: "\r. 0 \ r \ emeasure M A = ennreal r" using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto lemma (in finite_measure) bounded_measure: "measure M A \ measure M (space M)" using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_Diff: assumes sets: "A \ sets M" "B \ sets M" and "B \ A" shows "measure M (A - B) = measure M A - measure M B" using measure_Diff[OF _ assms] by simp lemma (in finite_measure) finite_measure_Union: assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" shows "measure M (A \ B) = measure M A + measure M B" using measure_Union[OF _ _ assms] by simp lemma (in finite_measure) finite_measure_finite_Union: assumes measurable: "finite S" "A`S \ sets M" "disjoint_family_on A S" shows "measure M (\i\S. A i) = (\i\S. measure M (A i))" using measure_finite_Union[OF assms] by simp lemma (in finite_measure) finite_measure_UNION: assumes A: "range A \ sets M" "disjoint_family A" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" using measure_UNION[OF A] by simp lemma (in finite_measure) finite_measure_mono: assumes "A \ B" "B \ sets M" shows "measure M A \ measure M B" using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_subadditive: assumes m: "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" using measure_subadditive[OF m] by simp lemma (in finite_measure) finite_measure_subadditive_finite: assumes "finite I" "A`I \ sets M" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" using measure_subadditive_finite[OF assms] by simp lemma (in finite_measure) finite_measure_subadditive_countably: "range A \ sets M \ summable (\i. measure M (A i)) \ measure M (\i. A i) \ (\i. measure M (A i))" by (rule measure_subadditive_countably) (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_sum_singleton: assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" shows "measure M S = (\x\S. measure M {x})" using measure_eq_sum_singleton[OF assms] by simp lemma (in finite_measure) finite_Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(\i. measure M (A i)) \ measure M (\i. A i)" using Lim_measure_incseq[OF A] by simp lemma (in finite_measure) finite_Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" shows "(\n. measure M (A n)) \ measure M (\i. A i)" using Lim_measure_decseq[OF A] by simp lemma (in finite_measure) finite_measure_compl: assumes S: "S \ sets M" shows "measure M (space M - S) = measure M (space M) - measure M S" using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp lemma (in finite_measure) finite_measure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "measure M A \ measure M B" using assms emeasure_mono_AE[OF imp B] by (simp add: emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure) lemma (in finite_measure) measure_increasing: "increasing M (measure M)" by (auto intro!: finite_measure_mono simp: increasing_def) lemma (in finite_measure) measure_zero_union: assumes "s \ sets M" "t \ sets M" "measure M t = 0" shows "measure M (s \ t) = measure M s" using assms proof - have "measure M (s \ t) \ measure M s" using finite_measure_subadditive[of s t] assms by auto moreover have "measure M (s \ t) \ measure M s" using assms by (blast intro: finite_measure_mono) ultimately show ?thesis by simp qed lemma (in finite_measure) measure_eq_compl: assumes "s \ sets M" "t \ sets M" assumes "measure M (space M - s) = measure M (space M - t)" shows "measure M s = measure M t" using assms finite_measure_compl by auto lemma (in finite_measure) measure_eq_bigunion_image: assumes "range f \ sets M" "range g \ sets M" assumes "disjoint_family f" "disjoint_family g" assumes "\ n :: nat. measure M (f n) = measure M (g n)" shows "measure M (\i. f i) = measure M (\i. g i)" using assms proof - have a: "(\ i. measure M (f i)) sums (measure M (\i. f i))" by (rule finite_measure_UNION[OF assms(1,3)]) have b: "(\ i. measure M (g i)) sums (measure M (\i. g i))" by (rule finite_measure_UNION[OF assms(2,4)]) show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp qed lemma (in finite_measure) measure_countably_zero: assumes "range c \ sets M" assumes "\ i. measure M (c i) = 0" shows "measure M (\i :: nat. c i) = 0" proof (rule antisym) show "measure M (\i :: nat. c i) \ 0" using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) qed simp lemma (in finite_measure) measure_space_inter: assumes events:"s \ sets M" "t \ sets M" assumes "measure M t = measure M (space M)" shows "measure M (s \ t) = measure M s" proof - have "measure M ((space M - s) \ (space M - t)) = measure M (space M - s)" using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union) also have "(space M - s) \ (space M - t) = space M - (s \ t)" by blast finally show "measure M (s \ t) = measure M s" using events by (auto intro!: measure_eq_compl[of "s \ t" s]) qed lemma (in finite_measure) measure_equiprobable_finite_unions: assumes s: "finite s" "\x. x \ s \ {x} \ sets M" assumes "\ x y. \x \ s; y \ s\ \ measure M {x} = measure M {y}" shows "measure M s = real (card s) * measure M {SOME x. x \ s}" proof cases assume "s \ {}" then have "\ x. x \ s" by blast from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ measure M {x} = measure M {SOME y. y \ s}" by blast have "measure M s = (\ x \ s. measure M {x})" using finite_measure_eq_sum_singleton[OF s] by simp also have "\ = (\ x \ s. measure M {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * measure M {(SOME x. x \ s)}" using sum_constant assms by simp finally show ?thesis by simp qed simp lemma (in finite_measure) measure_real_sum_image_fn: assumes "e \ sets M" assumes "\ x. x \ s \ e \ f x \ sets M" assumes "finite s" assumes disjoint: "\ x y. \x \ s ; y \ s ; x \ y\ \ f x \ f y = {}" assumes upper: "space M \ (\i \ s. f i)" shows "measure M e = (\ x \ s. measure M (e \ f x))" proof - have "e \ (\i\s. f i)" using \e \ sets M\ sets.sets_into_space upper by blast then have e: "e = (\i \ s. e \ f i)" by auto hence "measure M e = measure M (\i \ s. e \ f i)" by simp also have "\ = (\ x \ s. measure M (e \ f x))" proof (rule finite_measure_finite_Union) show "finite s" by fact show "(\i. e \ f i)`s \ sets M" using assms(2) by auto show "disjoint_family_on (\i. e \ f i) s" using disjoint by (auto simp: disjoint_family_on_def) qed finally show ?thesis . qed lemma (in finite_measure) measure_exclude: assumes "A \ sets M" "B \ sets M" assumes "measure M A = measure M (space M)" "A \ B = {}" shows "measure M B = 0" using measure_space_inter[of B A] assms by (auto simp: ac_simps) lemma (in finite_measure) finite_measure_distr: assumes f: "f \ measurable M M'" shows "finite_measure (distr M M' f)" proof (rule finite_measureI) have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) with f show "emeasure (distr M M' f) (space (distr M M' f)) \ \" by (auto simp: emeasure_distr) qed lemma emeasure_gfp[consumes 1, case_names cont measurable]: assumes sets[simp]: "\s. sets (M s) = sets N" assumes "\s. finite_measure (M s)" assumes cont: "inf_continuous F" "inf_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" assumes bound: "\P. f P \ f (\s. emeasure (M s) (space (M s)))" shows "emeasure (M s) {x\space N. gfp F x} = gfp f s" proof (subst gfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and P="Measurable.pred N", symmetric]) interpret finite_measure "M s" for s by fact fix C assume "decseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (INF i. C i) x}) = (INF i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding INF_apply by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) next show "f x \ (\s. emeasure (M s) {x \ space N. F top x})" for x using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter) qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont) subsection\<^marker>\tag unimportant\ \Counting space\ lemma strict_monoI_Suc: - assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" - unfolding strict_mono_def -proof safe - fix n m :: nat assume "n < m" then show "f n < f m" - by (induct m) (auto simp: less_Suc_eq intro: less_trans ord) -qed + assumes "(\n. f n < f (Suc n))" shows "strict_mono f" + by (simp add: assms strict_mono_Suc_iff) lemma emeasure_count_space: assumes "X \ A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \)" (is "_ = ?M X") unfolding count_space_def proof (rule emeasure_measure_of_sigma) show "X \ Pow A" using \X \ A\ by auto show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) show positive: "positive (Pow A) ?M" by (auto simp: positive_def) have additive: "additive (Pow A) ?M" by (auto simp: card_Un_disjoint additive_def) interpret ring_of_sets A "Pow A" by (rule ring_of_setsI) auto show "countably_additive (Pow A) ?M" unfolding countably_additive_iff_continuous_from_below[OF positive additive] proof safe fix F :: "nat \ 'a set" assume "incseq F" show "(\i. ?M (F i)) \ ?M (\i. F i)" proof cases assume "\i. \j\i. F i = F j" then obtain i where i: "\j\i. F i = F j" .. with \incseq F\ have "F j \ F i" for j by (cases "i \ j") (auto simp: incseq_def) then have eq: "(\i. F i) = F i" by auto with i show ?thesis by (auto intro!: Lim_transform_eventually[OF tendsto_const] eventually_sequentiallyI[where c=i]) next assume "\ (\i. \j\i. F i = F j)" then obtain f where f: "\i. i \ f i" "\i. F i \ F (f i)" by metis then have "\i. F i \ F (f i)" using \incseq F\ by (auto simp: incseq_def) with f have *: "\i. F i \ F (f i)" by auto have "incseq (\i. ?M (F i))" using \incseq F\ unfolding incseq_def by (auto simp: card_mono dest: finite_subset) then have "(\i. ?M (F i)) \ (SUP n. ?M (F n))" by (rule LIMSEQ_SUP) moreover have "(SUP n. ?M (F n)) = top" proof (rule ennreal_SUP_eq_top) fix n :: nat show "\k::nat\UNIV. of_nat n \ ?M (F k)" proof (induct n) case (Suc n) then obtain k where "of_nat n \ ?M (F k)" .. moreover have "finite (F k) \ finite (F (f k)) \ card (F k) < card (F (f k))" using \F k \ F (f k)\ by (simp add: psubset_card_mono) moreover have "finite (F (f k)) \ finite (F k)" using \k \ f k\ \incseq F\ by (auto simp: incseq_def dest: finite_subset) ultimately show ?case by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc) qed auto qed moreover have "inj (\n. F ((f ^^ n) 0))" by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *) then have 1: "infinite (range (\i. F ((f ^^ i) 0)))" by (rule range_inj_infinite) have "infinite (Pow (\i. F i))" by (rule infinite_super[OF _ 1]) auto then have "infinite (\i. F i)" by auto ultimately show ?thesis by (simp only:) simp qed qed qed lemma distr_bij_count_space: assumes f: "bij_betw f A B" shows "distr (count_space A) (count_space B) f = count_space B" proof (rule measure_eqI) have f': "f \ measurable (count_space A) (count_space B)" using f unfolding Pi_def bij_betw_def by auto fix X assume "X \ sets (distr (count_space A) (count_space B) f)" then have X: "X \ sets (count_space B)" by auto moreover from X have "f -` X \ A = the_inv_into A f ` X" using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric]) moreover have "inj_on (the_inv_into A f) B" using X f by (auto simp: bij_betw_def inj_on_the_inv_into) with X have "inj_on (the_inv_into A f) X" by (auto intro: subset_inj_on) ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X" using f unfolding emeasure_distr[OF f' X] by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD) qed simp lemma emeasure_count_space_finite[simp]: "X \ A \ finite X \ emeasure (count_space A) X = of_nat (card X)" using emeasure_count_space[of X A] by simp lemma emeasure_count_space_infinite[simp]: "X \ A \ infinite X \ emeasure (count_space A) X = \" using emeasure_count_space[of X A] by simp lemma measure_count_space: "measure (count_space A) X = (if X \ A then of_nat (card X) else 0)" by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat measure_zero_top measure_eq_emeasure_eq_ennreal) lemma emeasure_count_space_eq_0: "emeasure (count_space A) X = 0 \ (X \ A \ X = {})" proof cases assume X: "X \ A" then show ?thesis proof (intro iffI impI) assume "emeasure (count_space A) X = 0" with X show "X = {}" by (subst (asm) emeasure_count_space) (auto split: if_split_asm) qed simp qed (simp add: emeasure_notin_sets) lemma null_sets_count_space: "null_sets (count_space A) = { {} }" unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0) lemma AE_count_space: "(AE x in count_space A. P x) \ (\x\A. P x)" unfolding eventually_ae_filter by (auto simp add: null_sets_count_space) lemma sigma_finite_measure_count_space_countable: assumes A: "countable A" shows "sigma_finite_measure (count_space A)" proof qed (insert A, auto intro!: exI[of _ "(\a. {a}) ` A"]) lemma sigma_finite_measure_count_space: fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)" by (rule sigma_finite_measure_count_space_countable) auto lemma finite_measure_count_space: assumes [simp]: "finite A" shows "finite_measure (count_space A)" by rule simp lemma sigma_finite_measure_count_space_finite: assumes A: "finite A" shows "sigma_finite_measure (count_space A)" proof - interpret finite_measure "count_space A" using A by (rule finite_measure_count_space) show "sigma_finite_measure (count_space A)" .. qed subsection\<^marker>\tag unimportant\ \Measure restricted to space\ lemma emeasure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "emeasure (restrict_space M \) A = emeasure M A" proof (cases "A \ sets M") case True show ?thesis proof (rule emeasure_measure_of[OF restrict_space_def]) show "(\) \ ` sets M \ Pow (\ \ space M)" "A \ sets (restrict_space M \)" using \A \ \\ \A \ sets M\ sets.space_closed by (auto simp: sets_restrict_space) show "positive (sets (restrict_space M \)) (emeasure M)" by (auto simp: positive_def) show "countably_additive (sets (restrict_space M \)) (emeasure M)" proof (rule countably_additiveI) fix A :: "nat \ _" assume "range A \ sets (restrict_space M \)" "disjoint_family A" with assms have "\i. A i \ sets M" "\i. A i \ space M" "disjoint_family A" by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff dest: sets.sets_into_space)+ then show "(\i. emeasure M (A i)) = emeasure M (\i. A i)" by (subst suminf_emeasure) (auto simp: disjoint_family_subset) qed qed next case False with assms have "A \ sets (restrict_space M \)" by (simp add: sets_restrict_space_iff) with False show ?thesis by (simp add: emeasure_notin_sets) qed lemma measure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "measure (restrict_space M \) A = measure M A" using emeasure_restrict_space[OF assms] by (simp add: measure_def) lemma AE_restrict_space_iff: assumes "\ \ space M \ sets M" shows "(AE x in restrict_space M \. P x) \ (AE x in M. x \ \ \ P x)" proof - have ex_cong: "\P Q f. (\x. P x \ Q x) \ (\x. Q x \ P (f x)) \ (\x. P x) \ (\x. Q x)" by auto { fix X assume X: "X \ sets M" "emeasure M X = 0" then have "emeasure M (\ \ space M \ X) \ emeasure M X" by (intro emeasure_mono) auto then have "emeasure M (\ \ space M \ X) = 0" using X by (auto intro!: antisym) } with assms show ?thesis unfolding eventually_ae_filter by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff emeasure_restrict_space cong: conj_cong intro!: ex_cong[where f="\X. (\ \ space M) \ X"]) qed lemma restrict_restrict_space: assumes "A \ space M \ sets M" "B \ space M \ sets M" shows "restrict_space (restrict_space M A) B = restrict_space M (A \ B)" (is "?l = ?r") proof (rule measure_eqI[symmetric]) show "sets ?r = sets ?l" unfolding sets_restrict_space image_comp by (intro image_cong) auto next fix X assume "X \ sets (restrict_space M (A \ B))" then obtain Y where "Y \ sets M" "X = Y \ A \ B" by (auto simp: sets_restrict_space) with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X" by (subst (1 2) emeasure_restrict_space) (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps) qed lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \ B)" proof (rule measure_eqI) show "sets (restrict_space (count_space B) A) = sets (count_space (A \ B))" by (subst sets_restrict_space) auto moreover fix X assume "X \ sets (restrict_space (count_space B) A)" ultimately have "X \ A \ B" by auto then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \ B)) X" by (cases "finite X") (auto simp add: emeasure_restrict_space) qed lemma sigma_finite_measure_restrict_space: assumes "sigma_finite_measure M" and A: "A \ sets M" shows "sigma_finite_measure (restrict_space M A)" proof - interpret sigma_finite_measure M by fact from sigma_finite_countable obtain C where C: "countable C" "C \ sets M" "(\C) = space M" "\a\C. emeasure M a \ \" by blast let ?C = "(\) A ` C" from C have "countable ?C" "?C \ sets (restrict_space M A)" "(\?C) = space (restrict_space M A)" by(auto simp add: sets_restrict_space space_restrict_space) moreover { fix a assume "a \ ?C" then obtain a' where "a = A \ a'" "a' \ C" .. then have "emeasure (restrict_space M A) a \ emeasure M a'" using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono) also have "\ < \" using C(4)[rule_format, of a'] \a' \ C\ by (simp add: less_top) finally have "emeasure (restrict_space M A) a \ \" by simp } ultimately show ?thesis by unfold_locales (rule exI conjI|assumption|blast)+ qed lemma finite_measure_restrict_space: assumes "finite_measure M" and A: "A \ sets M" shows "finite_measure (restrict_space M A)" proof - interpret finite_measure M by fact show ?thesis by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space) qed lemma restrict_distr: assumes [measurable]: "f \ measurable M N" assumes [simp]: "\ \ space N \ sets N" and restrict: "f \ space M \ \" shows "restrict_space (distr M N f) \ = distr M (restrict_space N \) f" (is "?l = ?r") proof (rule measure_eqI) fix A assume "A \ sets (restrict_space (distr M N f) \)" with restrict show "emeasure ?l A = emeasure ?r A" by (subst emeasure_distr) (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr intro!: measurable_restrict_space2) qed (simp add: sets_restrict_space) lemma measure_eqI_restrict_generator: assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" assumes sets_eq: "sets M = sets N" and \: "\ \ sets M" assumes "sets (restrict_space M \) = sigma_sets \ E" assumes "sets (restrict_space N \) = sigma_sets \ E" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" assumes A: "countable A" "A \ {}" "A \ E" "\A = \" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof (rule measure_eqI) fix X assume X: "X \ sets M" then have "emeasure M X = emeasure (restrict_space M \) (X \ \)" using ae \ by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE) also have "restrict_space M \ = restrict_space N \" proof (rule measure_eqI_generator_eq) fix X assume "X \ E" then show "emeasure (restrict_space M \) X = emeasure (restrict_space N \) X" using E \ by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) next show "range (from_nat_into A) \ E" "(\i. from_nat_into A i) = \" using A by (auto cong del: SUP_cong_simp) next fix i have "emeasure (restrict_space M \) (from_nat_into A i) = emeasure M (from_nat_into A i)" using A \ by (subst emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into) with A show "emeasure (restrict_space M \) (from_nat_into A i) \ \" by (auto intro: from_nat_into) qed fact+ also have "emeasure (restrict_space N \) (X \ \) = emeasure N X" using X ae \ by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE) finally show "emeasure M X = emeasure N X" . qed fact subsection\<^marker>\tag unimportant\ \Null measure\ definition null_measure :: "'a measure \ 'a measure" where "null_measure M = sigma (space M) (sets M)" lemma space_null_measure[simp]: "space (null_measure M) = space M" by (simp add: null_measure_def) lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" by (simp add: null_measure_def) lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0" by (cases "X \ sets M", rule emeasure_measure_of) (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def dest: sets.sets_into_space) lemma measure_null_measure[simp]: "measure (null_measure M) X = 0" by (intro measure_eq_emeasure_eq_ennreal) auto lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M" by(rule measure_eqI) simp_all subsection \Scaling a measure\ definition\<^marker>\tag important\ scale_measure :: "ennreal \ 'a measure \ 'a measure" where "scale_measure r M = measure_of (space M) (sets M) (\A. r * emeasure M A)" lemma space_scale_measure: "space (scale_measure r M) = space M" by (simp add: scale_measure_def) lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M" by (simp add: scale_measure_def) lemma emeasure_scale_measure [simp]: "emeasure (scale_measure r M) A = r * emeasure M A" (is "_ = ?\ A") proof(cases "A \ sets M") case True show ?thesis unfolding scale_measure_def proof(rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" by (simp add: positive_def) show "countably_additive (sets M) ?\" proof (rule countably_additiveI) fix A :: "nat \ _" assume *: "range A \ sets M" "disjoint_family A" have "(\i. ?\ (A i)) = r * (\i. emeasure M (A i))" by simp also have "\ = ?\ (\i. A i)" using * by(simp add: suminf_emeasure) finally show "(\i. ?\ (A i)) = ?\ (\i. A i)" . qed qed(fact True) qed(simp add: emeasure_notin_sets) lemma scale_measure_1 [simp]: "scale_measure 1 M = M" by(rule measure_eqI) simp_all lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M" by(rule measure_eqI) simp_all lemma measure_scale_measure [simp]: "0 \ r \ measure (scale_measure r M) A = r * measure M A" using emeasure_scale_measure[of r M A] emeasure_eq_ennreal_measure[of M A] measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A] by (cases "emeasure (scale_measure r M) A = top") (auto simp del: emeasure_scale_measure simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric]) lemma scale_scale_measure [simp]: "scale_measure r (scale_measure r' M) = scale_measure (r * r') M" by (rule measure_eqI) (simp_all add: max_def mult.assoc) lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M" by (rule measure_eqI) simp_all subsection \Complete lattice structure on measures\ lemma (in finite_measure) finite_measure_Diff': "A \ sets M \ B \ sets M \ measure M (A - B) = measure M A - measure M (A \ B)" using finite_measure_Diff[of A "A \ B"] by (auto simp: Diff_Int) lemma (in finite_measure) finite_measure_Union': "A \ sets M \ B \ sets M \ measure M (A \ B) = measure M A + measure M (B - A)" using finite_measure_Union[of A "B - A"] by auto lemma finite_unsigned_Hahn_decomposition: assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M" shows "\Y\sets M. (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ Y = {} \ M X \ N X)" proof - interpret M: finite_measure M by fact interpret N: finite_measure N by fact define d where "d X = measure M X - measure N X" for X have [intro]: "bdd_above (d`sets M)" using sets.sets_into_space[of _ M] by (intro bdd_aboveI[where M="measure M (space M)"]) (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono) define \ where "\ = (SUP X\sets M. d X)" have le_\[intro]: "X \ sets M \ d X \ \" for X by (auto simp: \_def intro!: cSUP_upper) have "\f. \n. f n \ sets M \ d (f n) > \ - 1 / 2^n" proof (intro choice_iff[THEN iffD1] allI) fix n have "\X\sets M. \ - 1 / 2^n < d X" unfolding \_def by (intro less_cSUP_iff[THEN iffD1]) auto then show "\y. y \ sets M \ \ - 1 / 2 ^ n < d y" by auto qed then obtain E where [measurable]: "E n \ sets M" and E: "d (E n) > \ - 1 / 2^n" for n by auto define F where "F m n = (if m \ n then \i\{m..n}. E i else space M)" for m n have [measurable]: "m \ n \ F m n \ sets M" for m n by (auto simp: F_def) have 1: "\ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" if "m \ n" for m n using that proof (induct rule: dec_induct) case base with E[of m] show ?case by (simp add: F_def field_simps) next case (step i) have F_Suc: "F m (Suc i) = F m i \ E (Suc i)" using \m \ i\ by (auto simp: F_def le_Suc_eq) have "\ + (\ - 2 / 2^m + 1 / 2 ^ Suc i) \ (\ - 1 / 2^Suc i) + (\ - 2 / 2^m + 1 / 2^i)" by (simp add: field_simps) also have "\ \ d (E (Suc i)) + d (F m i)" using E[of "Suc i"] by (intro add_mono step) auto also have "\ = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff') also have "\ = d (E (Suc i) \ F m i) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union') also have "\ \ \ + d (F m (Suc i))" using \m \ i\ by auto finally show ?case by auto qed define F' where "F' m = (\i\{m..}. E i)" for m have F'_eq: "F' m = (\i. F m (i + m))" for m by (fastforce simp: le_iff_add[of m] F'_def F_def) have [measurable]: "F' m \ sets M" for m by (auto simp: F'_def) have \_le: "\ - 0 \ d (\m. F' m)" proof (rule LIMSEQ_le) show "(\n. \ - 2 / 2 ^ n) \ \ - 0" by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto have "incseq F'" by (auto simp: incseq_def F'_def) then show "(\m. d (F' m)) \ d (\m. F' m)" unfolding d_def by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto have "\ - 2 / 2 ^ m + 0 \ d (F' m)" for m proof (rule LIMSEQ_le) have *: "decseq (\n. F m (n + m))" by (auto simp: decseq_def F_def) show "(\n. d (F m n)) \ d (F' m)" unfolding d_def F'_eq by (rule LIMSEQ_offset[where k=m]) (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *) show "(\n. \ - 2 / 2 ^ m + 1 / 2 ^ n) \ \ - 2 / 2 ^ m + 0" by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto show "\N. \n\N. \ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" using 1[of m] by (intro exI[of _ m]) auto qed then show "\N. \n\N. \ - 2 / 2 ^ n \ d (F' n)" by auto qed show ?thesis proof (safe intro!: bexI[of _ "\m. F' m"]) fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m)" have "d (\m. F' m) - d X = d ((\m. F' m) - X)" using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff) also have "\ \ \" by auto finally have "0 \ d X" using \_le by auto then show "emeasure N X \ emeasure M X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) next fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m) = {}" then have "d (\m. F' m) + d X = d (X \ (\m. F' m))" by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union) also have "\ \ \" by auto finally have "d X \ 0" using \_le by auto then show "emeasure M X \ emeasure N X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) qed auto qed proposition unsigned_Hahn_decomposition: assumes [simp]: "sets N = sets M" and [measurable]: "A \ sets M" and [simp]: "emeasure M A \ top" "emeasure N A \ top" shows "\Y\sets M. Y \ A \ (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ A \ X \ Y = {} \ M X \ N X)" proof - have "\Y\sets (restrict_space M A). (\X\sets (restrict_space M A). X \ Y \ (restrict_space N A) X \ (restrict_space M A) X) \ (\X\sets (restrict_space M A). X \ Y = {} \ (restrict_space M A) X \ (restrict_space N A) X)" proof (rule finite_unsigned_Hahn_decomposition) show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)" by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI) qed (simp add: sets_restrict_space) with assms show ?thesis by (metis Int_subset_iff emeasure_restrict_space sets.Int_space_eq2 sets_restrict_space_iff space_restrict_space) qed text\<^marker>\tag important\ \ Define a lexicographical order on \<^type>\measure\, in the order space, sets and measure. The parts of the lexicographical order are point-wise ordered. \ instantiation measure :: (type) order_bot begin inductive less_eq_measure :: "'a measure \ 'a measure \ bool" where "space M \ space N \ less_eq_measure M N" | "space M = space N \ sets M \ sets N \ less_eq_measure M N" | "space M = space N \ sets M = sets N \ emeasure M \ emeasure N \ less_eq_measure M N" lemma le_measure_iff: "M \ N \ (if space M = space N then if sets M = sets N then emeasure M \ emeasure N else sets M \ sets N else space M \ space N)" by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros) definition\<^marker>\tag important\ less_measure :: "'a measure \ 'a measure \ bool" where "less_measure M N \ (M \ N \ \ N \ M)" definition\<^marker>\tag important\ bot_measure :: "'a measure" where "bot_measure = sigma {} {}" lemma shows space_bot[simp]: "space bot = {}" and sets_bot[simp]: "sets bot = {{}}" and emeasure_bot[simp]: "emeasure bot X = 0" by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma) instance proof standard show "bot \ a" for a :: "'a measure" by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def) qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI) end proposition le_measure: "sets M = sets N \ M \ N \ (\A\sets M. emeasure M A \ emeasure N A)" by (metis emeasure_neq_0_sets le_fun_def le_measure_iff order_class.order_eq_iff sets_eq_imp_space_eq) definition\<^marker>\tag important\ sup_measure' :: "'a measure \ 'a measure \ 'a measure" where "sup_measure' A B = measure_of (space A) (sets A) (\X. SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" lemma assumes [simp]: "sets B = sets A" shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A" and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A" using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def) lemma emeasure_sup_measure': assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \ sets A" shows "emeasure (sup_measure' A B) X = (SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" (is "_ = ?S X") proof - note sets_eq_imp_space_eq[OF sets_eq, simp] show ?thesis using sup_measure'_def proof (rule emeasure_measure_of) let ?d = "\X Y. emeasure A (X \ Y) + emeasure B (X \ - Y)" show "countably_additive (sets (sup_measure' A B)) (\X. SUP Y \ sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" proof (rule countably_additiveI, goal_cases) case (1 X) then have [measurable]: "\i. X i \ sets A" and "disjoint_family X" by auto have disjoint: "disjoint_family (\i. X i \ Y)" "disjoint_family (\i. X i - Y)" for Y using "1"(2) disjoint_family_subset by fastforce+ have "(\i. ?S (X i)) = (SUP Y\sets A. \i. ?d (X i) Y)" proof (rule ennreal_suminf_SUP_eq_directed) fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \ sets A" "b \ sets A" have "\c\sets A. c \ X i \ (\a\sets A. ?d (X i) a \ ?d (X i) c)" for i proof cases assume "emeasure A (X i) = top \ emeasure B (X i) = top" then show ?thesis by force next assume finite: "\ (emeasure A (X i) = top \ emeasure B (X i) = top)" then have "\Y\sets A. Y \ X i \ (\C\sets A. C \ Y \ B C \ A C) \ (\C\sets A. C \ X i \ C \ Y = {} \ A C \ B C)" using unsigned_Hahn_decomposition[of B A "X i"] by simp then obtain Y where [measurable]: "Y \ sets A" and [simp]: "Y \ X i" and B_le_A: "\C. C \ sets A \ C \ Y \ B C \ A C" and A_le_B: "\C. C \ sets A \ C \ X i \ C \ Y = {} \ A C \ B C" by auto show ?thesis proof (intro bexI ballI conjI) fix a assume [measurable]: "a \ sets A" have *: "(X i \ a \ Y \ (X i \ a - Y)) = X i \ a" "(X i - a) \ Y \ (X i - a - Y) = X i \ - a" for a Y by auto then have "?d (X i) a = (A (X i \ a \ Y) + A (X i \ a \ - Y)) + (B (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ a \ Y) + B (X i \ a \ - Y)) + (A (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ Y \ a) + A (X i \ Y \ - a)) + (B (X i \ - Y \ a) + B (X i \ - Y \ - a))" by (simp add: ac_simps) also have "\ \ A (X i \ Y) + B (X i \ - Y)" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *) finally show "?d (X i) a \ ?d (X i) Y" . qed auto qed then obtain C where [measurable]: "C i \ sets A" and "C i \ X i" and C: "\a. a \ sets A \ ?d (X i) a \ ?d (X i) (C i)" for i by metis have *: "X i \ (\i. C i) = X i \ C i" for i using \disjoint_family X\ \\i. C i \ X i\ by (simp add: disjoint_family_on_def disjoint_iff_not_equal set_eq_iff) (metis subsetD) then have **: "X i \ - (\i. C i) = X i \ - C i" for i by blast moreover have "(\i. C i) \ sets A" by fastforce ultimately show "\c\sets A. \i\J. ?d (X i) a \ ?d (X i) c \ ?d (X i) b \ ?d (X i) c" by (metis "*" C \a \ sets A\ \b \ sets A\) qed also have "\ = ?S (\i. X i)" proof - have "\Y. Y \ sets A \ (\i. emeasure A (X i \ Y) + emeasure B (X i \ -Y)) = emeasure A (\i. X i \ Y) + emeasure B (\i. X i \ -Y)" using disjoint by (auto simp flip: suminf_add Diff_eq simp add: image_subset_iff suminf_emeasure) then show ?thesis by force qed finally show "(\i. ?S (X i)) = ?S (\i. X i)" . qed qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const) qed lemma le_emeasure_sup_measure'1: assumes "sets B = sets A" "X \ sets A" shows "emeasure A X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms) lemma le_emeasure_sup_measure'2: assumes "sets B = sets A" "X \ sets A" shows "emeasure B X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms) lemma emeasure_sup_measure'_le2: assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \ sets C" assumes A: "\Y. Y \ X \ Y \ sets A \ emeasure A Y \ emeasure C Y" assumes B: "\Y. Y \ X \ Y \ sets A \ emeasure B Y \ emeasure C Y" shows "emeasure (sup_measure' A B) X \ emeasure C X" proof (subst emeasure_sup_measure') show "(SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y)) \ emeasure C X" unfolding \sets A = sets C\ proof (intro SUP_least) fix Y assume [measurable]: "Y \ sets C" have [simp]: "X \ Y \ (X - Y) = X" by auto have "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C (X \ Y) + emeasure C (X \ - Y)" by (intro add_mono A B) (auto simp: Diff_eq[symmetric]) also have "\ = emeasure C X" by (subst plus_emeasure) (auto simp: Diff_eq[symmetric]) finally show "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C X" . qed qed simp_all definition\<^marker>\tag important\ sup_lexord :: "'a \ 'a \ ('a \ 'b::order) \ 'a \ 'a \ 'a" where "sup_lexord A B k s c = (if k A = k B then c else if \ k A \ k B \ \ k B \ k A then s else if k B \ k A then A else B)" lemma sup_lexord: "(k A < k B \ P B) \ (k B < k A \ P A) \ (k A = k B \ P c) \ (\ k B \ k A \ \ k A \ k B \ P s) \ P (sup_lexord A B k s c)" by (auto simp: sup_lexord_def) lemmas le_sup_lexord = sup_lexord[where P="\a. c \ a" for c] lemma sup_lexord1: "k A = k B \ sup_lexord A B k s c = c" by (simp add: sup_lexord_def) lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c" by (auto simp: sup_lexord_def) lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \ \ sets x) = (\ \ sets x)" using sets.sigma_sets_subset[of \ x] by auto lemma sigma_le_iff: "\ \ Pow \ \ sigma \ \ \ x \ (\ \ space x \ (space x = \ \ \ \ sets x))" by (cases "\ = space x") (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def sigma_sets_superset_generator sigma_sets_le_sets_iff) instantiation measure :: (type) semilattice_sup begin definition\<^marker>\tag important\ sup_measure :: "'a measure \ 'a measure \ 'a measure" where "sup_measure A B = sup_lexord A B space (sigma (space A \ space B) {}) (sup_lexord A B sets (sigma (space A) (sets A \ sets B)) (sup_measure' A B))" instance proof fix x y z :: "'a measure" show "x \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume "space x = space y" then have *: "sets x \ sets y \ Pow (space x)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets x \ sets x \ sets y" by auto also have "\ \ sigma (space x) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "x \ sigma (space x) (sets x \ sets y)" by (simp add: space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "x \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "x \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'1) qed (auto intro: less_eq_measure.intros) show "y \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume **: "space x = space y" then have *: "sets x \ sets y \ Pow (space y)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets y \ sets x \ sets y" by auto also have "\ \ sigma (space y) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "y \ sigma (space x) (sets x \ sets y)" by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "y \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "y \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'2) qed (auto intro: less_eq_measure.intros) show "x \ y \ z \ y \ sup x z \ y" unfolding sup_measure_def proof (intro sup_lexord[where P="\x. x \ y"]) assume "x \ y" "z \ y" and [simp]: "space x = space z" "sets x = sets z" from \x \ y\ show "sup_measure' x z \ y" proof cases case 1 then show ?thesis by (intro less_eq_measure.intros(1)) simp next case 2 then show ?thesis by (intro less_eq_measure.intros(2)) simp_all next case 3 with \z \ y\ \x \ y\ show ?thesis by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2) qed next assume **: "x \ y" "z \ y" "space x = space z" "\ sets z \ sets x" "\ sets x \ sets z" then have *: "sets x \ sets z \ Pow (space x)" using sets.space_closed by auto show "sigma (space x) (sets x \ sets z) \ y" unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm) next assume "x \ y" "z \ y" "\ space z \ space x" "\ space x \ space z" then have "space x \ space y" "space z \ space y" by (auto simp: le_measure_iff split: if_split_asm) then show "sigma (space x \ space z) {} \ y" by (simp add: sigma_le_iff) qed qed end lemma space_empty_eq_bot: "space a = {} \ a = bot" using space_empty[of a] by (auto intro!: measure_eqI) lemma sets_eq_iff_bounded: "A \ B \ B \ C \ sets A = sets C \ sets B = sets A" by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm) lemma sets_sup: "sets A = sets M \ sets B = sets M \ sets (sup A B) = sets M" by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq) lemma le_measureD1: "A \ B \ space A \ space B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD2: "A \ B \ space A = space B \ sets A \ sets B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD3: "A \ B \ sets A = sets B \ emeasure A X \ emeasure B X" by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm) lemma UN_space_closed: "\(sets ` S) \ Pow (\(space ` S))" using sets.space_closed by auto definition\<^marker>\tag important\ Sup_lexord :: "('a \ 'b::complete_lattice) \ ('a set \ 'a) \ ('a set \ 'a) \ 'a set \ 'a" where "Sup_lexord k c s A = (let U = (SUP a\A. k a) in if \a\A. k a = U then c {a\A. k a = U} else s A)" lemma Sup_lexord: "(\a S. a \ A \ k a = (SUP a\A. k a) \ S = {a'\A. k a' = k a} \ P (c S)) \ ((\a. a \ A \ k a \ (SUP a\A. k a)) \ P (s A)) \ P (Sup_lexord k c s A)" by (auto simp: Sup_lexord_def Let_def) lemma Sup_lexord1: assumes A: "A \ {}" "(\a. a \ A \ k a = (\a\A. k a))" "P (c A)" shows "P (Sup_lexord k c s A)" unfolding Sup_lexord_def Let_def proof (clarsimp, safe) show "\a\A. k a \ (\x\A. k x) \ P (s A)" by (metis assms(1,2) ex_in_conv) next fix a assume "a \ A" "k a = (\x\A. k x)" then have "{a \ A. k a = (\x\A. k x)} = {a \ A. k a = k a}" by (metis A(2)[symmetric]) then show "P (c {a \ A. k a = (\x\A. k x)})" by (simp add: A(3)) qed instantiation measure :: (type) complete_lattice begin interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure" by standard (auto intro!: antisym) lemma sup_measure_F_mono': "finite J \ finite I \ sup_measure.F id I \ sup_measure.F id (I \ J)" proof (induction J rule: finite_induct) case empty then show ?case by simp next case (insert i J) show ?case proof cases assume "i \ I" with insert show ?thesis by (auto simp: insert_absorb) next assume "i \ I" have "sup_measure.F id I \ sup_measure.F id (I \ J)" by (intro insert) also have "\ \ sup_measure.F id (insert i (I \ J))" using insert \i \ I\ by (subst sup_measure.insert) auto finally show ?thesis by auto qed qed lemma sup_measure_F_mono: "finite I \ J \ I \ sup_measure.F id J \ sup_measure.F id I" using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1) lemma sets_sup_measure_F: "finite I \ I \ {} \ (\i. i \ I \ sets i = sets M) \ sets (sup_measure.F id I) = sets M" by (induction I rule: finite_ne_induct) (simp_all add: sets_sup) definition\<^marker>\tag important\ Sup_measure' :: "'a measure set \ 'a measure" where "Sup_measure' M = measure_of (\a\M. space a) (\a\M. sets a) (\X. (SUP P\{P. finite P \ P \ M }. sup_measure.F id P X))" lemma space_Sup_measure'2: "space (Sup_measure' M) = (\m\M. space m)" unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed]) lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\m\M. space m) (\m\M. sets m)" unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed]) lemma sets_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "sets (Sup_measure' M) = sets A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def) lemma space_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "space (Sup_measure' M) = space A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def ) lemma emeasure_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "X \ sets A" "M \ {}" shows "emeasure (Sup_measure' M) X = (SUP P\{P. finite P \ P \ M}. sup_measure.F id P X)" (is "_ = ?S X") using Sup_measure'_def proof (rule emeasure_measure_of) note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A" using \M \ {}\ by (simp_all add: Sup_measure'_def) let ?\ = "sup_measure.F id" show "countably_additive (sets (Sup_measure' M)) ?S" proof (rule countably_additiveI, goal_cases) case (1 F) then have **: "range F \ sets A" by (auto simp: *) show "(\i. ?S (F i)) = ?S (\i. F i)" proof (subst ennreal_suminf_SUP_eq_directed) fix i j and N :: "nat set" assume ij: "i \ {P. finite P \ P \ M}" "j \ {P. finite P \ P \ M}" have "(i \ {} \ sets (?\ i) = sets A) \ (j \ {} \ sets (?\ j) = sets A) \ (i \ {} \ j \ {} \ sets (?\ (i \ j)) = sets A)" using ij by (intro impI sets_sup_measure_F conjI) auto then have "?\ j (F n) \ ?\ (i \ j) (F n) \ ?\ i (F n) \ ?\ (i \ j) (F n)" for n using ij by (cases "i = {}"; cases "j = {}") (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F simp del: id_apply) with ij show "\k\{P. finite P \ P \ M}. \n\N. ?\ i (F n) \ ?\ k (F n) \ ?\ j (F n) \ ?\ k (F n)" by (safe intro!: bexI[of _ "i \ j"]) auto next show "(SUP P \ {P. finite P \ P \ M}. \n. ?\ P (F n)) = (SUP P \ {P. finite P \ P \ M}. ?\ P (\(F ` UNIV)))" proof (intro arg_cong [of _ _ Sup] image_cong refl) fix i assume i: "i \ {P. finite P \ P \ M}" show "(\n. ?\ i (F n)) = ?\ i (\(F ` UNIV))" proof cases assume "i \ {}" with i ** show ?thesis by (smt (verit, best) "1"(2) Measure_Space.sets_sup_measure_F assms(1) mem_Collect_eq subset_eq suminf_cong suminf_emeasure) qed simp qed qed qed show "positive (sets (Sup_measure' M)) ?S" by (auto simp: positive_def bot_ennreal[symmetric]) show "X \ sets (Sup_measure' M)" using assms * by auto qed (rule UN_space_closed) definition\<^marker>\tag important\ Sup_measure :: "'a measure set \ 'a measure" where "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure' (\U. sigma (\u\U. space u) (\u\U. sets u))) (\U. sigma (\u\U. space u) {})" definition\<^marker>\tag important\ Inf_measure :: "'a measure set \ 'a measure" where "Inf_measure A = Sup {x. \a\A. x \ a}" definition\<^marker>\tag important\ inf_measure :: "'a measure \ 'a measure \ 'a measure" where "inf_measure a b = Inf {a, b}" definition\<^marker>\tag important\ top_measure :: "'a measure" where "top_measure = Inf {}" instance proof note UN_space_closed [simp] show upper: "x \ Sup A" if x: "x \ A" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. x \ y"]) assume "\a. a \ A \ space a \ (\a\A. space a)" from this[OF \x \ A\] \x \ A\ show "x \ sigma (\a\A. space a) {}" by (intro less_eq_measure.intros) auto next fix a S assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and neq: "\aa. aa \ S \ sets aa \ (\a\S. sets a)" have sp_a: "space a = (\(space ` S))" using \a\A\ by (auto simp: S) show "x \ sigma (\(space ` S)) (\(sets ` S))" proof cases assume [simp]: "space x = space a" have "sets x \ (\a\S. sets a)" using \x\A\ neq[of x] by (auto simp: S) also have "\ \ sigma_sets (\x\S. space x) (\x\S. sets x)" by (rule sigma_sets_superset_generator) finally show ?thesis by (intro less_eq_measure.intros(2)) (simp_all add: sp_a) next assume "space x \ space a" moreover have "space x \ space a" unfolding a using \x\A\ by auto ultimately show ?thesis by (intro less_eq_measure.intros) (simp add: less_le sp_a) qed next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "x \ Sup_measure' S'" proof cases assume "x \ S" with \b \ S\ have "space x = space b" by (simp add: S) show ?thesis proof cases assume "x \ S'" show "x \ Sup_measure' S'" proof (intro le_measure[THEN iffD2] ballI) show "sets x = sets (Sup_measure' S')" using \x\S'\ * by (simp add: S') fix X assume "X \ sets x" show "emeasure x X \ emeasure (Sup_measure' S') X" proof (subst emeasure_Sup_measure'[OF _ \X \ sets x\]) show "emeasure x X \ (SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X)" using \x\S'\ by (intro SUP_upper2[where i="{x}"]) auto qed (insert \x\S'\ S', auto) qed next assume "x \ S'" then have "sets x \ sets b" using \x\S\ by (auto simp: S') moreover have "sets x \ sets b" using \x\S\ unfolding b by auto ultimately show ?thesis using * \x \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * \space x = space b\ less_le) qed next assume "x \ S" with \x\A\ \x \ S\ \space b = space a\ show ?thesis by (intro less_eq_measure.intros) (simp_all add: * less_le a SUP_upper S) qed qed show least: "Sup A \ x" if x: "\z. z \ A \ z \ x" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. y \ x"]) assume "\a. a \ A \ space a \ (\a\A. space a)" show "sigma (\(space ` A)) {} \ x" using x[THEN le_measureD1] by (subst sigma_le_iff) auto next fix a S assume "a \ A" "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" "\a. a \ S \ sets a \ (\a\S. sets a)" have "\(space ` S) \ space x" using S le_measureD1[OF x] by auto moreover have "\(space ` S) = space a" using \a\A\ S by auto then have "space x = \(space ` S) \ \(sets ` S) \ sets x" using \a \ A\ le_measureD2[OF x] by (auto simp: S) ultimately show "sigma (\(space ` S)) (\(sets ` S)) \ x" by (subst sigma_le_iff) simp_all next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "Sup_measure' S' \ x" proof cases assume "space x = space a" show ?thesis proof cases assume **: "sets x = sets b" show ?thesis proof (intro le_measure[THEN iffD2] ballI) show ***: "sets (Sup_measure' S') = sets x" by (simp add: * **) fix X assume "X \ sets (Sup_measure' S')" show "emeasure (Sup_measure' S') X \ emeasure x X" unfolding *** proof (subst emeasure_Sup_measure'[OF _ \X \ sets (Sup_measure' S')\]) show "(SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X) \ emeasure x X" proof (safe intro!: SUP_least) fix P assume P: "finite P" "P \ S'" show "emeasure (sup_measure.F id P) X \ emeasure x X" proof cases assume "P = {}" then show ?thesis by auto next assume "P \ {}" from P have "finite P" "P \ A" unfolding S' S by (simp_all add: subset_eq) then have "sup_measure.F id P \ x" by (induction P) (auto simp: x) moreover have "sets (sup_measure.F id P) = sets x" using \finite P\ \P \ {}\ \P \ S'\ \sets x = sets b\ by (intro sets_sup_measure_F) (auto simp: S') ultimately show "emeasure (sup_measure.F id P) X \ emeasure x X" by (rule le_measureD3) qed qed show "m \ S' \ sets m = sets (Sup_measure' S')" for m unfolding * by (simp add: S') qed fact qed next assume "sets x \ sets b" moreover have "sets b \ sets x" unfolding b S using x[THEN le_measureD2] \space x = space a\ by auto ultimately show "Sup_measure' S' \ x" using \space x = space a\ \b \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * S) qed next assume "space x \ space a" then have "space a < space x" using le_measureD1[OF x[OF \a\A\]] by auto then show "Sup_measure' S' \ x" by (intro less_eq_measure.intros) (simp add: * \space b = space a\) qed qed show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)" by (auto intro!: antisym least simp: top_measure_def) show lower: "x \ A \ Inf A \ x" for x :: "'a measure" and A unfolding Inf_measure_def by (intro least) auto show greatest: "(\z. z \ A \ x \ z) \ x \ Inf A" for x :: "'a measure" and A unfolding Inf_measure_def by (intro upper) auto show "inf x y \ x" "inf x y \ y" "x \ y \ x \ z \ x \ inf y z" for x y z :: "'a measure" by (auto simp: inf_measure_def intro!: lower greatest) qed end lemma sets_SUP: assumes "\x. x \ I \ sets (M x) = sets N" shows "I \ {} \ sets (SUP i\I. M i) = sets N" unfolding Sup_measure_def using assms assms[THEN sets_eq_imp_space_eq] sets_Sup_measure'[where A=N and M="M`I"] by (intro Sup_lexord1[where P="\x. sets x = sets N"]) auto lemma emeasure_SUP: assumes sets: "\i. i \ I \ sets (M i) = sets N" "X \ sets N" "I \ {}" shows "emeasure (SUP i\I. M i) X = (SUP J\{J. J \ {} \ finite J \ J \ I}. emeasure (SUP i\J. M i) X)" proof - interpret sup_measure: comm_monoid_set sup "bot :: 'b measure" by standard (auto intro!: antisym) have eq: "finite J \ sup_measure.F id J = (SUP i\J. i)" for J :: "'b measure set" by (induction J rule: finite_induct) auto have 1: "J \ {} \ J \ I \ sets (SUP x\J. M x) = sets N" for J by (intro sets_SUP sets) (auto ) from \I \ {}\ obtain i where "i\I" by auto have "Sup_measure' (M`I) X = (SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X)" using sets by (intro emeasure_Sup_measure') auto also have "Sup_measure' (M`I) = (SUP i\I. M i)" unfolding Sup_measure_def using \I \ {}\ sets sets(1)[THEN sets_eq_imp_space_eq] by (intro Sup_lexord1[where P="\x. _ = x"]) auto also have "(SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X) = (SUP J\{J. J \ {} \ finite J \ J \ I}. (SUP i\J. M i) X)" proof (intro SUP_eq) fix J assume "J \ {P. finite P \ P \ M`I}" then obtain J' where J': "J' \ I" "finite J'" and J: "J = M`J'" and "finite J" using finite_subset_image[of J M I] by auto show "\j\{J. J \ {} \ finite J \ J \ I}. sup_measure.F id J X \ (SUP i\j. M i) X" proof cases assume "J' = {}" with \i \ I\ show ?thesis by (auto simp add: J) next assume "J' \ {}" with J J' show ?thesis by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply) qed next fix J assume J: "J \ {P. P \ {} \ finite P \ P \ I}" show "\J'\{J. finite J \ J \ M`I}. (SUP i\J. M i) X \ sup_measure.F id J' X" using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply) qed finally show ?thesis . qed lemma emeasure_SUP_chain: assumes sets: "\i. i \ A \ sets (M i) = sets N" "X \ sets N" assumes ch: "Complete_Partial_Order.chain (\) (M ` A)" and "A \ {}" shows "emeasure (SUP i\A. M i) X = (SUP i\A. emeasure (M i) X)" proof (subst emeasure_SUP[OF sets \A \ {}\]) show "(SUP J\{J. J \ {} \ finite J \ J \ A}. emeasure (Sup (M ` J)) X) = (SUP i\A. emeasure (M i) X)" proof (rule SUP_eq) fix J assume "J \ {J. J \ {} \ finite J \ J \ A}" then have J: "Complete_Partial_Order.chain (\) (M ` J)" "finite J" "J \ {}" and "J \ A" using ch[THEN chain_subset, of "M`J"] by auto with in_chain_finite[OF J(1)] obtain j where "j \ J" "(SUP j\J. M j) = M j" by auto with \J \ A\ show "\j\A. emeasure (Sup (M ` J)) X \ emeasure (M j) X" by auto next fix j assume "j\A" then show "\i\{J. J \ {} \ finite J \ J \ A}. emeasure (M j) X \ emeasure (Sup (M ` i)) X" by (intro bexI[of _ "{j}"]) auto qed qed subsubsection\<^marker>\tag unimportant\ \Supremum of a set of \\\-algebras\ lemma space_Sup_eq_UN: "space (Sup M) = (\x\M. space x)" (is "?L=?R") proof show "?L \ ?R" using Sup_lexord[where P="\x. space x = _"] apply (clarsimp simp: Sup_measure_def) by (smt (verit) Sup_lexord_def UN_E mem_Collect_eq space_Sup_measure'2 space_measure_of_conv) qed (use Sup_upper le_measureD1 in fastforce) lemma sets_Sup_eq: assumes *: "\m. m \ M \ space m = X" and "M \ {}" shows "sets (Sup M) = sigma_sets X (\x\M. sets x)" unfolding Sup_measure_def proof (rule Sup_lexord1 [OF \M \ {}\]) show "sets (Sup_lexord sets Sup_measure' (\U. sigma (\ (space ` U)) (\ (sets ` U))) M) = sigma_sets X (\ (sets ` M))" apply (rule Sup_lexord) apply (metis (mono_tags, lifting) "*" empty_iff mem_Collect_eq sets.sigma_sets_eq sets_Sup_measure') by (metis "*" SUP_eq_const UN_space_closed assms(2) sets_measure_of) qed (use * in blast) lemma in_sets_Sup: "(\m. m \ M \ space m = X) \ m \ M \ A \ sets m \ A \ sets (Sup M)" by (subst sets_Sup_eq[where X=X]) auto lemma Sup_lexord_rel: assumes "\i. i \ I \ k (A i) = k (B i)" "R (c (A ` {a \ I. k (B a) = (SUP x\I. k (B x))})) (c (B ` {a \ I. k (B a) = (SUP x\I. k (B x))}))" "R (s (A`I)) (s (B`I))" shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))" proof - have "A ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ A ` I. k a = (SUP x\I. k (B x))}" using assms(1) by auto moreover have "B ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ B ` I. k a = (SUP x\I. k (B x))}" by auto ultimately show ?thesis using assms by (auto simp: Sup_lexord_def Let_def image_comp) qed lemma sets_SUP_cong: assumes eq: "\i. i \ I \ sets (M i) = sets (N i)" shows "sets (SUP i\I. M i) = sets (SUP i\I. N i)" unfolding Sup_measure_def using eq eq[THEN sets_eq_imp_space_eq] by (intro Sup_lexord_rel[where R="\x y. sets x = sets y"], simp_all add: sets_Sup_measure'2) lemma sets_Sup_in_sets: assumes "M \ {}" assumes "\m. m \ M \ space m = space N" assumes "\m. m \ M \ sets m \ sets N" shows "sets (Sup M) \ sets N" proof - have *: "\(space ` M) = space N" using assms by auto show ?thesis unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset) qed lemma measurable_Sup1: assumes m: "m \ M" and f: "f \ measurable m N" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable (Sup M) N" proof - have "space (Sup M) = space m" using m by (auto simp add: space_Sup_eq_UN dest: const_space) then show ?thesis using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space]) qed lemma measurable_Sup2: assumes M: "M \ {}" assumes f: "\m. m \ M \ f \ measurable N m" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable N (Sup M)" proof - from M obtain m where "m \ M" by auto have space_eq: "\n. n \ M \ space n = space m" by (intro const_space \m \ M\) have eq: "sets (sigma (\ (space ` M)) (\ (sets ` M))) = sets (Sup M)" by (metis M SUP_eq_const UN_space_closed sets_Sup_eq sets_measure_of space_eq) have "f \ measurable N (sigma (\m\M. space m) (\m\M. sets m))" proof (rule measurable_measure_of) show "f \ space N \ \(space ` M)" using measurable_space[OF f] M by auto qed (auto intro: measurable_sets f dest: sets.sets_into_space) also have "measurable N (sigma (\m\M. space m) (\m\M. sets m)) = measurable N (Sup M)" using eq measurable_cong_sets by blast finally show ?thesis . qed lemma measurable_SUP2: "I \ {} \ (\i. i \ I \ f \ measurable N (M i)) \ (\i j. i \ I \ j \ I \ space (M i) = space (M j)) \ f \ measurable N (SUP i\I. M i)" by (auto intro!: measurable_Sup2) lemma sets_Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" proof - { fix a m assume "a \ sigma_sets \ m" "m \ M" then have "a \ sigma_sets \ (\M)" by induction (auto intro: sigma_sets.intros(2-)) } then have "sigma_sets \ (\ (sigma_sets \ ` M)) = sigma_sets \ (\ M)" by (smt (verit, best) UN_iff Union_iff sigma_sets.Basic sigma_sets_eqI) then show "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" by (subst sets_Sup_eq) (fastforce simp add: M Union_least)+ qed lemma Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "(SUP m\M. sigma \ m) = (sigma \ (\M))" proof (intro antisym SUP_least) have *: "\M \ Pow \" using M by auto show "sigma \ (\M) \ (SUP m\M. sigma \ m)" proof (intro less_eq_measure.intros(3)) show "space (sigma \ (\M)) = space (SUP m\M. sigma \ m)" "sets (sigma \ (\M)) = sets (SUP m\M. sigma \ m)" by (auto simp add: M sets_Sup_sigma sets_eq_imp_space_eq space_measure_of_conv) qed (simp add: emeasure_sigma le_fun_def) fix m assume "m \ M" then show "sigma \ m \ sigma \ (\M)" by (subst sigma_le_iff) (auto simp add: M *) qed lemma SUP_sigma_sigma: "M \ {} \ (\m. m \ M \ f m \ Pow \) \ (SUP m\M. sigma \ (f m)) = sigma \ (\m\M. f m)" using Sup_sigma[of "f`M" \] by (auto simp: image_comp) lemma sets_vimage_Sup_eq: assumes *: "M \ {}" "f \ X \ Y" "\m. m \ M \ space m = Y" shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \ M. vimage_algebra X f m)" (is "?L = ?R") proof have "\m. m \ M \ f \ Sup (vimage_algebra X f ` M) \\<^sub>M m" using assms by (smt (verit, del_insts) Pi_iff imageE image_eqI measurable_Sup1 measurable_vimage_algebra1 space_vimage_algebra) then show "?L \ ?R" by (intro sets_image_in_sets measurable_Sup2) (simp_all add: space_Sup_eq_UN *) show "?R \ ?L" apply (intro sets_Sup_in_sets) apply (force simp add: * space_Sup_eq_UN sets_vimage_algebra2 intro: in_sets_Sup)+ done qed lemma restrict_space_eq_vimage_algebra': "sets (restrict_space M \) = sets (vimage_algebra (\ \ space M) (\x. x) M)" proof - have *: "{A \ (\ \ space M) |A. A \ sets M} = {A \ \ |A. A \ sets M}" using sets.sets_into_space[of _ M] by blast show ?thesis unfolding restrict_space_def by (subst sets_measure_of) (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets]) qed lemma sigma_le_sets: assumes [simp]: "A \ Pow X" shows "sets (sigma X A) \ sets N \ X \ sets N \ A \ sets N" proof have "X \ sigma_sets X A" "A \ sigma_sets X A" by (auto intro: sigma_sets_top) moreover assume "sets (sigma X A) \ sets N" ultimately show "X \ sets N \ A \ sets N" by auto next assume *: "X \ sets N \ A \ sets N" { fix Y assume "Y \ sigma_sets X A" from this * have "Y \ sets N" by induction auto } then show "sets (sigma X A) \ sets N" by auto qed lemma measurable_iff_sets: "f \ measurable M N \ (f \ space M \ space N \ sets (vimage_algebra (space M) f N) \ sets M)" unfolding measurable_def by (smt (verit, ccfv_threshold) mem_Collect_eq sets_vimage_algebra sigma_sets_le_sets_iff subset_eq) lemma sets_vimage_algebra_space: "X \ sets (vimage_algebra X f M)" using sets.top[of "vimage_algebra X f M"] by simp lemma measurable_mono: assumes N: "sets N' \ sets N" "space N = space N'" assumes M: "sets M \ sets M'" "space M = space M'" shows "measurable M N \ measurable M' N'" unfolding measurable_def proof safe fix f A assume "f \ space M \ space N" "A \ sets N'" moreover assume "\y\sets N. f -` y \ space M \ sets M" note this[THEN bspec, of A] ultimately show "f -` A \ space M' \ sets M'" using assms by auto qed (use N M in auto) lemma measurable_Sup_measurable: assumes f: "f \ space N \ A" shows "f \ measurable N (Sup {M. space M = A \ f \ measurable N M})" proof (rule measurable_Sup2) show "{M. space M = A \ f \ measurable N M} \ {}" using f unfolding ex_in_conv[symmetric] by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of) qed auto lemma (in sigma_algebra) sigma_sets_subset': assumes a: "a \ M" "\' \ M" shows "sigma_sets \' a \ M" proof show "x \ M" if x: "x \ sigma_sets \' a" for x using x by (induct rule: sigma_sets.induct) (use a in auto) qed lemma in_sets_SUP: "i \ I \ (\i. i \ I \ space (M i) = Y) \ X \ sets (M i) \ X \ sets (SUP i\I. M i)" by (intro in_sets_Sup[where X=Y]) auto lemma measurable_SUP1: "i \ I \ f \ measurable (M i) N \ (\m n. m \ I \ n \ I \ space (M m) = space (M n)) \ f \ measurable (SUP i\I. M i) N" by (auto intro: measurable_Sup1) lemma sets_image_in_sets': assumes X: "X \ sets N" assumes f: "\A. A \ sets M \ f -` A \ X \ sets N" shows "sets (vimage_algebra X f M) \ sets N" unfolding sets_vimage_algebra by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f) lemma mono_vimage_algebra: "sets M \ sets N \ sets (vimage_algebra X f M) \ sets (vimage_algebra X f N)" using sets.top[of "sigma X {f -` A \ X |A. A \ sets N}"] unfolding vimage_algebra_def by (smt (verit, del_insts) space_measure_of sigma_le_sets Pow_iff inf_le2 mem_Collect_eq subset_eq) lemma mono_restrict_space: "sets M \ sets N \ sets (restrict_space M X) \ sets (restrict_space N X)" unfolding sets_restrict_space by (rule image_mono) lemma sets_eq_bot: "sets M = {{}} \ M = bot" by (metis measure_eqI emeasure_empty sets_bot singletonD) lemma sets_eq_bot2: "{{}} = sets M \ M = bot" using sets_eq_bot[of M] by blast lemma (in finite_measure) countable_support: "countable {x. measure M {x} \ 0}" proof cases assume "measure M (space M) = 0" then show ?thesis by (metis (mono_tags, lifting) bounded_measure measure_le_0_iff Collect_empty_eq countable_empty) next let ?M = "measure M (space M)" and ?m = "\x. measure M {x}" assume "?M \ 0" then have *: "{x. ?m x \ 0} = (\n. {x. ?M / Suc n < ?m x})" using reals_Archimedean[of "?m x / ?M" for x] by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff) have **: "\n. finite {x. ?M / Suc n < ?m x}" proof (rule ccontr) fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X") then obtain X where "finite X" "card X = Suc (Suc n)" "X \ ?X" by (metis infinite_arbitrarily_large) then have *: "\x. x \ X \ ?M / Suc n \ ?m x" by auto { fix x assume "x \ X" from \?M \ 0\ *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) then have "{x} \ sets M" by (auto dest: measure_notin_sets) } note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" using \?M \ 0\ by (simp add: \card X = Suc (Suc n)\ field_simps less_le) also have "\ \ (\x\X. ?m x)" by (rule sum_mono) fact also have "\ = measure M (\x\X. {x})" using singleton_sets \finite X\ by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) finally have "?M < measure M (\x\X. {x})" . moreover have "measure M (\x\X. {x}) \ ?M" using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto ultimately show False by simp qed show ?thesis unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) qed 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,643 +1,684 @@ section\The binary product topology\ theory Product_Topology imports Function_Topology begin 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 +text \Missing the opposite direction. Does it hold? A converse is proved for proper maps, a stronger condition\ +lemma closed_map_prod: + assumes "closed_map (prod_topology X Y) (prod_topology X' Y') (\(x,y). (f x, g y))" + shows "topspace(prod_topology X Y) = {} \ closed_map X X' f \ closed_map Y Y' g" +proof (cases "topspace(prod_topology X Y) = {}") + case False + then have ne: "topspace X \ {}" "topspace Y \ {}" + by auto + have "closed_map X X' f" + unfolding closed_map_def + proof (intro strip) + fix C + assume "closedin X C" + show "closedin X' (f ` C)" + proof (cases "C={}") + case False + with assms have "closedin (prod_topology X' Y') ((\(x,y). (f x, g y)) ` (C \ topspace Y))" + by (simp add: \closedin X C\ closed_map_def closedin_prod_Times_iff) + with False ne show ?thesis + by (simp add: image_paired_Times closedin_Times closedin_prod_Times_iff) + qed auto + qed + moreover + have "closed_map Y Y' g" + unfolding closed_map_def + proof (intro strip) + fix C + assume "closedin Y C" + show "closedin Y' (g ` C)" + proof (cases "C={}") + case False + with assms have "closedin (prod_topology X' Y') ((\(x,y). (f x, g y)) ` (topspace X \ C))" + by (simp add: \closedin Y C\ closed_map_def closedin_prod_Times_iff) + with False ne show ?thesis + by (simp add: image_paired_Times closedin_Times closedin_prod_Times_iff) + qed auto + qed + ultimately show ?thesis + by (auto simp: False) +qed auto + 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) 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) 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 homeomorphic_space_prod_topology_swap: "(prod_topology X Y) homeomorphic_space (prod_topology Y X)" using homeomorphic_map_swap homeomorphic_space by blast 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 proposition connected_space_prod_topology: "connected_space(prod_topology X Y) \ topspace(prod_topology X Y) = {} \ connected_space X \ connected_space Y" (is "?lhs=?rhs") proof (cases "topspace(prod_topology X Y) = {}") case True then show ?thesis using connected_space_topspace_empty by blast next case False then have nonempty: "topspace X \ {}" "topspace Y \ {}" by force+ show ?thesis proof assume ?lhs then show ?rhs by (meson connected_space_quotient_map_image nonempty quotient_map_fst quotient_map_snd) next assume ?rhs then have conX: "connected_space X" and conY: "connected_space Y" using False by blast+ have False if "openin (prod_topology X Y) U" and "openin (prod_topology X Y) V" and UV: "topspace X \ topspace Y \ U \ V" "U \ V = {}" and "U \ {}" and "V \ {}" for U V proof - have Usub: "U \ topspace X \ topspace Y" and Vsub: "V \ topspace X \ topspace Y" using that by (metis openin_subset topspace_prod_topology)+ obtain a b where ab: "(a,b) \ U" and a: "a \ topspace X" and b: "b \ topspace Y" using \U \ {}\ Usub by auto have "\ topspace X \ topspace Y \ U" using Usub Vsub \U \ V = {}\ \V \ {}\ by auto then obtain x y where x: "x \ topspace X" and y: "y \ topspace Y" and "(x,y) \ U" by blast have oX: "openin X {x \ topspace X. (x,y) \ U}" "openin X {x \ topspace X. (x,y) \ V}" and oY: "openin Y {y \ topspace Y. (a,y) \ U}" "openin Y {y \ topspace Y. (a,y) \ V}" by (force intro: openin_continuous_map_preimage [where Y = "prod_topology X Y"] simp: that continuous_map_pairwise o_def x y a)+ have 1: "topspace Y \ {y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V}" using a that(3) by auto have 2: "{y \ topspace Y. (a,y) \ U} \ {y \ topspace Y. (a,y) \ V} = {}" using that(4) by auto have 3: "{y \ topspace Y. (a,y) \ U} \ {}" using ab b by auto have 4: "{y \ topspace Y. (a,y) \ V} \ {}" proof - show ?thesis using connected_spaceD [OF conX oX] UV \(x,y) \ U\ a x y disjoint_iff_not_equal by blast qed show ?thesis using connected_spaceD [OF conY oY 1 2 3 4] by auto qed then show ?lhs unfolding connected_space_def topspace_prod_topology by blast qed qed lemma connectedin_Times: "connectedin (prod_topology X Y) (S \ T) \ S = {} \ T = {} \ connectedin X S \ connectedin Y T" by (force simp: connectedin_def subtopology_Times connected_space_prod_topology) 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,750 +1,756 @@ section\T1 and Hausdorff spaces\ theory T1_Spaces 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 continuous_closed_imp_proper_map: + "\compact_space X; t1_space Y; continuous_map X Y f; closed_map X Y f\ \ proper_map X Y f" + unfolding proper_map_def + by (smt (verit) closedin_compact_space closedin_continuous_map_preimage + Collect_cong singleton_iff t1_space_closedin_singleton) + lemma t1_space_euclidean: "t1_space (euclidean :: 'a::metric_space topology)" by (simp add: t1_space_closedin_singleton) 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 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 by (metis Hausdorff_space_compact_sets Hausdorff_space_def compactin_discrete_topology equalityE openin_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 infinite_perfect_set: "\Hausdorff_space X; S \ X derived_set_of S; S \ {}\ \ infinite S" using derived_set_of_finite by blast 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_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 lemma Hausdorff_space_closed_neighbourhood: "Hausdorff_space X \ (\x \ topspace X. \U C. openin X U \ closedin X C \ Hausdorff_space(subtopology X C) \ x \ U \ U \ C)" (is "_ = ?rhs") proof assume R: ?rhs show "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" obtain T C where *: "openin X T" "closedin X C" "x \ T" "T \ C" and C: "Hausdorff_space (subtopology X C)" by (meson R \x \ topspace X\) show "\U V. openin X U \ openin X V \ x \ U \ y \ V \ disjnt U V" proof (cases "y \ C") case True with * C obtain U V where U: "openin (subtopology X C) U" and V: "openin (subtopology X C) V" and "x \ U" "y \ V" "disjnt U V" unfolding Hausdorff_space_def by (smt (verit, best) \x \ y\ closedin_subset subsetD topspace_subtopology_subset) then obtain U' V' where UV': "U = U' \ C" "openin X U'" "V = V' \ C" "openin X V'" by (meson openin_subtopology) have "disjnt (T \ U') V'" using \disjnt U V\ UV' \T \ C\ by (force simp: disjnt_iff) with \T \ C\ have "disjnt (T \ U') (V' \ (topspace X - C))" unfolding disjnt_def by blast moreover have "openin X (T \ U')" by (simp add: \openin X T\ \openin X U'\ openin_Int) moreover have "openin X (V' \ (topspace X - C))" using \closedin X C\ \openin X V'\ by auto ultimately show ?thesis using UV' \x \ T\ \x \ U\ \y \ V\ by blast next case False with * y show ?thesis by (force simp: closedin_def disjnt_def) qed qed qed fastforce end diff --git a/src/HOL/Nat.thy b/src/HOL/Nat.thy --- a/src/HOL/Nat.thy +++ b/src/HOL/Nat.thy @@ -1,2572 +1,2583 @@ (* Title: HOL/Nat.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Natural numbers\ theory Nat imports Inductive Typedef Fun Rings begin subsection \Type \ind\\ typedecl ind axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \ ind" \ \The axiom of infinity in 2 parts:\ where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \ x = y" and Suc_Rep_not_Zero_Rep: "Suc_Rep x \ Zero_Rep" subsection \Type nat\ text \Type definition\ inductive Nat :: "ind \ bool" where Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ Nat (Suc_Rep i)" typedef nat = "{n. Nat n}" morphisms Rep_Nat Abs_Nat using Nat.Zero_RepI by auto lemma Nat_Rep_Nat: "Nat (Rep_Nat n)" using Rep_Nat by simp lemma Nat_Abs_Nat_inverse: "Nat n \ Rep_Nat (Abs_Nat n) = n" using Abs_Nat_inverse by simp lemma Nat_Abs_Nat_inject: "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" using Abs_Nat_inject by simp instantiation nat :: zero begin definition Zero_nat_def: "0 = Abs_Nat Zero_Rep" instance .. end definition Suc :: "nat \ nat" where "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" lemma Suc_not_Zero: "Suc m \ 0" by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat) lemma Zero_not_Suc: "0 \ Suc m" by (rule not_sym) (rule Suc_not_Zero) lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \ x = y" by (rule iffI, rule Suc_Rep_inject) simp_all lemma nat_induct0: assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" proof - have "P (Abs_Nat (Rep_Nat n))" using assms unfolding Zero_nat_def Suc_def by (iprover intro: Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst]) then show ?thesis by (simp add: Rep_Nat_inverse) qed free_constructors case_nat for "0 :: nat" | Suc pred where "pred (0 :: nat) = (0 :: nat)" proof atomize_elim fix n show "n = 0 \ (\m. n = Suc m)" by (induction n rule: nat_induct0) auto next fix n m show "(Suc n = Suc m) = (n = m)" by (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) next fix n show "0 \ Suc n" by (simp add: Suc_not_Zero) qed \ \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "0 :: nat" Suc by (erule nat_induct0) auto setup \Sign.parent_path\ \ \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "nat"\ declare old.nat.inject[iff del] and old.nat.distinct(1)[simp del, induct_simp del] lemmas induct = old.nat.induct lemmas inducts = old.nat.inducts lemmas rec = old.nat.rec lemmas simps = nat.inject nat.distinct nat.case nat.rec setup \Sign.parent_path\ abbreviation rec_nat :: "'a \ (nat \ 'a \ 'a) \ nat \ 'a" where "rec_nat \ old.rec_nat" declare nat.sel[code del] hide_const (open) Nat.pred \ \hide everything related to the selector\ hide_fact nat.case_eq_if nat.collapse nat.expand nat.sel nat.exhaust_sel nat.split_sel nat.split_sel_asm lemma nat_exhaust [case_names 0 Suc, cases type: nat]: "(y = 0 \ P) \ (\nat. y = Suc nat \ P) \ P" \ \for backward compatibility -- names of variables differ\ by (rule old.nat.exhaust) lemma nat_induct [case_names 0 Suc, induct type: nat]: fixes n assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" \ \for backward compatibility -- names of variables differ\ using assms by (rule nat.induct) hide_fact nat_exhaust nat_induct0 ML \ val nat_basic_lfp_sugar = let val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}} end; \ setup \ let fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; in BNF_LFP_Rec_Sugar.register_lfp_rec_extension {nested_simps = [], special_endgame_tac = K (K (K (K no_tac))), is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of, rewrite_nested_rec_call = NONE} end \ text \Injectiveness and distinctness lemmas\ lemma inj_Suc [simp]: "inj_on Suc N" by (simp add: inj_on_def) lemma bij_betw_Suc [simp]: "bij_betw Suc M N \ Suc ` M = N" by (simp add: bij_betw_def) lemma Suc_neq_Zero: "Suc m = 0 \ R" by (rule notE) (rule Suc_not_Zero) lemma Zero_neq_Suc: "0 = Suc m \ R" by (rule Suc_neq_Zero) (erule sym) lemma Suc_inject: "Suc x = Suc y \ x = y" by (rule inj_Suc [THEN injD]) lemma n_not_Suc_n: "n \ Suc n" by (induct n) simp_all lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" and "\x y. P x y \ P (Suc x) (Suc y)" shows "P m n" proof (induct n arbitrary: m) case 0 show ?case by (rule assms(1)) next case (Suc n) show ?case proof (induct m) case 0 show ?case by (rule assms(2)) next case (Suc m) from \P m n\ show ?case by (rule assms(3)) qed qed subsection \Arithmetic operators\ instantiation nat :: comm_monoid_diff begin primrec plus_nat where add_0: "0 + n = (n::nat)" | add_Suc: "Suc m + n = Suc (m + n)" lemma add_0_right [simp]: "m + 0 = m" for m :: nat by (induct m) simp_all lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)" by (induct m) simp_all declare add_0 [code] lemma add_Suc_shift [code]: "Suc m + n = m + Suc n" by simp primrec minus_nat where diff_0 [code]: "m - 0 = (m::nat)" | diff_Suc: "m - Suc n = (case m - n of 0 \ 0 | Suc k \ k)" declare diff_Suc [simp del] lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat by (induct n) (simp_all add: diff_Suc) lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n" by (induct n) (simp_all add: diff_Suc) instance proof fix n m q :: nat show "(n + m) + q = n + (m + q)" by (induct n) simp_all show "n + m = m + n" by (induct n) simp_all show "m + n - m = n" by (induct m) simp_all show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc) show "0 + n = n" by simp show "0 - n = 0" by simp qed end hide_fact (open) add_0 add_0_right diff_0 instantiation nat :: comm_semiring_1_cancel begin definition One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0::nat)" | mult_Suc: "Suc m * n = n + (m * n)" lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat by (induct m) (simp_all add: add.assoc) instance proof fix k n m q :: nat show "0 \ (1::nat)" by simp show "1 * n = n" by simp show "n * m = m * n" by (induct n) simp_all show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib) show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib) show "k * (m - n) = (k * m) - (k * n)" by (induct m n rule: diff_induct) simp_all qed end subsubsection \Addition\ text \Reasoning about \m + 0 = 0\, etc.\ lemma add_is_0 [iff]: "m + n = 0 \ m = 0 \ n = 0" for m n :: nat by (cases m) simp_all lemma add_is_1: "m + n = Suc 0 \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (cases m) simp_all lemma one_is_add: "Suc 0 = m + n \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (rule trans, rule eq_commute, rule add_is_1) lemma add_eq_self_zero: "m + n = m \ n = 0" for m n :: nat by (induct m) simp_all lemma plus_1_eq_Suc: "plus 1 = Suc" by (simp add: fun_eq_iff) lemma Suc_eq_plus1: "Suc n = n + 1" by simp lemma Suc_eq_plus1_left: "Suc n = 1 + n" by simp subsubsection \Difference\ lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_add) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" by simp subsubsection \Multiplication\ lemma mult_is_0 [simp]: "m * n = 0 \ m = 0 \ n = 0" for m n :: nat by (induct m) auto lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \ m = Suc 0 \ n = Suc 0" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (induct n) auto qed lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \ m = Suc 0 \ n = Suc 0" by (simp add: eq_commute flip: mult_eq_1_iff) lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" and nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" for m n :: nat by auto lemma mult_cancel1 [simp]: "k * m = k * n \ m = n \ k = 0" for k m n :: nat proof - have "k \ 0 \ k * m = k * n \ m = n" proof (induct n arbitrary: m) case 0 then show "m = 0" by simp next case (Suc n) then show "m = Suc n" by (cases m) (simp_all add: eq_commute [of 0]) qed then show ?thesis by auto qed lemma mult_cancel2 [simp]: "m * k = n * k \ m = n \ k = 0" for k m n :: nat by (simp add: mult.commute) lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \ m = n" by (subst mult_cancel1) simp subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ instantiation nat :: linorder begin primrec less_eq_nat where "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] lemma le0 [iff]: "0 \ n" for n :: nat by (simp add: less_eq_nat.simps) lemma [code]: "0 \ n \ True" for n :: nat by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" lemma Suc_le_mono [iff]: "Suc n \ Suc m \ n \ m" by (simp add: less_eq_nat.simps(2)) lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. lemma le_0_eq [iff]: "n \ 0 \ n = 0" for n :: nat by (induct n) (simp_all add: less_eq_nat.simps(2)) lemma not_less0 [iff]: "\ n < 0" for n :: nat by (simp add: less_eq_Suc_le) lemma less_nat_zero_code [code]: "n < 0 \ False" for n :: nat by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" by (simp add: less_eq_Suc_le) lemma less_Suc_eq_le [code]: "m < Suc n \ m \ n" by (simp add: less_eq_Suc_le) lemma Suc_less_eq2: "Suc n < m \ (\m'. m = Suc m' \ n < m')" by (cases m) auto lemma le_SucI: "m \ n \ m \ Suc n" by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits) lemma Suc_leD: "Suc m \ n \ m \ n" by (cases n) (auto intro: le_SucI) lemma less_SucI: "m < n \ m < Suc n" by (simp add: less_eq_Suc_le) (erule Suc_leD) lemma Suc_lessD: "Suc m < n \ m < n" by (simp add: less_eq_Suc_le) (erule Suc_leD) instance proof fix n m q :: nat show "n < m \ n \ m \ \ m \ n" proof (induct n arbitrary: m) case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le) next case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le) qed show "n \ n" by (induct n) simp_all then show "n = m" if "n \ m" and "m \ n" using that by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) show "n \ q" if "n \ m" and "m \ q" using that proof (induct n arbitrary: m q) case 0 show ?case by simp next case (Suc n) then show ?case by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits) qed show "n \ m \ m \ n" by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) qed end instantiation nat :: order_bot begin definition bot_nat :: nat where "bot_nat = 0" instance by standard (simp add: bot_nat_def) end instance nat :: no_top by standard (auto intro: less_Suc_eq_le [THEN iffD2]) subsubsection \Introduction properties\ lemma lessI [iff]: "n < Suc n" by (simp add: less_Suc_eq_le) lemma zero_less_Suc [iff]: "0 < Suc n" by (simp add: less_Suc_eq_le) subsubsection \Elimination properties\ lemma less_not_refl: "\ n < n" for n :: nat by (rule order_less_irrefl) lemma less_not_refl2: "n < m \ m \ n" for m n :: nat by (rule not_sym) (rule less_imp_neq) lemma less_not_refl3: "s < t \ s \ t" for s t :: nat by (rule less_imp_neq) lemma less_irrefl_nat: "n < n \ R" for n :: nat by (rule notE, rule less_not_refl) lemma less_zeroE: "n < 0 \ R" for n :: nat by (rule notE) (rule not_less0) lemma less_Suc_eq: "m < Suc n \ m < n \ m = n" unfolding less_Suc_eq_le le_less .. lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)" by (simp add: less_Suc_eq) lemma less_one [iff]: "n < 1 \ n = 0" for n :: nat unfolding One_nat_def by (rule less_Suc0) lemma Suc_mono: "m < n \ Suc m < Suc n" by simp text \"Less than" is antisymmetric, sort of.\ lemma less_antisym: "\ n < m \ n < Suc m \ m = n" unfolding not_less less_Suc_eq_le by (rule antisym) lemma nat_neq_iff: "m \ n \ m < n \ n < m" for m n :: nat by (rule linorder_neq_iff) subsubsection \Inductive (?) properties\ lemma Suc_lessI: "m < n \ Suc m \ n \ Suc m < n" unfolding less_eq_Suc_le [of m] le_less by simp lemma lessE: assumes major: "i < k" and 1: "k = Suc i \ P" and 2: "\j. i < j \ k = Suc j \ P" shows P proof - from major have "\j. i \ j \ k = Suc j" unfolding less_eq_Suc_le by (induct k) simp_all then have "(\j. i < j \ k = Suc j) \ k = Suc i" by (auto simp add: less_le) with 1 2 show P by auto qed lemma less_SucE: assumes major: "m < Suc n" and less: "m < n \ P" and eq: "m = n \ P" shows P proof (rule major [THEN lessE]) show "Suc n = Suc m \ P" using eq by blast show "\j. \m < j; Suc n = Suc j\ \ P" by (blast intro: less) qed lemma Suc_lessE: assumes major: "Suc i < k" and minor: "\j. i < j \ k = Suc j \ P" shows P proof (rule major [THEN lessE]) show "k = Suc (Suc i) \ P" using lessI minor by iprover show "\j. \Suc i < j; k = Suc j\ \ P" using Suc_lessD minor by iprover qed lemma Suc_less_SucD: "Suc m < Suc n \ m < n" by simp lemma less_trans_Suc: assumes le: "i < j" shows "j < k \ Suc i < k" proof (induct k) case 0 then show ?case by simp next case (Suc k) with le show ?case by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) lemma not_less_eq_eq: "\ m \ n \ Suc n \ m" by (simp only: not_le Suc_le_eq) text \Properties of "less than or equal".\ lemma le_imp_less_Suc: "m \ n \ m < Suc n" by (simp only: less_Suc_eq_le) lemma Suc_n_not_le_n: "\ Suc n \ n" by (simp add: not_le less_Suc_eq_le) lemma le_Suc_eq: "m \ Suc n \ m \ n \ m = Suc n" by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq) lemma le_SucE: "m \ Suc n \ (m \ n \ R) \ (m = Suc n \ R) \ R" by (drule le_Suc_eq [THEN iffD1], iprover+) lemma Suc_leI: "m < n \ Suc m \ n" by (simp only: Suc_le_eq) text \Stronger version of \Suc_leD\.\ lemma Suc_le_lessD: "Suc m \ n \ m < n" by (simp only: Suc_le_eq) lemma less_imp_le_nat: "m < n \ m \ n" for m n :: nat unfolding less_eq_Suc_le by (rule Suc_leD) text \For instance, \(Suc m < Suc n) = (Suc m \ n) = (m < n)\\ lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq text \Equivalence of \m \ n\ and \m < n \ m = n\\ lemma less_or_eq_imp_le: "m < n \ m = n \ m \ n" for m n :: nat unfolding le_less . lemma le_eq_less_or_eq: "m \ n \ m < n \ m = n" for m n :: nat by (rule le_less) text \Useful with \blast\.\ lemma eq_imp_le: "m = n \ m \ n" for m n :: nat by auto lemma le_refl: "n \ n" for n :: nat by simp lemma le_trans: "i \ j \ j \ k \ i \ k" for i j k :: nat by (rule order_trans) lemma le_antisym: "m \ n \ n \ m \ m = n" for m n :: nat by (rule antisym) lemma nat_less_le: "m < n \ m \ n \ m \ n" for m n :: nat by (rule less_le) lemma le_neq_implies_less: "m \ n \ m \ n \ m < n" for m n :: nat unfolding less_le .. lemma nat_le_linear: "m \ n \ n \ m" for m n :: nat by (rule linear) lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat] lemma le_less_Suc_eq: "m \ n \ n < Suc m \ n = m" unfolding less_Suc_eq_le by auto lemma not_less_less_Suc_eq: "\ n < m \ n < Suc m \ n = m" unfolding not_less by (rule le_less_Suc_eq) lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq lemma not0_implies_Suc: "n \ 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr0_implies_Suc: "n > 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr_implies_not0: "m < n \ n \ 0" for m n :: nat by (cases n) simp_all lemma neq0_conv[iff]: "n \ 0 \ 0 < n" for n :: nat by (cases n) simp_all text \This theorem is useful with \blast\\ lemma gr0I: "(n = 0 \ False) \ 0 < n" for n :: nat by (rule neq0_conv[THEN iffD1]) iprover lemma gr0_conv_Suc: "0 < n \ (\m. n = Suc m)" by (fast intro: not0_implies_Suc) lemma not_gr0 [iff]: "\ 0 < n \ n = 0" for n :: nat using neq0_conv by blast lemma Suc_le_D: "Suc n \ m' \ \m. m' = Suc m" by (induct m') simp_all text \Useful in certain inductive arguments\ lemma less_Suc_eq_0_disj: "m < Suc n \ m = 0 \ (\j. m = Suc j \ j < n)" by (cases m) simp_all lemma All_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma All_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) lemma Ex_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma Ex_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) text \@{term mono} (non-strict) doesn't imply increasing, as the function could be constant\ lemma strict_mono_imp_increasing: fixes n::nat assumes "strict_mono f" shows "f n \ n" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case unfolding not_less_eq_eq [symmetric] using Suc_n_not_le_n assms order_trans strict_mono_less_eq by blast qed subsubsection \Monotonicity of Addition\ lemma Suc_pred [simp]: "n > 0 \ Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) lemma Suc_diff_1 [simp]: "0 < n \ Suc (n - 1) = n" unfolding One_nat_def by (rule Suc_pred) lemma nat_add_left_cancel_le [simp]: "k + m \ k + n \ m \ n" for k m n :: nat by (induct k) simp_all lemma nat_add_left_cancel_less [simp]: "k + m < k + n \ m < n" for k m n :: nat by (induct k) simp_all lemma add_gr_0 [iff]: "m + n > 0 \ m > 0 \ n > 0" for m n :: nat by (auto dest: gr0_implies_Suc) text \strict, in 1st argument\ lemma add_less_mono1: "i < j \ i + k < j + k" for i j k :: nat by (induct k) simp_all text \strict, in both arguments\ lemma add_less_mono: fixes i j k l :: nat assumes "i < j" "k < l" shows "i + k < j + l" proof - have "i + k < j + k" by (simp add: add_less_mono1 assms) also have "... < j + l" using \i < j\ by (induction j) (auto simp: assms) finally show ?thesis . qed lemma less_imp_Suc_add: "m < n \ \k. n = Suc (m + k)" proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: order_le_less) (blast elim!: less_SucE intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric]) qed lemma le_Suc_ex: "k \ l \ (\n. l = k + n)" for k l :: nat by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add) lemma less_natE: assumes \m < n\ obtains q where \n = Suc (m + q)\ using assms by (auto dest: less_imp_Suc_add intro: that) text \strict, in 1st argument; proof is by induction on \k > 0\\ lemma mult_less_mono2: fixes i j :: nat assumes "i < j" and "0 < k" shows "k * i < k * j" using \0 < k\ proof (induct k) case 0 then show ?case by simp next case (Suc k) with \i < j\ show ?case by (cases k) (simp_all add: add_less_mono) qed text \Addition is the inverse of subtraction: if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma nat_le_iff_add: "m \ n \ (\k. n = m + k)" for m n :: nat using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex) text \The naturals form an ordered \semidom\ and a \dioid\.\ instance nat :: linordered_semidom proof fix m n q :: nat show "0 < (1::nat)" by simp show "m \ n \ q + m \ q + n" by simp show "m < n \ 0 < q \ q * m < q * n" by (simp add: mult_less_mono2) show "m \ 0 \ n \ 0 \ m * n \ 0" by simp show "n \ m \ (m - n) + n = m" by (simp add: add_diff_inverse_nat add.commute linorder_not_less) qed instance nat :: dioid by standard (rule nat_le_iff_add) declare le0[simp del] \ \This is now @{thm zero_le}\ declare le_0_eq[simp del] \ \This is now @{thm le_zero_eq}\ declare not_less0[simp del] \ \This is now @{thm not_less_zero}\ declare not_gr0[simp del] \ \This is now @{thm not_gr_zero}\ instance nat :: ordered_cancel_comm_monoid_add .. instance nat :: ordered_cancel_comm_monoid_diff .. subsubsection \\<^term>\min\ and \<^term>\max\\ global_interpretation bot_nat_0: ordering_top \(\)\ \(>)\ \0::nat\ by standard simp global_interpretation max_nat: semilattice_neutr_order max \0::nat\ \(\)\ \(>)\ by standard (simp add: max_def) lemma mono_Suc: "mono Suc" by (rule monoI) simp lemma min_0L [simp]: "min 0 n = 0" for n :: nat by (rule min_absorb1) simp lemma min_0R [simp]: "min n 0 = 0" for n :: nat by (rule min_absorb2) simp lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)" by (simp add: mono_Suc min_of_mono) lemma min_Suc1: "min (Suc n) m = (case m of 0 \ 0 | Suc m' \ Suc(min n m'))" by (simp split: nat.split) lemma min_Suc2: "min m (Suc n) = (case m of 0 \ 0 | Suc m' \ Suc(min m' n))" by (simp split: nat.split) lemma max_0L [simp]: "max 0 n = n" for n :: nat by (fact max_nat.left_neutral) lemma max_0R [simp]: "max n 0 = n" for n :: nat by (fact max_nat.right_neutral) lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)" by (simp add: mono_Suc max_of_mono) lemma max_Suc1: "max (Suc n) m = (case m of 0 \ Suc n | Suc m' \ Suc (max n m'))" by (simp split: nat.split) lemma max_Suc2: "max m (Suc n) = (case m of 0 \ Suc n | Suc m' \ Suc (max m' n))" by (simp split: nat.split) lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat by (simp add: max_def) lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat by (simp add: max_def) lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ instance nat :: wellorder proof fix P and n :: nat assume step: "(\m. m < n \ P m) \ P n" for n :: nat have "\q. q \ n \ P q" proof (induct n) case (0 n) have "P 0" by (rule step) auto with 0 show ?case by auto next case (Suc m n) then have "n \ m \ n = Suc m" by (simp add: le_Suc_eq) then show ?case proof assume "n \ m" then show "P n" by (rule Suc(1)) next assume n: "n = Suc m" show "P n" by (rule step) (rule Suc(1), simp add: n le_simps) qed qed then show "P n" by auto qed lemma Least_eq_0[simp]: "P 0 \ Least P = 0" for P :: "nat \ bool" by (rule Least_equality[OF _ le0]) lemma Least_Suc: assumes "P n" "\ P 0" shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))" proof (cases n) case (Suc m) show ?thesis proof (rule antisym) show "(LEAST x. P x) \ Suc (LEAST x. P (Suc x))" using assms Suc by (force intro: LeastI Least_le) have \
: "P (LEAST x. P x)" by (blast intro: LeastI assms) show "Suc (LEAST m. P (Suc m)) \ (LEAST n. P n)" proof (cases "(LEAST n. P n)") case 0 then show ?thesis using \
by (simp add: assms) next case Suc with \
show ?thesis by (auto simp: Least_le) qed qed qed (use assms in auto) lemma Least_Suc2: "P n \ Q m \ \ P 0 \ \k. P (Suc k) = Q k \ Least P = Suc (Least Q)" by (erule (1) Least_Suc [THEN ssubst]) simp lemma ex_least_nat_le: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\k\n. (\i P i) \ P k" proof (cases n) case (Suc m) with assms show ?thesis by (blast intro: Least_le LeastI_ex dest: not_less_Least) qed (use assms in auto) lemma ex_least_nat_less: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\ki\k. \ P i) \ P (Suc k)" proof (cases n) case (Suc m) then obtain k where k: "k \ n" "\i P i" "P k" using ex_least_nat_le [OF assms] by blast show ?thesis by (cases k) (use assms k less_eq_Suc_le in auto) qed (use assms in auto) lemma nat_less_induct: fixes P :: "nat \ bool" assumes "\n. \m. m < n \ P m \ P n" shows "P n" using assms less_induct by blast lemma measure_induct_rule [case_names less]: fixes f :: "'a \ 'b::wellorder" assumes step: "\x. (\y. f y < f x \ P y) \ P x" shows "P a" by (induct m \ "f a" arbitrary: a rule: less_induct) (auto intro: step) text \old style induction rules:\ lemma measure_induct: fixes f :: "'a \ 'b::wellorder" shows "(\x. \y. f y < f x \ P y \ P x) \ P a" by (rule measure_induct_rule [of f P a]) iprover lemma full_nat_induct: assumes step: "\n. (\m. Suc m \ n \ P m) \ P n" shows "P n" by (rule less_induct) (auto intro: step simp:le_simps) text\An induction rule for establishing binary relations\ lemma less_Suc_induct [consumes 1]: assumes less: "i < j" and step: "\i. P i (Suc i)" and trans: "\i j k. i < j \ j < k \ P i j \ P j k \ P i k" shows "P i j" proof - from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add) have "P i (Suc (i + k))" proof (induct k) case 0 show ?case by (simp add: step) next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp then have "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed then show "P i j" by (simp add: j) qed text \ The method of infinite descent, frequently used in number theory. Provided by Roelof Oosterhuis. \P n\ is true for all natural numbers if \<^item> case ``0'': given \n = 0\ prove \P n\ \<^item> case ``smaller'': given \n > 0\ and \\ P n\ prove there exists a smaller natural number \m\ such that \\ P m\. \ lemma infinite_descent: "(\n. \ P n \ \m P m) \ P n" for P :: "nat \ bool" \ \compact version without explicit base case\ by (induct n rule: less_induct) auto lemma infinite_descent0 [case_names 0 smaller]: fixes P :: "nat \ bool" assumes "P 0" and "\n. n > 0 \ \ P n \ \m. m < n \ \ P m" shows "P n" proof (rule infinite_descent) fix n show "\ P n \ \m P m" using assms by (cases "n > 0") auto qed text \ Infinite descent using a mapping to \nat\: \P x\ is true for all \x \ D\ if there exists a \V \ D \ nat\ and \<^item> case ``0'': given \V x = 0\ prove \P x\ \<^item> ``smaller'': given \V x > 0\ and \\ P x\ prove there exists a \y \ D\ such that \V y < V x\ and \\ P y\. \ corollary infinite_descent0_measure [case_names 0 smaller]: fixes V :: "'a \ nat" assumes 1: "\x. V x = 0 \ P x" and 2: "\x. V x > 0 \ \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent0) case 0 with 1 show "P x" by auto next case (smaller n) then obtain x where *: "V x = n " and "V x > 0 \ \ P x" by auto with 2 obtain y where "V y < V x \ \ P y" by auto with * obtain m where "m = V y \ m < n \ \ P y" by auto then show ?case by auto qed ultimately show "P x" by auto qed text \Again, without explicit base case:\ lemma infinite_descent_measure: fixes V :: "'a \ nat" assumes "\x. \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - from assms obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof - have "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x using assms and that by auto then show "\x. V x = n \ P x" by (induct n rule: infinite_descent, auto) qed ultimately show "P x" by auto qed text \A (clumsy) way of lifting \<\ monotonicity to \\\ monotonicity\ lemma less_mono_imp_le_mono: fixes f :: "nat \ nat" and i j :: nat assumes "\i j::nat. i < j \ f i < f j" and "i \ j" shows "f i \ f j" using assms by (auto simp add: order_le_less) text \non-strict, in 1st argument\ lemma add_le_mono1: "i \ j \ i + k \ j + k" for i j k :: nat by (rule add_right_mono) text \non-strict, in both arguments\ lemma add_le_mono: "i \ j \ k \ l \ i + k \ j + l" for i j k l :: nat by (rule add_mono) lemma le_add2: "n \ m + n" for m n :: nat by simp lemma le_add1: "n \ n + m" for m n :: nat by simp lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) lemma less_add_Suc2: "i < Suc (m + i)" by (rule le_less_trans, rule le_add2, rule lessI) lemma less_iff_Suc_add: "m < n \ (\k. n = Suc (m + k))" by (iprover intro!: less_add_Suc1 less_imp_Suc_add) lemma trans_le_add1: "i \ j \ i \ j + m" for i j m :: nat by (rule le_trans, assumption, rule le_add1) lemma trans_le_add2: "i \ j \ i \ m + j" for i j m :: nat by (rule le_trans, assumption, rule le_add2) lemma trans_less_add1: "i < j \ i < j + m" for i j m :: nat by (rule less_le_trans, assumption, rule le_add1) lemma trans_less_add2: "i < j \ i < m + j" for i j m :: nat by (rule less_le_trans, assumption, rule le_add2) lemma add_lessD1: "i + j < k \ i < k" for i j k :: nat by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1) lemma not_add_less1 [iff]: "\ i + j < i" for i j :: nat by simp lemma not_add_less2 [iff]: "\ j + i < i" for i j :: nat by simp lemma add_leD1: "m + k \ n \ m \ n" for k m n :: nat by (rule order_trans [of _ "m + k"]) (simp_all add: le_add1) lemma add_leD2: "m + k \ n \ k \ n" for k m n :: nat by (force simp add: add.commute dest: add_leD1) lemma add_leE: "m + k \ n \ (m \ n \ k \ n \ R) \ R" for k m n :: nat by (blast dest: add_leD1 add_leD2) text \needs \\k\ for \ac_simps\ to work\ lemma less_add_eq_less: "\k. k < l \ m + l = k + n \ m < n" for l m n :: nat by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps) subsubsection \More results about difference\ lemma Suc_diff_le: "n \ m \ Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all lemma diff_less_Suc: "m - n < Suc m" by (induct m n rule: diff_induct) (auto simp: less_Suc_eq) lemma diff_le_self [simp]: "m - n \ m" for m n :: nat by (induct m n rule: diff_induct) (simp_all add: le_SucI) lemma less_imp_diff_less: "j < k \ j - n < k" for j k n :: nat by (rule le_less_trans, rule diff_le_self) lemma diff_Suc_less [simp]: "0 < n \ n - Suc i < n" by (cases n) (auto simp add: le_simps) lemma diff_add_assoc: "k \ j \ (i + j) - k = i + (j - k)" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc) lemma add_diff_assoc [simp]: "k \ j \ i + (j - k) = i + j - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc) lemma diff_add_assoc2: "k \ j \ (j + i) - k = (j - k) + i" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc2) lemma add_diff_assoc2 [simp]: "k \ j \ j - k + i = j + i - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc2) lemma le_imp_diff_is_add: "i \ j \ (j - i = k) = (j = k + i)" for i j k :: nat by auto lemma diff_is_0_eq [simp]: "m - n = 0 \ m \ n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma diff_is_0_eq' [simp]: "m \ n \ m - n = 0" for m n :: nat by (rule iffD2, rule diff_is_0_eq) lemma zero_less_diff [simp]: "0 < n - m \ m < n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma less_imp_add_positive: assumes "i < j" shows "\k::nat. 0 < k \ i + k = j" proof from assms show "0 < j - i \ i + (j - i) = j" by (simp add: order_less_imp_le) qed text \a nice rewrite for bounded subtraction\ lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat by (simp add: max_def not_le order_less_imp_le) lemma nat_diff_split: "P (a - b) \ (a < b \ P 0) \ (\d. a = b + d \ P d)" for a b :: nat \ \elimination of \-\ on \nat\\ by (cases "a < b") (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym]) lemma nat_diff_split_asm: "P (a - b) \ \ (a < b \ \ P 0 \ (\d. a = b + d \ \ P d))" for a b :: nat \ \elimination of \-\ on \nat\ in assumptions\ by (auto split: nat_diff_split) lemma Suc_pred': "0 < n \ n = Suc(n - 1)" by simp lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat by (cases m) simp_all lemma Suc_diff_eq_diff_pred: "0 < n \ Suc m - n = m - (n - 1)" by (cases n) simp_all lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (cases m) simp_all lemma Let_Suc [simp]: "Let (Suc n) f \ f (Suc n)" by (fact Let_def) subsubsection \Monotonicity of multiplication\ lemma mult_le_mono1: "i \ j \ i * k \ j * k" for i j k :: nat by (simp add: mult_right_mono) lemma mult_le_mono2: "i \ j \ k * i \ k * j" for i j k :: nat by (simp add: mult_left_mono) text \\\\ monotonicity, BOTH arguments\ lemma mult_le_mono: "i \ j \ k \ l \ i * k \ j * l" for i j k l :: nat by (simp add: mult_mono) lemma mult_less_mono1: "i < j \ 0 < k \ i * k < j * k" for i j k :: nat by (simp add: mult_strict_right_mono) text \Differs from the standard \zero_less_mult_iff\ in that there are no negative numbers.\ lemma nat_0_less_mult_iff [simp]: "0 < m * n \ 0 < m \ 0 < n" for m n :: nat proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma one_le_mult_iff [simp]: "Suc 0 \ m * n \ Suc 0 \ m \ Suc 0 \ n" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma mult_less_cancel2 [simp]: "m * k < n * k \ 0 < k \ m < n" for k m n :: nat proof (intro iffI conjI) assume m: "m * k < n * k" then show "0 < k" by (cases k) auto show "m < n" proof (cases k) case 0 then show ?thesis using m by auto next case (Suc k') then show ?thesis using m by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1) qed next assume "0 < k \ m < n" then show "m * k < n * k" by (blast intro: mult_less_mono1) qed lemma mult_less_cancel1 [simp]: "k * m < k * n \ 0 < k \ m < n" for k m n :: nat by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "k * m \ k * n \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma mult_le_cancel2 [simp]: "m * k \ n * k \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \ m < n" by (subst mult_less_cancel1) simp lemma Suc_mult_le_cancel1: "Suc k * m \ Suc k * n \ m \ n" by (subst mult_le_cancel1) simp lemma le_square: "m \ m * m" for m :: nat by (cases m) (auto intro: le_add1) lemma le_cube: "m \ m * (m * m)" for m :: nat by (cases m) (auto intro: le_add1) text \Lemma for \gcd\\ lemma mult_eq_self_implies_10: fixes m n :: nat assumes "m = m * n" shows "n = 1 \ m = 0" proof (rule disjCI) assume "m \ 0" show "n = 1" proof (cases n "1::nat" rule: linorder_cases) case greater show ?thesis using assms mult_less_mono2 [OF greater, of m] \m \ 0\ by auto qed (use assms \m \ 0\ in auto) qed lemma mono_times_nat: fixes n :: nat assumes "n > 0" shows "mono (times n)" proof fix m q :: nat assume "m \ q" with assms show "n * m \ n * q" by simp qed text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin definition "(inf :: nat \ nat \ nat) = min" definition "(sup :: nat \ nat \ nat) = max" instance by intro_classes (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def intro: order_less_imp_le antisym elim!: order_trans order_less_trans) end subsection \Natural operation of natural numbers on functions\ text \ We use the same logical constant for the power operations on functions and relations, in order to share the same syntax. \ consts compow :: "nat \ 'a \ 'a" abbreviation compower :: "'a \ nat \ 'a" (infixr "^^" 80) where "f ^^ n \ compow n f" notation (latex output) compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \\f ^^ n = f \ \ \ f\, the \n\-fold composition of \f\\ overloading funpow \ "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f \ funpow n f" end lemma funpow_0 [simp]: "(f ^^ 0) x = x" by simp lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \ f" proof (induct n) case 0 then show ?case by simp next fix n assume "f ^^ Suc n = f ^^ n \ f" then show "f ^^ Suc (Suc n) = f ^^ Suc n \ f" by (simp add: o_assoc) qed lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right text \For code generation.\ context begin qualified definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where funpow_code_def [code_abbrev]: "funpow = compow" lemma [code]: "funpow (Suc n) f = f \ funpow n f" "funpow 0 f = id" by (simp_all add: funpow_code_def) end lemma funpow_add: "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \ 'a" by (induct n) (simp_all add: funpow_add) lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp also have "\ = (f ^^ n \ f ^^ 1) x" by (simp only: funpow_add) also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \ 'a" by (induct n) simp_all lemma Suc_funpow[simp]: "Suc ^^ n = ((+) n)" by (induct n) simp_all lemma id_funpow[simp]: "id ^^ n = id" by (induct n) simp_all lemma funpow_mono: "mono f \ A \ B \ (f ^^ n) A \ (f ^^ n) B" for f :: "'a \ ('a::order)" by (induct n arbitrary: A B) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def) lemma funpow_mono2: assumes "mono f" and "i \ j" and "x \ y" and "x \ f x" shows "(f ^^ i) x \ (f ^^ j) y" using assms(2,3) proof (induct j arbitrary: y) case 0 then show ?case by simp next case (Suc j) show ?case proof(cases "i = Suc j") case True with assms(1) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right monoD funpow_mono) next case False with assms(1,4) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right le_eq_less_or_eq less_Suc_eq_le) (simp add: Suc.hyps monoD order_subst1) qed qed lemma inj_fn[simp]: fixes f::"'a \ 'a" assumes "inj f" shows "inj (f^^n)" proof (induction n) case Suc thus ?case using inj_compose[OF assms Suc.IH] by (simp del: comp_apply) qed simp lemma surj_fn[simp]: fixes f::"'a \ 'a" assumes "surj f" shows "surj (f^^n)" proof (induction n) case Suc thus ?case by (simp add: comp_surj[OF Suc.IH assms] del: comp_apply) qed simp lemma bij_fn[simp]: fixes f::"'a \ 'a" assumes "bij f" shows "bij (f^^n)" by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]]) lemma bij_betw_funpow: \<^marker>\contributor \Lars Noschinski\\ assumes "bij_betw f S S" shows "bij_betw (f ^^ n) S S" proof (induct n) case 0 then show ?case by (auto simp: id_def[symmetric]) next case (Suc n) then show ?case unfolding funpow.simps using assms by (rule bij_betw_trans) qed subsection \Kleene iteration\ lemma Kleene_iter_lpfp: fixes f :: "'a::order_bot \ 'a" assumes "mono f" and "f p \ p" shows "(f ^^ k) bot \ p" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma lfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) bot = (f ^^ k) bot" shows "lfp f = (f ^^ k) bot" proof (rule antisym) show "lfp f \ (f ^^ k) bot" proof (rule lfp_lowerbound) show "f ((f ^^ k) bot) \ (f ^^ k) bot" using assms(2) by simp qed show "(f ^^ k) bot \ lfp f" using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp qed lemma mono_pow: "mono f \ mono (f ^^ n)" for f :: "'a \ 'a::complete_lattice" by (induct n) (auto simp: mono_def) lemma lfp_funpow: assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f" proof (rule antisym) show "lfp f \ lfp (f ^^ Suc n)" proof (rule lfp_lowerbound) have "f (lfp (f ^^ Suc n)) = lfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: lfp_rolling f mono_pow comp_def) then show "f (lfp (f ^^ Suc n)) \ lfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (lfp f) = lfp f" for n by (induct n) (auto intro: f lfp_fixpoint) then show "lfp (f ^^ Suc n) \ lfp f" by (intro lfp_lowerbound) (simp del: funpow.simps) qed lemma gfp_funpow: assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f" proof (rule antisym) show "gfp f \ gfp (f ^^ Suc n)" proof (rule gfp_upperbound) have "f (gfp (f ^^ Suc n)) = gfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: gfp_rolling f mono_pow comp_def) then show "f (gfp (f ^^ Suc n)) \ gfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (gfp f) = gfp f" for n by (induct n) (auto intro: f gfp_fixpoint) then show "gfp (f ^^ Suc n) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed lemma Kleene_iter_gpfp: fixes f :: "'a::order_top \ 'a" assumes "mono f" and "p \ f p" shows "p \ (f ^^ k) top" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma gfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) top = (f ^^ k) top" shows "gfp f = (f ^^ k) top" (is "?lhs = ?rhs") proof (rule antisym) have "?rhs \ f ?rhs" using assms(2) by simp then show "?rhs \ ?lhs" by (rule gfp_upperbound) show "?lhs \ ?rhs" using Kleene_iter_gpfp[OF assms(1)] gfp_unfold[OF assms(1)] by simp qed subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin definition of_nat :: "nat \ 'a" where "of_nat n = (plus 1 ^^ n) 0" lemma of_nat_simps [simp]: shows of_nat_0: "of_nat 0 = 0" and of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m" by (simp_all add: of_nat_def) lemma of_nat_1 [simp]: "of_nat 1 = 1" by (simp add: of_nat_def) lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n" by (induct m) (simp_all add: ac_simps) lemma of_nat_mult [simp]: "of_nat (m * n) = of_nat m * of_nat n" by (induct m) (simp_all add: ac_simps distrib_right) lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x" by (induct x) (simp_all add: algebra_simps) primrec of_nat_aux :: "('a \ 'a) \ nat \ 'a \ 'a" where "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \ \tail recursive\ lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "\i. of_nat_aux (\i. i + 1) n (i + 1) = of_nat_aux (\i. i + 1) n i + 1" by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp with Suc show ?case by (simp add: add.commute) qed lemma of_nat_of_bool [simp]: "of_nat (of_bool P) = of_bool P" by auto end declare of_nat_code [code] context semiring_1_cancel begin lemma of_nat_diff: \of_nat (m - n) = of_nat m - of_nat n\ if \n \ m\ proof - from that obtain q where \m = n + q\ by (blast dest: le_Suc_ex) then show ?thesis by simp qed end text \Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.\ class semiring_char_0 = semiring_1 + assumes inj_of_nat: "inj of_nat" begin lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \ m = n" by (auto intro: inj_of_nat injD) text \Special cases where either operand is zero\ lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \ 0 = n" by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0]) lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \ m = 0" by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0]) lemma of_nat_1_eq_iff [simp]: "1 = of_nat n \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_eq_1_iff [simp]: "of_nat n = 1 \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_neq_0 [simp]: "of_nat (Suc n) \ 0" unfolding of_nat_eq_0_iff by simp lemma of_nat_0_neq [simp]: "0 \ of_nat (Suc n)" unfolding of_nat_0_eq_iff by simp end class ring_char_0 = ring_1 + semiring_char_0 context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" by (induct n) simp_all lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" by (auto simp: le_iff_add intro!: add_increasing2) lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" proof(induct m n rule: diff_induct) case (1 m) then show ?case by auto next case (2 n) then show ?case by (simp add: add_pos_nonneg) next case (3 m n) then show ?case by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) lemma less_imp_of_nat_less: "m < n \ of_nat m < of_nat n" by simp lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: order.eq_iff) text \Special cases where either operand is zero\ lemma of_nat_le_0_iff [simp]: "of_nat m \ 0 \ m = 0" by (rule of_nat_le_iff [of _ 0, simplified]) lemma of_nat_0_less_iff [simp]: "0 < of_nat n \ 0 < n" by (rule of_nat_less_iff [of 0, simplified]) end context linordered_nonzero_semiring begin lemma of_nat_max: "of_nat (max x y) = max (of_nat x) (of_nat y)" by (auto simp: max_def ord_class.max_def) lemma of_nat_min: "of_nat (min x y) = min (of_nat x) (of_nat y)" by (auto simp: min_def ord_class.min_def) end context linordered_semidom begin subclass linordered_nonzero_semiring .. subclass semiring_char_0 .. end context linordered_idom begin lemma abs_of_nat [simp]: "\of_nat n\ = of_nat n" by (simp add: abs_if) lemma sgn_of_nat [simp]: "sgn (of_nat n) = of_bool (n > 0)" by simp end lemma of_nat_id [simp]: "of_nat n = n" by (induct n) simp_all lemma of_nat_eq_id [simp]: "of_nat = id" by (auto simp add: fun_eq_iff) subsection \The set of natural numbers\ context semiring_1 begin definition Nats :: "'a set" ("\") where "\ = range of_nat" lemma of_nat_in_Nats [simp]: "of_nat n \ \" by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ \" using of_nat_0 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_1 [simp]: "1 \ \" using of_nat_1 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" unfolding Nats_def using of_nat_add [symmetric] by (blast intro: range_eqI) lemma Nats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" unfolding Nats_def using of_nat_mult [symmetric] by (blast intro: range_eqI) lemma Nats_cases [cases set: Nats]: assumes "x \ \" obtains (of_nat) n where "x = of_nat n" unfolding Nats_def proof - from \x \ \\ have "x \ range of_nat" unfolding Nats_def . then obtain n where "x = of_nat n" .. then show thesis .. qed lemma Nats_induct [case_names of_nat, induct set: Nats]: "x \ \ \ (\n. P (of_nat n)) \ P x" by (rule Nats_cases) auto end lemma Nats_diff [simp]: fixes a:: "'a::linordered_idom" assumes "a \ \" "b \ \" "b \ a" shows "a - b \ \" proof - obtain i where i: "a = of_nat i" using Nats_cases assms by blast obtain j where j: "b = of_nat j" using Nats_cases assms by blast have "j \ i" using \b \ a\ i j of_nat_le_iff by blast then have *: "of_nat i - of_nat j = (of_nat (i-j) :: 'a)" by (simp add: of_nat_diff) then show ?thesis by (simp add: * i j) qed subsection \Further arithmetic facts concerning the natural numbers\ lemma subst_equals: assumes "t = s" and "u = t" shows "u = s" using assms(2,1) by (rule trans) locale nat_arith begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma suc1: "A == k + a \ Suc A \ k + Suc a" by (simp only: add_Suc_right) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/nat_arith.ML\ simproc_setup nateq_cancel_sums ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") = \fn phi => try o Nat_Arith.cancel_eq_conv\ simproc_setup natless_cancel_sums ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") = \fn phi => try o Nat_Arith.cancel_less_conv\ simproc_setup natle_cancel_sums ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "Suc m \ n" | "m \ Suc n") = \fn phi => try o Nat_Arith.cancel_le_conv\ simproc_setup natdiff_cancel_sums ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") = \fn phi => try o Nat_Arith.cancel_diff_conv\ context order begin lemma lift_Suc_mono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_antimono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_mono_less: assumes mono: "\n. f n < f (Suc n)" and "n < n'" shows "f n < f n'" using \n < n'\ by (induct n n' rule: less_Suc_induct) (auto intro: mono) lemma lift_Suc_mono_less_iff: "(\n. f n < f (Suc n)) \ f n < f m \ n < m" by (blast intro: less_asym' lift_Suc_mono_less [of f] dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1]) end lemma mono_iff_le_Suc: "mono f \ (\n. f n \ f (Suc n))" unfolding mono_def by (auto intro: lift_Suc_mono_le [of f]) lemma antimono_iff_le_Suc: "antimono f \ (\n. f (Suc n) \ f n)" unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f]) +lemma strict_mono_Suc_iff: "strict_mono f \ (\n. f n < f (Suc n))" +proof (intro iffI strict_monoI) + assume *: "\n. f n < f (Suc n)" + fix m n :: nat assume "m < n" + thus "f m < f n" + by (induction rule: less_Suc_induct) (use * in auto) +qed (auto simp: strict_mono_def) + +lemma strict_mono_add: "strict_mono (\n::'a::linordered_semidom. n + k)" + by (auto simp: strict_mono_def) + lemma mono_nat_linear_lb: fixes f :: "nat \ nat" assumes "\m n. m < n \ f m < f n" shows "f m + k \ f (m + k)" proof (induct k) case 0 then show ?case by simp next case (Suc k) then have "Suc (f m + k) \ Suc (f (m + k))" by simp also from assms [of "m + k" "Suc (m + k)"] have "Suc (f (m + k)) \ f (Suc (m + k))" by (simp add: Suc_le_eq) finally show ?case by simp qed text \Subtraction laws, mostly by Clemens Ballarin\ lemma diff_less_mono: fixes a b c :: nat assumes "a < b" and "c \ a" shows "a - c < b - c" proof - from assms obtain d e where "b = c + (d + e)" and "a = c + e" and "d > 0" by (auto dest!: le_Suc_ex less_imp_Suc_add simp add: ac_simps) then show ?thesis by simp qed lemma less_diff_conv: "i < j - k \ i + k < j" for i j k :: nat by (cases "k \ j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex) lemma less_diff_conv2: "k \ j \ j - k < i \ j < i + k" for j k i :: nat by (auto dest: le_Suc_ex) lemma le_diff_conv: "j - k \ i \ j \ i + k" for j k i :: nat by (cases "k \ j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex) lemma diff_diff_cancel [simp]: "i \ n \ n - (n - i) = i" for i n :: nat by (auto dest: le_Suc_ex) lemma diff_less [simp]: "0 < n \ 0 < m \ m - n < m" for i n :: nat by (auto dest: less_imp_Suc_add) text \Simplification of relational expressions involving subtraction\ lemma diff_diff_eq: "k \ m \ k \ n \ m - k - (n - k) = m - n" for m n k :: nat by (auto dest!: le_Suc_ex) hide_fact (open) diff_diff_eq lemma eq_diff_iff: "k \ m \ k \ n \ m - k = n - k \ m = n" for m n k :: nat by (auto dest: le_Suc_ex) lemma less_diff_iff: "k \ m \ k \ n \ m - k < n - k \ m < n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff: "k \ m \ k \ n \ m - k \ n - k \ m \ n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff': "a \ c \ b \ c \ c - a \ c - b \ b \ a" for a b c :: nat by (force dest: le_Suc_ex) text \(Anti)Monotonicity of subtraction -- by Stephan Merz\ lemma diff_le_mono: "m \ n \ m - l \ n - l" for m n l :: nat by (auto dest: less_imp_le less_imp_Suc_add split: nat_diff_split) lemma diff_le_mono2: "m \ n \ l - n \ l - m" for m n l :: nat by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split: nat_diff_split) lemma diff_less_mono2: "m < n \ m < l \ l - n < l - m" for m n l :: nat by (auto dest: less_imp_Suc_add split: nat_diff_split) lemma diffs0_imp_equal: "m - n = 0 \ n - m = 0 \ m = n" for m n :: nat by (simp split: nat_diff_split) lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat by (cases m n rule: le_cases) (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono) lemma inj_on_diff_nat: fixes k :: nat assumes "\n. n \ N \ k \ n" shows "inj_on (\n. n - k) N" proof (rule inj_onI) fix x y assume a: "x \ N" "y \ N" "x - k = y - k" with assms have "x - k + k = y - k + k" by auto with a assms show "x = y" by (auto simp add: eq_diff_iff) qed text \Rewriting to pull differences out\ lemma diff_diff_right [simp]: "k \ j \ i - (j - k) = i + k - j" for i j k :: nat by (fact diff_diff_right) lemma diff_Suc_diff_eq1 [simp]: assumes "k \ j" shows "i - Suc (j - k) = i + k - Suc j" proof - from assms have *: "Suc (j - k) = Suc j - k" by (simp add: Suc_diff_le) from assms have "k \ Suc j" by (rule order_trans) simp with diff_diff_right [of k "Suc j" i] * show ?thesis by simp qed lemma diff_Suc_diff_eq2 [simp]: assumes "k \ j" shows "Suc (j - k) - i = Suc j - (k + i)" proof - from assms obtain n where "j = k + n" by (auto dest: le_Suc_ex) moreover have "Suc n - i = (k + Suc n) - (k + i)" using add_diff_cancel_left [of k "Suc n" i] by simp ultimately show ?thesis by simp qed lemma Suc_diff_Suc: assumes "n < m" shows "Suc (m - Suc n) = m - n" proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp qed lemma one_less_mult: "Suc 0 < n \ Suc 0 < m \ Suc 0 < m * n" using less_1_mult [of n m] by (simp add: ac_simps) lemma n_less_m_mult_n: "0 < n \ Suc 0 < m \ n < m * n" using mult_strict_right_mono [of 1 m n] by simp lemma n_less_n_mult_m: "0 < n \ Suc 0 < m \ n < n * m" using mult_strict_left_mono [of 1 m n] by simp text \Induction starting beyond zero\ lemma nat_induct_at_least [consumes 1, case_names base Suc]: "P n" if "n \ m" "P m" "\n. n \ m \ P n \ P (Suc n)" proof - define q where "q = n - m" with \n \ m\ have "n = m + q" by simp moreover have "P (m + q)" by (induction q) (use that in simp_all) ultimately show "P n" by simp qed lemma nat_induct_non_zero [consumes 1, case_names 1 Suc]: "P n" if "n > 0" "P 1" "\n. n > 0 \ P n \ P (Suc n)" proof - from \n > 0\ have "n \ 1" by (cases n) simp_all moreover note \P 1\ moreover have "\n. n \ 1 \ P n \ P (Suc n)" using \\n. n > 0 \ P n \ P (Suc n)\ by (simp add: Suc_le_eq) ultimately show "P n" by (rule nat_induct_at_least) qed text \Specialized induction principles that work "backwards":\ lemma inc_induct [consumes 1, case_names base step]: assumes less: "i \ j" and base: "P j" and step: "\n. i \ n \ n < j \ P (Suc n) \ P n" shows "P i" using less step proof (induct "j - i" arbitrary: i) case (0 i) then have "i = j" by simp with base show ?case by simp next case (Suc d n) from Suc.hyps have "n \ j" by auto with Suc have "n < j" by (simp add: less_le) from \Suc d = j - n\ have "d + 1 = j - n" by simp then have "d + 1 - 1 = j - n - 1" by simp then have "d = j - n - 1" by simp then have "d = j - (n + 1)" by (simp add: diff_diff_eq) then have "d = j - Suc n" by simp moreover from \n < j\ have "Suc n \ j" by (simp add: Suc_le_eq) ultimately have "P (Suc n)" proof (rule Suc.hyps) fix q assume "Suc n \ q" then have "n \ q" by (simp add: Suc_le_eq less_imp_le) moreover assume "q < j" moreover assume "P (Suc q)" ultimately show "P q" by (rule Suc.prems) qed with order_refl \n < j\ show "P n" by (rule Suc.prems) qed lemma strict_inc_induct [consumes 1, case_names base step]: assumes less: "i < j" and base: "\i. j = Suc i \ P i" and step: "\i. i < j \ P (Suc i) \ P i" shows "P i" using less proof (induct "j - i - 1" arbitrary: i) case (0 i) from \i < j\ obtain n where "j = i + n" and "n > 0" by (auto dest!: less_imp_Suc_add) with 0 have "j = Suc i" by (auto intro: order_antisym simp add: Suc_le_eq) with base show ?case by simp next case (Suc d i) from \Suc d = j - i - 1\ have *: "Suc d = j - Suc i" by (simp add: diff_diff_add) then have "Suc d - 1 = j - Suc i - 1" by simp then have "d = j - Suc i - 1" by simp moreover from * have "j - Suc i \ 0" by auto then have "Suc i < j" by (simp add: not_le) ultimately have "P (Suc i)" by (rule Suc.hyps) with \i < j\ show "P i" by (rule step) qed lemma zero_induct_lemma: "P k \ (\n. P (Suc n) \ P n) \ P (k - i)" using inc_induct[of "k - i" k P, simplified] by blast lemma zero_induct: "P k \ (\n. P (Suc n) \ P n) \ P 0" using inc_induct[of 0 k P] by blast text \Further induction rule similar to @{thm inc_induct}.\ lemma dec_induct [consumes 1, case_names base step]: "i \ j \ P i \ (\n. i \ n \ n < j \ P n \ P (Suc n)) \ P j" proof (induct j arbitrary: i) case 0 then show ?case by simp next case (Suc j) from Suc.prems consider "i \ j" | "i = Suc j" by (auto simp add: le_Suc_eq) then show ?case proof cases case 1 moreover have "j < Suc j" by simp moreover have "P j" using \i \ j\ \P i\ proof (rule Suc.hyps) fix q assume "i \ q" moreover assume "q < j" then have "q < Suc j" by (simp add: less_Suc_eq) moreover assume "P q" ultimately show "P (Suc q)" by (rule Suc.prems) qed ultimately show "P (Suc j)" by (rule Suc.prems) next case 2 with \P i\ show "P (Suc j)" by simp qed qed lemma transitive_stepwise_le: assumes "m \ n" "\x. R x x" "\x y z. R x y \ R y z \ R x z" and "\n. R n (Suc n)" shows "R m n" using \m \ n\ by (induction rule: dec_induct) (use assms in blast)+ subsubsection \Greatest operator\ lemma ex_has_greatest_nat: "P (k::nat) \ \y. P y \ y \ b \ \x. P x \ (\y. P y \ y \ x)" proof (induction "b-k" arbitrary: b k rule: less_induct) case less show ?case proof cases assume "\n>k. P n" then obtain n where "n>k" "P n" by blast have "n \ b" using \P n\ less.prems(2) by auto hence "b-n < b-k" by(rule diff_less_mono2[OF \k less_le_trans[OF \k]]) from less.hyps[OF this \P n\ less.prems(2)] show ?thesis . next assume "\ (\n>k. P n)" hence "\y. P y \ y \ k" by (auto simp: not_less) thus ?thesis using less.prems(1) by auto qed qed lemma fixes k::nat assumes "P k" and minor: "\y. P y \ y \ b" shows GreatestI_nat: "P (Greatest P)" and Greatest_le_nat: "k \ Greatest P" proof - obtain x where "P x" "\y. P y \ y \ x" using assms ex_has_greatest_nat by blast with \P k\ show "P (Greatest P)" "k \ Greatest P" using GreatestI2_order by blast+ qed lemma GreatestI_ex_nat: "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" by (blast intro: GreatestI_nat) subsection \Monotonicity of \funpow\\ lemma funpow_increasing: "m \ n \ mono f \ (f ^^ n) \ \ (f ^^ m) \" for f :: "'a::{lattice,order_top} \ 'a" by (induct rule: inc_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma funpow_decreasing: "m \ n \ mono f \ (f ^^ m) \ \ (f ^^ n) \" for f :: "'a::{lattice,order_bot} \ 'a" by (induct rule: dec_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma mono_funpow: "mono Q \ mono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_bot} \ 'a" by (auto intro!: funpow_decreasing simp: mono_def) lemma antimono_funpow: "mono Q \ antimono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_top} \ 'a" by (auto intro!: funpow_increasing simp: antimono_def) subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def) lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \ m = Suc 0" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \ m = 1" for m :: nat by (simp add: dvd_def) lemma dvd_antisym: "m dvd n \ n dvd m \ m = n" for m n :: nat unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) lemma dvd_diff_nat [simp]: "k dvd m \ k dvd n \ k dvd (m - n)" for k m n :: nat unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric]) lemma dvd_diffD: fixes k m n :: nat assumes "k dvd m - n" "k dvd n" "n \ m" shows "k dvd m" proof - have "k dvd n + (m - n)" using assms by (blast intro: dvd_add) with assms show ?thesis by simp qed lemma dvd_diffD1: "k dvd m - n \ k dvd m \ n \ m \ k dvd n" for k m n :: nat by (drule_tac m = m in dvd_diff_nat) auto lemma dvd_mult_cancel: fixes m n k :: nat assumes "k * m dvd k * n" and "0 < k" shows "m dvd n" proof - from assms(1) obtain q where "k * n = (k * m) * q" .. then have "k * n = k * (m * q)" by (simp add: ac_simps) with \0 < k\ have "n = m * q" by (auto simp add: mult_left_cancel) then show ?thesis .. qed lemma dvd_mult_cancel1: fixes m n :: nat assumes "0 < m" shows "m * n dvd m \ n = 1" proof assume "m * n dvd m" then have "m * n dvd m * 1" by simp then have "n dvd 1" by (iprover intro: assms dvd_mult_cancel) then show "n = 1" by auto qed auto lemma dvd_mult_cancel2: "0 < m \ n * m dvd m \ n = 1" for m n :: nat using dvd_mult_cancel1 [of m n] by (simp add: ac_simps) lemma dvd_imp_le: "k dvd n \ 0 < n \ k \ n" for k n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma nat_dvd_not_less: "0 < m \ m < n \ \ n dvd m" for m n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma less_eq_dvd_minus: fixes m n :: nat assumes "m \ n" shows "m dvd n \ m dvd n - m" proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. then show ?thesis by (simp add: add.commute [of m]) qed lemma dvd_minus_self: "m dvd n - m \ n < m \ m dvd n" for m n :: nat by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le) lemma dvd_minus_add: fixes m n q r :: nat assumes "q \ n" "q \ r * m" shows "m dvd n - q \ m dvd n + (r * m - q)" proof - have "m dvd n - q \ m dvd r * m + (n - q)" using dvd_add_times_triv_left_iff [of m r] by simp also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed subsection \Aliasses\ lemma nat_mult_1: "1 * n = n" for n :: nat by (fact mult_1_left) lemma nat_mult_1_right: "n * 1 = n" for n :: nat by (fact mult_1_right) lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat by (fact left_diff_distrib') lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat by (fact right_diff_distrib') (*Used in AUTO2 and Groups.le_diff_conv2 (with variables renamed) doesn't work for some reason*) lemma le_diff_conv2: "k \ j \ (i \ j - k) = (i + k \ j)" for i j k :: nat by (fact le_diff_conv2) lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat by (fact diff_cancel) lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat by (fact diff_diff_add) lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat by (fact add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = 0" for m n :: nat by (fact diff_add_zero) lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat by (fact distrib_left) lemmas nat_distrib = add_mult_distrib distrib_left diff_mult_distrib diff_mult_distrib2 subsection \Size of a datatype value\ class size = fixes size :: "'a \ nat" \ \see further theory \Wellfounded\\ instantiation nat :: size begin definition size_nat where [simp, code]: "size (n::nat) = n" instance .. end lemmas size_nat = size_nat_def lemma size_neq_size_imp_neq: "size x \ size y \ x \ y" by (erule contrapos_nn) (rule arg_cong) subsection \Code module namespace\ code_identifier code_module Nat \ (SML) Arith and (OCaml) Arith and (Haskell) Arith hide_const (open) of_nat_aux end diff --git a/src/HOL/Topological_Spaces.thy b/src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy +++ b/src/HOL/Topological_Spaces.thy @@ -1,3970 +1,3959 @@ (* Title: HOL/Topological_Spaces.thy Author: Brian Huffman Author: Johannes Hölzl *) section \Topological Spaces\ theory Topological_Spaces imports Main begin named_theorems continuous_intros "structural introduction rules for continuity" subsection \Topological space\ class "open" = fixes "open" :: "'a set \ bool" class topological_space = "open" + assumes open_UNIV [simp, intro]: "open UNIV" assumes open_Int [intro]: "open S \ open T \ open (S \ T)" assumes open_Union [intro]: "\S\K. open S \ open (\K)" begin definition closed :: "'a set \ bool" where "closed S \ open (- S)" lemma open_empty [continuous_intros, intro, simp]: "open {}" using open_Union [of "{}"] by simp lemma open_Un [continuous_intros, intro]: "open S \ open T \ open (S \ T)" using open_Union [of "{S, T}"] by simp lemma open_UN [continuous_intros, intro]: "\x\A. open (B x) \ open (\x\A. B x)" using open_Union [of "B ` A"] by simp lemma open_Inter [continuous_intros, intro]: "finite S \ \T\S. open T \ open (\S)" by (induction set: finite) auto lemma open_INT [continuous_intros, intro]: "finite A \ \x\A. open (B x) \ open (\x\A. B x)" using open_Inter [of "B ` A"] by simp lemma openI: assumes "\x. x \ S \ \T. open T \ x \ T \ T \ S" shows "open S" proof - have "open (\{T. open T \ T \ S})" by auto moreover have "\{T. open T \ T \ S} = S" by (auto dest!: assms) ultimately show "open S" by simp qed lemma open_subopen: "open S \ (\x\S. \T. open T \ x \ T \ T \ S)" by (auto intro: openI) lemma closed_empty [continuous_intros, intro, simp]: "closed {}" unfolding closed_def by simp lemma closed_Un [continuous_intros, intro]: "closed S \ closed T \ closed (S \ T)" unfolding closed_def by auto lemma closed_UNIV [continuous_intros, intro, simp]: "closed UNIV" unfolding closed_def by simp lemma closed_Int [continuous_intros, intro]: "closed S \ closed T \ closed (S \ T)" unfolding closed_def by auto lemma closed_INT [continuous_intros, intro]: "\x\A. closed (B x) \ closed (\x\A. B x)" unfolding closed_def by auto lemma closed_Inter [continuous_intros, intro]: "\S\K. closed S \ closed (\K)" unfolding closed_def uminus_Inf by auto lemma closed_Union [continuous_intros, intro]: "finite S \ \T\S. closed T \ closed (\S)" by (induct set: finite) auto lemma closed_UN [continuous_intros, intro]: "finite A \ \x\A. closed (B x) \ closed (\x\A. B x)" using closed_Union [of "B ` A"] by simp lemma open_closed: "open S \ closed (- S)" by (simp add: closed_def) lemma closed_open: "closed S \ open (- S)" by (rule closed_def) lemma open_Diff [continuous_intros, intro]: "open S \ closed T \ open (S - T)" by (simp add: closed_open Diff_eq open_Int) lemma closed_Diff [continuous_intros, intro]: "closed S \ open T \ closed (S - T)" by (simp add: open_closed Diff_eq closed_Int) lemma open_Compl [continuous_intros, intro]: "closed S \ open (- S)" by (simp add: closed_open) lemma closed_Compl [continuous_intros, intro]: "open S \ closed (- S)" by (simp add: open_closed) lemma open_Collect_neg: "closed {x. P x} \ open {x. \ P x}" unfolding Collect_neg_eq by (rule open_Compl) lemma open_Collect_conj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \ Q x}" using open_Int[OF assms] by (simp add: Int_def) lemma open_Collect_disj: assumes "open {x. P x}" "open {x. Q x}" shows "open {x. P x \ Q x}" using open_Un[OF assms] by (simp add: Un_def) lemma open_Collect_ex: "(\i. open {x. P i x}) \ open {x. \i. P i x}" using open_UN[of UNIV "\i. {x. P i x}"] unfolding Collect_ex_eq by simp lemma open_Collect_imp: "closed {x. P x} \ open {x. Q x} \ open {x. P x \ Q x}" unfolding imp_conv_disj by (intro open_Collect_disj open_Collect_neg) lemma open_Collect_const: "open {x. P}" by (cases P) auto lemma closed_Collect_neg: "open {x. P x} \ closed {x. \ P x}" unfolding Collect_neg_eq by (rule closed_Compl) lemma closed_Collect_conj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \ Q x}" using closed_Int[OF assms] by (simp add: Int_def) lemma closed_Collect_disj: assumes "closed {x. P x}" "closed {x. Q x}" shows "closed {x. P x \ Q x}" using closed_Un[OF assms] by (simp add: Un_def) lemma closed_Collect_all: "(\i. closed {x. P i x}) \ closed {x. \i. P i x}" using closed_INT[of UNIV "\i. {x. P i x}"] by (simp add: Collect_all_eq) lemma closed_Collect_imp: "open {x. P x} \ closed {x. Q x} \ closed {x. P x \ Q x}" unfolding imp_conv_disj by (intro closed_Collect_disj closed_Collect_neg) lemma closed_Collect_const: "closed {x. P}" by (cases P) auto end subsection \Hausdorff and other separation properties\ class t0_space = topological_space + assumes t0_space: "x \ y \ \U. open U \ \ (x \ U \ y \ U)" class t1_space = topological_space + assumes t1_space: "x \ y \ \U. open U \ x \ U \ y \ U" instance t1_space \ t0_space by standard (fast dest: t1_space) context t1_space begin lemma separation_t1: "x \ y \ (\U. open U \ x \ U \ y \ U)" using t1_space[of x y] by blast lemma closed_singleton [iff]: "closed {a}" proof - let ?T = "\{S. open S \ a \ S}" have "open ?T" by (simp add: open_Union) also have "?T = - {a}" by (auto simp add: set_eq_iff separation_t1) finally show "closed {a}" by (simp only: closed_def) qed lemma closed_insert [continuous_intros, simp]: assumes "closed S" shows "closed (insert a S)" proof - from closed_singleton assms have "closed ({a} \ S)" by (rule closed_Un) then show "closed (insert a S)" by simp qed lemma finite_imp_closed: "finite S \ closed S" by (induct pred: finite) simp_all end text \T2 spaces are also known as Hausdorff spaces.\ class t2_space = topological_space + assumes hausdorff: "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" instance t2_space \ t1_space by standard (fast dest: hausdorff) lemma (in t2_space) separation_t2: "x \ y \ (\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {})" using hausdorff [of x y] by blast lemma (in t0_space) separation_t0: "x \ y \ (\U. open U \ \ (x \ U \ y \ U))" using t0_space [of x y] by blast text \A classical separation axiom for topological space, the T3 axiom -- also called regularity: if a point is not in a closed set, then there are open sets separating them.\ class t3_space = t2_space + assumes t3_space: "closed S \ y \ S \ \U V. open U \ open V \ y \ U \ S \ V \ U \ V = {}" text \A classical separation axiom for topological space, the T4 axiom -- also called normality: if two closed sets are disjoint, then there are open sets separating them.\ class t4_space = t2_space + assumes t4_space: "closed S \ closed T \ S \ T = {} \ \U V. open U \ open V \ S \ U \ T \ V \ U \ V = {}" text \T4 is stronger than T3, and weaker than metric.\ instance t4_space \ t3_space proof fix S and y::'a assume "closed S" "y \ S" then show "\U V. open U \ open V \ y \ U \ S \ V \ U \ V = {}" using t4_space[of "{y}" S] by auto qed text \A perfect space is a topological space with no isolated points.\ class perfect_space = topological_space + assumes not_open_singleton: "\ open {x}" lemma (in perfect_space) UNIV_not_singleton: "UNIV \ {x}" for x::'a by (metis (no_types) open_UNIV not_open_singleton) subsection \Generators for toplogies\ inductive generate_topology :: "'a set set \ 'a set \ bool" for S :: "'a set set" where UNIV: "generate_topology S UNIV" | Int: "generate_topology S (a \ b)" if "generate_topology S a" and "generate_topology S b" | UN: "generate_topology S (\K)" if "(\k. k \ K \ generate_topology S k)" | Basis: "generate_topology S s" if "s \ S" hide_fact (open) UNIV Int UN Basis lemma generate_topology_Union: "(\k. k \ I \ generate_topology S (K k)) \ generate_topology S (\k\I. K k)" using generate_topology.UN [of "K ` I"] by auto lemma topological_space_generate_topology: "class.topological_space (generate_topology S)" by standard (auto intro: generate_topology.intros) subsection \Order topologies\ class order_topology = order + "open" + assumes open_generated_order: "open = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" begin subclass topological_space unfolding open_generated_order by (rule topological_space_generate_topology) lemma open_greaterThan [continuous_intros, simp]: "open {a <..}" unfolding open_generated_order by (auto intro: generate_topology.Basis) lemma open_lessThan [continuous_intros, simp]: "open {..< a}" unfolding open_generated_order by (auto intro: generate_topology.Basis) lemma open_greaterThanLessThan [continuous_intros, simp]: "open {a <..< b}" unfolding greaterThanLessThan_eq by (simp add: open_Int) end class linorder_topology = linorder + order_topology lemma closed_atMost [continuous_intros, simp]: "closed {..a}" for a :: "'a::linorder_topology" by (simp add: closed_open) lemma closed_atLeast [continuous_intros, simp]: "closed {a..}" for a :: "'a::linorder_topology" by (simp add: closed_open) lemma closed_atLeastAtMost [continuous_intros, simp]: "closed {a..b}" for a b :: "'a::linorder_topology" proof - have "{a .. b} = {a ..} \ {.. b}" by auto then show ?thesis by (simp add: closed_Int) qed lemma (in order) less_separate: assumes "x < y" shows "\a b. x \ {..< a} \ y \ {b <..} \ {..< a} \ {b <..} = {}" proof (cases "\z. x < z \ z < y") case True then obtain z where "x < z \ z < y" .. then have "x \ {..< z} \ y \ {z <..} \ {z <..} \ {..< z} = {}" by auto then show ?thesis by blast next case False with \x < y\ have "x \ {..< y}" "y \ {x <..}" "{x <..} \ {..< y} = {}" by auto then show ?thesis by blast qed instance linorder_topology \ t2_space proof fix x y :: 'a show "x \ y \ \U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" using less_separate [of x y] less_separate [of y x] by (elim neqE; metis open_lessThan open_greaterThan Int_commute) qed lemma (in linorder_topology) open_right: assumes "open S" "x \ S" and gt_ex: "x < y" shows "\b>x. {x ..< b} \ S" using assms unfolding open_generated_order proof induct case UNIV then show ?case by blast next case (Int A B) then obtain a b where "a > x" "{x ..< a} \ A" "b > x" "{x ..< b} \ B" by auto then show ?case by (auto intro!: exI[of _ "min a b"]) next case UN then show ?case by blast next case Basis then show ?case by (fastforce intro: exI[of _ y] gt_ex) qed lemma (in linorder_topology) open_left: assumes "open S" "x \ S" and lt_ex: "y < x" shows "\b S" using assms unfolding open_generated_order proof induction case UNIV then show ?case by blast next case (Int A B) then obtain a b where "a < x" "{a <.. x} \ A" "b < x" "{b <.. x} \ B" by auto then show ?case by (auto intro!: exI[of _ "max a b"]) next case UN then show ?case by blast next case Basis then show ?case by (fastforce intro: exI[of _ y] lt_ex) qed subsection \Setup some topologies\ subsubsection \Boolean is an order topology\ class discrete_topology = topological_space + assumes open_discrete: "\A. open A" instance discrete_topology < t2_space proof fix x y :: 'a assume "x \ y" then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by (intro exI[of _ "{_}"]) (auto intro!: open_discrete) qed instantiation bool :: linorder_topology begin definition open_bool :: "bool set \ bool" where "open_bool = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_bool_def) end instance bool :: discrete_topology proof fix A :: "bool set" have *: "{False <..} = {True}" "{..< True} = {False}" by auto have "A = UNIV \ A = {} \ A = {False <..} \ A = {..< True}" using subset_UNIV[of A] unfolding UNIV_bool * by blast then show "open A" by auto qed instantiation nat :: linorder_topology begin definition open_nat :: "nat set \ bool" where "open_nat = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_nat_def) end instance nat :: discrete_topology proof fix A :: "nat set" have "open {n}" for n :: nat proof (cases n) case 0 moreover have "{0} = {..<1::nat}" by auto ultimately show ?thesis by auto next case (Suc n') then have "{n} = {.. {n' <..}" by auto with Suc show ?thesis by (auto intro: open_lessThan open_greaterThan) qed then have "open (\a\A. {a})" by (intro open_UN) auto then show "open A" by simp qed instantiation int :: linorder_topology begin definition open_int :: "int set \ bool" where "open_int = generate_topology (range (\a. {..< a}) \ range (\a. {a <..}))" instance by standard (rule open_int_def) end instance int :: discrete_topology proof fix A :: "int set" have "{.. {i-1 <..} = {i}" for i :: int by auto then have "open {i}" for i :: int using open_Int[OF open_lessThan[of "i + 1"] open_greaterThan[of "i - 1"]] by auto then have "open (\a\A. {a})" by (intro open_UN) auto then show "open A" by simp qed subsubsection \Topological filters\ definition (in topological_space) nhds :: "'a \ 'a filter" where "nhds a = (INF S\{S. open S \ a \ S}. principal S)" definition (in topological_space) at_within :: "'a \ 'a set \ 'a filter" ("at (_)/ within (_)" [1000, 60] 60) where "at a within s = inf (nhds a) (principal (s - {a}))" abbreviation (in topological_space) at :: "'a \ 'a filter" ("at") where "at x \ at x within (CONST UNIV)" abbreviation (in order_topology) at_right :: "'a \ 'a filter" where "at_right x \ at x within {x <..}" abbreviation (in order_topology) at_left :: "'a \ 'a filter" where "at_left x \ at x within {..< x}" lemma (in topological_space) nhds_generated_topology: "open = generate_topology T \ nhds x = (INF S\{S\T. x \ S}. principal S)" unfolding nhds_def proof (safe intro!: antisym INF_greatest) fix S assume "generate_topology T S" "x \ S" then show "(INF S\{S \ T. x \ S}. principal S) \ principal S" by induct (auto intro: INF_lower order_trans simp: inf_principal[symmetric] simp del: inf_principal) qed (auto intro!: INF_lower intro: generate_topology.intros) lemma (in topological_space) eventually_nhds: "eventually P (nhds a) \ (\S. open S \ a \ S \ (\x\S. P x))" unfolding nhds_def by (subst eventually_INF_base) (auto simp: eventually_principal) lemma eventually_eventually: "eventually (\y. eventually P (nhds y)) (nhds x) = eventually P (nhds x)" by (auto simp: eventually_nhds) lemma (in topological_space) eventually_nhds_in_open: "open s \ x \ s \ eventually (\y. y \ s) (nhds x)" by (subst eventually_nhds) blast lemma (in topological_space) eventually_nhds_x_imp_x: "eventually P (nhds x) \ P x" by (subst (asm) eventually_nhds) blast lemma (in topological_space) nhds_neq_bot [simp]: "nhds a \ bot" by (simp add: trivial_limit_def eventually_nhds) lemma (in t1_space) t1_space_nhds: "x \ y \ (\\<^sub>F x in nhds x. x \ y)" by (drule t1_space) (auto simp: eventually_nhds) lemma (in topological_space) nhds_discrete_open: "open {x} \ nhds x = principal {x}" by (auto simp: nhds_def intro!: antisym INF_greatest INF_lower2[of "{x}"]) lemma (in discrete_topology) nhds_discrete: "nhds x = principal {x}" by (simp add: nhds_discrete_open open_discrete) lemma (in discrete_topology) at_discrete: "at x within S = bot" unfolding at_within_def nhds_discrete by simp lemma (in discrete_topology) tendsto_discrete: "filterlim (f :: 'b \ 'a) (nhds y) F \ eventually (\x. f x = y) F" by (auto simp: nhds_discrete filterlim_principal) lemma (in topological_space) at_within_eq: "at x within s = (INF S\{S. open S \ x \ S}. principal (S \ s - {x}))" unfolding nhds_def at_within_def by (subst INF_inf_const2[symmetric]) (auto simp: Diff_Int_distrib) lemma (in topological_space) eventually_at_filter: "eventually P (at a within s) \ eventually (\x. x \ a \ x \ s \ P x) (nhds a)" by (simp add: at_within_def eventually_inf_principal imp_conjL[symmetric] conj_commute) lemma (in topological_space) at_le: "s \ t \ at x within s \ at x within t" unfolding at_within_def by (intro inf_mono) auto lemma (in topological_space) eventually_at_topological: "eventually P (at a within s) \ (\S. open S \ a \ S \ (\x\S. x \ a \ x \ s \ P x))" by (simp add: eventually_nhds eventually_at_filter) lemma eventually_at_in_open: assumes "open A" "x \ A" shows "eventually (\y. y \ A - {x}) (at x)" using assms eventually_at_topological by blast lemma eventually_at_in_open': assumes "open A" "x \ A" shows "eventually (\y. y \ A) (at x)" using assms eventually_at_topological by blast lemma (in topological_space) at_within_open: "a \ S \ open S \ at a within S = at a" unfolding filter_eq_iff eventually_at_topological by (metis open_Int Int_iff UNIV_I) lemma (in topological_space) at_within_open_NO_MATCH: "a \ s \ open s \ NO_MATCH UNIV s \ at a within s = at a" by (simp only: at_within_open) lemma (in topological_space) at_within_open_subset: "a \ S \ open S \ S \ T \ at a within T = at a" by (metis at_le at_within_open dual_order.antisym subset_UNIV) lemma (in topological_space) at_within_nhd: assumes "x \ S" "open S" "T \ S - {x} = U \ S - {x}" shows "at x within T = at x within U" unfolding filter_eq_iff eventually_at_filter proof (intro allI eventually_subst) have "eventually (\x. x \ S) (nhds x)" using \x \ S\ \open S\ by (auto simp: eventually_nhds) then show "\\<^sub>F n in nhds x. (n \ x \ n \ T \ P n) = (n \ x \ n \ U \ P n)" for P by eventually_elim (insert \T \ S - {x} = U \ S - {x}\, blast) qed lemma (in topological_space) at_within_empty [simp]: "at a within {} = bot" unfolding at_within_def by simp lemma (in topological_space) at_within_union: "at x within (S \ T) = sup (at x within S) (at x within T)" unfolding filter_eq_iff eventually_sup eventually_at_filter by (auto elim!: eventually_rev_mp) lemma (in topological_space) at_eq_bot_iff: "at a = bot \ open {a}" unfolding trivial_limit_def eventually_at_topological by (metis UNIV_I empty_iff is_singletonE is_singletonI' singleton_iff) lemma (in t1_space) eventually_neq_at_within: "eventually (\w. w \ x) (at z within A)" by (smt (verit, ccfv_threshold) eventually_True eventually_at_topological separation_t1) lemma (in perfect_space) at_neq_bot [simp]: "at a \ bot" by (simp add: at_eq_bot_iff not_open_singleton) lemma (in order_topology) nhds_order: "nhds x = inf (INF a\{x <..}. principal {..< a}) (INF a\{..< x}. principal {a <..})" proof - have 1: "{S \ range lessThan \ range greaterThan. x \ S} = (\a. {..< a}) ` {x <..} \ (\a. {a <..}) ` {..< x}" by auto show ?thesis by (simp only: nhds_generated_topology[OF open_generated_order] INF_union 1 INF_image comp_def) qed lemma (in topological_space) filterlim_at_within_If: assumes "filterlim f G (at x within (A \ {x. P x}))" and "filterlim g G (at x within (A \ {x. \P x}))" shows "filterlim (\x. if P x then f x else g x) G (at x within A)" proof (rule filterlim_If) note assms(1) also have "at x within (A \ {x. P x}) = inf (nhds x) (principal (A \ Collect P - {x}))" by (simp add: at_within_def) also have "A \ Collect P - {x} = (A - {x}) \ Collect P" by blast also have "inf (nhds x) (principal \) = inf (at x within A) (principal (Collect P))" by (simp add: at_within_def inf_assoc) finally show "filterlim f G (inf (at x within A) (principal (Collect P)))" . next note assms(2) also have "at x within (A \ {x. \ P x}) = inf (nhds x) (principal (A \ {x. \ P x} - {x}))" by (simp add: at_within_def) also have "A \ {x. \ P x} - {x} = (A - {x}) \ {x. \ P x}" by blast also have "inf (nhds x) (principal \) = inf (at x within A) (principal {x. \ P x})" by (simp add: at_within_def inf_assoc) finally show "filterlim g G (inf (at x within A) (principal {x. \ P x}))" . qed lemma (in topological_space) filterlim_at_If: assumes "filterlim f G (at x within {x. P x})" and "filterlim g G (at x within {x. \P x})" shows "filterlim (\x. if P x then f x else g x) G (at x)" using assms by (intro filterlim_at_within_If) simp_all lemma (in linorder_topology) at_within_order: assumes "UNIV \ {x}" shows "at x within s = inf (INF a\{x <..}. principal ({..< a} \ s - {x})) (INF a\{..< x}. principal ({a <..} \ s - {x}))" proof (cases "{x <..} = {}" "{..< x} = {}" rule: case_split [case_product case_split]) case True_True have "UNIV = {..< x} \ {x} \ {x <..}" by auto with assms True_True show ?thesis by auto qed (auto simp del: inf_principal simp: at_within_def nhds_order Int_Diff inf_principal[symmetric] INF_inf_const2 inf_sup_aci[where 'a="'a filter"]) lemma (in linorder_topology) at_left_eq: "y < x \ at_left x = (INF a\{..< x}. principal {a <..< x})" by (subst at_within_order) (auto simp: greaterThan_Int_greaterThan greaterThanLessThan_eq[symmetric] min.absorb2 INF_constant intro!: INF_lower2 inf_absorb2) lemma (in linorder_topology) eventually_at_left: "y < x \ eventually P (at_left x) \ (\by>b. y < x \ P y)" unfolding at_left_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) lemma (in linorder_topology) at_right_eq: "x < y \ at_right x = (INF a\{x <..}. principal {x <..< a})" by (subst at_within_order) (auto simp: lessThan_Int_lessThan greaterThanLessThan_eq[symmetric] max.absorb2 INF_constant Int_commute intro!: INF_lower2 inf_absorb1) lemma (in linorder_topology) eventually_at_right: "x < y \ eventually P (at_right x) \ (\b>x. \y>x. y < b \ P y)" unfolding at_right_eq by (subst eventually_INF_base) (auto simp: eventually_principal Ball_def) lemma eventually_at_right_less: "\\<^sub>F y in at_right (x::'a::{linorder_topology, no_top}). x < y" using gt_ex[of x] eventually_at_right[of x] by auto lemma trivial_limit_at_right_top: "at_right (top::_::{order_top,linorder_topology}) = bot" by (auto simp: filter_eq_iff eventually_at_topological) lemma trivial_limit_at_left_bot: "at_left (bot::_::{order_bot,linorder_topology}) = bot" by (auto simp: filter_eq_iff eventually_at_topological) lemma trivial_limit_at_left_real [simp]: "\ trivial_limit (at_left x)" for x :: "'a::{no_bot,dense_order,linorder_topology}" using lt_ex [of x] by safe (auto simp add: trivial_limit_def eventually_at_left dest: dense) lemma trivial_limit_at_right_real [simp]: "\ trivial_limit (at_right x)" for x :: "'a::{no_top,dense_order,linorder_topology}" using gt_ex[of x] by safe (auto simp add: trivial_limit_def eventually_at_right dest: dense) lemma (in linorder_topology) at_eq_sup_left_right: "at x = sup (at_left x) (at_right x)" by (auto simp: eventually_at_filter filter_eq_iff eventually_sup elim: eventually_elim2 eventually_mono) lemma (in linorder_topology) eventually_at_split: "eventually P (at x) \ eventually P (at_left x) \ eventually P (at_right x)" by (subst at_eq_sup_left_right) (simp add: eventually_sup) lemma (in order_topology) eventually_at_leftI: assumes "\x. x \ {a<.. P x" "a < b" shows "eventually P (at_left b)" using assms unfolding eventually_at_topological by (intro exI[of _ "{a<..}"]) auto lemma (in order_topology) eventually_at_rightI: assumes "\x. x \ {a<.. P x" "a < b" shows "eventually P (at_right a)" using assms unfolding eventually_at_topological by (intro exI[of _ "{.. (\S. open S \ x \ S \ (\x. f x \ S \ P x))" unfolding eventually_filtercomap eventually_nhds by auto lemma eventually_filtercomap_at_topological: "eventually P (filtercomap f (at A within B)) \ (\S. open S \ A \ S \ (\x. f x \ S \ B - {A} \ P x))" (is "?lhs = ?rhs") unfolding at_within_def filtercomap_inf eventually_inf_principal filtercomap_principal eventually_filtercomap_nhds eventually_principal by blast lemma eventually_at_right_field: "eventually P (at_right x) \ (\b>x. \y>x. y < b \ P y)" for x :: "'a::{linordered_field, linorder_topology}" using linordered_field_no_ub[rule_format, of x] by (auto simp: eventually_at_right) lemma eventually_at_left_field: "eventually P (at_left x) \ (\by>b. y < x \ P y)" for x :: "'a::{linordered_field, linorder_topology}" using linordered_field_no_lb[rule_format, of x] by (auto simp: eventually_at_left) lemma filtermap_nhds_eq_imp_filtermap_at_eq: assumes "filtermap f (nhds z) = nhds (f z)" assumes "eventually (\x. f x = f z \ x = z) (at z)" shows "filtermap f (at z) = at (f z)" proof (rule filter_eqI) fix P :: "'a \ bool" have "eventually P (filtermap f (at z)) \ (\\<^sub>F x in nhds z. x \ z \ P (f x))" by (simp add: eventually_filtermap eventually_at_filter) also have "\ \ (\\<^sub>F x in nhds z. f x \ f z \ P (f x))" by (rule eventually_cong [OF assms(2)[unfolded eventually_at_filter]]) auto also have "\ \ (\\<^sub>F x in filtermap f (nhds z). x \ f z \ P x)" by (simp add: eventually_filtermap) also have "filtermap f (nhds z) = nhds (f z)" by (rule assms) also have "(\\<^sub>F x in nhds (f z). x \ f z \ P x) \ (\\<^sub>F x in at (f z). P x)" by (simp add: eventually_at_filter) finally show "eventually P (filtermap f (at z)) = eventually P (at (f z))" . qed subsubsection \Tendsto\ abbreviation (in topological_space) tendsto :: "('b \ 'a) \ 'a \ 'b filter \ bool" (infixr "\" 55) where "(f \ l) F \ filterlim f (nhds l) F" definition (in t2_space) Lim :: "'f filter \ ('f \ 'a) \ 'a" where "Lim A f = (THE l. (f \ l) A)" lemma (in topological_space) tendsto_eq_rhs: "(f \ x) F \ x = y \ (f \ y) F" by simp named_theorems tendsto_intros "introduction rules for tendsto" setup \ Global_Theory.add_thms_dynamic (\<^binding>\tendsto_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\tendsto_intros\ |> map_filter (try (fn thm => @{thm tendsto_eq_rhs} OF [thm]))) \ context topological_space begin lemma tendsto_def: "(f \ l) F \ (\S. open S \ l \ S \ eventually (\x. f x \ S) F)" unfolding nhds_def filterlim_INF filterlim_principal by auto lemma tendsto_cong: "(f \ c) F \ (g \ c) F" if "eventually (\x. f x = g x) F" by (rule filterlim_cong [OF refl refl that]) lemma tendsto_mono: "F \ F' \ (f \ l) F' \ (f \ l) F" unfolding tendsto_def le_filter_def by fast lemma tendsto_ident_at [tendsto_intros, simp, intro]: "((\x. x) \ a) (at a within s)" by (auto simp: tendsto_def eventually_at_topological) lemma tendsto_const [tendsto_intros, simp, intro]: "((\x. k) \ k) F" by (simp add: tendsto_def) lemma filterlim_at: "(LIM x F. f x :> at b within s) \ eventually (\x. f x \ s \ f x \ b) F \ (f \ b) F" by (simp add: at_within_def filterlim_inf filterlim_principal conj_commute) lemma (in -) assumes "filterlim f (nhds L) F" shows tendsto_imp_filterlim_at_right: "eventually (\x. f x > L) F \ filterlim f (at_right L) F" and tendsto_imp_filterlim_at_left: "eventually (\x. f x < L) F \ filterlim f (at_left L) F" using assms by (auto simp: filterlim_at elim: eventually_mono) lemma filterlim_at_withinI: assumes "filterlim f (nhds c) F" assumes "eventually (\x. f x \ A - {c}) F" shows "filterlim f (at c within A) F" using assms by (simp add: filterlim_at) lemma filterlim_atI: assumes "filterlim f (nhds c) F" assumes "eventually (\x. f x \ c) F" shows "filterlim f (at c) F" using assms by (intro filterlim_at_withinI) simp_all lemma topological_tendstoI: "(\S. open S \ l \ S \ eventually (\x. f x \ S) F) \ (f \ l) F" by (auto simp: tendsto_def) lemma topological_tendstoD: "(f \ l) F \ open S \ l \ S \ eventually (\x. f x \ S) F" by (auto simp: tendsto_def) lemma tendsto_bot [simp]: "(f \ a) bot" by (simp add: tendsto_def) lemma tendsto_eventually: "eventually (\x. f x = l) net \ ((\x. f x) \ l) net" by (rule topological_tendstoI) (auto elim: eventually_mono) (* Contributed by Dominique Unruh *) lemma tendsto_principal_singleton[simp]: shows "(f \ f x) (principal {x})" unfolding tendsto_def eventually_principal by simp end lemma (in topological_space) filterlim_within_subset: "filterlim f l (at x within S) \ T \ S \ filterlim f l (at x within T)" by (blast intro: filterlim_mono at_le) lemmas tendsto_within_subset = filterlim_within_subset lemma (in order_topology) order_tendsto_iff: "(f \ x) F \ (\lx. l < f x) F) \ (\u>x. eventually (\x. f x < u) F)" by (auto simp: nhds_order filterlim_inf filterlim_INF filterlim_principal) lemma (in order_topology) order_tendstoI: "(\a. a < y \ eventually (\x. a < f x) F) \ (\a. y < a \ eventually (\x. f x < a) F) \ (f \ y) F" by (auto simp: order_tendsto_iff) lemma (in order_topology) order_tendstoD: assumes "(f \ y) F" shows "a < y \ eventually (\x. a < f x) F" and "y < a \ eventually (\x. f x < a) F" using assms by (auto simp: order_tendsto_iff) lemma (in linorder_topology) tendsto_max[tendsto_intros]: assumes X: "(X \ x) net" and Y: "(Y \ y) net" shows "((\x. max (X x) (Y x)) \ max x y) net" proof (rule order_tendstoI) fix a assume "a < max x y" then show "eventually (\x. a < max (X x) (Y x)) net" using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] by (auto simp: less_max_iff_disj elim: eventually_mono) next fix a assume "max x y < a" then show "eventually (\x. max (X x) (Y x) < a) net" using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] by (auto simp: eventually_conj_iff) qed lemma (in linorder_topology) tendsto_min[tendsto_intros]: assumes X: "(X \ x) net" and Y: "(Y \ y) net" shows "((\x. min (X x) (Y x)) \ min x y) net" proof (rule order_tendstoI) fix a assume "a < min x y" then show "eventually (\x. a < min (X x) (Y x)) net" using order_tendstoD(1)[OF X, of a] order_tendstoD(1)[OF Y, of a] by (auto simp: eventually_conj_iff) next fix a assume "min x y < a" then show "eventually (\x. min (X x) (Y x) < a) net" using order_tendstoD(2)[OF X, of a] order_tendstoD(2)[OF Y, of a] by (auto simp: min_less_iff_disj elim: eventually_mono) qed lemma (in order_topology) assumes "a < b" shows at_within_Icc_at_right: "at a within {a..b} = at_right a" and at_within_Icc_at_left: "at b within {a..b} = at_left b" using order_tendstoD(2)[OF tendsto_ident_at assms, of "{a<..}"] using order_tendstoD(1)[OF tendsto_ident_at assms, of "{.. x < b \ at x within {a..b} = at x" by (rule at_within_open_subset[where S="{a<.. bot" and "(f \ a) F" and "(f \ b) F" shows "a = b" proof (rule ccontr) assume "a \ b" obtain U V where "open U" "open V" "a \ U" "b \ V" "U \ V = {}" using hausdorff [OF \a \ b\] by fast have "eventually (\x. f x \ U) F" using \(f \ a) F\ \open U\ \a \ U\ by (rule topological_tendstoD) moreover have "eventually (\x. f x \ V) F" using \(f \ b) F\ \open V\ \b \ V\ by (rule topological_tendstoD) ultimately have "eventually (\x. False) F" proof eventually_elim case (elim x) then have "f x \ U \ V" by simp with \U \ V = {}\ show ?case by simp qed with \\ trivial_limit F\ show "False" by (simp add: trivial_limit_def) qed lemma (in t2_space) tendsto_const_iff: fixes a b :: 'a assumes "\ trivial_limit F" shows "((\x. a) \ b) F \ a = b" by (auto intro!: tendsto_unique [OF assms tendsto_const]) lemma (in t2_space) tendsto_unique': assumes "F \ bot" shows "\\<^sub>\\<^sub>1l. (f \ l) F" using Uniq_def assms local.tendsto_unique by fastforce lemma Lim_in_closed_set: assumes "closed S" "eventually (\x. f(x) \ S) F" "F \ bot" "(f \ l) F" shows "l \ S" proof (rule ccontr) assume "l \ S" with \closed S\ have "open (- S)" "l \ - S" by (simp_all add: open_Compl) with assms(4) have "eventually (\x. f x \ - S) F" by (rule topological_tendstoD) with assms(2) have "eventually (\x. False) F" by (rule eventually_elim2) simp with assms(3) show "False" by (simp add: eventually_False) qed lemma (in t3_space) nhds_closed: assumes "x \ A" and "open A" shows "\A'. x \ A' \ closed A' \ A' \ A \ eventually (\y. y \ A') (nhds x)" proof - from assms have "\U V. open U \ open V \ x \ U \ - A \ V \ U \ V = {}" by (intro t3_space) auto then obtain U V where UV: "open U" "open V" "x \ U" "-A \ V" "U \ V = {}" by auto have "eventually (\y. y \ U) (nhds x)" using \open U\ and \x \ U\ by (intro eventually_nhds_in_open) hence "eventually (\y. y \ -V) (nhds x)" by eventually_elim (use UV in auto) with UV show ?thesis by (intro exI[of _ "-V"]) auto qed lemma (in order_topology) increasing_tendsto: assumes bdd: "eventually (\n. f n \ l) F" and en: "\x. x < l \ eventually (\n. x < f n) F" shows "(f \ l) F" using assms by (intro order_tendstoI) (auto elim!: eventually_mono) lemma (in order_topology) decreasing_tendsto: assumes bdd: "eventually (\n. l \ f n) F" and en: "\x. l < x \ eventually (\n. f n < x) F" shows "(f \ l) F" using assms by (intro order_tendstoI) (auto elim!: eventually_mono) lemma (in order_topology) tendsto_sandwich: assumes ev: "eventually (\n. f n \ g n) net" "eventually (\n. g n \ h n) net" assumes lim: "(f \ c) net" "(h \ c) net" shows "(g \ c) net" proof (rule order_tendstoI) fix a show "a < c \ eventually (\x. a < g x) net" using order_tendstoD[OF lim(1), of a] ev by (auto elim: eventually_elim2) next fix a show "c < a \ eventually (\x. g x < a) net" using order_tendstoD[OF lim(2), of a] ev by (auto elim: eventually_elim2) qed lemma (in t1_space) limit_frequently_eq: assumes "F \ bot" and "frequently (\x. f x = c) F" and "(f \ d) F" shows "d = c" proof (rule ccontr) assume "d \ c" from t1_space[OF this] obtain U where "open U" "d \ U" "c \ U" by blast with assms have "eventually (\x. f x \ U) F" unfolding tendsto_def by blast then have "eventually (\x. f x \ c) F" by eventually_elim (insert \c \ U\, blast) with assms(2) show False unfolding frequently_def by contradiction qed lemma (in t1_space) tendsto_imp_eventually_ne: assumes "(f \ c) F" "c \ c'" shows "eventually (\z. f z \ c') F" proof (cases "F=bot") case True thus ?thesis by auto next case False show ?thesis proof (rule ccontr) assume "\ eventually (\z. f z \ c') F" then have "frequently (\z. f z = c') F" by (simp add: frequently_def) from limit_frequently_eq[OF False this \(f \ c) F\] and \c \ c'\ show False by contradiction qed qed lemma (in linorder_topology) tendsto_le: assumes F: "\ trivial_limit F" and x: "(f \ x) F" and y: "(g \ y) F" and ev: "eventually (\x. g x \ f x) F" shows "y \ x" proof (rule ccontr) assume "\ y \ x" with less_separate[of x y] obtain a b where xy: "x < a" "b < y" "{.. {b<..} = {}" by (auto simp: not_le) then have "eventually (\x. f x < a) F" "eventually (\x. b < g x) F" using x y by (auto intro: order_tendstoD) with ev have "eventually (\x. False) F" by eventually_elim (insert xy, fastforce) with F show False by (simp add: eventually_False) qed lemma (in linorder_topology) tendsto_lowerbound: assumes x: "(f \ x) F" and ev: "eventually (\i. a \ f i) F" and F: "\ trivial_limit F" shows "a \ x" using F x tendsto_const ev by (rule tendsto_le) lemma (in linorder_topology) tendsto_upperbound: assumes x: "(f \ x) F" and ev: "eventually (\i. a \ f i) F" and F: "\ trivial_limit F" shows "a \ x" by (rule tendsto_le [OF F tendsto_const x ev]) lemma filterlim_at_within_not_equal: fixes f::"'a \ 'b::t2_space" assumes "filterlim f (at a within s) F" shows "eventually (\w. f w\s \ f w \b) F" proof (cases "a=b") case True then show ?thesis using assms by (simp add: filterlim_at) next case False from hausdorff[OF this] obtain U V where UV:"open U" "open V" "a \ U" "b \ V" "U \ V = {}" by auto have "(f \ a) F" using assms filterlim_at by auto then have "\\<^sub>F x in F. f x \ U" using UV unfolding tendsto_def by auto moreover have "\\<^sub>F x in F. f x \ s \ f x\a" using assms filterlim_at by auto ultimately show ?thesis apply eventually_elim using UV by auto qed subsubsection \Rules about \<^const>\Lim\\ lemma tendsto_Lim: "\ trivial_limit net \ (f \ l) net \ Lim net f = l" unfolding Lim_def using tendsto_unique [of net f] by auto lemma Lim_ident_at: "\ trivial_limit (at x within s) \ Lim (at x within s) (\x. x) = x" by (rule tendsto_Lim[OF _ tendsto_ident_at]) auto lemma Lim_cong: assumes "eventually (\x. f x = g x) F" "F = G" shows "Lim F f = Lim G g" proof (cases "(\c. (f \ c) F) \ F \ bot") case True then obtain c where c: "(f \ c) F" by blast hence "Lim F f = c" using True by (intro tendsto_Lim) auto moreover have "(f \ c) F \ (g \ c) G" using assms by (intro filterlim_cong) auto with True c assms have "Lim G g = c" by (intro tendsto_Lim) auto ultimately show ?thesis by simp next case False show ?thesis proof (cases "F = bot") case True thus ?thesis using assms by (auto simp: Topological_Spaces.Lim_def) next case False have "(f \ c) F \ (g \ c) G" for c using assms by (intro filterlim_cong) auto thus ?thesis by (auto simp: Topological_Spaces.Lim_def) qed qed lemma eventually_Lim_ident_at: "(\\<^sub>F y in at x within X. P (Lim (at x within X) (\x. x)) y) \ (\\<^sub>F y in at x within X. P x y)" for x::"'a::t2_space" by (cases "at x within X = bot") (auto simp: Lim_ident_at) lemma filterlim_at_bot_at_right: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" and bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" and Q: "eventually Q (at_right a)" and bound: "\b. Q b \ a < b" and P: "eventually P at_bot" shows "filterlim f at_bot (at_right a)" proof - from P obtain x where x: "\y. y \ x \ P y" unfolding eventually_at_bot_linorder by auto show ?thesis proof (intro filterlim_at_bot_le[THEN iffD2] allI impI) fix z assume "z \ x" with x have "P z" by auto have "eventually (\x. x \ g z) (at_right a)" using bound[OF bij(2)[OF \P z\]] unfolding eventually_at_right[OF bound[OF bij(2)[OF \P z\]]] by (auto intro!: exI[of _ "g z"]) with Q show "eventually (\x. f x \ z) (at_right a)" by eventually_elim (metis bij \P z\ mono) qed qed lemma filterlim_at_top_at_left: fixes f :: "'a::linorder_topology \ 'b::linorder" assumes mono: "\x y. Q x \ Q y \ x \ y \ f x \ f y" and bij: "\x. P x \ f (g x) = x" "\x. P x \ Q (g x)" and Q: "eventually Q (at_left a)" and bound: "\b. Q b \ b < a" and P: "eventually P at_top" shows "filterlim f at_top (at_left a)" proof - from P obtain x where x: "\y. x \ y \ P y" unfolding eventually_at_top_linorder by auto show ?thesis proof (intro filterlim_at_top_ge[THEN iffD2] allI impI) fix z assume "x \ z" with x have "P z" by auto have "eventually (\x. g z \ x) (at_left a)" using bound[OF bij(2)[OF \P z\]] unfolding eventually_at_left[OF bound[OF bij(2)[OF \P z\]]] by (auto intro!: exI[of _ "g z"]) with Q show "eventually (\x. z \ f x) (at_left a)" by eventually_elim (metis bij \P z\ mono) qed qed lemma filterlim_split_at: "filterlim f F (at_left x) \ filterlim f F (at_right x) \ filterlim f F (at x)" for x :: "'a::linorder_topology" by (subst at_eq_sup_left_right) (rule filterlim_sup) lemma filterlim_at_split: "filterlim f F (at x) \ filterlim f F (at_left x) \ filterlim f F (at_right x)" for x :: "'a::linorder_topology" by (subst at_eq_sup_left_right) (simp add: filterlim_def filtermap_sup) lemma eventually_nhds_top: fixes P :: "'a :: {order_top,linorder_topology} \ bool" and b :: 'a assumes "b < top" shows "eventually P (nhds top) \ (\bz. b < z \ P z))" unfolding eventually_nhds proof safe fix S :: "'a set" assume "open S" "top \ S" note open_left[OF this \b < top\] moreover assume "\s\S. P s" ultimately show "\bz>b. P z" by (auto simp: subset_eq Ball_def) next fix b assume "b < top" "\z>b. P z" then show "\S. open S \ top \ S \ (\xa\S. P xa)" by (intro exI[of _ "{b <..}"]) auto qed lemma tendsto_at_within_iff_tendsto_nhds: "(g \ g l) (at l within S) \ (g \ g l) (inf (nhds l) (principal S))" unfolding tendsto_def eventually_at_filter eventually_inf_principal by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) subsection \Limits on sequences\ abbreviation (in topological_space) LIMSEQ :: "[nat \ 'a, 'a] \ bool" ("((_)/ \ (_))" [60, 60] 60) where "X \ L \ (X \ L) sequentially" abbreviation (in t2_space) lim :: "(nat \ 'a) \ 'a" where "lim X \ Lim sequentially X" definition (in topological_space) convergent :: "(nat \ 'a) \ bool" where "convergent X = (\L. X \ L)" lemma lim_def: "lim X = (THE L. X \ L)" unfolding Lim_def .. lemma lim_explicit: "f \ f0 \ (\S. open S \ f0 \ S \ (\N. \n\N. f n \ S))" unfolding tendsto_def eventually_sequentially by auto subsection \Monotone sequences and subsequences\ text \ Definition of monotonicity. The use of disjunction here complicates proofs considerably. One alternative is to add a Boolean argument to indicate the direction. Another is to develop the notions of increasing and decreasing first. \ definition monoseq :: "(nat \ 'a::order) \ bool" where "monoseq X \ (\m. \n\m. X m \ X n) \ (\m. \n\m. X n \ X m)" abbreviation incseq :: "(nat \ 'a::order) \ bool" where "incseq X \ mono X" lemma incseq_def: "incseq X \ (\m. \n\m. X n \ X m)" unfolding mono_def .. abbreviation decseq :: "(nat \ 'a::order) \ bool" where "decseq X \ antimono X" lemma decseq_def: "decseq X \ (\m. \n\m. X n \ X m)" unfolding antimono_def .. subsubsection \Definition of subsequence.\ (* For compatibility with the old "subseq" *) lemma strict_mono_leD: "strict_mono r \ m \ n \ r m \ r n" by (erule (1) monoD [OF strict_mono_mono]) lemma strict_mono_id: "strict_mono id" by (simp add: strict_mono_def) lemma incseq_SucI: "(\n. X n \ X (Suc n)) \ incseq X" using lift_Suc_mono_le[of X] by (auto simp: incseq_def) lemma incseqD: "incseq f \ i \ j \ f i \ f j" by (auto simp: incseq_def) lemma incseq_SucD: "incseq A \ A i \ A (Suc i)" using incseqD[of A i "Suc i"] by auto lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" by (auto intro: incseq_SucI dest: incseq_SucD) lemma incseq_const[simp, intro]: "incseq (\x. k)" unfolding incseq_def by auto lemma decseq_SucI: "(\n. X (Suc n) \ X n) \ decseq X" using order.lift_Suc_mono_le[OF dual_order, of X] by (auto simp: decseq_def) lemma decseqD: "decseq f \ i \ j \ f j \ f i" by (auto simp: decseq_def) lemma decseq_SucD: "decseq A \ A (Suc i) \ A i" using decseqD[of A i "Suc i"] by auto lemma decseq_Suc_iff: "decseq f \ (\n. f (Suc n) \ f n)" by (auto intro: decseq_SucI dest: decseq_SucD) lemma decseq_const[simp, intro]: "decseq (\x. k)" unfolding decseq_def by auto lemma monoseq_iff: "monoseq X \ incseq X \ decseq X" unfolding monoseq_def incseq_def decseq_def .. lemma monoseq_Suc: "monoseq X \ (\n. X n \ X (Suc n)) \ (\n. X (Suc n) \ X n)" unfolding monoseq_iff incseq_Suc_iff decseq_Suc_iff .. lemma monoI1: "\m. \n \ m. X m \ X n \ monoseq X" by (simp add: monoseq_def) lemma monoI2: "\m. \n \ m. X n \ X m \ monoseq X" by (simp add: monoseq_def) lemma mono_SucI1: "\n. X n \ X (Suc n) \ monoseq X" by (simp add: monoseq_Suc) lemma mono_SucI2: "\n. X (Suc n) \ X n \ monoseq X" by (simp add: monoseq_Suc) lemma monoseq_minus: fixes a :: "nat \ 'a::ordered_ab_group_add" assumes "monoseq a" shows "monoseq (\ n. - a n)" proof (cases "\m. \n \ m. a m \ a n") case True then have "\m. \n \ m. - a n \ - a m" by auto then show ?thesis by (rule monoI2) next case False then have "\m. \n \ m. - a m \ - a n" using \monoseq a\[unfolded monoseq_def] by auto then show ?thesis by (rule monoI1) qed subsubsection \Subsequence (alternative definition, (e.g. Hoskins)\ -lemma strict_mono_Suc_iff: "strict_mono f \ (\n. f n < f (Suc n))" -proof (intro iffI strict_monoI) - assume *: "\n. f n < f (Suc n)" - fix m n :: nat assume "m < n" - thus "f m < f n" - by (induction rule: less_Suc_induct) (use * in auto) -qed (auto simp: strict_mono_def) - -lemma strict_mono_add: "strict_mono (\n::'a::linordered_semidom. n + k)" - by (auto simp: strict_mono_def) - text \For any sequence, there is a monotonic subsequence.\ lemma seq_monosub: fixes s :: "nat \ 'a::linorder" shows "\f. strict_mono f \ monoseq (\n. (s (f n)))" proof (cases "\n. \p>n. \m\p. s m \ s p") case True then have "\f. \n. (\m\f n. s m \ s (f n)) \ f n < f (Suc n)" by (intro dependent_nat_choice) (auto simp: conj_commute) then obtain f :: "nat \ nat" where f: "strict_mono f" and mono: "\n m. f n \ m \ s m \ s (f n)" by (auto simp: strict_mono_Suc_iff) then have "incseq f" unfolding strict_mono_Suc_iff incseq_Suc_iff by (auto intro: less_imp_le) then have "monoseq (\n. s (f n))" by (auto simp add: incseq_def intro!: mono monoI2) with f show ?thesis by auto next case False then obtain N where N: "p > N \ \m>p. s p < s m" for p by (force simp: not_le le_less) have "\f. \n. N < f n \ f n < f (Suc n) \ s (f n) \ s (f (Suc n))" proof (intro dependent_nat_choice) fix x assume "N < x" with N[of x] show "\y>N. x < y \ s x \ s y" by (auto intro: less_trans) qed auto then show ?thesis by (auto simp: monoseq_iff incseq_Suc_iff strict_mono_Suc_iff) qed lemma seq_suble: assumes sf: "strict_mono (f :: nat \ nat)" shows "n \ f n" proof (induct n) case 0 show ?case by simp next case (Suc n) with sf [unfolded strict_mono_Suc_iff, rule_format, of n] have "n < f (Suc n)" by arith then show ?case by arith qed lemma eventually_subseq: "strict_mono r \ eventually P sequentially \ eventually (\n. P (r n)) sequentially" unfolding eventually_sequentially by (metis seq_suble le_trans) lemma not_eventually_sequentiallyD: assumes "\ eventually P sequentially" shows "\r::nat\nat. strict_mono r \ (\n. \ P (r n))" proof - from assms have "\n. \m\n. \ P m" unfolding eventually_sequentially by (simp add: not_less) then obtain r where "\n. r n \ n" "\n. \ P (r n)" by (auto simp: choice_iff) then show ?thesis by (auto intro!: exI[of _ "\n. r (((Suc \ r) ^^ Suc n) 0)"] simp: less_eq_Suc_le strict_mono_Suc_iff) qed lemma sequentially_offset: assumes "eventually (\i. P i) sequentially" shows "eventually (\i. P (i + k)) sequentially" using assms by (rule eventually_sequentially_seg [THEN iffD2]) lemma seq_offset_neg: "(f \ l) sequentially \ ((\i. f(i - k)) \ l) sequentially" apply (erule filterlim_compose) apply (simp add: filterlim_def le_sequentially eventually_filtermap eventually_sequentially, arith) done lemma filterlim_subseq: "strict_mono f \ filterlim f sequentially sequentially" unfolding filterlim_iff by (metis eventually_subseq) lemma strict_mono_o: "strict_mono r \ strict_mono s \ strict_mono (r \ s)" unfolding strict_mono_def by simp lemma strict_mono_compose: "strict_mono r \ strict_mono s \ strict_mono (\x. r (s x))" using strict_mono_o[of r s] by (simp add: o_def) lemma incseq_imp_monoseq: "incseq X \ monoseq X" by (simp add: incseq_def monoseq_def) lemma decseq_imp_monoseq: "decseq X \ monoseq X" by (simp add: decseq_def monoseq_def) lemma decseq_eq_incseq: "decseq X = incseq (\n. - X n)" for X :: "nat \ 'a::ordered_ab_group_add" by (simp add: decseq_def incseq_def) lemma INT_decseq_offset: assumes "decseq F" shows "(\i. F i) = (\i\{n..}. F i)" proof safe fix x i assume x: "x \ (\i\{n..}. F i)" show "x \ F i" proof cases from x have "x \ F n" by auto also assume "i \ n" with \decseq F\ have "F n \ F i" unfolding decseq_def by simp finally show ?thesis . qed (insert x, simp) qed auto lemma LIMSEQ_const_iff: "(\n. k) \ l \ k = l" for k l :: "'a::t2_space" using trivial_limit_sequentially by (rule tendsto_const_iff) lemma LIMSEQ_SUP: "incseq X \ X \ (SUP i. X i :: 'a::{complete_linorder,linorder_topology})" by (intro increasing_tendsto) (auto simp: SUP_upper less_SUP_iff incseq_def eventually_sequentially intro: less_le_trans) lemma LIMSEQ_INF: "decseq X \ X \ (INF i. X i :: 'a::{complete_linorder,linorder_topology})" by (intro decreasing_tendsto) (auto simp: INF_lower INF_less_iff decseq_def eventually_sequentially intro: le_less_trans) lemma LIMSEQ_ignore_initial_segment: "f \ a \ (\n. f (n + k)) \ a" unfolding tendsto_def by (subst eventually_sequentially_seg[where k=k]) lemma LIMSEQ_offset: "(\n. f (n + k)) \ a \ f \ a" unfolding tendsto_def by (subst (asm) eventually_sequentially_seg[where k=k]) lemma LIMSEQ_Suc: "f \ l \ (\n. f (Suc n)) \ l" by (drule LIMSEQ_ignore_initial_segment [where k="Suc 0"]) simp lemma LIMSEQ_imp_Suc: "(\n. f (Suc n)) \ l \ f \ l" by (rule LIMSEQ_offset [where k="Suc 0"]) simp lemma LIMSEQ_lessThan_iff_atMost: shows "(\n. f {.. x \ (\n. f {..n}) \ x" apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp only: lessThan_Suc_atMost) done lemma (in t2_space) LIMSEQ_Uniq: "\\<^sub>\\<^sub>1l. X \ l" by (simp add: tendsto_unique') lemma (in t2_space) LIMSEQ_unique: "X \ a \ X \ b \ a = b" using trivial_limit_sequentially by (rule tendsto_unique) lemma LIMSEQ_le_const: "X \ x \ \N. \n\N. a \ X n \ a \ x" for a x :: "'a::linorder_topology" by (simp add: eventually_at_top_linorder tendsto_lowerbound) lemma LIMSEQ_le: "X \ x \ Y \ y \ \N. \n\N. X n \ Y n \ x \ y" for x y :: "'a::linorder_topology" using tendsto_le[of sequentially Y y X x] by (simp add: eventually_sequentially) lemma LIMSEQ_le_const2: "X \ x \ \N. \n\N. X n \ a \ x \ a" for a x :: "'a::linorder_topology" by (rule LIMSEQ_le[of X x "\n. a"]) auto lemma Lim_bounded: "f \ l \ \n\M. f n \ C \ l \ C" for l :: "'a::linorder_topology" by (intro LIMSEQ_le_const2) auto lemma Lim_bounded2: fixes f :: "nat \ 'a::linorder_topology" assumes lim:"f \ l" and ge: "\n\N. f n \ C" shows "l \ C" using ge by (intro tendsto_le[OF trivial_limit_sequentially lim tendsto_const]) (auto simp: eventually_sequentially) lemma lim_mono: fixes X Y :: "nat \ 'a::linorder_topology" assumes "\n. N \ n \ X n \ Y n" and "X \ x" and "Y \ y" shows "x \ y" using assms(1) by (intro LIMSEQ_le[OF assms(2,3)]) auto lemma Sup_lim: fixes a :: "'a::{complete_linorder,linorder_topology}" assumes "\n. b n \ s" and "b \ a" shows "a \ Sup s" by (metis Lim_bounded assms complete_lattice_class.Sup_upper) lemma Inf_lim: fixes a :: "'a::{complete_linorder,linorder_topology}" assumes "\n. b n \ s" and "b \ a" shows "Inf s \ a" by (metis Lim_bounded2 assms complete_lattice_class.Inf_lower) lemma SUP_Lim: fixes X :: "nat \ 'a::{complete_linorder,linorder_topology}" assumes inc: "incseq X" and l: "X \ l" shows "(SUP n. X n) = l" using LIMSEQ_SUP[OF inc] tendsto_unique[OF trivial_limit_sequentially l] by simp lemma INF_Lim: fixes X :: "nat \ 'a::{complete_linorder,linorder_topology}" assumes dec: "decseq X" and l: "X \ l" shows "(INF n. X n) = l" using LIMSEQ_INF[OF dec] tendsto_unique[OF trivial_limit_sequentially l] by simp lemma convergentD: "convergent X \ \L. X \ L" by (simp add: convergent_def) lemma convergentI: "X \ L \ convergent X" by (auto simp add: convergent_def) lemma convergent_LIMSEQ_iff: "convergent X \ X \ lim X" by (auto intro: theI LIMSEQ_unique simp add: convergent_def lim_def) lemma convergent_const: "convergent (\n. c)" by (rule convergentI) (rule tendsto_const) lemma monoseq_le: "monoseq a \ a \ x \ (\n. a n \ x) \ (\m. \n\m. a m \ a n) \ (\n. x \ a n) \ (\m. \n\m. a n \ a m)" for x :: "'a::linorder_topology" by (metis LIMSEQ_le_const LIMSEQ_le_const2 decseq_def incseq_def monoseq_iff) lemma LIMSEQ_subseq_LIMSEQ: "X \ L \ strict_mono f \ (X \ f) \ L" unfolding comp_def by (rule filterlim_compose [of X, OF _ filterlim_subseq]) lemma convergent_subseq_convergent: "convergent X \ strict_mono f \ convergent (X \ f)" by (auto simp: convergent_def intro: LIMSEQ_subseq_LIMSEQ) lemma limI: "X \ L \ lim X = L" by (rule tendsto_Lim) (rule trivial_limit_sequentially) lemma lim_le: "convergent f \ (\n. f n \ x) \ lim f \ x" for x :: "'a::linorder_topology" using LIMSEQ_le_const2[of f "lim f" x] by (simp add: convergent_LIMSEQ_iff) lemma lim_const [simp]: "lim (\m. a) = a" by (simp add: limI) subsubsection \Increasing and Decreasing Series\ lemma incseq_le: "incseq X \ X \ L \ X n \ L" for L :: "'a::linorder_topology" by (metis incseq_def LIMSEQ_le_const) lemma decseq_ge: "decseq X \ X \ L \ L \ X n" for L :: "'a::linorder_topology" by (metis decseq_def LIMSEQ_le_const2) subsection \First countable topologies\ class first_countable_topology = topological_space + assumes first_countable_basis: "\A::nat \ 'a set. (\i. x \ A i \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" lemma (in first_countable_topology) countable_basis_at_decseq: obtains A :: "nat \ 'a set" where "\i. open (A i)" "\i. x \ (A i)" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" proof atomize_elim from first_countable_basis[of x] obtain A :: "nat \ 'a set" where nhds: "\i. open (A i)" "\i. x \ A i" and incl: "\S. open S \ x \ S \ \i. A i \ S" by auto define F where "F n = (\i\n. A i)" for n show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially)" proof (safe intro!: exI[of _ F]) fix i show "open (F i)" using nhds(1) by (auto simp: F_def) show "x \ F i" using nhds(2) by (auto simp: F_def) next fix S assume "open S" "x \ S" from incl[OF this] obtain i where "F i \ S" unfolding F_def by auto moreover have "\j. i \ j \ F j \ F i" by (simp add: Inf_superset_mono F_def image_mono) ultimately show "eventually (\i. F i \ S) sequentially" by (auto simp: eventually_sequentially) qed qed lemma (in first_countable_topology) nhds_countable: obtains X :: "nat \ 'a set" where "decseq X" "\n. open (X n)" "\n. x \ X n" "nhds x = (INF n. principal (X n))" proof - from first_countable_basis obtain A :: "nat \ 'a set" where *: "\n. x \ A n" "\n. open (A n)" "\S. open S \ x \ S \ \i. A i \ S" by metis show thesis proof show "decseq (\n. \i\n. A i)" by (simp add: antimono_iff_le_Suc atMost_Suc) show "x \ (\i\n. A i)" "\n. open (\i\n. A i)" for n using * by auto with * show "nhds x = (INF n. principal (\i\n. A i))" unfolding nhds_def apply (intro INF_eq) apply fastforce apply blast done qed qed lemma (in first_countable_topology) countable_basis: obtains A :: "nat \ 'a set" where "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" proof atomize_elim obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. x \ A i" "\S. open S \ x \ S \ eventually (\i. A i \ S) sequentially" by (rule countable_basis_at_decseq) blast have "eventually (\n. F n \ S) sequentially" if "\n. F n \ A n" "open S" "x \ S" for F S using *(3)[of S] that by (auto elim: eventually_mono simp: subset_eq) with * show "\A. (\i. open (A i)) \ (\i. x \ A i) \ (\F. (\n. F n \ A n) \ F \ x)" by (intro exI[of _ A]) (auto simp: tendsto_def) qed lemma (in first_countable_topology) sequentially_imp_eventually_nhds_within: assumes "\f. (\n. f n \ s) \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (inf (nhds a) (principal s))" proof (rule ccontr) obtain A :: "nat \ 'a set" where *: "\i. open (A i)" "\i. a \ A i" "\F. \n. F n \ A n \ F \ a" by (rule countable_basis) blast assume "\ ?thesis" with * have "\F. \n. F n \ s \ F n \ A n \ \ P (F n)" unfolding eventually_inf_principal eventually_nhds by (intro choice) fastforce then obtain F where F: "\n. F n \ s" and "\n. F n \ A n" and F': "\n. \ P (F n)" by blast with * have "F \ a" by auto then have "eventually (\n. P (F n)) sequentially" using assms F by simp then show False by (simp add: F') qed lemma (in first_countable_topology) eventually_nhds_within_iff_sequentially: "eventually P (inf (nhds a) (principal s)) \ (\f. (\n. f n \ s) \ f \ a \ eventually (\n. P (f n)) sequentially)" proof (safe intro!: sequentially_imp_eventually_nhds_within) assume "eventually P (inf (nhds a) (principal s))" then obtain S where "open S" "a \ S" "\x\S. x \ s \ P x" by (auto simp: eventually_inf_principal eventually_nhds) moreover fix f assume "\n. f n \ s" "f \ a" ultimately show "eventually (\n. P (f n)) sequentially" by (auto dest!: topological_tendstoD elim: eventually_mono) qed lemma (in first_countable_topology) eventually_nhds_iff_sequentially: "eventually P (nhds a) \ (\f. f \ a \ eventually (\n. P (f n)) sequentially)" using eventually_nhds_within_iff_sequentially[of P a UNIV] by simp (*Thanks to Sébastien Gouëzel*) lemma Inf_as_limit: fixes A::"'a::{linorder_topology, first_countable_topology, complete_linorder} set" assumes "A \ {}" shows "\u. (\n. u n \ A) \ u \ Inf A" proof (cases "Inf A \ A") case True show ?thesis by (rule exI[of _ "\n. Inf A"], auto simp add: True) next case False obtain y where "y \ A" using assms by auto then have "Inf A < y" using False Inf_lower less_le by auto obtain F :: "nat \ 'a set" where F: "\i. open (F i)" "\i. Inf A \ F i" "\u. (\n. u n \ F n) \ u \ Inf A" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ F n \ z \ A)" have "\z. z \ U \ z \ A" if "Inf A \ U" "open U" for U proof - obtain b where "b > Inf A" "{Inf A .. U" using open_right[OF \open U\ \Inf A \ U\ \Inf A < y\] by auto obtain z where "z < b" "z \ A" using \Inf A < b\ Inf_less_iff by auto then have "z \ {Inf A ..z \ A\ \{Inf A .. U\ by auto qed then have *: "u n \ F n \ u n \ A" for n using \Inf A \ F n\ \open (F n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ Inf A" using F(3) by simp then show ?thesis using * by auto qed lemma tendsto_at_iff_sequentially: "(f \ a) (at x within s) \ (\X. (\i. X i \ s - {x}) \ X \ x \ ((f \ X) \ a))" for f :: "'a::first_countable_topology \ _" unfolding filterlim_def[of _ "nhds a"] le_filter_def eventually_filtermap at_within_def eventually_nhds_within_iff_sequentially comp_def by metis lemma approx_from_above_dense_linorder: fixes x::"'a::{dense_linorder, linorder_topology, first_countable_topology}" assumes "x < y" shows "\u. (\n. u n > x) \ (u \ x)" proof - obtain A :: "nat \ 'a set" where A: "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ A n \ z > x)" have "\z. z \ U \ x < z" if "x \ U" "open U" for U using open_right[OF \open U\ \x \ U\ \x < y\] by (meson atLeastLessThan_iff dense less_imp_le subset_eq) then have *: "u n \ A n \ x < u n" for n using \x \ A n\ \open (A n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ x" using A(3) by simp then show ?thesis using * by auto qed lemma approx_from_below_dense_linorder: fixes x::"'a::{dense_linorder, linorder_topology, first_countable_topology}" assumes "x > y" shows "\u. (\n. u n < x) \ (u \ x)" proof - obtain A :: "nat \ 'a set" where A: "\i. open (A i)" "\i. x \ A i" "\F. (\n. F n \ A n) \ F \ x" by (metis first_countable_topology_class.countable_basis) define u where "u = (\n. SOME z. z \ A n \ z < x)" have "\z. z \ U \ z < x" if "x \ U" "open U" for U using open_left[OF \open U\ \x \ U\ \x > y\] by (meson dense greaterThanAtMost_iff less_imp_le subset_eq) then have *: "u n \ A n \ u n < x" for n using \x \ A n\ \open (A n)\ unfolding u_def by (metis (no_types, lifting) someI_ex) then have "u \ x" using A(3) by simp then show ?thesis using * by auto qed subsection \Function limit at a point\ abbreviation LIM :: "('a::topological_space \ 'b::topological_space) \ 'a \ 'b \ bool" ("((_)/ \(_)/\ (_))" [60, 0, 60] 60) where "f \a\ L \ (f \ L) (at a)" lemma tendsto_within_open: "a \ S \ open S \ (f \ l) (at a within S) \ (f \a\ l)" by (simp add: tendsto_def at_within_open[where S = S]) lemma tendsto_within_open_NO_MATCH: "a \ S \ NO_MATCH UNIV S \ open S \ (f \ l)(at a within S) \ (f \ l)(at a)" for f :: "'a::topological_space \ 'b::topological_space" using tendsto_within_open by blast lemma LIM_const_not_eq[tendsto_intros]: "k \ L \ \ (\x. k) \a\ L" for a :: "'a::perfect_space" and k L :: "'b::t2_space" by (simp add: tendsto_const_iff) lemmas LIM_not_zero = LIM_const_not_eq [where L = 0] lemma LIM_const_eq: "(\x. k) \a\ L \ k = L" for a :: "'a::perfect_space" and k L :: "'b::t2_space" by (simp add: tendsto_const_iff) lemma LIM_unique: "f \a\ L \ f \a\ M \ L = M" for a :: "'a::perfect_space" and L M :: "'b::t2_space" using at_neq_bot by (rule tendsto_unique) lemma LIM_Uniq: "\\<^sub>\\<^sub>1L::'b::t2_space. f \a\ L" for a :: "'a::perfect_space" by (auto simp add: Uniq_def LIM_unique) text \Limits are equal for functions equal except at limit point.\ lemma LIM_equal: "\x. x \ a \ f x = g x \ (f \a\ l) \ (g \a\ l)" by (simp add: tendsto_def eventually_at_topological) lemma LIM_cong: "a = b \ (\x. x \ b \ f x = g x) \ l = m \ (f \a\ l) \ (g \b\ m)" by (simp add: LIM_equal) lemma tendsto_cong_limit: "(f \ l) F \ k = l \ (f \ k) F" by simp lemma tendsto_at_iff_tendsto_nhds: "g \l\ g l \ (g \ g l) (nhds l)" unfolding tendsto_def eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma tendsto_compose: "g \l\ g l \ (f \ l) F \ ((\x. g (f x)) \ g l) F" unfolding tendsto_at_iff_tendsto_nhds by (rule filterlim_compose[of g]) lemma tendsto_compose_eventually: "g \l\ m \ (f \ l) F \ eventually (\x. f x \ l) F \ ((\x. g (f x)) \ m) F" by (rule filterlim_compose[of g _ "at l"]) (auto simp add: filterlim_at) lemma LIM_compose_eventually: assumes "f \a\ b" and "g \b\ c" and "eventually (\x. f x \ b) (at a)" shows "(\x. g (f x)) \a\ c" using assms(2,1,3) by (rule tendsto_compose_eventually) lemma tendsto_compose_filtermap: "((g \ f) \ T) F \ (g \ T) (filtermap f F)" by (simp add: filterlim_def filtermap_filtermap comp_def) lemma tendsto_compose_at: assumes f: "(f \ y) F" and g: "(g \ z) (at y)" and fg: "eventually (\w. f w = y \ g y = z) F" shows "((g \ f) \ z) F" proof - have "(\\<^sub>F a in F. f a \ y) \ g y = z" using fg by force moreover have "(g \ z) (filtermap f F) \ \ (\\<^sub>F a in F. f a \ y)" by (metis (no_types) filterlim_atI filterlim_def tendsto_mono f g) ultimately show ?thesis by (metis (no_types) f filterlim_compose filterlim_filtermap g tendsto_at_iff_tendsto_nhds tendsto_compose_filtermap) qed lemma tendsto_nhds_iff: "(f \ (c :: 'a :: t1_space)) (nhds x) \ f \x\ c \ f x = c" proof safe assume lim: "(f \ c) (nhds x)" show "f x = c" proof (rule ccontr) assume "f x \ c" hence "c \ f x" by auto then obtain A where A: "open A" "c \ A" "f x \ A" by (subst (asm) separation_t1) auto with lim obtain B where "open B" "x \ B" "\x. x \ B \ f x \ A" unfolding tendsto_def eventually_nhds by metis with \f x \ A\ show False by blast qed show "(f \ c) (at x)" using lim by (rule filterlim_mono) (auto simp: at_within_def) next assume "f \x\ f x" "c = f x" thus "(f \ f x) (nhds x)" unfolding tendsto_def eventually_at_filter by (fast elim: eventually_mono) qed subsubsection \Relation of \LIM\ and \LIMSEQ\\ lemma (in first_countable_topology) sequentially_imp_eventually_within: "(\f. (\n. f n \ s \ f n \ a) \ f \ a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a within s)" unfolding at_within_def by (intro sequentially_imp_eventually_nhds_within) auto lemma (in first_countable_topology) sequentially_imp_eventually_at: "(\f. (\n. f n \ a) \ f \ a \ eventually (\n. P (f n)) sequentially) \ eventually P (at a)" using sequentially_imp_eventually_within [where s=UNIV] by simp lemma LIMSEQ_SEQ_conv: "(\S. (\n. S n \ a) \ S \ a \ (\n. X (S n)) \ L) \ X \a\ L" (is "?lhs=?rhs") for a :: "'a::first_countable_topology" and L :: "'b::topological_space" proof assume ?lhs then show ?rhs by (simp add: sequentially_imp_eventually_within tendsto_def) next assume ?rhs then show ?lhs using tendsto_compose_eventually eventuallyI by blast qed lemma sequentially_imp_eventually_at_left: fixes a :: "'a::{linorder_topology,first_countable_topology}" assumes b[simp]: "b < a" and *: "\f. (\n. b < f n) \ (\n. f n < a) \ incseq f \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (at_left a)" proof (safe intro!: sequentially_imp_eventually_within) fix X assume X: "\n. X n \ {..< a} \ X n \ a" "X \ a" show "eventually (\n. P (X n)) sequentially" proof (rule ccontr) assume neg: "\ ?thesis" have "\s. \n. (\ P (X (s n)) \ b < X (s n)) \ (X (s n) \ X (s (Suc n)) \ Suc (s n) \ s (Suc n))" (is "\s. ?P s") proof (rule dependent_nat_choice) have "\ eventually (\n. b < X n \ P (X n)) sequentially" by (intro not_eventually_impI neg order_tendstoD(1) [OF X(2) b]) then show "\x. \ P (X x) \ b < X x" by (auto dest!: not_eventuallyD) next fix x n have "\ eventually (\n. Suc x \ n \ b < X n \ X x < X n \ P (X n)) sequentially" using X by (intro not_eventually_impI order_tendstoD(1)[OF X(2)] eventually_ge_at_top neg) auto then show "\n. (\ P (X n) \ b < X n) \ (X x \ X n \ Suc x \ n)" by (auto dest!: not_eventuallyD) qed then obtain s where "?P s" .. with X have "b < X (s n)" and "X (s n) < a" and "incseq (\n. X (s n))" and "(\n. X (s n)) \ a" and "\ P (X (s n))" for n by (auto simp: strict_mono_Suc_iff Suc_le_eq incseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \X \ a\, unfolded comp_def]) from *[OF this(1,2,3,4)] this(5) show False by auto qed qed lemma tendsto_at_left_sequentially: fixes a b :: "'b::{linorder_topology,first_countable_topology}" assumes "b < a" assumes *: "\S. (\n. S n < a) \ (\n. b < S n) \ incseq S \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at_left a)" using assms by (simp add: tendsto_def [where l=L] sequentially_imp_eventually_at_left) lemma sequentially_imp_eventually_at_right: fixes a b :: "'a::{linorder_topology,first_countable_topology}" assumes b[simp]: "a < b" assumes *: "\f. (\n. a < f n) \ (\n. f n < b) \ decseq f \ f \ a \ eventually (\n. P (f n)) sequentially" shows "eventually P (at_right a)" proof (safe intro!: sequentially_imp_eventually_within) fix X assume X: "\n. X n \ {a <..} \ X n \ a" "X \ a" show "eventually (\n. P (X n)) sequentially" proof (rule ccontr) assume neg: "\ ?thesis" have "\s. \n. (\ P (X (s n)) \ X (s n) < b) \ (X (s (Suc n)) \ X (s n) \ Suc (s n) \ s (Suc n))" (is "\s. ?P s") proof (rule dependent_nat_choice) have "\ eventually (\n. X n < b \ P (X n)) sequentially" by (intro not_eventually_impI neg order_tendstoD(2) [OF X(2) b]) then show "\x. \ P (X x) \ X x < b" by (auto dest!: not_eventuallyD) next fix x n have "\ eventually (\n. Suc x \ n \ X n < b \ X n < X x \ P (X n)) sequentially" using X by (intro not_eventually_impI order_tendstoD(2)[OF X(2)] eventually_ge_at_top neg) auto then show "\n. (\ P (X n) \ X n < b) \ (X n \ X x \ Suc x \ n)" by (auto dest!: not_eventuallyD) qed then obtain s where "?P s" .. with X have "a < X (s n)" and "X (s n) < b" and "decseq (\n. X (s n))" and "(\n. X (s n)) \ a" and "\ P (X (s n))" for n by (auto simp: strict_mono_Suc_iff Suc_le_eq decseq_Suc_iff intro!: LIMSEQ_subseq_LIMSEQ[OF \X \ a\, unfolded comp_def]) from *[OF this(1,2,3,4)] this(5) show False by auto qed qed lemma tendsto_at_right_sequentially: fixes a :: "_ :: {linorder_topology, first_countable_topology}" assumes "a < b" and *: "\S. (\n. a < S n) \ (\n. S n < b) \ decseq S \ S \ a \ (\n. X (S n)) \ L" shows "(X \ L) (at_right a)" using assms by (simp add: tendsto_def [where l=L] sequentially_imp_eventually_at_right) subsection \Continuity\ subsubsection \Continuity on a set\ definition continuous_on :: "'a set \ ('a::topological_space \ 'b::topological_space) \ bool" where "continuous_on s f \ (\x\s. (f \ f x) (at x within s))" lemma continuous_on_cong [cong]: "s = t \ (\x. x \ t \ f x = g x) \ continuous_on s f \ continuous_on t g" unfolding continuous_on_def by (intro ball_cong filterlim_cong) (auto simp: eventually_at_filter) lemma continuous_on_cong_simp: "s = t \ (\x. x \ t =simp=> f x = g x) \ continuous_on s f \ continuous_on t g" unfolding simp_implies_def by (rule continuous_on_cong) lemma continuous_on_topological: "continuous_on s f \ (\x\s. \B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" unfolding continuous_on_def tendsto_def eventually_at_topological by metis lemma continuous_on_open_invariant: "continuous_on s f \ (\B. open B \ (\A. open A \ A \ s = f -` B \ s))" proof safe fix B :: "'b set" assume "continuous_on s f" "open B" then have "\x\f -` B \ s. (\A. open A \ x \ A \ s \ A \ f -` B)" by (auto simp: continuous_on_topological subset_eq Ball_def imp_conjL) then obtain A where "\x\f -` B \ s. open (A x) \ x \ A x \ s \ A x \ f -` B" unfolding bchoice_iff .. then show "\A. open A \ A \ s = f -` B \ s" by (intro exI[of _ "\x\f -` B \ s. A x"]) auto next assume B: "\B. open B \ (\A. open A \ A \ s = f -` B \ s)" show "continuous_on s f" unfolding continuous_on_topological proof safe fix x B assume "x \ s" "open B" "f x \ B" with B obtain A where A: "open A" "A \ s = f -` B \ s" by auto with \x \ s\ \f x \ B\ show "\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)" by (intro exI[of _ A]) auto qed qed lemma continuous_on_open_vimage: "open s \ continuous_on s f \ (\B. open B \ open (f -` B \ s))" unfolding continuous_on_open_invariant by (metis open_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) corollary continuous_imp_open_vimage: assumes "continuous_on s f" "open s" "open B" "f -` B \ s" shows "open (f -` B)" by (metis assms continuous_on_open_vimage le_iff_inf) corollary open_vimage[continuous_intros]: assumes "open s" and "continuous_on UNIV f" shows "open (f -` s)" using assms by (simp add: continuous_on_open_vimage [OF open_UNIV]) lemma continuous_on_closed_invariant: "continuous_on s f \ (\B. closed B \ (\A. closed A \ A \ s = f -` B \ s))" proof - have *: "(\A. P A \ Q (- A)) \ (\A. P A) \ (\A. Q A)" for P Q :: "'b set \ bool" by (metis double_compl) show ?thesis unfolding continuous_on_open_invariant by (intro *) (auto simp: open_closed[symmetric]) qed lemma continuous_on_closed_vimage: "closed s \ continuous_on s f \ (\B. closed B \ closed (f -` B \ s))" unfolding continuous_on_closed_invariant by (metis closed_Int Int_absorb Int_commute[of s] Int_assoc[of _ _ s]) corollary closed_vimage_Int[continuous_intros]: assumes "closed s" and "continuous_on t f" and t: "closed t" shows "closed (f -` s \ t)" using assms by (simp add: continuous_on_closed_vimage [OF t]) corollary closed_vimage[continuous_intros]: assumes "closed s" and "continuous_on UNIV f" shows "closed (f -` s)" using closed_vimage_Int [OF assms] by simp lemma continuous_on_empty [simp]: "continuous_on {} f" by (simp add: continuous_on_def) lemma continuous_on_sing [simp]: "continuous_on {x} f" by (simp add: continuous_on_def at_within_def) lemma continuous_on_open_Union: "(\s. s \ S \ open s) \ (\s. s \ S \ continuous_on s f) \ continuous_on (\S) f" unfolding continuous_on_def by safe (metis open_Union at_within_open UnionI) lemma continuous_on_open_UN: "(\s. s \ S \ open (A s)) \ (\s. s \ S \ continuous_on (A s) f) \ continuous_on (\s\S. A s) f" by (rule continuous_on_open_Union) auto lemma continuous_on_open_Un: "open s \ open t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" using continuous_on_open_Union [of "{s,t}"] by auto lemma continuous_on_closed_Un: "closed s \ closed t \ continuous_on s f \ continuous_on t f \ continuous_on (s \ t) f" by (auto simp add: continuous_on_closed_vimage closed_Un Int_Un_distrib) lemma continuous_on_closed_Union: assumes "finite I" "\i. i \ I \ closed (U i)" "\i. i \ I \ continuous_on (U i) f" shows "continuous_on (\ i \ I. U i) f" using assms by (induction I) (auto intro!: continuous_on_closed_Un) lemma continuous_on_If: assumes closed: "closed s" "closed t" and cont: "continuous_on s f" "continuous_on t g" and P: "\x. x \ s \ \ P x \ f x = g x" "\x. x \ t \ P x \ f x = g x" shows "continuous_on (s \ t) (\x. if P x then f x else g x)" (is "continuous_on _ ?h") proof- from P have "\x\s. f x = ?h x" "\x\t. g x = ?h x" by auto with cont have "continuous_on s ?h" "continuous_on t ?h" by simp_all with closed show ?thesis by (rule continuous_on_closed_Un) qed lemma continuous_on_cases: "closed s \ closed t \ continuous_on s f \ continuous_on t g \ \x. (x\s \ \ P x) \ (x \ t \ P x) \ f x = g x \ continuous_on (s \ t) (\x. if P x then f x else g x)" by (rule continuous_on_If) auto lemma continuous_on_id[continuous_intros,simp]: "continuous_on s (\x. x)" unfolding continuous_on_def by fast lemma continuous_on_id'[continuous_intros,simp]: "continuous_on s id" unfolding continuous_on_def id_def by fast lemma continuous_on_const[continuous_intros,simp]: "continuous_on s (\x. c)" unfolding continuous_on_def by auto lemma continuous_on_subset: "continuous_on s f \ t \ s \ continuous_on t f" unfolding continuous_on_def by (metis subset_eq tendsto_within_subset) lemma continuous_on_compose[continuous_intros]: "continuous_on s f \ continuous_on (f ` s) g \ continuous_on s (g \ f)" unfolding continuous_on_topological by simp metis lemma continuous_on_compose2: "continuous_on t g \ continuous_on s f \ f ` s \ t \ continuous_on s (\x. g (f x))" using continuous_on_compose[of s f g] continuous_on_subset by (force simp add: comp_def) lemma continuous_on_generate_topology: assumes *: "open = generate_topology X" and **: "\B. B \ X \ \C. open C \ C \ A = f -` B \ A" shows "continuous_on A f" unfolding continuous_on_open_invariant proof safe fix B :: "'a set" assume "open B" then show "\C. open C \ C \ A = f -` B \ A" unfolding * proof induct case (UN K) then obtain C where "\k. k \ K \ open (C k)" "\k. k \ K \ C k \ A = f -` k \ A" by metis then show ?case by (intro exI[of _ "\k\K. C k"]) blast qed (auto intro: **) qed lemma continuous_onI_mono: fixes f :: "'a::linorder_topology \ 'b::{dense_order,linorder_topology}" assumes "open (f`A)" and mono: "\x y. x \ A \ y \ A \ x \ y \ f x \ f y" shows "continuous_on A f" proof (rule continuous_on_generate_topology[OF open_generated_order], safe) have monoD: "\x y. x \ A \ y \ A \ f x < f y \ x < y" by (auto simp: not_le[symmetric] mono) have "\x. x \ A \ f x < b \ a < x" if a: "a \ A" and fa: "f a < b" for a b proof - obtain y where "f a < y" "{f a ..< y} \ f`A" using open_right[OF \open (f`A)\, of "f a" b] a fa by auto obtain z where z: "f a < z" "z < min b y" using dense[of "f a" "min b y"] \f a < y\ \f a < b\ by auto then obtain c where "z = f c" "c \ A" using \{f a ..< y} \ f`A\[THEN subsetD, of z] by (auto simp: less_imp_le) with a z show ?thesis by (auto intro!: exI[of _ c] simp: monoD) qed then show "\C. open C \ C \ A = f -` {.. A" for b by (intro exI[of _ "(\x\{x\A. f x < b}. {..< x})"]) (auto intro: le_less_trans[OF mono] less_imp_le) have "\x. x \ A \ b < f x \ x < a" if a: "a \ A" and fa: "b < f a" for a b proof - note a fa moreover obtain y where "y < f a" "{y <.. f a} \ f`A" using open_left[OF \open (f`A)\, of "f a" b] a fa by auto then obtain z where z: "max b y < z" "z < f a" using dense[of "max b y" "f a"] \y < f a\ \b < f a\ by auto then obtain c where "z = f c" "c \ A" using \{y <.. f a} \ f`A\[THEN subsetD, of z] by (auto simp: less_imp_le) with a z show ?thesis by (auto intro!: exI[of _ c] simp: monoD) qed then show "\C. open C \ C \ A = f -` {b <..} \ A" for b by (intro exI[of _ "(\x\{x\A. b < f x}. {x <..})"]) (auto intro: less_le_trans[OF _ mono] less_imp_le) qed lemma continuous_on_IccI: "\(f \ f a) (at_right a); (f \ f b) (at_left b); (\x. a < x \ x < b \ f \x\ f x); a < b\ \ continuous_on {a .. b} f" for a::"'a::linorder_topology" using at_within_open[of _ "{a<.. f a) (at_right a)" and continuous_on_Icc_at_leftD: "(f \ f b) (at_left b)" using assms by (auto simp: at_within_Icc_at_right at_within_Icc_at_left continuous_on_def dest: bspec[where x=a] bspec[where x=b]) lemma continuous_on_discrete [simp]: "continuous_on A (f :: 'a :: discrete_topology \ _)" by (auto simp: continuous_on_def at_discrete) lemma continuous_on_of_nat [continuous_intros]: assumes "continuous_on A f" shows "continuous_on A (\n. of_nat (f n))" using continuous_on_compose[OF assms continuous_on_discrete[of _ of_nat]] by (simp add: o_def) lemma continuous_on_of_int [continuous_intros]: assumes "continuous_on A f" shows "continuous_on A (\n. of_int (f n))" using continuous_on_compose[OF assms continuous_on_discrete[of _ of_int]] by (simp add: o_def) subsubsection \Continuity at a point\ definition continuous :: "'a::t2_space filter \ ('a \ 'b::topological_space) \ bool" where "continuous F f \ (f \ f (Lim F (\x. x))) F" lemma continuous_bot[continuous_intros, simp]: "continuous bot f" unfolding continuous_def by auto lemma continuous_trivial_limit: "trivial_limit net \ continuous net f" by simp lemma continuous_within: "continuous (at x within s) f \ (f \ f x) (at x within s)" by (cases "trivial_limit (at x within s)") (auto simp add: Lim_ident_at continuous_def) lemma continuous_within_topological: "continuous (at x within s) f \ (\B. open B \ f x \ B \ (\A. open A \ x \ A \ (\y\s. y \ A \ f y \ B)))" unfolding continuous_within tendsto_def eventually_at_topological by metis lemma continuous_within_compose[continuous_intros]: "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ continuous (at x within s) (g \ f)" by (simp add: continuous_within_topological) metis lemma continuous_within_compose2: "continuous (at x within s) f \ continuous (at (f x) within f ` s) g \ continuous (at x within s) (\x. g (f x))" using continuous_within_compose[of x s f g] by (simp add: comp_def) lemma continuous_at: "continuous (at x) f \ f \x\ f x" using continuous_within[of x UNIV f] by simp lemma continuous_ident[continuous_intros, simp]: "continuous (at x within S) (\x. x)" unfolding continuous_within by (rule tendsto_ident_at) lemma continuous_id[continuous_intros, simp]: "continuous (at x within S) id" by (simp add: id_def) lemma continuous_const[continuous_intros, simp]: "continuous F (\x. c)" unfolding continuous_def by (rule tendsto_const) lemma continuous_on_eq_continuous_within: "continuous_on s f \ (\x\s. continuous (at x within s) f)" unfolding continuous_on_def continuous_within .. lemma continuous_discrete [simp]: "continuous (at x within A) (f :: 'a :: discrete_topology \ _)" by (auto simp: continuous_def at_discrete) abbreviation isCont :: "('a::t2_space \ 'b::topological_space) \ 'a \ bool" where "isCont f a \ continuous (at a) f" lemma isCont_def: "isCont f a \ f \a\ f a" by (rule continuous_at) lemma isContD: "isCont f x \ f \x\ f x" by (simp add: isCont_def) lemma isCont_cong: assumes "eventually (\x. f x = g x) (nhds x)" shows "isCont f x \ isCont g x" proof - from assms have [simp]: "f x = g x" by (rule eventually_nhds_x_imp_x) from assms have "eventually (\x. f x = g x) (at x)" by (auto simp: eventually_at_filter elim!: eventually_mono) with assms have "isCont f x \ isCont g x" unfolding isCont_def by (intro filterlim_cong) (auto elim!: eventually_mono) with assms show ?thesis by simp qed lemma continuous_at_imp_continuous_at_within: "isCont f x \ continuous (at x within s) f" by (auto intro: tendsto_mono at_le simp: continuous_at continuous_within) lemma continuous_on_eq_continuous_at: "open s \ continuous_on s f \ (\x\s. isCont f x)" by (simp add: continuous_on_def continuous_at at_within_open[of _ s]) lemma continuous_within_open: "a \ A \ open A \ continuous (at a within A) f \ isCont f a" by (simp add: at_within_open_NO_MATCH) lemma continuous_at_imp_continuous_on: "\x\s. isCont f x \ continuous_on s f" by (auto intro: continuous_at_imp_continuous_at_within simp: continuous_on_eq_continuous_within) lemma isCont_o2: "isCont f a \ isCont g (f a) \ isCont (\x. g (f x)) a" unfolding isCont_def by (rule tendsto_compose) lemma continuous_at_compose[continuous_intros]: "isCont f a \ isCont g (f a) \ isCont (g \ f) a" unfolding o_def by (rule isCont_o2) lemma isCont_tendsto_compose: "isCont g l \ (f \ l) F \ ((\x. g (f x)) \ g l) F" unfolding isCont_def by (rule tendsto_compose) lemma continuous_on_tendsto_compose: assumes f_cont: "continuous_on s f" and g: "(g \ l) F" and l: "l \ s" and ev: "\\<^sub>Fx in F. g x \ s" shows "((\x. f (g x)) \ f l) F" proof - from f_cont l have f: "(f \ f l) (at l within s)" by (simp add: continuous_on_def) have i: "((\x. if g x = l then f l else f (g x)) \ f l) F" by (rule filterlim_If) (auto intro!: filterlim_compose[OF f] eventually_conj tendsto_mono[OF _ g] simp: filterlim_at eventually_inf_principal eventually_mono[OF ev]) show ?thesis by (rule filterlim_cong[THEN iffD1[OF _ i]]) auto qed lemma continuous_within_compose3: "isCont g (f x) \ continuous (at x within s) f \ continuous (at x within s) (\x. g (f x))" using continuous_at_imp_continuous_at_within continuous_within_compose2 by blast lemma at_within_isCont_imp_nhds: fixes f:: "'a:: {t2_space,perfect_space} \ 'b:: t2_space" assumes "\\<^sub>F w in at z. f w = g w" "isCont f z" "isCont g z" shows "\\<^sub>F w in nhds z. f w = g w" proof - have "g \z\ f z" using assms isContD tendsto_cong by blast moreover have "g \z\ g z" using \isCont g z\ using isCont_def by blast ultimately have "f z=g z" using LIM_unique by auto moreover have "\\<^sub>F x in nhds z. x \ z \ f x = g x" using assms unfolding eventually_at_filter by auto ultimately show ?thesis by (auto elim:eventually_mono) qed lemma filtermap_nhds_open_map': assumes cont: "isCont f a" and "open A" "a \ A" and open_map: "\S. open S \ S \ A \ open (f ` S)" shows "filtermap f (nhds a) = nhds (f a)" unfolding filter_eq_iff proof safe fix P assume "eventually P (filtermap f (nhds a))" then obtain S where S: "open S" "a \ S" "\x\S. P (f x)" by (auto simp: eventually_filtermap eventually_nhds) show "eventually P (nhds (f a))" unfolding eventually_nhds proof (rule exI [of _ "f ` (A \ S)"], safe) show "open (f ` (A \ S))" using S by (intro open_Int assms) auto show "f a \ f ` (A \ S)" using assms S by auto show "P (f x)" if "x \ A" "x \ S" for x using S that by auto qed qed (metis filterlim_iff tendsto_at_iff_tendsto_nhds isCont_def eventually_filtermap cont) lemma filtermap_nhds_open_map: assumes cont: "isCont f a" and open_map: "\S. open S \ open (f`S)" shows "filtermap f (nhds a) = nhds (f a)" using cont filtermap_nhds_open_map' open_map by blast lemma continuous_at_split: "continuous (at x) f \ continuous (at_left x) f \ continuous (at_right x) f" for x :: "'a::linorder_topology" by (simp add: continuous_within filterlim_at_split) lemma continuous_on_max [continuous_intros]: fixes f g :: "'a::topological_space \ 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. max (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_max) lemma continuous_on_min [continuous_intros]: fixes f g :: "'a::topological_space \ 'b::linorder_topology" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. min (f x) (g x))" by (auto simp: continuous_on_def intro!: tendsto_min) lemma continuous_max [continuous_intros]: fixes f :: "'a::t2_space \ 'b::linorder_topology" shows "\continuous F f; continuous F g\ \ continuous F (\x. (max (f x) (g x)))" by (simp add: tendsto_max continuous_def) lemma continuous_min [continuous_intros]: fixes f :: "'a::t2_space \ 'b::linorder_topology" shows "\continuous F f; continuous F g\ \ continuous F (\x. (min (f x) (g x)))" by (simp add: tendsto_min continuous_def) text \ The following open/closed Collect lemmas are ported from Sébastien Gouëzel's \Ergodic_Theory\. \ lemma open_Collect_neq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "open {x. f x \ g x}" proof (rule openI) fix t assume "t \ {x. f x \ g x}" then obtain U V where *: "open U" "open V" "f t \ U" "g t \ V" "U \ V = {}" by (auto simp add: separation_t2) with open_vimage[OF \open U\ f] open_vimage[OF \open V\ g] show "\T. open T \ t \ T \ T \ {x. f x \ g x}" by (intro exI[of _ "f -` U \ g -` V"]) auto qed lemma closed_Collect_eq: fixes f g :: "'a::topological_space \ 'b::t2_space" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "closed {x. f x = g x}" using open_Collect_neq[OF f g] by (simp add: closed_def Collect_neg_eq) lemma open_Collect_less: fixes f g :: "'a::topological_space \ 'b::linorder_topology" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "open {x. f x < g x}" proof (rule openI) fix t assume t: "t \ {x. f x < g x}" show "\T. open T \ t \ T \ T \ {x. f x < g x}" proof (cases "\z. f t < z \ z < g t") case True then obtain z where "f t < z \ z < g t" by blast then show ?thesis using open_vimage[OF _ f, of "{..< z}"] open_vimage[OF _ g, of "{z <..}"] by (intro exI[of _ "f -` {.. g -` {z<..}"]) auto next case False then have *: "{g t ..} = {f t <..}" "{..< g t} = {.. f t}" using t by (auto intro: leI) show ?thesis using open_vimage[OF _ f, of "{..< g t}"] open_vimage[OF _ g, of "{f t <..}"] t apply (intro exI[of _ "f -` {..< g t} \ g -` {f t<..}"]) apply (simp add: open_Int) apply (auto simp add: *) done qed qed lemma closed_Collect_le: fixes f g :: "'a :: topological_space \ 'b::linorder_topology" assumes f: "continuous_on UNIV f" and g: "continuous_on UNIV g" shows "closed {x. f x \ g x}" using open_Collect_less [OF g f] by (simp add: closed_def Collect_neg_eq[symmetric] not_le) subsubsection \Open-cover compactness\ context topological_space begin definition compact :: "'a set \ bool" where compact_eq_Heine_Borel: (* This name is used for backwards compatibility *) "compact S \ (\C. (\c\C. open c) \ S \ \C \ (\D\C. finite D \ S \ \D))" lemma compactI: assumes "\C. \t\C. open t \ s \ \C \ \C'. C' \ C \ finite C' \ s \ \C'" shows "compact s" unfolding compact_eq_Heine_Borel using assms by metis lemma compact_empty[simp]: "compact {}" by (auto intro!: compactI) lemma compactE: (*related to COMPACT_IMP_HEINE_BOREL in HOL Light*) assumes "compact S" "S \ \\" "\B. B \ \ \ open B" obtains \' where "\' \ \" "finite \'" "S \ \\'" by (meson assms compact_eq_Heine_Borel) lemma compactE_image: assumes "compact S" and opn: "\T. T \ C \ open (f T)" and S: "S \ (\c\C. f c)" obtains C' where "C' \ C" and "finite C'" and "S \ (\c\C'. f c)" apply (rule compactE[OF \compact S\ S]) using opn apply force by (metis finite_subset_image) lemma compact_Int_closed [intro]: assumes "compact S" and "closed T" shows "compact (S \ T)" proof (rule compactI) fix C assume C: "\c\C. open c" assume cover: "S \ T \ \C" from C \closed T\ have "\c\C \ {- T}. open c" by auto moreover from cover have "S \ \(C \ {- T})" by auto ultimately have "\D\C \ {- T}. finite D \ S \ \D" using \compact S\ unfolding compact_eq_Heine_Borel by auto then obtain D where "D \ C \ {- T} \ finite D \ S \ \D" .. then show "\D\C. finite D \ S \ T \ \D" by (intro exI[of _ "D - {-T}"]) auto qed lemma compact_diff: "\compact S; open T\ \ compact(S - T)" by (simp add: Diff_eq compact_Int_closed open_closed) lemma inj_setminus: "inj_on uminus (A::'a set set)" by (auto simp: inj_on_def) subsection \Finite intersection property\ lemma compact_fip: "compact U \ (\A. (\a\A. closed a) \ (\B \ A. finite B \ U \ \B \ {}) \ U \ \A \ {})" (is "_ \ ?R") proof (safe intro!: compact_eq_Heine_Borel[THEN iffD2]) fix A assume "compact U" assume A: "\a\A. closed a" "U \ \A = {}" assume fin: "\B \ A. finite B \ U \ \B \ {}" from A have "(\a\uminus`A. open a) \ U \ \(uminus`A)" by auto with \compact U\ obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B)" unfolding compact_eq_Heine_Borel by (metis subset_image_iff) with fin[THEN spec, of B] show False by (auto dest: finite_imageD intro: inj_setminus) next fix A assume ?R assume "\a\A. open a" "U \ \A" then have "U \ \(uminus`A) = {}" "\a\uminus`A. closed a" by auto with \?R\ obtain B where "B \ A" "finite (uminus`B)" "U \ \(uminus`B) = {}" by (metis subset_image_iff) then show "\T\A. finite T \ U \ \T" by (auto intro!: exI[of _ B] inj_setminus dest: finite_imageD) qed lemma compact_imp_fip: assumes "compact S" and "\T. T \ F \ closed T" and "\F'. finite F' \ F' \ F \ S \ (\F') \ {}" shows "S \ (\F) \ {}" using assms unfolding compact_fip by auto lemma compact_imp_fip_image: assumes "compact s" and P: "\i. i \ I \ closed (f i)" and Q: "\I'. finite I' \ I' \ I \ (s \ (\i\I'. f i) \ {})" shows "s \ (\i\I. f i) \ {}" proof - from P have "\i \ f ` I. closed i" by blast moreover have "\A. finite A \ A \ f ` I \ (s \ (\A) \ {})" by (metis Q finite_subset_image) ultimately show "s \ (\(f ` I)) \ {}" by (metis \compact s\ compact_imp_fip) qed end lemma (in t2_space) compact_imp_closed: assumes "compact s" shows "closed s" unfolding closed_def proof (rule openI) fix y assume "y \ - s" let ?C = "\x\s. {u. open u \ x \ u \ eventually (\y. y \ u) (nhds y)}" have "s \ \?C" proof fix x assume "x \ s" with \y \ - s\ have "x \ y" by clarsimp then have "\u v. open u \ open v \ x \ u \ y \ v \ u \ v = {}" by (rule hausdorff) with \x \ s\ show "x \ \?C" unfolding eventually_nhds by auto qed then obtain D where "D \ ?C" and "finite D" and "s \ \D" by (rule compactE [OF \compact s\]) auto from \D \ ?C\ have "\x\D. eventually (\y. y \ x) (nhds y)" by auto with \finite D\ have "eventually (\y. y \ \D) (nhds y)" by (simp add: eventually_ball_finite) with \s \ \D\ have "eventually (\y. y \ s) (nhds y)" by (auto elim!: eventually_mono) then show "\t. open t \ y \ t \ t \ - s" by (simp add: eventually_nhds subset_eq) qed lemma compact_continuous_image: assumes f: "continuous_on s f" and s: "compact s" shows "compact (f ` s)" proof (rule compactI) fix C assume "\c\C. open c" and cover: "f`s \ \C" with f have "\c\C. \A. open A \ A \ s = f -` c \ s" unfolding continuous_on_open_invariant by blast then obtain A where A: "\c\C. open (A c) \ A c \ s = f -` c \ s" unfolding bchoice_iff .. with cover have "\c. c \ C \ open (A c)" "s \ (\c\C. A c)" by (fastforce simp add: subset_eq set_eq_iff)+ from compactE_image[OF s this] obtain D where "D \ C" "finite D" "s \ (\c\D. A c)" . with A show "\D \ C. finite D \ f`s \ \D" by (intro exI[of _ D]) (fastforce simp add: subset_eq set_eq_iff)+ qed lemma continuous_on_inv: fixes f :: "'a::topological_space \ 'b::t2_space" assumes "continuous_on s f" and "compact s" and "\x\s. g (f x) = x" shows "continuous_on (f ` s) g" unfolding continuous_on_topological proof (clarsimp simp add: assms(3)) fix x :: 'a and B :: "'a set" assume "x \ s" and "open B" and "x \ B" have 1: "\x\s. f x \ f ` (s - B) \ x \ s - B" using assms(3) by (auto, metis) have "continuous_on (s - B) f" using \continuous_on s f\ Diff_subset by (rule continuous_on_subset) moreover have "compact (s - B)" using \open B\ and \compact s\ unfolding Diff_eq by (intro compact_Int_closed closed_Compl) ultimately have "compact (f ` (s - B))" by (rule compact_continuous_image) then have "closed (f ` (s - B))" by (rule compact_imp_closed) then have "open (- f ` (s - B))" by (rule open_Compl) moreover have "f x \ - f ` (s - B)" using \x \ s\ and \x \ B\ by (simp add: 1) moreover have "\y\s. f y \ - f ` (s - B) \ y \ B" by (simp add: 1) ultimately show "\A. open A \ f x \ A \ (\y\s. f y \ A \ y \ B)" by fast qed lemma continuous_on_inv_into: fixes f :: "'a::topological_space \ 'b::t2_space" assumes s: "continuous_on s f" "compact s" and f: "inj_on f s" shows "continuous_on (f ` s) (the_inv_into s f)" by (rule continuous_on_inv[OF s]) (auto simp: the_inv_into_f_f[OF f]) lemma (in linorder_topology) compact_attains_sup: assumes "compact S" "S \ {}" shows "\s\S. \t\S. t \ s" proof (rule classical) assume "\ (\s\S. \t\S. t \ s)" then obtain t where t: "\s\S. t s \ S" and "\s\S. s < t s" by (metis not_le) then have "\s. s\S \ open {..< t s}" "S \ (\s\S. {..< t s})" by auto with \compact S\ obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {..< t s})" by (metis compactE_image) with \S \ {}\ have Max: "Max (t`C) \ t`C" and "\s\t`C. s \ Max (t`C)" by (auto intro!: Max_in) with C have "S \ {..< Max (t`C)}" by (auto intro: less_le_trans simp: subset_eq) with t Max \C \ S\ show ?thesis by fastforce qed lemma (in linorder_topology) compact_attains_inf: assumes "compact S" "S \ {}" shows "\s\S. \t\S. s \ t" proof (rule classical) assume "\ (\s\S. \t\S. s \ t)" then obtain t where t: "\s\S. t s \ S" and "\s\S. t s < s" by (metis not_le) then have "\s. s\S \ open {t s <..}" "S \ (\s\S. {t s <..})" by auto with \compact S\ obtain C where "C \ S" "finite C" and C: "S \ (\s\C. {t s <..})" by (metis compactE_image) with \S \ {}\ have Min: "Min (t`C) \ t`C" and "\s\t`C. Min (t`C) \ s" by (auto intro!: Min_in) with C have "S \ {Min (t`C) <..}" by (auto intro: le_less_trans simp: subset_eq) with t Min \C \ S\ show ?thesis by fastforce qed lemma continuous_attains_sup: fixes f :: "'a::topological_space \ 'b::linorder_topology" shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f y \ f x)" using compact_attains_sup[of "f ` s"] compact_continuous_image[of s f] by auto lemma continuous_attains_inf: fixes f :: "'a::topological_space \ 'b::linorder_topology" shows "compact s \ s \ {} \ continuous_on s f \ (\x\s. \y\s. f x \ f y)" using compact_attains_inf[of "f ` s"] compact_continuous_image[of s f] by auto subsection \Connectedness\ context topological_space begin definition "connected S \ \ (\A B. open A \ open B \ S \ A \ B \ A \ B \ S = {} \ A \ S \ {} \ B \ S \ {})" lemma connectedI: "(\A B. open A \ open B \ A \ U \ {} \ B \ U \ {} \ A \ B \ U = {} \ U \ A \ B \ False) \ connected U" by (auto simp: connected_def) lemma connected_empty [simp]: "connected {}" by (auto intro!: connectedI) lemma connected_sing [simp]: "connected {x}" by (auto intro!: connectedI) lemma connectedD: "connected A \ open U \ open V \ U \ V \ A = {} \ A \ U \ V \ U \ A = {} \ V \ A = {}" by (auto simp: connected_def) end lemma connected_closed: "connected s \ \ (\A B. closed A \ closed B \ s \ A \ B \ A \ B \ s = {} \ A \ s \ {} \ B \ s \ {})" apply (simp add: connected_def del: ex_simps, safe) apply (drule_tac x="-A" in spec) apply (drule_tac x="-B" in spec) apply (fastforce simp add: closed_def [symmetric]) apply (drule_tac x="-A" in spec) apply (drule_tac x="-B" in spec) apply (fastforce simp add: open_closed [symmetric]) done lemma connected_closedD: "\connected s; A \ B \ s = {}; s \ A \ B; closed A; closed B\ \ A \ s = {} \ B \ s = {}" by (simp add: connected_closed) lemma connected_Union: assumes cs: "\s. s \ S \ connected s" and ne: "\S \ {}" shows "connected(\S)" proof (rule connectedI) fix A B assume A: "open A" and B: "open B" and Alap: "A \ \S \ {}" and Blap: "B \ \S \ {}" and disj: "A \ B \ \S = {}" and cover: "\S \ A \ B" have disjs:"\s. s \ S \ A \ B \ s = {}" using disj by auto obtain sa where sa: "sa \ S" "A \ sa \ {}" using Alap by auto obtain sb where sb: "sb \ S" "B \ sb \ {}" using Blap by auto obtain x where x: "\s. s \ S \ x \ s" using ne by auto then have "x \ \S" using \sa \ S\ by blast then have "x \ A \ x \ B" using cover by auto then show False using cs [unfolded connected_def] by (metis A B IntI Sup_upper sa sb disjs x cover empty_iff subset_trans) qed lemma connected_Un: "connected s \ connected t \ s \ t \ {} \ connected (s \ t)" using connected_Union [of "{s,t}"] by auto lemma connected_diff_open_from_closed: assumes st: "s \ t" and tu: "t \ u" and s: "open s" and t: "closed t" and u: "connected u" and ts: "connected (t - s)" shows "connected(u - s)" proof (rule connectedI) fix A B assume AB: "open A" "open B" "A \ (u - s) \ {}" "B \ (u - s) \ {}" and disj: "A \ B \ (u - s) = {}" and cover: "u - s \ A \ B" then consider "A \ (t - s) = {}" | "B \ (t - s) = {}" using st ts tu connectedD [of "t-s" "A" "B"] by auto then show False proof cases case 1 then have "(A - t) \ (B \ s) \ u = {}" using disj st by auto moreover have "u \ (A - t) \ (B \ s)" using 1 cover by auto ultimately show False using connectedD [of u "A - t" "B \ s"] AB s t 1 u by auto next case 2 then have "(A \ s) \ (B - t) \ u = {}" using disj st by auto moreover have "u \ (A \ s) \ (B - t)" using 2 cover by auto ultimately show False using connectedD [of u "A \ s" "B - t"] AB s t 2 u by auto qed qed lemma connected_iff_const: fixes S :: "'a::topological_space set" shows "connected S \ (\P::'a \ bool. continuous_on S P \ (\c. \s\S. P s = c))" proof safe fix P :: "'a \ bool" assume "connected S" "continuous_on S P" then have "\b. \A. open A \ A \ S = P -` {b} \ S" unfolding continuous_on_open_invariant by (simp add: open_discrete) from this[of True] this[of False] obtain t f where "open t" "open f" and *: "f \ S = P -` {False} \ S" "t \ S = P -` {True} \ S" by meson then have "t \ S = {} \ f \ S = {}" by (intro connectedD[OF \connected S\]) auto then show "\c. \s\S. P s = c" proof (rule disjE) assume "t \ S = {}" then show ?thesis unfolding * by (intro exI[of _ False]) auto next assume "f \ S = {}" then show ?thesis unfolding * by (intro exI[of _ True]) auto qed next assume P: "\P::'a \ bool. continuous_on S P \ (\c. \s\S. P s = c)" show "connected S" proof (rule connectedI) fix A B assume *: "open A" "open B" "A \ S \ {}" "B \ S \ {}" "A \ B \ S = {}" "S \ A \ B" have "continuous_on S (\x. x \ A)" unfolding continuous_on_open_invariant proof safe fix C :: "bool set" have "C = UNIV \ C = {True} \ C = {False} \ C = {}" using subset_UNIV[of C] unfolding UNIV_bool by auto with * show "\T. open T \ T \ S = (\x. x \ A) -` C \ S" by (intro exI[of _ "(if True \ C then A else {}) \ (if False \ C then B else {})"]) auto qed from P[rule_format, OF this] obtain c where "\s. s \ S \ (s \ A) = c" by blast with * show False by (cases c) auto qed qed lemma connectedD_const: "connected S \ continuous_on S P \ \c. \s\S. P s = c" for P :: "'a::topological_space \ bool" by (auto simp: connected_iff_const) lemma connectedI_const: "(\P::'a::topological_space \ bool. continuous_on S P \ \c. \s\S. P s = c) \ connected S" by (auto simp: connected_iff_const) lemma connected_local_const: assumes "connected A" "a \ A" "b \ A" and *: "\a\A. eventually (\b. f a = f b) (at a within A)" shows "f a = f b" proof - obtain S where S: "\a. a \ A \ a \ S a" "\a. a \ A \ open (S a)" "\a x. a \ A \ x \ S a \ x \ A \ f a = f x" using * unfolding eventually_at_topological by metis let ?P = "\b\{b\A. f a = f b}. S b" and ?N = "\b\{b\A. f a \ f b}. S b" have "?P \ A = {} \ ?N \ A = {}" using \connected A\ S \a\A\ by (intro connectedD) (auto, metis) then show "f a = f b" proof assume "?N \ A = {}" then have "\x\A. f a = f x" using S(1) by auto with \b\A\ show ?thesis by auto next assume "?P \ A = {}" then show ?thesis using \a \ A\ S(1)[of a] by auto qed qed lemma (in linorder_topology) connectedD_interval: assumes "connected U" and xy: "x \ U" "y \ U" and "x \ z" "z \ y" shows "z \ U" proof - have eq: "{.. {z<..} = - {z}" by auto have "\ connected U" if "z \ U" "x < z" "z < y" using xy that apply (simp only: connected_def simp_thms) apply (rule_tac exI[of _ "{..< z}"]) apply (rule_tac exI[of _ "{z <..}"]) apply (auto simp add: eq) done with assms show "z \ U" by (metis less_le) qed lemma (in linorder_topology) not_in_connected_cases: assumes conn: "connected S" assumes nbdd: "x \ S" assumes ne: "S \ {}" obtains "bdd_above S" "\y. y \ S \ x \ y" | "bdd_below S" "\y. y \ S \ x \ y" proof - obtain s where "s \ S" using ne by blast { assume "s \ x" have "False" if "x \ y" "y \ S" for y using connectedD_interval[OF conn \s \ S\ \y \ S\ \s \ x\ \x \ y\] \x \ S\ by simp then have wit: "y \ S \ x \ y" for y using le_cases by blast then have "bdd_above S" by (rule local.bdd_aboveI) note this wit } moreover { assume "x \ s" have "False" if "x \ y" "y \ S" for y using connectedD_interval[OF conn \y \ S\ \s \ S\ \x \ y\ \s \ x\ ] \x \ S\ by simp then have wit: "y \ S \ x \ y" for y using le_cases by blast then have "bdd_below S" by (rule bdd_belowI) note this wit } ultimately show ?thesis by (meson le_cases that) qed lemma connected_continuous_image: assumes *: "continuous_on s f" and "connected s" shows "connected (f ` s)" proof (rule connectedI_const) fix P :: "'b \ bool" assume "continuous_on (f ` s) P" then have "continuous_on s (P \ f)" by (rule continuous_on_compose[OF *]) from connectedD_const[OF \connected s\ this] show "\c. \s\f ` s. P s = c" by auto qed lemma connected_Un_UN: assumes "connected A" "\X. X \ B \ connected X" "\X. X \ B \ A \ X \ {}" shows "connected (A \ \B)" proof (rule connectedI_const) fix f :: "'a \ bool" assume f: "continuous_on (A \ \B) f" have "connected A" "continuous_on A f" by (auto intro: assms continuous_on_subset[OF f(1)]) from connectedD_const[OF this] obtain c where c: "\x. x \ A \ f x = c" by metis have "f x = c" if "x \ X" "X \ B" for x X proof - have "connected X" "continuous_on X f" using that by (auto intro: assms continuous_on_subset[OF f]) from connectedD_const[OF this] obtain c' where c': "\x. x \ X \ f x = c'" by metis from assms(3) and that obtain y where "y \ A \ X" by auto with c[of y] c'[of y] c'[of x] that show ?thesis by auto qed with c show "\c. \x\A \ \ B. f x = c" by (intro exI[of _ c]) auto qed section \Linear Continuum Topologies\ class linear_continuum_topology = linorder_topology + linear_continuum begin lemma Inf_notin_open: assumes A: "open A" and bnd: "\a\A. x < a" shows "Inf A \ A" proof assume "Inf A \ A" then obtain b where "b < Inf A" "{b <.. Inf A} \ A" using open_left[of A "Inf A" x] assms by auto with dense[of b "Inf A"] obtain c where "c < Inf A" "c \ A" by (auto simp: subset_eq) then show False using cInf_lower[OF \c \ A\] bnd by (metis not_le less_imp_le bdd_belowI) qed lemma Sup_notin_open: assumes A: "open A" and bnd: "\a\A. a < x" shows "Sup A \ A" proof assume "Sup A \ A" with assms obtain b where "Sup A < b" "{Sup A ..< b} \ A" using open_right[of A "Sup A" x] by auto with dense[of "Sup A" b] obtain c where "Sup A < c" "c \ A" by (auto simp: subset_eq) then show False using cSup_upper[OF \c \ A\] bnd by (metis less_imp_le not_le bdd_aboveI) qed end instance linear_continuum_topology \ perfect_space proof fix x :: 'a obtain y where "x < y \ y < x" using ex_gt_or_lt [of x] .. with Inf_notin_open[of "{x}" y] Sup_notin_open[of "{x}" y] show "\ open {x}" by auto qed lemma connectedI_interval: fixes U :: "'a :: linear_continuum_topology set" assumes *: "\x y z. x \ U \ y \ U \ x \ z \ z \ y \ z \ U" shows "connected U" proof (rule connectedI) { fix A B assume "open A" "open B" "A \ B \ U = {}" "U \ A \ B" fix x y assume "x < y" "x \ A" "y \ B" "x \ U" "y \ U" let ?z = "Inf (B \ {x <..})" have "x \ ?z" "?z \ y" using \y \ B\ \x < y\ by (auto intro: cInf_lower cInf_greatest) with \x \ U\ \y \ U\ have "?z \ U" by (rule *) moreover have "?z \ B \ {x <..}" using \open B\ by (intro Inf_notin_open) auto ultimately have "?z \ A" using \x \ ?z\ \A \ B \ U = {}\ \x \ A\ \U \ A \ B\ by auto have "\b\B. b \ A \ b \ U" if "?z < y" proof - obtain a where "?z < a" "{?z ..< a} \ A" using open_right[OF \open A\ \?z \ A\ \?z < y\] by auto moreover obtain b where "b \ B" "x < b" "b < min a y" using cInf_less_iff[of "B \ {x <..}" "min a y"] \?z < a\ \?z < y\ \x < y\ \y \ B\ by auto moreover have "?z \ b" using \b \ B\ \x < b\ by (intro cInf_lower) auto moreover have "b \ U" using \x \ ?z\ \?z \ b\ \b < min a y\ by (intro *[OF \x \ U\ \y \ U\]) (auto simp: less_imp_le) ultimately show ?thesis by (intro bexI[of _ b]) auto qed then have False using \?z \ y\ \?z \ A\ \y \ B\ \y \ U\ \A \ B \ U = {}\ unfolding le_less by blast } note not_disjoint = this fix A B assume AB: "open A" "open B" "U \ A \ B" "A \ B \ U = {}" moreover assume "A \ U \ {}" then obtain x where x: "x \ U" "x \ A" by auto moreover assume "B \ U \ {}" then obtain y where y: "y \ U" "y \ B" by auto moreover note not_disjoint[of B A y x] not_disjoint[of A B x y] ultimately show False by (cases x y rule: linorder_cases) auto qed lemma connected_iff_interval: "connected U \ (\x\U. \y\U. \z. x \ z \ z \ y \ z \ U)" for U :: "'a::linear_continuum_topology set" by (auto intro: connectedI_interval dest: connectedD_interval) lemma connected_UNIV[simp]: "connected (UNIV::'a::linear_continuum_topology set)" by (simp add: connected_iff_interval) lemma connected_Ioi[simp]: "connected {a<..}" for a :: "'a::linear_continuum_topology" by (auto simp: connected_iff_interval) lemma connected_Ici[simp]: "connected {a..}" for a :: "'a::linear_continuum_topology" by (auto simp: connected_iff_interval) lemma connected_Iio[simp]: "connected {.. A" "b \ A" shows "{a <..< b} \ A" using connectedD_interval[OF assms] by (simp add: subset_eq Ball_def less_imp_le) lemma connected_contains_Icc: fixes A :: "'a::linorder_topology set" assumes "connected A" "a \ A" "b \ A" shows "{a..b} \ A" proof fix x assume "x \ {a..b}" then have "x = a \ x = b \ x \ {a<.. A" using assms connected_contains_Ioo[of A a b] by auto qed subsection \Intermediate Value Theorem\ lemma IVT': fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes y: "f a \ y" "y \ f b" "a \ b" and *: "continuous_on {a .. b} f" shows "\x. a \ x \ x \ b \ f x = y" proof - have "connected {a..b}" unfolding connected_iff_interval by auto from connected_continuous_image[OF * this, THEN connectedD_interval, of "f a" "f b" y] y show ?thesis by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) qed lemma IVT2': fixes f :: "'a :: linear_continuum_topology \ 'b :: linorder_topology" assumes y: "f b \ y" "y \ f a" "a \ b" and *: "continuous_on {a .. b} f" shows "\x. a \ x \ x \ b \ f x = y" proof - have "connected {a..b}" unfolding connected_iff_interval by auto from connected_continuous_image[OF * this, THEN connectedD_interval, of "f b" "f a" y] y show ?thesis by (auto simp add: atLeastAtMost_def atLeast_def atMost_def) qed lemma IVT: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" shows "f a \ y \ y \ f b \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" by (rule IVT') (auto intro: continuous_at_imp_continuous_on) lemma IVT2: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" shows "f b \ y \ y \ f a \ a \ b \ (\x. a \ x \ x \ b \ isCont f x) \ \x. a \ x \ x \ b \ f x = y" by (rule IVT2') (auto intro: continuous_at_imp_continuous_on) lemma continuous_inj_imp_mono: fixes f :: "'a::linear_continuum_topology \ 'b::linorder_topology" assumes x: "a < x" "x < b" and cont: "continuous_on {a..b} f" and inj: "inj_on f {a..b}" shows "(f a < f x \ f x < f b) \ (f b < f x \ f x < f a)" proof - note I = inj_on_eq_iff[OF inj] { assume "f x < f a" "f x < f b" then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f x < f s" using IVT'[of f x "min (f a) (f b)" b] IVT2'[of f x "min (f a) (f b)" a] x by (auto simp: continuous_on_subset[OF cont] less_imp_le) with x I have False by auto } moreover { assume "f a < f x" "f b < f x" then obtain s t where "x \ s" "s \ b" "a \ t" "t \ x" "f s = f t" "f s < f x" using IVT'[of f a "max (f a) (f b)" x] IVT2'[of f b "max (f a) (f b)" x] x by (auto simp: continuous_on_subset[OF cont] less_imp_le) with x I have False by auto } ultimately show ?thesis using I[of a x] I[of x b] x less_trans[OF x] by (auto simp add: le_less less_imp_neq neq_iff) qed lemma continuous_at_Sup_mono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "mono f" and cont: "continuous (at_left (Sup S)) f" and S: "S \ {}" "bdd_above S" shows "f (Sup S) = (SUP s\S. f s)" proof (rule antisym) have f: "(f \ f (Sup S)) (at_left (Sup S))" using cont unfolding continuous_within . show "f (Sup S) \ (SUP s\S. f s)" proof cases assume "Sup S \ S" then show ?thesis by (rule cSUP_upper) (auto intro: bdd_above_image_mono S \mono f\) next assume "Sup S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Sup S \ S\ S have "s < Sup S" unfolding less_le by (blast intro: cSup_upper) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(1)[OF f, of "SUP s\S. f s"] obtain b where "b < Sup S" and *: "\y. b < y \ y < Sup S \ (SUP s\S. f s) < f y" by (auto simp: not_le eventually_at_left[OF \s < Sup S\]) with \S \ {}\ obtain c where "c \ S" "b < c" using less_cSupD[of S b] by auto with \Sup S \ S\ S have "c < Sup S" unfolding less_le by (blast intro: cSup_upper) from *[OF \b < c\ \c < Sup S\] cSUP_upper[OF \c \ S\ bdd_above_image_mono[of f]] show False by (auto simp: assms) qed qed qed (intro cSUP_least \mono f\[THEN monoD] cSup_upper S) lemma continuous_at_Sup_antimono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "antimono f" and cont: "continuous (at_left (Sup S)) f" and S: "S \ {}" "bdd_above S" shows "f (Sup S) = (INF s\S. f s)" proof (rule antisym) have f: "(f \ f (Sup S)) (at_left (Sup S))" using cont unfolding continuous_within . show "(INF s\S. f s) \ f (Sup S)" proof cases assume "Sup S \ S" then show ?thesis by (intro cINF_lower) (auto intro: bdd_below_image_antimono S \antimono f\) next assume "Sup S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Sup S \ S\ S have "s < Sup S" unfolding less_le by (blast intro: cSup_upper) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(2)[OF f, of "INF s\S. f s"] obtain b where "b < Sup S" and *: "\y. b < y \ y < Sup S \ f y < (INF s\S. f s)" by (auto simp: not_le eventually_at_left[OF \s < Sup S\]) with \S \ {}\ obtain c where "c \ S" "b < c" using less_cSupD[of S b] by auto with \Sup S \ S\ S have "c < Sup S" unfolding less_le by (blast intro: cSup_upper) from *[OF \b < c\ \c < Sup S\] cINF_lower[OF bdd_below_image_antimono, of f S c] \c \ S\ show False by (auto simp: assms) qed qed qed (intro cINF_greatest \antimono f\[THEN antimonoD] cSup_upper S) lemma continuous_at_Inf_mono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "mono f" and cont: "continuous (at_right (Inf S)) f" and S: "S \ {}" "bdd_below S" shows "f (Inf S) = (INF s\S. f s)" proof (rule antisym) have f: "(f \ f (Inf S)) (at_right (Inf S))" using cont unfolding continuous_within . show "(INF s\S. f s) \ f (Inf S)" proof cases assume "Inf S \ S" then show ?thesis by (rule cINF_lower[rotated]) (auto intro: bdd_below_image_mono S \mono f\) next assume "Inf S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Inf S \ S\ S have "Inf S < s" unfolding less_le by (blast intro: cInf_lower) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(2)[OF f, of "INF s\S. f s"] obtain b where "Inf S < b" and *: "\y. Inf S < y \ y < b \ f y < (INF s\S. f s)" by (auto simp: not_le eventually_at_right[OF \Inf S < s\]) with \S \ {}\ obtain c where "c \ S" "c < b" using cInf_lessD[of S b] by auto with \Inf S \ S\ S have "Inf S < c" unfolding less_le by (blast intro: cInf_lower) from *[OF \Inf S < c\ \c < b\] cINF_lower[OF bdd_below_image_mono[of f] \c \ S\] show False by (auto simp: assms) qed qed qed (intro cINF_greatest \mono f\[THEN monoD] cInf_lower \bdd_below S\ \S \ {}\) lemma continuous_at_Inf_antimono: fixes f :: "'a::{linorder_topology,conditionally_complete_linorder} \ 'b::{linorder_topology,conditionally_complete_linorder}" assumes "antimono f" and cont: "continuous (at_right (Inf S)) f" and S: "S \ {}" "bdd_below S" shows "f (Inf S) = (SUP s\S. f s)" proof (rule antisym) have f: "(f \ f (Inf S)) (at_right (Inf S))" using cont unfolding continuous_within . show "f (Inf S) \ (SUP s\S. f s)" proof cases assume "Inf S \ S" then show ?thesis by (rule cSUP_upper) (auto intro: bdd_above_image_antimono S \antimono f\) next assume "Inf S \ S" from \S \ {}\ obtain s where "s \ S" by auto with \Inf S \ S\ S have "Inf S < s" unfolding less_le by (blast intro: cInf_lower) show ?thesis proof (rule ccontr) assume "\ ?thesis" with order_tendstoD(1)[OF f, of "SUP s\S. f s"] obtain b where "Inf S < b" and *: "\y. Inf S < y \ y < b \ (SUP s\S. f s) < f y" by (auto simp: not_le eventually_at_right[OF \Inf S < s\]) with \S \ {}\ obtain c where "c \ S" "c < b" using cInf_lessD[of S b] by auto with \Inf S \ S\ S have "Inf S < c" unfolding less_le by (blast intro: cInf_lower) from *[OF \Inf S < c\ \c < b\] cSUP_upper[OF \c \ S\ bdd_above_image_antimono[of f]] show False by (auto simp: assms) qed qed qed (intro cSUP_least \antimono f\[THEN antimonoD] cInf_lower S) subsection \Uniform spaces\ class uniformity = fixes uniformity :: "('a \ 'a) filter" begin abbreviation uniformity_on :: "'a set \ ('a \ 'a) filter" where "uniformity_on s \ inf uniformity (principal (s\s))" end lemma uniformity_Abort: "uniformity = Filter.abstract_filter (\u. Code.abort (STR ''uniformity is not executable'') (\u. uniformity))" by simp class open_uniformity = "open" + uniformity + assumes open_uniformity: "\U. open U \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" begin subclass topological_space by standard (force elim: eventually_mono eventually_elim2 simp: split_beta' open_uniformity)+ end class uniform_space = open_uniformity + assumes uniformity_refl: "eventually E uniformity \ E (x, x)" and uniformity_sym: "eventually E uniformity \ eventually (\(x, y). E (y, x)) uniformity" and uniformity_trans: "eventually E uniformity \ \D. eventually D uniformity \ (\x y z. D (x, y) \ D (y, z) \ E (x, z))" begin lemma uniformity_bot: "uniformity \ bot" using uniformity_refl by auto lemma uniformity_trans': "eventually E uniformity \ eventually (\((x, y), (y', z)). y = y' \ E (x, z)) (uniformity \\<^sub>F uniformity)" by (drule uniformity_trans) (auto simp add: eventually_prod_same) lemma uniformity_transE: assumes "eventually E uniformity" obtains D where "eventually D uniformity" "\x y z. D (x, y) \ D (y, z) \ E (x, z)" using uniformity_trans [OF assms] by auto lemma eventually_nhds_uniformity: "eventually P (nhds x) \ eventually (\(x', y). x' = x \ P y) uniformity" (is "_ \ ?N P x") unfolding eventually_nhds proof safe assume *: "?N P x" have "?N (?N P) x" if "?N P x" for x proof - from that obtain D where ev: "eventually D uniformity" and D: "D (a, b) \ D (b, c) \ case (a, c) of (x', y) \ x' = x \ P y" for a b c by (rule uniformity_transE) simp from ev show ?thesis by eventually_elim (insert ev D, force elim: eventually_mono split: prod.split) qed then have "open {x. ?N P x}" by (simp add: open_uniformity) then show "\S. open S \ x \ S \ (\x\S. P x)" by (intro exI[of _ "{x. ?N P x}"]) (auto dest: uniformity_refl simp: *) qed (force simp add: open_uniformity elim: eventually_mono) subsubsection \Totally bounded sets\ definition totally_bounded :: "'a set \ bool" where "totally_bounded S \ (\E. eventually E uniformity \ (\X. finite X \ (\s\S. \x\X. E (x, s))))" lemma totally_bounded_empty[iff]: "totally_bounded {}" by (auto simp add: totally_bounded_def) lemma totally_bounded_subset: "totally_bounded S \ T \ S \ totally_bounded T" by (fastforce simp add: totally_bounded_def) lemma totally_bounded_Union[intro]: assumes M: "finite M" "\S. S \ M \ totally_bounded S" shows "totally_bounded (\M)" unfolding totally_bounded_def proof safe fix E assume "eventually E uniformity" with M obtain X where "\S\M. finite (X S) \ (\s\S. \x\X S. E (x, s))" by (metis totally_bounded_def) with \finite M\ show "\X. finite X \ (\s\\M. \x\X. E (x, s))" by (intro exI[of _ "\S\M. X S"]) force qed subsubsection \Cauchy filter\ definition cauchy_filter :: "'a filter \ bool" where "cauchy_filter F \ F \\<^sub>F F \ uniformity" definition Cauchy :: "(nat \ 'a) \ bool" where Cauchy_uniform: "Cauchy X = cauchy_filter (filtermap X sequentially)" lemma Cauchy_uniform_iff: "Cauchy X \ (\P. eventually P uniformity \ (\N. \n\N. \m\N. P (X n, X m)))" unfolding Cauchy_uniform cauchy_filter_def le_filter_def eventually_prod_same eventually_filtermap eventually_sequentially proof safe let ?U = "\P. eventually P uniformity" { fix P assume "?U P" "\P. ?U P \ (\Q. (\N. \n\N. Q (X n)) \ (\x y. Q x \ Q y \ P (x, y)))" then obtain Q N where "\n. n \ N \ Q (X n)" "\x y. Q x \ Q y \ P (x, y)" by metis then show "\N. \n\N. \m\N. P (X n, X m)" by blast next fix P assume "?U P" and P: "\P. ?U P \ (\N. \n\N. \m\N. P (X n, X m))" then obtain Q where "?U Q" and Q: "\x y z. Q (x, y) \ Q (y, z) \ P (x, z)" by (auto elim: uniformity_transE) then have "?U (\x. Q x \ (\(x, y). Q (y, x)) x)" unfolding eventually_conj_iff by (simp add: uniformity_sym) from P[rule_format, OF this] obtain N where N: "\n m. n \ N \ m \ N \ Q (X n, X m) \ Q (X m, X n)" by auto show "\Q. (\N. \n\N. Q (X n)) \ (\x y. Q x \ Q y \ P (x, y))" proof (safe intro!: exI[of _ "\x. \n\N. Q (x, X n) \ Q (X n, x)"] exI[of _ N] N) fix x y assume "\n\N. Q (x, X n) \ Q (X n, x)" "\n\N. Q (y, X n) \ Q (X n, y)" then have "Q (x, X N)" "Q (X N, y)" by auto then show "P (x, y)" by (rule Q) qed } qed lemma nhds_imp_cauchy_filter: assumes *: "F \ nhds x" shows "cauchy_filter F" proof - have "F \\<^sub>F F \ nhds x \\<^sub>F nhds x" by (intro prod_filter_mono *) also have "\ \ uniformity" unfolding le_filter_def eventually_nhds_uniformity eventually_prod_same proof safe fix P assume "eventually P uniformity" then obtain Ql where ev: "eventually Ql uniformity" and "Ql (x, y) \ Ql (y, z) \ P (x, z)" for x y z by (rule uniformity_transE) simp with ev[THEN uniformity_sym] show "\Q. eventually (\(x', y). x' = x \ Q y) uniformity \ (\x y. Q x \ Q y \ P (x, y))" by (rule_tac exI[of _ "\y. Ql (y, x) \ Ql (x, y)"]) (fastforce elim: eventually_elim2) qed finally show ?thesis by (simp add: cauchy_filter_def) qed lemma LIMSEQ_imp_Cauchy: "X \ x \ Cauchy X" unfolding Cauchy_uniform filterlim_def by (intro nhds_imp_cauchy_filter) lemma Cauchy_subseq_Cauchy: assumes "Cauchy X" "strict_mono f" shows "Cauchy (X \ f)" unfolding Cauchy_uniform comp_def filtermap_filtermap[symmetric] cauchy_filter_def by (rule order_trans[OF _ \Cauchy X\[unfolded Cauchy_uniform cauchy_filter_def]]) (intro prod_filter_mono filtermap_mono filterlim_subseq[OF \strict_mono f\, unfolded filterlim_def]) lemma convergent_Cauchy: "convergent X \ Cauchy X" unfolding convergent_def by (erule exE, erule LIMSEQ_imp_Cauchy) definition complete :: "'a set \ bool" where complete_uniform: "complete S \ (\F \ principal S. F \ bot \ cauchy_filter F \ (\x\S. F \ nhds x))" lemma (in uniform_space) cauchy_filter_complete_converges: assumes "cauchy_filter F" "complete A" "F \ principal A" "F \ bot" shows "\c. F \ nhds c" using assms unfolding complete_uniform by blast end subsubsection \Uniformly continuous functions\ definition uniformly_continuous_on :: "'a set \ ('a::uniform_space \ 'b::uniform_space) \ bool" where uniformly_continuous_on_uniformity: "uniformly_continuous_on s f \ (LIM (x, y) (uniformity_on s). (f x, f y) :> uniformity)" lemma uniformly_continuous_onD: "uniformly_continuous_on s f \ eventually E uniformity \ eventually (\(x, y). x \ s \ y \ s \ E (f x, f y)) uniformity" by (simp add: uniformly_continuous_on_uniformity filterlim_iff eventually_inf_principal split_beta' mem_Times_iff imp_conjL) lemma uniformly_continuous_on_const[continuous_intros]: "uniformly_continuous_on s (\x. c)" by (auto simp: uniformly_continuous_on_uniformity filterlim_iff uniformity_refl) lemma uniformly_continuous_on_id[continuous_intros]: "uniformly_continuous_on s (\x. x)" by (auto simp: uniformly_continuous_on_uniformity filterlim_def) lemma uniformly_continuous_on_compose: "uniformly_continuous_on s g \ uniformly_continuous_on (g`s) f \ uniformly_continuous_on s (\x. f (g x))" using filterlim_compose[of "\(x, y). (f x, f y)" uniformity "uniformity_on (g`s)" "\(x, y). (g x, g y)" "uniformity_on s"] by (simp add: split_beta' uniformly_continuous_on_uniformity filterlim_inf filterlim_principal eventually_inf_principal mem_Times_iff) lemma uniformly_continuous_imp_continuous: assumes f: "uniformly_continuous_on s f" shows "continuous_on s f" by (auto simp: filterlim_iff eventually_at_filter eventually_nhds_uniformity continuous_on_def elim: eventually_mono dest!: uniformly_continuous_onD[OF f]) section \Product Topology\ subsection \Product is a topological space\ instantiation prod :: (topological_space, topological_space) topological_space begin definition open_prod_def[code del]: "open (S :: ('a \ 'b) set) \ (\x\S. \A B. open A \ open B \ x \ A \ B \ A \ B \ S)" lemma open_prod_elim: assumes "open S" and "x \ S" obtains A B where "open A" and "open B" and "x \ A \ B" and "A \ B \ S" using assms unfolding open_prod_def by fast lemma open_prod_intro: assumes "\x. x \ S \ \A B. open A \ open B \ x \ A \ B \ A \ B \ S" shows "open S" using assms unfolding open_prod_def by fast instance proof show "open (UNIV :: ('a \ 'b) set)" unfolding open_prod_def by auto next fix S T :: "('a \ 'b) set" assume "open S" "open T" show "open (S \ T)" proof (rule open_prod_intro) fix x assume x: "x \ S \ T" from x have "x \ S" by simp obtain Sa Sb where A: "open Sa" "open Sb" "x \ Sa \ Sb" "Sa \ Sb \ S" using \open S\ and \x \ S\ by (rule open_prod_elim) from x have "x \ T" by simp obtain Ta Tb where B: "open Ta" "open Tb" "x \ Ta \ Tb" "Ta \ Tb \ T" using \open T\ and \x \ T\ by (rule open_prod_elim) let ?A = "Sa \ Ta" and ?B = "Sb \ Tb" have "open ?A \ open ?B \ x \ ?A \ ?B \ ?A \ ?B \ S \ T" using A B by (auto simp add: open_Int) then show "\A B. open A \ open B \ x \ A \ B \ A \ B \ S \ T" by fast qed next fix K :: "('a \ 'b) set set" assume "\S\K. open S" then show "open (\K)" unfolding open_prod_def by fast qed end declare [[code abort: "open :: ('a::topological_space \ 'b::topological_space) set \ bool"]] lemma open_Times: "open S \ open T \ open (S \ T)" unfolding open_prod_def by auto lemma fst_vimage_eq_Times: "fst -` S = S \ UNIV" by auto lemma snd_vimage_eq_Times: "snd -` S = UNIV \ S" by auto lemma open_vimage_fst: "open S \ open (fst -` S)" by (simp add: fst_vimage_eq_Times open_Times) lemma open_vimage_snd: "open S \ open (snd -` S)" by (simp add: snd_vimage_eq_Times open_Times) lemma closed_vimage_fst: "closed S \ closed (fst -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_fst) lemma closed_vimage_snd: "closed S \ closed (snd -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_snd) lemma closed_Times: "closed S \ closed T \ closed (S \ T)" proof - have "S \ T = (fst -` S) \ (snd -` T)" by auto then show "closed S \ closed T \ closed (S \ T)" by (simp add: closed_vimage_fst closed_vimage_snd closed_Int) qed lemma subset_fst_imageI: "A \ B \ S \ y \ B \ A \ fst ` S" unfolding image_def subset_eq by force lemma subset_snd_imageI: "A \ B \ S \ x \ A \ B \ snd ` S" unfolding image_def subset_eq by force lemma open_image_fst: assumes "open S" shows "open (fst ` S)" proof (rule openI) fix x assume "x \ fst ` S" then obtain y where "(x, y) \ S" by auto then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" using \open S\ unfolding open_prod_def by auto from \A \ B \ S\ \y \ B\ have "A \ fst ` S" by (rule subset_fst_imageI) with \open A\ \x \ A\ have "open A \ x \ A \ A \ fst ` S" by simp then show "\T. open T \ x \ T \ T \ fst ` S" .. qed lemma open_image_snd: assumes "open S" shows "open (snd ` S)" proof (rule openI) fix y assume "y \ snd ` S" then obtain x where "(x, y) \ S" by auto then obtain A B where "open A" "open B" "x \ A" "y \ B" "A \ B \ S" using \open S\ unfolding open_prod_def by auto from \A \ B \ S\ \x \ A\ have "B \ snd ` S" by (rule subset_snd_imageI) with \open B\ \y \ B\ have "open B \ y \ B \ B \ snd ` S" by simp then show "\T. open T \ y \ T \ T \ snd ` S" .. qed lemma nhds_prod: "nhds (a, b) = nhds a \\<^sub>F nhds b" unfolding nhds_def proof (subst prod_filter_INF, auto intro!: antisym INF_greatest simp: principal_prod_principal) fix S T assume "open S" "a \ S" "open T" "b \ T" then show "(INF x \ {S. open S \ (a, b) \ S}. principal x) \ principal (S \ T)" by (intro INF_lower) (auto intro!: open_Times) next fix S' assume "open S'" "(a, b) \ S'" then obtain S T where "open S" "a \ S" "open T" "b \ T" "S \ T \ S'" by (auto elim: open_prod_elim) then show "(INF x \ {S. open S \ a \ S}. INF y \ {S. open S \ b \ S}. principal (x \ y)) \ principal S'" by (auto intro!: INF_lower2) qed subsubsection \Continuity of operations\ lemma tendsto_fst [tendsto_intros]: assumes "(f \ a) F" shows "((\x. fst (f x)) \ fst a) F" proof (rule topological_tendstoI) fix S assume "open S" and "fst a \ S" then have "open (fst -` S)" and "a \ fst -` S" by (simp_all add: open_vimage_fst) with assms have "eventually (\x. f x \ fst -` S) F" by (rule topological_tendstoD) then show "eventually (\x. fst (f x) \ S) F" by simp qed lemma tendsto_snd [tendsto_intros]: assumes "(f \ a) F" shows "((\x. snd (f x)) \ snd a) F" proof (rule topological_tendstoI) fix S assume "open S" and "snd a \ S" then have "open (snd -` S)" and "a \ snd -` S" by (simp_all add: open_vimage_snd) with assms have "eventually (\x. f x \ snd -` S) F" by (rule topological_tendstoD) then show "eventually (\x. snd (f x) \ S) F" by simp qed lemma tendsto_Pair [tendsto_intros]: assumes "(f \ a) F" and "(g \ b) F" shows "((\x. (f x, g x)) \ (a, b)) F" unfolding nhds_prod using assms by (rule filterlim_Pair) lemma continuous_fst[continuous_intros]: "continuous F f \ continuous F (\x. fst (f x))" unfolding continuous_def by (rule tendsto_fst) lemma continuous_snd[continuous_intros]: "continuous F f \ continuous F (\x. snd (f x))" unfolding continuous_def by (rule tendsto_snd) lemma continuous_Pair[continuous_intros]: "continuous F f \ continuous F g \ continuous F (\x. (f x, g x))" unfolding continuous_def by (rule tendsto_Pair) lemma continuous_on_fst[continuous_intros]: "continuous_on s f \ continuous_on s (\x. fst (f x))" unfolding continuous_on_def by (auto intro: tendsto_fst) lemma continuous_on_snd[continuous_intros]: "continuous_on s f \ continuous_on s (\x. snd (f x))" unfolding continuous_on_def by (auto intro: tendsto_snd) lemma continuous_on_Pair[continuous_intros]: "continuous_on s f \ continuous_on s g \ continuous_on s (\x. (f x, g x))" unfolding continuous_on_def by (auto intro: tendsto_Pair) lemma continuous_on_swap[continuous_intros]: "continuous_on A prod.swap" by (simp add: prod.swap_def continuous_on_fst continuous_on_snd continuous_on_Pair continuous_on_id) lemma continuous_on_swap_args: assumes "continuous_on (A\B) (\(x,y). d x y)" shows "continuous_on (B\A) (\(x,y). d y x)" proof - have "(\(x,y). d y x) = (\(x,y). d x y) \ prod.swap" by force then show ?thesis by (metis assms continuous_on_compose continuous_on_swap product_swap) qed lemma isCont_fst [simp]: "isCont f a \ isCont (\x. fst (f x)) a" by (fact continuous_fst) lemma isCont_snd [simp]: "isCont f a \ isCont (\x. snd (f x)) a" by (fact continuous_snd) lemma isCont_Pair [simp]: "\isCont f a; isCont g a\ \ isCont (\x. (f x, g x)) a" by (fact continuous_Pair) lemma continuous_on_compose_Pair: assumes f: "continuous_on (Sigma A B) (\(a, b). f a b)" assumes g: "continuous_on C g" assumes h: "continuous_on C h" assumes subset: "\c. c \ C \ g c \ A" "\c. c \ C \ h c \ B (g c)" shows "continuous_on C (\c. f (g c) (h c))" using continuous_on_compose2[OF f continuous_on_Pair[OF g h]] subset by auto subsubsection \Connectedness of products\ proposition connected_Times: assumes S: "connected S" and T: "connected T" shows "connected (S \ T)" proof (rule connectedI_const) fix P::"'a \ 'b \ bool" assume P[THEN continuous_on_compose2, continuous_intros]: "continuous_on (S \ T) P" have "continuous_on S (\s. P (s, t))" if "t \ T" for t by (auto intro!: continuous_intros that) from connectedD_const[OF S this] obtain c1 where c1: "\s t. t \ T \ s \ S \ P (s, t) = c1 t" by metis moreover have "continuous_on T (\t. P (s, t))" if "s \ S" for s by (auto intro!: continuous_intros that) from connectedD_const[OF T this] obtain c2 where "\s t. t \ T \ s \ S \ P (s, t) = c2 s" by metis ultimately show "\c. \s\S \ T. P s = c" by auto qed corollary connected_Times_eq [simp]: "connected (S \ T) \ S = {} \ T = {} \ connected S \ connected T" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof cases assume "S \ {} \ T \ {}" moreover have "connected (fst ` (S \ T))" "connected (snd ` (S \ T))" using continuous_on_fst continuous_on_snd continuous_on_id by (blast intro: connected_continuous_image [OF _ L])+ ultimately show ?thesis by auto qed auto qed (auto simp: connected_Times) subsubsection \Separation axioms\ instance prod :: (t0_space, t0_space) t0_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U. open U \ (x \ U) \ (y \ U)" by (fast dest: t0_space elim: open_vimage_fst open_vimage_snd) qed instance prod :: (t1_space, t1_space) t1_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U. open U \ x \ U \ y \ U" by (fast dest: t1_space elim: open_vimage_fst open_vimage_snd) qed instance prod :: (t2_space, t2_space) t2_space proof fix x y :: "'a \ 'b" assume "x \ y" then have "fst x \ fst y \ snd x \ snd y" by (simp add: prod_eq_iff) then show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" by (fast dest: hausdorff elim: open_vimage_fst open_vimage_snd) qed lemma isCont_swap[continuous_intros]: "isCont prod.swap a" using continuous_on_eq_continuous_within continuous_on_swap by blast lemma open_diagonal_complement: "open {(x,y) |x y. x \ (y::('a::t2_space))}" proof - have "open {(x, y). x \ (y::'a)}" unfolding split_def by (intro open_Collect_neq continuous_intros) also have "{(x, y). x \ (y::'a)} = {(x, y) |x y. x \ (y::'a)}" by auto finally show ?thesis . qed lemma closed_diagonal: "closed {y. \ x::('a::t2_space). y = (x,x)}" proof - have "{y. \ x::'a. y = (x,x)} = UNIV - {(x,y) | x y. x \ y}" by auto then show ?thesis using open_diagonal_complement closed_Diff by auto qed lemma open_superdiagonal: "open {(x,y) | x y. x > (y::'a::{linorder_topology})}" proof - have "open {(x, y). x > (y::'a)}" unfolding split_def by (intro open_Collect_less continuous_intros) also have "{(x, y). x > (y::'a)} = {(x, y) |x y. x > (y::'a)}" by auto finally show ?thesis . qed lemma closed_subdiagonal: "closed {(x,y) | x y. x \ (y::'a::{linorder_topology})}" proof - have "{(x,y) | x y. x \ (y::'a)} = UNIV - {(x,y) | x y. x > (y::'a)}" by auto then show ?thesis using open_superdiagonal closed_Diff by auto qed lemma open_subdiagonal: "open {(x,y) | x y. x < (y::'a::{linorder_topology})}" proof - have "open {(x, y). x < (y::'a)}" unfolding split_def by (intro open_Collect_less continuous_intros) also have "{(x, y). x < (y::'a)} = {(x, y) |x y. x < (y::'a)}" by auto finally show ?thesis . qed lemma closed_superdiagonal: "closed {(x,y) | x y. x \ (y::('a::{linorder_topology}))}" proof - have "{(x,y) | x y. x \ (y::'a)} = UNIV - {(x,y) | x y. x < y}" by auto then show ?thesis using open_subdiagonal closed_Diff by auto qed end