| author | wenzelm | 
| Tue, 15 Apr 2025 16:53:07 +0200 | |
| changeset 82545 | 0d955ab17466 | 
| parent 78336 | 6bae28577994 | 
| permissions | -rw-r--r-- | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 1 | section\<open>T1 and Hausdorff spaces\<close> | 
| 69874 | 2 | |
| 3 | theory T1_Spaces | |
| 71200 
3548d54ce3ee
split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
 immler parents: 
70196diff
changeset | 4 | imports Product_Topology | 
| 69874 | 5 | begin | 
| 6 | ||
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 7 | section\<open>T1 spaces with equivalences to many naturally "nice" properties. \<close> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 8 | |
| 69874 | 9 | definition t1_space where | 
| 10 | "t1_space X \<equiv> \<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x\<noteq>y \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> y \<notin> U)" | |
| 11 | ||
| 12 | lemma t1_space_expansive: | |
| 13 | "\<lbrakk>topspace Y = topspace X; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> \<Longrightarrow> t1_space X \<Longrightarrow> t1_space Y" | |
| 14 | by (metis t1_space_def) | |
| 15 | ||
| 16 | lemma t1_space_alt: | |
| 17 | "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x\<noteq>y \<longrightarrow> (\<exists>U. closedin X U \<and> x \<in> U \<and> y \<notin> U))" | |
| 18 | by (metis DiffE DiffI closedin_def openin_closedin_eq t1_space_def) | |
| 19 | ||
| 78336 | 20 | lemma t1_space_empty [iff]: "t1_space trivial_topology" | 
| 69874 | 21 | by (simp add: t1_space_def) | 
| 22 | ||
| 23 | lemma t1_space_derived_set_of_singleton: | |
| 24 |   "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. X derived_set_of {x} = {})"
 | |
| 25 | apply (simp add: t1_space_def derived_set_of_def, safe) | |
| 26 | apply (metis openin_topspace) | |
| 27 | by force | |
| 28 | ||
| 29 | lemma t1_space_derived_set_of_finite: | |
| 30 |    "t1_space X \<longleftrightarrow> (\<forall>S. finite S \<longrightarrow> X derived_set_of S = {})"
 | |
| 31 | proof (intro iffI allI impI) | |
| 32 | fix S :: "'a set" | |
| 33 | assume "finite S" | |
| 34 |   then have fin: "finite ((\<lambda>x. {x}) ` (topspace X \<inter> S))"
 | |
| 35 | by blast | |
| 36 | assume "t1_space X" | |
| 37 |   then have "X derived_set_of (\<Union>x \<in> topspace X \<inter> S. {x}) = {}"
 | |
| 38 | unfolding derived_set_of_Union [OF fin] | |
| 39 | by (auto simp: t1_space_derived_set_of_singleton) | |
| 40 |   then have "X derived_set_of (topspace X \<inter> S) = {}"
 | |
| 41 | by simp | |
| 42 |   then show "X derived_set_of S = {}"
 | |
| 43 | by simp | |
| 44 | qed (auto simp: t1_space_derived_set_of_singleton) | |
| 45 | ||
| 46 | lemma t1_space_closedin_singleton: | |
| 47 |    "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. closedin X {x})"
 | |
| 48 | apply (rule iffI) | |
| 49 | apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_singleton) | |
| 50 | using t1_space_alt by auto | |
| 51 | ||
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 52 | lemma continuous_closed_imp_proper_map: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 53 | "\<lbrakk>compact_space X; t1_space Y; continuous_map X Y f; closed_map X Y f\<rbrakk> \<Longrightarrow> proper_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 54 | unfolding proper_map_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 55 | by (smt (verit) closedin_compact_space closedin_continuous_map_preimage | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 56 | Collect_cong singleton_iff t1_space_closedin_singleton) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
77943diff
changeset | 57 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77935diff
changeset | 58 | lemma t1_space_euclidean: "t1_space (euclidean :: 'a::metric_space topology)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77935diff
changeset | 59 | by (simp add: t1_space_closedin_singleton) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77935diff
changeset | 60 | |
| 69874 | 61 | lemma closedin_t1_singleton: | 
| 62 |    "\<lbrakk>t1_space X; a \<in> topspace X\<rbrakk> \<Longrightarrow> closedin X {a}"
 | |
| 63 | by (simp add: t1_space_closedin_singleton) | |
| 64 | ||
| 65 | lemma t1_space_closedin_finite: | |
| 66 | "t1_space X \<longleftrightarrow> (\<forall>S. finite S \<and> S \<subseteq> topspace X \<longrightarrow> closedin X S)" | |
| 67 | apply (rule iffI) | |
| 68 | apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_finite) | |
| 69 | by (simp add: t1_space_closedin_singleton) | |
| 70 | ||
| 71 | lemma closure_of_singleton: | |
| 72 |    "t1_space X \<Longrightarrow> X closure_of {a} = (if a \<in> topspace X then {a} else {})"
 | |
| 73 | by (simp add: closure_of_eq t1_space_closedin_singleton closure_of_eq_empty_gen) | |
| 74 | ||
| 75 | lemma separated_in_singleton: | |
| 76 | assumes "t1_space X" | |
| 77 |   shows "separatedin X {a} S \<longleftrightarrow> a \<in> topspace X \<and> S \<subseteq> topspace X \<and> (a \<notin> X closure_of S)"
 | |
| 78 |         "separatedin X S {a} \<longleftrightarrow> a \<in> topspace X \<and> S \<subseteq> topspace X \<and> (a \<notin> X closure_of S)"
 | |
| 79 | unfolding separatedin_def | |
| 80 | using assms closure_of closure_of_singleton by fastforce+ | |
| 81 | ||
| 82 | lemma t1_space_openin_delete: | |
| 83 |    "t1_space X \<longleftrightarrow> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (U - {x}))"
 | |
| 84 | apply (rule iffI) | |
| 85 | apply (meson closedin_t1_singleton in_mono openin_diff openin_subset) | |
| 86 | by (simp add: closedin_def t1_space_closedin_singleton) | |
| 87 | ||
| 88 | lemma t1_space_openin_delete_alt: | |
| 89 |    "t1_space X \<longleftrightarrow> (\<forall>U x. openin X U \<longrightarrow> openin X (U - {x}))"
 | |
| 90 | by (metis Diff_empty Diff_insert0 t1_space_openin_delete) | |
| 91 | ||
| 92 | ||
| 93 | lemma t1_space_singleton_Inter_open: | |
| 94 |       "t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<Inter>{U. openin X U \<and> x \<in> U} = {x})"  (is "?P=?Q")
 | |
| 95 | and t1_space_Inter_open_supersets: | |
| 96 |      "t1_space X \<longleftrightarrow> (\<forall>S. S \<subseteq> topspace X \<longrightarrow> \<Inter>{U. openin X U \<and> S \<subseteq> U} = S)" (is "?P=?R")
 | |
| 97 | proof - | |
| 98 | have "?R \<Longrightarrow> ?Q" | |
| 99 | apply clarify | |
| 100 |     apply (drule_tac x="{x}" in spec, simp)
 | |
| 101 | done | |
| 102 | moreover have "?Q \<Longrightarrow> ?P" | |
| 103 | apply (clarsimp simp add: t1_space_def) | |
| 104 | apply (drule_tac x=x in bspec) | |
| 105 | apply (simp_all add: set_eq_iff) | |
| 106 | by (metis (no_types, lifting)) | |
| 107 | moreover have "?P \<Longrightarrow> ?R" | |
| 108 | proof (clarsimp simp add: t1_space_closedin_singleton, rule subset_antisym) | |
| 109 | fix S | |
| 110 |     assume S: "\<forall>x\<in>topspace X. closedin X {x}" "S \<subseteq> topspace X"
 | |
| 111 |     then show "\<Inter> {U. openin X U \<and> S \<subseteq> U} \<subseteq> S"
 | |
| 112 | apply clarsimp | |
| 113 | by (metis Diff_insert_absorb Set.set_insert closedin_def openin_topspace subset_insert) | |
| 114 | qed force | |
| 115 | ultimately show "?P=?Q" "?P=?R" | |
| 116 | by auto | |
| 117 | qed | |
| 118 | ||
| 119 | lemma t1_space_derived_set_of_infinite_openin: | |
| 120 | "t1_space X \<longleftrightarrow> | |
| 121 | (\<forall>S. X derived_set_of S = | |
| 122 |              {x \<in> topspace X. \<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite(S \<inter> U)})"
 | |
| 123 | (is "_ = ?rhs") | |
| 124 | proof | |
| 125 | assume "t1_space X" | |
| 126 | show ?rhs | |
| 127 | proof safe | |
| 128 | fix S x U | |
| 129 | assume "x \<in> X derived_set_of S" "x \<in> U" "openin X U" "finite (S \<inter> U)" | |
| 130 | with \<open>t1_space X\<close> show "False" | |
| 131 | apply (simp add: t1_space_derived_set_of_finite) | |
| 132 | by (metis IntI empty_iff empty_subsetI inf_commute openin_Int_derived_set_of_subset subset_antisym) | |
| 133 | next | |
| 134 | fix S x | |
| 135 |     have eq: "(\<exists>y. (y \<noteq> x) \<and> y \<in> S \<and> y \<in> T) \<longleftrightarrow> ~((S \<inter> T) \<subseteq> {x})" for x S T
 | |
| 136 | by blast | |
| 137 | assume "x \<in> topspace X" "\<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite (S \<inter> U)" | |
| 138 | then show "x \<in> X derived_set_of S" | |
| 139 | apply (clarsimp simp add: derived_set_of_def eq) | |
| 140 | by (meson finite.emptyI finite.insertI finite_subset) | |
| 141 | qed (auto simp: in_derived_set_of) | |
| 142 | qed (auto simp: t1_space_derived_set_of_singleton) | |
| 143 | ||
| 144 | lemma finite_t1_space_imp_discrete_topology: | |
| 145 | "\<lbrakk>topspace X = U; finite U; t1_space X\<rbrakk> \<Longrightarrow> X = discrete_topology U" | |
| 146 | by (metis discrete_topology_unique_derived_set t1_space_derived_set_of_finite) | |
| 147 | ||
| 148 | lemma t1_space_subtopology: "t1_space X \<Longrightarrow> t1_space(subtopology X U)" | |
| 149 | by (simp add: derived_set_of_subtopology t1_space_derived_set_of_finite) | |
| 150 | ||
| 151 | lemma closedin_derived_set_of_gen: | |
| 152 | "t1_space X \<Longrightarrow> closedin X (X derived_set_of S)" | |
| 153 | apply (clarsimp simp add: in_derived_set_of closedin_contains_derived_set derived_set_of_subset_topspace) | |
| 154 | by (metis DiffD2 insert_Diff insert_iff t1_space_openin_delete) | |
| 155 | ||
| 156 | lemma derived_set_of_derived_set_subset_gen: | |
| 157 | "t1_space X \<Longrightarrow> X derived_set_of (X derived_set_of S) \<subseteq> X derived_set_of S" | |
| 158 | by (meson closedin_contains_derived_set closedin_derived_set_of_gen) | |
| 159 | ||
| 160 | lemma subtopology_eq_discrete_topology_gen_finite: | |
| 161 | "\<lbrakk>t1_space X; finite S\<rbrakk> \<Longrightarrow> subtopology X S = discrete_topology(topspace X \<inter> S)" | |
| 162 | by (simp add: subtopology_eq_discrete_topology_gen t1_space_derived_set_of_finite) | |
| 163 | ||
| 164 | lemma subtopology_eq_discrete_topology_finite: | |
| 165 | "\<lbrakk>t1_space X; S \<subseteq> topspace X; finite S\<rbrakk> | |
| 166 | \<Longrightarrow> subtopology X S = discrete_topology S" | |
| 167 | by (simp add: subtopology_eq_discrete_topology_eq t1_space_derived_set_of_finite) | |
| 168 | ||
| 169 | lemma t1_space_closed_map_image: | |
| 170 | "\<lbrakk>closed_map X Y f; f ` (topspace X) = topspace Y; t1_space X\<rbrakk> \<Longrightarrow> t1_space Y" | |
| 171 | by (metis closed_map_def finite_subset_image t1_space_closedin_finite) | |
| 172 | ||
| 173 | lemma homeomorphic_t1_space: "X homeomorphic_space Y \<Longrightarrow> (t1_space X \<longleftrightarrow> t1_space Y)" | |
| 174 | apply (clarsimp simp add: homeomorphic_space_def) | |
| 175 | by (meson homeomorphic_eq_everything_map homeomorphic_maps_map t1_space_closed_map_image) | |
| 176 | ||
| 177 | proposition t1_space_product_topology: | |
| 178 | "t1_space (product_topology X I) | |
| 78336 | 179 | \<longleftrightarrow> (product_topology X I) = trivial_topology \<or> (\<forall>i \<in> I. t1_space (X i))" | 
| 180 | proof (cases "(product_topology X I) = trivial_topology") | |
| 69874 | 181 | case True | 
| 182 | then show ?thesis | |
| 78336 | 183 | using True t1_space_empty by force | 
| 69874 | 184 | next | 
| 185 | case False | |
| 186 | then obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))" | |
| 78336 | 187 | using discrete_topology_unique by (fastforce iff: null_topspace_iff_trivial) | 
| 69874 | 188 | have "t1_space (product_topology X I) \<longleftrightarrow> (\<forall>i\<in>I. t1_space (X i))" | 
| 189 | proof (intro iffI ballI) | |
| 190 | show "t1_space (X i)" if "t1_space (product_topology X I)" and "i \<in> I" for i | |
| 191 | proof - | |
| 192 |       have clo: "\<And>h. h \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<Longrightarrow> closedin (product_topology X I) {h}"
 | |
| 193 | using that by (simp add: t1_space_closedin_singleton) | |
| 194 | show ?thesis | |
| 195 | unfolding t1_space_closedin_singleton | |
| 196 | proof clarify | |
| 197 |         show "closedin (X i) {xi}" if "xi \<in> topspace (X i)" for xi
 | |
| 198 | using clo [of "\<lambda>j \<in> I. if i=j then xi else f j"] f that \<open>i \<in> I\<close> | |
| 199 | by (fastforce simp add: closedin_product_topology_singleton) | |
| 200 | qed | |
| 201 | qed | |
| 202 | next | |
| 203 | next | |
| 204 | show "t1_space (product_topology X I)" if "\<forall>i\<in>I. t1_space (X i)" | |
| 205 | using that | |
| 206 | by (simp add: t1_space_closedin_singleton Ball_def PiE_iff closedin_product_topology_singleton) | |
| 207 | qed | |
| 208 | then show ?thesis | |
| 209 | using False by blast | |
| 210 | qed | |
| 211 | ||
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 212 | lemma t1_space_prod_topology: | 
| 78336 | 213 | "t1_space(prod_topology X Y) \<longleftrightarrow> (prod_topology X Y) = trivial_topology \<or> t1_space X \<and> t1_space Y" | 
| 214 | proof (cases "(prod_topology X Y) = trivial_topology") | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 215 | case True then show ?thesis | 
| 78336 | 216 | by auto | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 217 | next | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 218 | case False | 
| 78336 | 219 |   have eq: "{(x,y)} = {x} \<times> {y}" for x::'a and y::'b
 | 
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 220 | by simp | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 221 | have "t1_space (prod_topology X Y) \<longleftrightarrow> (t1_space X \<and> t1_space Y)" | 
| 78336 | 222 | using False | 
| 223 | apply(simp add: t1_space_closedin_singleton closedin_prod_Times_iff eq | |
| 224 | del: insert_Times_insert flip: null_topspace_iff_trivial ex_in_conv) | |
| 225 | by blast | |
| 69939 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 226 | with False show ?thesis | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 227 | by simp | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 228 | qed | 
| 
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 229 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 230 | subsection\<open>Hausdorff Spaces\<close> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 231 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 232 | definition Hausdorff_space | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 233 | where | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 234 | "Hausdorff_space X \<equiv> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 235 | \<forall>x y. x \<in> topspace X \<and> y \<in> topspace X \<and> (x \<noteq> y) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 236 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 237 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 238 | lemma Hausdorff_space_expansive: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 239 | "\<lbrakk>Hausdorff_space X; topspace X = topspace Y; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> \<Longrightarrow> Hausdorff_space Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 240 | by (metis Hausdorff_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 241 | |
| 78336 | 242 | lemma Hausdorff_space_topspace_empty [iff]: "Hausdorff_space trivial_topology" | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 243 | by (simp add: Hausdorff_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 244 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 245 | lemma Hausdorff_imp_t1_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 246 | "Hausdorff_space X \<Longrightarrow> t1_space X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 247 | by (metis Hausdorff_space_def disjnt_iff t1_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 248 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 249 | lemma closedin_derived_set_of: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 250 | "Hausdorff_space X \<Longrightarrow> closedin X (X derived_set_of S)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 251 | by (simp add: Hausdorff_imp_t1_space closedin_derived_set_of_gen) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 252 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 253 | lemma t1_or_Hausdorff_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 254 | "t1_space X \<or> Hausdorff_space X \<longleftrightarrow> t1_space X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 255 | using Hausdorff_imp_t1_space by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 256 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 257 | lemma Hausdorff_space_sing_Inter_opens: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 258 |    "\<lbrakk>Hausdorff_space X; a \<in> topspace X\<rbrakk> \<Longrightarrow> \<Inter>{u. openin X u \<and> a \<in> u} = {a}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 259 | using Hausdorff_imp_t1_space t1_space_singleton_Inter_open by force | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 260 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 261 | lemma Hausdorff_space_subtopology: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 262 | assumes "Hausdorff_space X" shows "Hausdorff_space(subtopology X S)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 263 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 264 | have *: "disjnt U V \<Longrightarrow> disjnt (S \<inter> U) (S \<inter> V)" for U V | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 265 | by (simp add: disjnt_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 266 | from assms show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 267 | apply (simp add: Hausdorff_space_def openin_subtopology_alt) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 268 | apply (fast intro: * elim!: all_forward) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 269 | done | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 270 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 271 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 272 | lemma Hausdorff_space_compact_separation: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 273 | assumes X: "Hausdorff_space X" and S: "compactin X S" and T: "compactin X T" and "disjnt S T" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 274 | obtains U V where "openin X U" "openin X V" "S \<subseteq> U" "T \<subseteq> V" "disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 275 | proof (cases "S = {}")
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 276 | case True | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 277 | then show thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 278 | by (metis \<open>compactin X T\<close> compactin_subset_topspace disjnt_empty1 empty_subsetI openin_empty openin_topspace that) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 279 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 280 | case False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 281 | have "\<forall>x \<in> S. \<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> T \<subseteq> V \<and> disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 282 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 283 | fix a | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 284 | assume "a \<in> S" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 285 | then have "a \<notin> T" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 286 | by (meson assms(4) disjnt_iff) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 287 | have a: "a \<in> topspace X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 288 | using S \<open>a \<in> S\<close> compactin_subset_topspace by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 289 | show "\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> T \<subseteq> V \<and> disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 290 |     proof (cases "T = {}")
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 291 | case True | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 292 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 293 | using a disjnt_empty2 openin_empty by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 294 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 295 | case False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 296 |       have "\<forall>x \<in> topspace X - {a}. \<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> a \<in> V \<and> disjnt U V"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 297 | using X a by (simp add: Hausdorff_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 298 |       then obtain U V where UV: "\<forall>x \<in> topspace X - {a}. openin X (U x) \<and> openin X (V x) \<and> x \<in> U x \<and> a \<in> V x \<and> disjnt (U x) (V x)"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 299 | by metis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 300 | with \<open>a \<notin> T\<close> compactin_subset_topspace [OF T] | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 301 | have Topen: "\<forall>W \<in> U ` T. openin X W" and Tsub: "T \<subseteq> \<Union> (U ` T)" | 
| 75455 
91c16c5ad3e9
tidied auto / simp with null arguments
 paulson <lp15@cam.ac.uk> parents: 
71200diff
changeset | 302 | by auto | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 303 | then obtain \<F> where \<F>: "finite \<F>" "\<F> \<subseteq> U ` T" and "T \<subseteq> \<Union>\<F>" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 304 | using T unfolding compactin_def by meson | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 305 | then obtain F where F: "finite F" "F \<subseteq> T" "\<F> = U ` F" and SUF: "T \<subseteq> \<Union>(U ` F)" and "a \<notin> F" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 306 | using finite_subset_image [OF \<F>] \<open>a \<notin> T\<close> by (metis subsetD) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 307 | have U: "\<And>x. \<lbrakk>x \<in> topspace X; x \<noteq> a\<rbrakk> \<Longrightarrow> openin X (U x)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 308 | and V: "\<And>x. \<lbrakk>x \<in> topspace X; x \<noteq> a\<rbrakk> \<Longrightarrow> openin X (V x)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 309 | and disj: "\<And>x. \<lbrakk>x \<in> topspace X; x \<noteq> a\<rbrakk> \<Longrightarrow> disjnt (U x) (V x)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 310 | using UV by blast+ | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 311 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 312 | proof (intro exI conjI) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 313 |         have "F \<noteq> {}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 314 | using False SUF by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 315 | with \<open>a \<notin> F\<close> show "openin X (\<Inter>(V ` F))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 316 | using F compactin_subset_topspace [OF T] by (force intro: V) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 317 | show "openin X (\<Union>(U ` F))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 318 | using F Topen Tsub by (force intro: U) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 319 | show "disjnt (\<Inter>(V ` F)) (\<Union>(U ` F))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 320 | using disj | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 321 | apply (auto simp: disjnt_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 322 | using \<open>F \<subseteq> T\<close> \<open>a \<notin> F\<close> compactin_subset_topspace [OF T] by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 323 | show "a \<in> (\<Inter>(V ` F))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 324 | using \<open>F \<subseteq> T\<close> T UV \<open>a \<notin> T\<close> compactin_subset_topspace by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 325 | qed (auto simp: SUF) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 326 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 327 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 328 | then obtain U V where UV: "\<forall>x \<in> S. openin X (U x) \<and> openin X (V x) \<and> x \<in> U x \<and> T \<subseteq> V x \<and> disjnt (U x) (V x)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 329 | by metis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 330 | then have "S \<subseteq> \<Union> (U ` S)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 331 | by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 332 | moreover have "\<forall>W \<in> U ` S. openin X W" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 333 | using UV by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 334 | ultimately obtain I where I: "S \<subseteq> \<Union> (U ` I)" "I \<subseteq> S" "finite I" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 335 | by (metis S compactin_def finite_subset_image) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 336 | show thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 337 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 338 | show "openin X (\<Union>(U ` I))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 339 | using \<open>I \<subseteq> S\<close> UV by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 340 | show "openin X (\<Inter> (V ` I))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 341 | using False UV \<open>I \<subseteq> S\<close> \<open>S \<subseteq> \<Union> (U ` I)\<close> \<open>finite I\<close> by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 342 | show "disjnt (\<Union>(U ` I)) (\<Inter> (V ` I))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 343 | by simp (meson UV \<open>I \<subseteq> S\<close> disjnt_subset2 in_mono le_INF_iff order_refl) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 344 | qed (use UV I in auto) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 345 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 346 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 347 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 348 | lemma Hausdorff_space_compact_sets: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 349 | "Hausdorff_space X \<longleftrightarrow> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 350 | (\<forall>S T. compactin X S \<and> compactin X T \<and> disjnt S T | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 351 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 352 | (is "?lhs = ?rhs") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 353 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 354 | assume ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 355 | then show ?rhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 356 | by (meson Hausdorff_space_compact_separation) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 357 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 358 | assume R [rule_format]: ?rhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 359 | show ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 360 | proof (clarsimp simp add: Hausdorff_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 361 | fix x y | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 362 | assume "x \<in> topspace X" "y \<in> topspace X" "x \<noteq> y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 363 | then show "\<exists>U. openin X U \<and> (\<exists>V. openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 364 |       using R [of "{x}" "{y}"] by auto
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 365 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 366 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 367 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 368 | lemma compactin_imp_closedin: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 369 | assumes X: "Hausdorff_space X" and S: "compactin X S" shows "closedin X S" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 370 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 371 | have "S \<subseteq> topspace X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 372 | by (simp add: assms compactin_subset_topspace) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 373 | moreover | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 374 | have "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> topspace X - S" if "x \<in> topspace X" "x \<notin> S" for x | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 375 |     using Hausdorff_space_compact_separation [OF X _ S, of "{x}"] that
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 376 | apply (simp add: disjnt_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 377 | by (metis Diff_mono Diff_triv openin_subset) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 378 | ultimately show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 379 | using closedin_def openin_subopen by force | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 380 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 381 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 382 | lemma closedin_Hausdorff_singleton: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 383 |    "\<lbrakk>Hausdorff_space X; x \<in> topspace X\<rbrakk> \<Longrightarrow> closedin X {x}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 384 | by (simp add: Hausdorff_imp_t1_space closedin_t1_singleton) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 385 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 386 | lemma closedin_Hausdorff_sing_eq: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 387 |    "Hausdorff_space X \<Longrightarrow> closedin X {x} \<longleftrightarrow> x \<in> topspace X"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 388 | by (meson closedin_Hausdorff_singleton closedin_subset insert_subset) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 389 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 390 | lemma Hausdorff_space_discrete_topology [simp]: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 391 | "Hausdorff_space (discrete_topology U)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 392 | unfolding Hausdorff_space_def | 
| 77935 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 393 | by (metis Hausdorff_space_compact_sets Hausdorff_space_def compactin_discrete_topology equalityE openin_discrete_topology) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 394 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 395 | lemma compactin_Int: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 396 | "\<lbrakk>Hausdorff_space X; compactin X S; compactin X T\<rbrakk> \<Longrightarrow> compactin X (S \<inter> T)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 397 | by (simp add: closed_Int_compactin compactin_imp_closedin) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 398 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 399 | lemma finite_topspace_imp_discrete_topology: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 400 | "\<lbrakk>topspace X = U; finite U; Hausdorff_space X\<rbrakk> \<Longrightarrow> X = discrete_topology U" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 401 | using Hausdorff_imp_t1_space finite_t1_space_imp_discrete_topology by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 402 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 403 | lemma derived_set_of_finite: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 404 |    "\<lbrakk>Hausdorff_space X; finite S\<rbrakk> \<Longrightarrow> X derived_set_of S = {}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 405 | using Hausdorff_imp_t1_space t1_space_derived_set_of_finite by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 406 | |
| 77935 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 407 | lemma infinite_perfect_set: | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 408 |    "\<lbrakk>Hausdorff_space X; S \<subseteq> X derived_set_of S; S \<noteq> {}\<rbrakk> \<Longrightarrow> infinite S"
 | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 409 | using derived_set_of_finite by blast | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 410 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 411 | lemma derived_set_of_singleton: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 412 |    "Hausdorff_space X \<Longrightarrow> X derived_set_of {x} = {}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 413 | by (simp add: derived_set_of_finite) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 414 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 415 | lemma closedin_Hausdorff_finite: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 416 | "\<lbrakk>Hausdorff_space X; S \<subseteq> topspace X; finite S\<rbrakk> \<Longrightarrow> closedin X S" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 417 | by (simp add: compactin_imp_closedin finite_imp_compactin_eq) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 418 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 419 | lemma open_in_Hausdorff_delete: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 420 |    "\<lbrakk>Hausdorff_space X; openin X S\<rbrakk> \<Longrightarrow> openin X (S - {x})"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 421 | using Hausdorff_imp_t1_space t1_space_openin_delete_alt by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 422 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 423 | lemma closedin_Hausdorff_finite_eq: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 424 | "\<lbrakk>Hausdorff_space X; finite S\<rbrakk> \<Longrightarrow> closedin X S \<longleftrightarrow> S \<subseteq> topspace X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 425 | by (meson closedin_Hausdorff_finite closedin_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 426 | |
| 70019 
095dce9892e8
A few results in Algebra, and bits for Analysis
 paulson <lp15@cam.ac.uk> parents: 
69994diff
changeset | 427 | lemma derived_set_of_infinite_openin: | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 428 | "Hausdorff_space X | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 429 | \<Longrightarrow> X derived_set_of S = | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 430 |             {x \<in> topspace X. \<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite(S \<inter> U)}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 431 | using Hausdorff_imp_t1_space t1_space_derived_set_of_infinite_openin by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 432 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 433 | lemma Hausdorff_space_discrete_compactin: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 434 | "Hausdorff_space X | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 435 |         \<Longrightarrow> S \<inter> X derived_set_of S = {} \<and> compactin X S \<longleftrightarrow> S \<subseteq> topspace X \<and> finite S"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 436 | using derived_set_of_finite discrete_compactin_eq_finite by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 437 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 438 | lemma Hausdorff_space_finite_topspace: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 439 |    "Hausdorff_space X \<Longrightarrow> X derived_set_of (topspace X) = {} \<and> compact_space X \<longleftrightarrow> finite(topspace X)"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 440 | using derived_set_of_finite discrete_compact_space_eq_finite by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 441 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 442 | lemma derived_set_of_derived_set_subset: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 443 | "Hausdorff_space X \<Longrightarrow> X derived_set_of (X derived_set_of S) \<subseteq> X derived_set_of S" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 444 | by (simp add: Hausdorff_imp_t1_space derived_set_of_derived_set_subset_gen) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 445 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 446 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 447 | lemma Hausdorff_space_injective_preimage: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 448 | assumes "Hausdorff_space Y" and cmf: "continuous_map X Y f" and "inj_on f (topspace X)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 449 | shows "Hausdorff_space X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 450 | unfolding Hausdorff_space_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 451 | proof clarify | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 452 | fix x y | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 453 | assume x: "x \<in> topspace X" and y: "y \<in> topspace X" and "x \<noteq> y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 454 | then obtain U V where "openin Y U" "openin Y V" "f x \<in> U" "f y \<in> V" "disjnt U V" | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 455 | using assms | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 456 | by (smt (verit, ccfv_threshold) Hausdorff_space_def continuous_map image_subset_iff inj_on_def) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 457 | show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 458 | proof (intro exI conjI) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 459 |     show "openin X {x \<in> topspace X. f x \<in> U}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 460 | using \<open>openin Y U\<close> cmf continuous_map by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 461 |     show "openin X {x \<in> topspace X. f x \<in> V}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 462 | using \<open>openin Y V\<close> cmf openin_continuous_map_preimage by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 463 |     show "disjnt {x \<in> topspace X. f x \<in> U} {x \<in> topspace X. f x \<in> V}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 464 | using \<open>disjnt U V\<close> by (auto simp add: disjnt_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 465 | qed (use x \<open>f x \<in> U\<close> y \<open>f y \<in> V\<close> in auto) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 466 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 467 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 468 | lemma homeomorphic_Hausdorff_space: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 469 | "X homeomorphic_space Y \<Longrightarrow> Hausdorff_space X \<longleftrightarrow> Hausdorff_space Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 470 | unfolding homeomorphic_space_def homeomorphic_maps_map | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 471 | by (auto simp: homeomorphic_eq_everything_map Hausdorff_space_injective_preimage) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 472 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 473 | lemma Hausdorff_space_retraction_map_image: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 474 | "\<lbrakk>retraction_map X Y r; Hausdorff_space X\<rbrakk> \<Longrightarrow> Hausdorff_space Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 475 | unfolding retraction_map_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 476 | using Hausdorff_space_subtopology homeomorphic_Hausdorff_space retraction_maps_section_image2 by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 477 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 478 | lemma compact_Hausdorff_space_optimal: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 479 | assumes eq: "topspace Y = topspace X" and XY: "\<And>U. openin X U \<Longrightarrow> openin Y U" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 480 | and "Hausdorff_space X" "compact_space Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 481 | shows "Y = X" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 482 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 483 | have "\<And>U. closedin X U \<Longrightarrow> closedin Y U" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 484 | using XY using topology_finer_closedin [OF eq] | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 485 | by metis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 486 | have "openin Y S = openin X S" for S | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 487 | by (metis XY assms(3) assms(4) closedin_compact_space compactin_contractive compactin_imp_closedin eq openin_closedin_eq) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 488 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 489 | by (simp add: topology_eq) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 490 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 491 | |
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 492 | lemma continuous_map_imp_closed_graph: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 493 | assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 494 | shows "closedin (prod_topology X Y) ((\<lambda>x. (x,f x)) ` topspace X)" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 495 | unfolding closedin_def | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 496 | proof | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 497 | show "(\<lambda>x. (x, f x)) ` topspace X \<subseteq> topspace (prod_topology X Y)" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 498 | using continuous_map_def f by fastforce | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 499 | show "openin (prod_topology X Y) (topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X)" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 500 | unfolding openin_prod_topology_alt | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 501 | proof (intro allI impI) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 502 | show "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 503 | if "(x,y) \<in> topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 504 | for x y | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 505 | proof - | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 506 | have "x \<in> topspace X" and y: "y \<in> topspace Y" "y \<noteq> f x" | 
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 507 | using that by auto | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 508 | then have "f x \<in> topspace Y" | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 509 | using continuous_map_image_subset_topspace f by blast | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 510 | then obtain U V where UV: "openin Y U" "openin Y V" "f x \<in> U" "y \<in> V" "disjnt U V" | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 511 | using Y y Hausdorff_space_def by metis | 
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 512 | show ?thesis | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 513 | proof (intro exI conjI) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 514 |         show "openin X {x \<in> topspace X. f x \<in> U}"
 | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 515 | using \<open>openin Y U\<close> f openin_continuous_map_preimage by blast | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 516 |         show "{x \<in> topspace X. f x \<in> U} \<times> V \<subseteq> topspace (prod_topology X Y) - (\<lambda>x. (x, f x)) ` topspace X"
 | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 517 | using UV by (auto simp: disjnt_iff dest: openin_subset) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 518 | qed (use UV \<open>x \<in> topspace X\<close> in auto) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 519 | qed | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 520 | qed | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 521 | qed | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 522 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 523 | lemma continuous_imp_closed_map: | 
| 70196 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 524 | "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y\<rbrakk> \<Longrightarrow> closed_map X Y f" | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 525 | by (meson closed_map_def closedin_compact_space compactin_imp_closedin image_compactin) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 526 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 527 | lemma continuous_imp_quotient_map: | 
| 70196 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 528 | "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y; f ` (topspace X) = topspace Y\<rbrakk> | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 529 | \<Longrightarrow> quotient_map X Y f" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 530 | by (simp add: continuous_imp_closed_map continuous_closed_imp_quotient_map) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 531 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 532 | lemma continuous_imp_homeomorphic_map: | 
| 70196 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 533 | "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y; | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 534 | f ` (topspace X) = topspace Y; inj_on f (topspace X)\<rbrakk> | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 535 | \<Longrightarrow> homeomorphic_map X Y f" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 536 | by (simp add: continuous_imp_closed_map bijective_closed_imp_homeomorphic_map) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 537 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 538 | lemma continuous_imp_embedding_map: | 
| 70196 
b7ef9090feed
Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
 paulson <lp15@cam.ac.uk> parents: 
70178diff
changeset | 539 | "\<lbrakk>continuous_map X Y f; compact_space X; Hausdorff_space Y; inj_on f (topspace X)\<rbrakk> | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 540 | \<Longrightarrow> embedding_map X Y f" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 541 | by (simp add: continuous_imp_closed_map injective_closed_imp_embedding_map) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 542 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 543 | lemma continuous_inverse_map: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 544 | assumes "compact_space X" "Hausdorff_space Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 545 | and cmf: "continuous_map X Y f" and gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g(f x) = x" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 546 | and Sf: "S \<subseteq> f ` (topspace X)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 547 | shows "continuous_map (subtopology Y S) X g" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 548 | proof (rule continuous_map_from_subtopology_mono [OF _ \<open>S \<subseteq> f ` (topspace X)\<close>]) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 549 | show "continuous_map (subtopology Y (f ` (topspace X))) X g" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 550 | unfolding continuous_map_closedin | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 551 | proof (intro conjI ballI allI impI) | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 552 | show "g \<in> topspace (subtopology Y (f ` topspace X)) \<rightarrow> topspace X" | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 553 | using gf by auto | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 554 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 555 | fix C | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 556 | assume C: "closedin X C" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 557 | show "closedin (subtopology Y (f ` topspace X)) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 558 |            {x \<in> topspace (subtopology Y (f ` topspace X)). g x \<in> C}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 559 | proof (rule compactin_imp_closedin) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 560 | show "Hausdorff_space (subtopology Y (f ` topspace X))" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 561 | using Hausdorff_space_subtopology [OF \<open>Hausdorff_space Y\<close>] by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 562 | have "compactin Y (f ` C)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 563 | using C cmf image_compactin closedin_compact_space [OF \<open>compact_space X\<close>] by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 564 |       moreover have "{x \<in> topspace Y. x \<in> f ` topspace X \<and> g x \<in> C} = f ` C"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 565 | using closedin_subset [OF C] cmf by (auto simp: gf continuous_map_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 566 |       ultimately have "compactin Y {x \<in> topspace Y. x \<in> f ` topspace X \<and> g x \<in> C}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 567 | by simp | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 568 | then show "compactin (subtopology Y (f ` topspace X)) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 569 |               {x \<in> topspace (subtopology Y (f ` topspace X)). g x \<in> C}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 570 | by (auto simp add: compactin_subtopology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 571 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 572 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 573 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 574 | |
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 575 | lemma closed_map_paired_continuous_map_right: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 576 | "\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk> \<Longrightarrow> closed_map X (prod_topology X Y) (\<lambda>x. (x,f x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 577 | by (simp add: continuous_map_imp_closed_graph embedding_map_graph embedding_imp_closed_map) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 578 | |
| 70178 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 579 | lemma closed_map_paired_continuous_map_left: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 580 | assumes f: "continuous_map X Y f" and Y: "Hausdorff_space Y" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 581 | shows "closed_map X (prod_topology Y X) (\<lambda>x. (f x,x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 582 | proof - | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 583 | have eq: "(\<lambda>x. (f x,x)) = (\<lambda>(a,b). (b,a)) \<circ> (\<lambda>x. (x,f x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 584 | by auto | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 585 | show ?thesis | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 586 | unfolding eq | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 587 | proof (rule closed_map_compose) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 588 | show "closed_map X (prod_topology X Y) (\<lambda>x. (x, f x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 589 | using Y closed_map_paired_continuous_map_right f by blast | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 590 | show "closed_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(a, b). (b, a))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 591 | by (metis homeomorphic_map_swap homeomorphic_imp_closed_map) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 592 | qed | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 593 | qed | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 594 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 595 | lemma proper_map_paired_continuous_map_right: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 596 | "\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk> | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 597 | \<Longrightarrow> proper_map X (prod_topology X Y) (\<lambda>x. (x,f x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 598 | using closed_injective_imp_proper_map closed_map_paired_continuous_map_right | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 599 | by (metis (mono_tags, lifting) Pair_inject inj_onI) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 600 | |
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 601 | lemma proper_map_paired_continuous_map_left: | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 602 | "\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk> | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 603 | \<Longrightarrow> proper_map X (prod_topology Y X) (\<lambda>x. (f x,x))" | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 604 | using closed_injective_imp_proper_map closed_map_paired_continuous_map_left | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 605 | by (metis (mono_tags, lifting) Pair_inject inj_onI) | 
| 
4900351361b0
Lindelöf spaces and supporting material
 paulson <lp15@cam.ac.uk> parents: 
70019diff
changeset | 606 | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 607 | lemma Hausdorff_space_prod_topology: | 
| 78336 | 608 | "Hausdorff_space(prod_topology X Y) \<longleftrightarrow> (prod_topology X Y) = trivial_topology \<or> Hausdorff_space X \<and> Hausdorff_space Y" | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 609 | (is "?lhs = ?rhs") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 610 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 611 | assume ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 612 | then show ?rhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 613 | by (rule topological_property_of_prod_component) (auto simp: Hausdorff_space_subtopology homeomorphic_Hausdorff_space) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 614 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 615 | assume R: ?rhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 616 | show ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 617 |   proof (cases "(topspace X \<times> topspace Y) = {}")
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 618 | case False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 619 |     with R have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}" and X: "Hausdorff_space X" and Y: "Hausdorff_space Y"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 620 | by auto | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 621 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 622 | unfolding Hausdorff_space_def | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 623 | proof clarify | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 624 | fix x y x' y' | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 625 | assume xy: "(x, y) \<in> topspace (prod_topology X Y)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 626 | and xy': "(x',y') \<in> topspace (prod_topology X Y)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 627 | and *: "\<nexists>U V. openin (prod_topology X Y) U \<and> openin (prod_topology X Y) V | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 628 | \<and> (x, y) \<in> U \<and> (x', y') \<in> V \<and> disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 629 | have False if "x \<noteq> x' \<or> y \<noteq> y'" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 630 | using that | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 631 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 632 | assume "x \<noteq> x'" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 633 | then obtain U V where "openin X U" "openin X V" "x \<in> U" "x' \<in> V" "disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 634 | by (metis Hausdorff_space_def X mem_Sigma_iff topspace_prod_topology xy xy') | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 635 | let ?U = "U \<times> topspace Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 636 | let ?V = "V \<times> topspace Y" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 637 | have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 638 | by (simp_all add: openin_prod_Times_iff \<open>openin X U\<close> \<open>openin X V\<close>) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 639 | moreover have "disjnt ?U ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 640 | by (simp add: \<open>disjnt U V\<close>) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 641 | ultimately show False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 642 | using * \<open>x \<in> U\<close> \<open>x' \<in> V\<close> xy xy' by (metis SigmaD2 SigmaI topspace_prod_topology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 643 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 644 | assume "y \<noteq> y'" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 645 | then obtain U V where "openin Y U" "openin Y V" "y \<in> U" "y' \<in> V" "disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 646 | by (metis Hausdorff_space_def Y mem_Sigma_iff topspace_prod_topology xy xy') | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 647 | let ?U = "topspace X \<times> U" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 648 | let ?V = "topspace X \<times> V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 649 | have "openin (prod_topology X Y) ?U" "openin (prod_topology X Y) ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 650 | by (simp_all add: openin_prod_Times_iff \<open>openin Y U\<close> \<open>openin Y V\<close>) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 651 | moreover have "disjnt ?U ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 652 | by (simp add: \<open>disjnt U V\<close>) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 653 | ultimately show False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 654 | using "*" \<open>y \<in> U\<close> \<open>y' \<in> V\<close> xy xy' by (metis SigmaD1 SigmaI topspace_prod_topology) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 655 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 656 | then show "x = x' \<and> y = y'" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 657 | by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 658 | qed | 
| 78336 | 659 | qed force | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 660 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 661 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 662 | |
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 663 | lemma Hausdorff_space_product_topology: | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 664 |    "Hausdorff_space (product_topology X I) \<longleftrightarrow> (\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {} \<or> (\<forall>i \<in> I. Hausdorff_space (X i))"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 665 | (is "?lhs = ?rhs") | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 666 | proof | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 667 | assume ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 668 | then show ?rhs | 
| 78336 | 669 | by (simp add: Hausdorff_space_subtopology PiE_eq_empty_iff homeomorphic_Hausdorff_space | 
| 670 | topological_property_of_product_component) | |
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 671 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 672 | assume R: ?rhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 673 | show ?lhs | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 674 |   proof (cases "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) = {}")
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 675 | case True | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 676 | then show ?thesis | 
| 78336 | 677 | by (simp add: Hausdorff_space_def) | 
| 69994 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 678 | next | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 679 | case False | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 680 | have "\<exists>U V. openin (product_topology X I) U \<and> openin (product_topology X I) V \<and> f \<in> U \<and> g \<in> V \<and> disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 681 | if f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" and g: "g \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" and "f \<noteq> g" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 682 | for f g :: "'a \<Rightarrow> 'b" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 683 | proof - | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 684 | obtain m where "f m \<noteq> g m" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 685 | using \<open>f \<noteq> g\<close> by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 686 | then have "m \<in> I" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 687 | using f g by fastforce | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 688 | then have "Hausdorff_space (X m)" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 689 | using False that R by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 690 | then obtain U V where U: "openin (X m) U" and V: "openin (X m) V" and "f m \<in> U" "g m \<in> V" "disjnt U V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 691 | by (metis Hausdorff_space_def PiE_mem \<open>f m \<noteq> g m\<close> \<open>m \<in> I\<close> f g) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 692 | show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 693 | proof (intro exI conjI) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 694 |         let ?U = "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<inter> {x. x m \<in> U}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 695 |         let ?V = "(\<Pi>\<^sub>E i\<in>I. topspace(X i)) \<inter> {x. x m \<in> V}"
 | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 696 | show "openin (product_topology X I) ?U" "openin (product_topology X I) ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 697 | using \<open>m \<in> I\<close> U V | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 698 | by (force simp add: openin_product_topology intro: arbitrary_union_of_inc relative_to_inc finite_intersection_of_inc)+ | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 699 | show "f \<in> ?U" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 700 | using \<open>f m \<in> U\<close> f by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 701 | show "g \<in> ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 702 | using \<open>g m \<in> V\<close> g by blast | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 703 | show "disjnt ?U ?V" | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 704 | using \<open>disjnt U V\<close> by (auto simp: PiE_def Pi_def disjnt_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 705 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 706 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 707 | then show ?thesis | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 708 | by (simp add: Hausdorff_space_def) | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 709 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 710 | qed | 
| 
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
 paulson <lp15@cam.ac.uk> parents: 
69986diff
changeset | 711 | |
| 77935 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 712 | lemma Hausdorff_space_closed_neighbourhood: | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 713 | "Hausdorff_space X \<longleftrightarrow> | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 714 | (\<forall>x \<in> topspace X. \<exists>U C. openin X U \<and> closedin X C \<and> | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 715 | Hausdorff_space(subtopology X C) \<and> x \<in> U \<and> U \<subseteq> C)" (is "_ = ?rhs") | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 716 | proof | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 717 | assume R: ?rhs | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 718 | show "Hausdorff_space X" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 719 | unfolding Hausdorff_space_def | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 720 | proof clarify | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 721 | fix x y | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 722 | assume x: "x \<in> topspace X" and y: "y \<in> topspace X" and "x \<noteq> y" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 723 | obtain T C where *: "openin X T" "closedin X C" "x \<in> T" "T \<subseteq> C" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 724 | and C: "Hausdorff_space (subtopology X C)" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 725 | by (meson R \<open>x \<in> topspace X\<close>) | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 726 | show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 727 | proof (cases "y \<in> C") | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 728 | case True | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 729 | with * C obtain U V where U: "openin (subtopology X C) U" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 730 | and V: "openin (subtopology X C) V" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 731 | and "x \<in> U" "y \<in> V" "disjnt U V" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 732 | unfolding Hausdorff_space_def | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 733 | by (smt (verit, best) \<open>x \<noteq> y\<close> closedin_subset subsetD topspace_subtopology_subset) | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 734 | then obtain U' V' where UV': "U = U' \<inter> C" "openin X U'" "V = V' \<inter> C" "openin X V'" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 735 | by (meson openin_subtopology) | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 736 | have "disjnt (T \<inter> U') V'" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 737 | using \<open>disjnt U V\<close> UV' \<open>T \<subseteq> C\<close> by (force simp: disjnt_iff) | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 738 | with \<open>T \<subseteq> C\<close> have "disjnt (T \<inter> U') (V' \<union> (topspace X - C))" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 739 | unfolding disjnt_def by blast | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 740 | moreover | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 741 | have "openin X (T \<inter> U')" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 742 | by (simp add: \<open>openin X T\<close> \<open>openin X U'\<close> openin_Int) | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 743 | moreover have "openin X (V' \<union> (topspace X - C))" | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 744 | using \<open>closedin X C\<close> \<open>openin X V'\<close> by auto | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 745 | ultimately show ?thesis | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 746 | using UV' \<open>x \<in> T\<close> \<open>x \<in> U\<close> \<open>y \<in> V\<close> by blast | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 747 | next | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 748 | case False | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 749 | with * y show ?thesis | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 750 | by (force simp: closedin_def disjnt_def) | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 751 | qed | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 752 | qed | 
| 
7f240b0dabd9
More new theorems, and a necessary correction
 paulson <lp15@cam.ac.uk> parents: 
75455diff
changeset | 753 | qed fastforce | 
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77935diff
changeset | 754 | |
| 69874 | 755 | end |