author | haftmann |
Mon, 13 Sep 2021 14:18:24 +0000 | |
changeset 74309 | 42523fbf643b |
parent 71200 | 3548d54ce3ee |
child 75455 | 91c16c5ad3e9 |
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:
69986
diff
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:
70196
diff
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:
69986
diff
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:
69986
diff
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 |
||
20 |
lemma t1_space_empty: "topspace X = {} \<Longrightarrow> t1_space X" |
|
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 |
||
52 |
lemma closedin_t1_singleton: |
|
53 |
"\<lbrakk>t1_space X; a \<in> topspace X\<rbrakk> \<Longrightarrow> closedin X {a}" |
|
54 |
by (simp add: t1_space_closedin_singleton) |
|
55 |
||
56 |
lemma t1_space_closedin_finite: |
|
57 |
"t1_space X \<longleftrightarrow> (\<forall>S. finite S \<and> S \<subseteq> topspace X \<longrightarrow> closedin X S)" |
|
58 |
apply (rule iffI) |
|
59 |
apply (simp add: closedin_contains_derived_set t1_space_derived_set_of_finite) |
|
60 |
by (simp add: t1_space_closedin_singleton) |
|
61 |
||
62 |
lemma closure_of_singleton: |
|
63 |
"t1_space X \<Longrightarrow> X closure_of {a} = (if a \<in> topspace X then {a} else {})" |
|
64 |
by (simp add: closure_of_eq t1_space_closedin_singleton closure_of_eq_empty_gen) |
|
65 |
||
66 |
lemma separated_in_singleton: |
|
67 |
assumes "t1_space X" |
|
68 |
shows "separatedin X {a} S \<longleftrightarrow> a \<in> topspace X \<and> S \<subseteq> topspace X \<and> (a \<notin> X closure_of S)" |
|
69 |
"separatedin X S {a} \<longleftrightarrow> a \<in> topspace X \<and> S \<subseteq> topspace X \<and> (a \<notin> X closure_of S)" |
|
70 |
unfolding separatedin_def |
|
71 |
using assms closure_of closure_of_singleton by fastforce+ |
|
72 |
||
73 |
lemma t1_space_openin_delete: |
|
74 |
"t1_space X \<longleftrightarrow> (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (U - {x}))" |
|
75 |
apply (rule iffI) |
|
76 |
apply (meson closedin_t1_singleton in_mono openin_diff openin_subset) |
|
77 |
by (simp add: closedin_def t1_space_closedin_singleton) |
|
78 |
||
79 |
lemma t1_space_openin_delete_alt: |
|
80 |
"t1_space X \<longleftrightarrow> (\<forall>U x. openin X U \<longrightarrow> openin X (U - {x}))" |
|
81 |
by (metis Diff_empty Diff_insert0 t1_space_openin_delete) |
|
82 |
||
83 |
||
84 |
lemma t1_space_singleton_Inter_open: |
|
85 |
"t1_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<Inter>{U. openin X U \<and> x \<in> U} = {x})" (is "?P=?Q") |
|
86 |
and t1_space_Inter_open_supersets: |
|
87 |
"t1_space X \<longleftrightarrow> (\<forall>S. S \<subseteq> topspace X \<longrightarrow> \<Inter>{U. openin X U \<and> S \<subseteq> U} = S)" (is "?P=?R") |
|
88 |
proof - |
|
89 |
have "?R \<Longrightarrow> ?Q" |
|
90 |
apply clarify |
|
91 |
apply (drule_tac x="{x}" in spec, simp) |
|
92 |
done |
|
93 |
moreover have "?Q \<Longrightarrow> ?P" |
|
94 |
apply (clarsimp simp add: t1_space_def) |
|
95 |
apply (drule_tac x=x in bspec) |
|
96 |
apply (simp_all add: set_eq_iff) |
|
97 |
by (metis (no_types, lifting)) |
|
98 |
moreover have "?P \<Longrightarrow> ?R" |
|
99 |
proof (clarsimp simp add: t1_space_closedin_singleton, rule subset_antisym) |
|
100 |
fix S |
|
101 |
assume S: "\<forall>x\<in>topspace X. closedin X {x}" "S \<subseteq> topspace X" |
|
102 |
then show "\<Inter> {U. openin X U \<and> S \<subseteq> U} \<subseteq> S" |
|
103 |
apply clarsimp |
|
104 |
by (metis Diff_insert_absorb Set.set_insert closedin_def openin_topspace subset_insert) |
|
105 |
qed force |
|
106 |
ultimately show "?P=?Q" "?P=?R" |
|
107 |
by auto |
|
108 |
qed |
|
109 |
||
110 |
lemma t1_space_derived_set_of_infinite_openin: |
|
111 |
"t1_space X \<longleftrightarrow> |
|
112 |
(\<forall>S. X derived_set_of S = |
|
113 |
{x \<in> topspace X. \<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite(S \<inter> U)})" |
|
114 |
(is "_ = ?rhs") |
|
115 |
proof |
|
116 |
assume "t1_space X" |
|
117 |
show ?rhs |
|
118 |
proof safe |
|
119 |
fix S x U |
|
120 |
assume "x \<in> X derived_set_of S" "x \<in> U" "openin X U" "finite (S \<inter> U)" |
|
121 |
with \<open>t1_space X\<close> show "False" |
|
122 |
apply (simp add: t1_space_derived_set_of_finite) |
|
123 |
by (metis IntI empty_iff empty_subsetI inf_commute openin_Int_derived_set_of_subset subset_antisym) |
|
124 |
next |
|
125 |
fix S x |
|
126 |
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 |
|
127 |
by blast |
|
128 |
assume "x \<in> topspace X" "\<forall>U. x \<in> U \<and> openin X U \<longrightarrow> infinite (S \<inter> U)" |
|
129 |
then show "x \<in> X derived_set_of S" |
|
130 |
apply (clarsimp simp add: derived_set_of_def eq) |
|
131 |
by (meson finite.emptyI finite.insertI finite_subset) |
|
132 |
qed (auto simp: in_derived_set_of) |
|
133 |
qed (auto simp: t1_space_derived_set_of_singleton) |
|
134 |
||
135 |
lemma finite_t1_space_imp_discrete_topology: |
|
136 |
"\<lbrakk>topspace X = U; finite U; t1_space X\<rbrakk> \<Longrightarrow> X = discrete_topology U" |
|
137 |
by (metis discrete_topology_unique_derived_set t1_space_derived_set_of_finite) |
|
138 |
||
139 |
lemma t1_space_subtopology: "t1_space X \<Longrightarrow> t1_space(subtopology X U)" |
|
140 |
by (simp add: derived_set_of_subtopology t1_space_derived_set_of_finite) |
|
141 |
||
142 |
lemma closedin_derived_set_of_gen: |
|
143 |
"t1_space X \<Longrightarrow> closedin X (X derived_set_of S)" |
|
144 |
apply (clarsimp simp add: in_derived_set_of closedin_contains_derived_set derived_set_of_subset_topspace) |
|
145 |
by (metis DiffD2 insert_Diff insert_iff t1_space_openin_delete) |
|
146 |
||
147 |
lemma derived_set_of_derived_set_subset_gen: |
|
148 |
"t1_space X \<Longrightarrow> X derived_set_of (X derived_set_of S) \<subseteq> X derived_set_of S" |
|
149 |
by (meson closedin_contains_derived_set closedin_derived_set_of_gen) |
|
150 |
||
151 |
lemma subtopology_eq_discrete_topology_gen_finite: |
|
152 |
"\<lbrakk>t1_space X; finite S\<rbrakk> \<Longrightarrow> subtopology X S = discrete_topology(topspace X \<inter> S)" |
|
153 |
by (simp add: subtopology_eq_discrete_topology_gen t1_space_derived_set_of_finite) |
|
154 |
||
155 |
lemma subtopology_eq_discrete_topology_finite: |
|
156 |
"\<lbrakk>t1_space X; S \<subseteq> topspace X; finite S\<rbrakk> |
|
157 |
\<Longrightarrow> subtopology X S = discrete_topology S" |
|
158 |
by (simp add: subtopology_eq_discrete_topology_eq t1_space_derived_set_of_finite) |
|
159 |
||
160 |
lemma t1_space_closed_map_image: |
|
161 |
"\<lbrakk>closed_map X Y f; f ` (topspace X) = topspace Y; t1_space X\<rbrakk> \<Longrightarrow> t1_space Y" |
|
162 |
by (metis closed_map_def finite_subset_image t1_space_closedin_finite) |
|
163 |
||
164 |
lemma homeomorphic_t1_space: "X homeomorphic_space Y \<Longrightarrow> (t1_space X \<longleftrightarrow> t1_space Y)" |
|
165 |
apply (clarsimp simp add: homeomorphic_space_def) |
|
166 |
by (meson homeomorphic_eq_everything_map homeomorphic_maps_map t1_space_closed_map_image) |
|
167 |
||
168 |
proposition t1_space_product_topology: |
|
169 |
"t1_space (product_topology X I) |
|
170 |
\<longleftrightarrow> topspace(product_topology X I) = {} \<or> (\<forall>i \<in> I. t1_space (X i))" |
|
171 |
proof (cases "topspace(product_topology X I) = {}") |
|
172 |
case True |
|
173 |
then show ?thesis |
|
174 |
using True t1_space_empty by blast |
|
175 |
next |
|
176 |
case False |
|
177 |
then obtain f where f: "f \<in> (\<Pi>\<^sub>E i\<in>I. topspace(X i))" |
|
178 |
by fastforce |
|
179 |
have "t1_space (product_topology X I) \<longleftrightarrow> (\<forall>i\<in>I. t1_space (X i))" |
|
180 |
proof (intro iffI ballI) |
|
181 |
show "t1_space (X i)" if "t1_space (product_topology X I)" and "i \<in> I" for i |
|
182 |
proof - |
|
183 |
have clo: "\<And>h. h \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<Longrightarrow> closedin (product_topology X I) {h}" |
|
184 |
using that by (simp add: t1_space_closedin_singleton) |
|
185 |
show ?thesis |
|
186 |
unfolding t1_space_closedin_singleton |
|
187 |
proof clarify |
|
188 |
show "closedin (X i) {xi}" if "xi \<in> topspace (X i)" for xi |
|
189 |
using clo [of "\<lambda>j \<in> I. if i=j then xi else f j"] f that \<open>i \<in> I\<close> |
|
190 |
by (fastforce simp add: closedin_product_topology_singleton) |
|
191 |
qed |
|
192 |
qed |
|
193 |
next |
|
194 |
next |
|
195 |
show "t1_space (product_topology X I)" if "\<forall>i\<in>I. t1_space (X i)" |
|
196 |
using that |
|
197 |
by (simp add: t1_space_closedin_singleton Ball_def PiE_iff closedin_product_topology_singleton) |
|
198 |
qed |
|
199 |
then show ?thesis |
|
200 |
using False by blast |
|
201 |
qed |
|
202 |
||
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
203 |
lemma t1_space_prod_topology: |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
204 |
"t1_space(prod_topology X Y) \<longleftrightarrow> topspace(prod_topology X Y) = {} \<or> t1_space X \<and> t1_space Y" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
205 |
proof (cases "topspace (prod_topology X Y) = {}") |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
206 |
case True then show ?thesis |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
207 |
by (auto simp: t1_space_empty) |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
208 |
next |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
209 |
case False |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
210 |
have eq: "{(x,y)} = {x} \<times> {y}" for x y |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
211 |
by simp |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
212 |
have "t1_space (prod_topology X Y) \<longleftrightarrow> (t1_space X \<and> t1_space Y)" |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
213 |
using False |
69986
f2d327275065
generalised homotopic_with to topologies; homotopic_with_canon is the old version
paulson <lp15@cam.ac.uk>
parents:
69939
diff
changeset
|
214 |
by (force simp: t1_space_closedin_singleton closedin_prod_Times_iff eq simp del: insert_Times_insert) |
69939
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
215 |
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:
69874
diff
changeset
|
216 |
by simp |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
217 |
qed |
812ce526da33
new material on topology: products, etc. Some renamings, esp continuous_on_topo -> continuous_map
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
218 |
|
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
219 |
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:
69986
diff
changeset
|
220 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
221 |
definition Hausdorff_space |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
222 |
where |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
223 |
"Hausdorff_space X \<equiv> |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
224 |
\<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:
69986
diff
changeset
|
225 |
\<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:
69986
diff
changeset
|
226 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
227 |
lemma Hausdorff_space_expansive: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
228 |
"\<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:
69986
diff
changeset
|
229 |
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:
69986
diff
changeset
|
230 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
231 |
lemma Hausdorff_space_topspace_empty: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
232 |
"topspace X = {} \<Longrightarrow> Hausdorff_space X" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
233 |
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:
69986
diff
changeset
|
234 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
235 |
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:
69986
diff
changeset
|
236 |
"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:
69986
diff
changeset
|
237 |
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:
69986
diff
changeset
|
238 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
239 |
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:
69986
diff
changeset
|
240 |
"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:
69986
diff
changeset
|
241 |
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:
69986
diff
changeset
|
242 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
243 |
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:
69986
diff
changeset
|
244 |
"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:
69986
diff
changeset
|
245 |
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:
69986
diff
changeset
|
246 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
247 |
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:
69986
diff
changeset
|
248 |
"\<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:
69986
diff
changeset
|
249 |
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:
69986
diff
changeset
|
250 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
251 |
lemma Hausdorff_space_subtopology: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
252 |
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:
69986
diff
changeset
|
253 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
254 |
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:
69986
diff
changeset
|
255 |
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:
69986
diff
changeset
|
256 |
from assms show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
257 |
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:
69986
diff
changeset
|
258 |
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:
69986
diff
changeset
|
259 |
done |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
260 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
261 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
262 |
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:
69986
diff
changeset
|
263 |
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:
69986
diff
changeset
|
264 |
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:
69986
diff
changeset
|
265 |
proof (cases "S = {}") |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
266 |
case True |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
267 |
then show thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
268 |
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:
69986
diff
changeset
|
269 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
270 |
case False |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
271 |
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:
69986
diff
changeset
|
272 |
proof |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
273 |
fix a |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
274 |
assume "a \<in> S" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
275 |
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:
69986
diff
changeset
|
276 |
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:
69986
diff
changeset
|
277 |
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:
69986
diff
changeset
|
278 |
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:
69986
diff
changeset
|
279 |
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:
69986
diff
changeset
|
280 |
proof (cases "T = {}") |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
281 |
case True |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
282 |
then show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
283 |
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:
69986
diff
changeset
|
284 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
285 |
case False |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
286 |
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:
69986
diff
changeset
|
287 |
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:
69986
diff
changeset
|
288 |
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:
69986
diff
changeset
|
289 |
by metis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
290 |
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:
69986
diff
changeset
|
291 |
have Topen: "\<forall>W \<in> U ` T. openin X W" and Tsub: "T \<subseteq> \<Union> (U ` T)" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
292 |
by (auto simp: ) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
293 |
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:
69986
diff
changeset
|
294 |
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:
69986
diff
changeset
|
295 |
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:
69986
diff
changeset
|
296 |
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:
69986
diff
changeset
|
297 |
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:
69986
diff
changeset
|
298 |
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:
69986
diff
changeset
|
299 |
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:
69986
diff
changeset
|
300 |
using UV by blast+ |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
301 |
show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
302 |
proof (intro exI conjI) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
303 |
have "F \<noteq> {}" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
304 |
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:
69986
diff
changeset
|
305 |
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:
69986
diff
changeset
|
306 |
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:
69986
diff
changeset
|
307 |
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:
69986
diff
changeset
|
308 |
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:
69986
diff
changeset
|
309 |
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:
69986
diff
changeset
|
310 |
using disj |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
311 |
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:
69986
diff
changeset
|
312 |
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:
69986
diff
changeset
|
313 |
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:
69986
diff
changeset
|
314 |
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:
69986
diff
changeset
|
315 |
qed (auto simp: SUF) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
316 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
317 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
318 |
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:
69986
diff
changeset
|
319 |
by metis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
320 |
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:
69986
diff
changeset
|
321 |
by auto |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
322 |
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:
69986
diff
changeset
|
323 |
using UV by blast |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
324 |
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:
69986
diff
changeset
|
325 |
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:
69986
diff
changeset
|
326 |
show thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
327 |
proof |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
328 |
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:
69986
diff
changeset
|
329 |
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:
69986
diff
changeset
|
330 |
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:
69986
diff
changeset
|
331 |
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:
69986
diff
changeset
|
332 |
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:
69986
diff
changeset
|
333 |
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:
69986
diff
changeset
|
334 |
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:
69986
diff
changeset
|
335 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
336 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
337 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
338 |
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:
69986
diff
changeset
|
339 |
"Hausdorff_space X \<longleftrightarrow> |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
340 |
(\<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:
69986
diff
changeset
|
341 |
\<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:
69986
diff
changeset
|
342 |
(is "?lhs = ?rhs") |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
343 |
proof |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
344 |
assume ?lhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
345 |
then show ?rhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
346 |
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:
69986
diff
changeset
|
347 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
348 |
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:
69986
diff
changeset
|
349 |
show ?lhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
350 |
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:
69986
diff
changeset
|
351 |
fix x y |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
352 |
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:
69986
diff
changeset
|
353 |
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:
69986
diff
changeset
|
354 |
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:
69986
diff
changeset
|
355 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
356 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
357 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
358 |
lemma compactin_imp_closedin: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
359 |
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:
69986
diff
changeset
|
360 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
361 |
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:
69986
diff
changeset
|
362 |
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:
69986
diff
changeset
|
363 |
moreover |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
364 |
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:
69986
diff
changeset
|
365 |
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:
69986
diff
changeset
|
366 |
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:
69986
diff
changeset
|
367 |
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:
69986
diff
changeset
|
368 |
ultimately show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
369 |
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:
69986
diff
changeset
|
370 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
371 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
372 |
lemma closedin_Hausdorff_singleton: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
373 |
"\<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:
69986
diff
changeset
|
374 |
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:
69986
diff
changeset
|
375 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
376 |
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:
69986
diff
changeset
|
377 |
"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:
69986
diff
changeset
|
378 |
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:
69986
diff
changeset
|
379 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
380 |
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:
69986
diff
changeset
|
381 |
"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:
69986
diff
changeset
|
382 |
unfolding Hausdorff_space_def |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
383 |
apply safe |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
384 |
by (metis discrete_topology_unique_alt disjnt_empty2 disjnt_insert2 insert_iff mk_disjoint_insert topspace_discrete_topology) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
385 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
386 |
lemma compactin_Int: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
387 |
"\<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:
69986
diff
changeset
|
388 |
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:
69986
diff
changeset
|
389 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
390 |
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:
69986
diff
changeset
|
391 |
"\<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:
69986
diff
changeset
|
392 |
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:
69986
diff
changeset
|
393 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
394 |
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:
69986
diff
changeset
|
395 |
"\<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:
69986
diff
changeset
|
396 |
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:
69986
diff
changeset
|
397 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
398 |
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:
69986
diff
changeset
|
399 |
"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:
69986
diff
changeset
|
400 |
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:
69986
diff
changeset
|
401 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
402 |
lemma closedin_Hausdorff_finite: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
403 |
"\<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:
69986
diff
changeset
|
404 |
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:
69986
diff
changeset
|
405 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
406 |
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:
69986
diff
changeset
|
407 |
"\<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:
69986
diff
changeset
|
408 |
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:
69986
diff
changeset
|
409 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
410 |
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:
69986
diff
changeset
|
411 |
"\<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:
69986
diff
changeset
|
412 |
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:
69986
diff
changeset
|
413 |
|
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69994
diff
changeset
|
414 |
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:
69986
diff
changeset
|
415 |
"Hausdorff_space X |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
416 |
\<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:
69986
diff
changeset
|
417 |
{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:
69986
diff
changeset
|
418 |
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:
69986
diff
changeset
|
419 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
420 |
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:
69986
diff
changeset
|
421 |
"Hausdorff_space X |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
422 |
\<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:
69986
diff
changeset
|
423 |
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:
69986
diff
changeset
|
424 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
425 |
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:
69986
diff
changeset
|
426 |
"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:
69986
diff
changeset
|
427 |
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:
69986
diff
changeset
|
428 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
429 |
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:
69986
diff
changeset
|
430 |
"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:
69986
diff
changeset
|
431 |
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:
69986
diff
changeset
|
432 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
433 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
434 |
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:
69986
diff
changeset
|
435 |
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:
69986
diff
changeset
|
436 |
shows "Hausdorff_space X" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
437 |
unfolding Hausdorff_space_def |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
438 |
proof clarify |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
439 |
fix x y |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
440 |
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:
69986
diff
changeset
|
441 |
then obtain U V where "openin Y U" "openin Y V" "f x \<in> U" "f 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:
69986
diff
changeset
|
442 |
using assms unfolding Hausdorff_space_def continuous_map_def by (meson inj_onD) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
443 |
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:
69986
diff
changeset
|
444 |
proof (intro exI conjI) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
445 |
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:
69986
diff
changeset
|
446 |
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:
69986
diff
changeset
|
447 |
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:
69986
diff
changeset
|
448 |
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:
69986
diff
changeset
|
449 |
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:
69986
diff
changeset
|
450 |
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:
69986
diff
changeset
|
451 |
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:
69986
diff
changeset
|
452 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
453 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
454 |
lemma homeomorphic_Hausdorff_space: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
455 |
"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:
69986
diff
changeset
|
456 |
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:
69986
diff
changeset
|
457 |
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:
69986
diff
changeset
|
458 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
459 |
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:
69986
diff
changeset
|
460 |
"\<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:
69986
diff
changeset
|
461 |
unfolding retraction_map_def |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
462 |
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:
69986
diff
changeset
|
463 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
464 |
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:
69986
diff
changeset
|
465 |
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:
69986
diff
changeset
|
466 |
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:
69986
diff
changeset
|
467 |
shows "Y = X" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
468 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
469 |
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:
69986
diff
changeset
|
470 |
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:
69986
diff
changeset
|
471 |
by metis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
472 |
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:
69986
diff
changeset
|
473 |
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:
69986
diff
changeset
|
474 |
then show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
475 |
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:
69986
diff
changeset
|
476 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
477 |
|
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
478 |
lemma continuous_map_imp_closed_graph: |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
479 |
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:
70019
diff
changeset
|
480 |
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:
70019
diff
changeset
|
481 |
unfolding closedin_def |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
482 |
proof |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
483 |
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:
70019
diff
changeset
|
484 |
using continuous_map_def f by fastforce |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
485 |
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:
70019
diff
changeset
|
486 |
unfolding openin_prod_topology_alt |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
487 |
proof (intro allI impI) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
488 |
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:
70019
diff
changeset
|
489 |
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:
70019
diff
changeset
|
490 |
for x y |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
491 |
proof - |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
492 |
have "x \<in> topspace X" "y \<in> topspace Y" "y \<noteq> f x" |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
493 |
using that by auto |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
494 |
moreover have "f x \<in> topspace Y" |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
495 |
by (meson \<open>x \<in> topspace X\<close> continuous_map_def f) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
496 |
ultimately obtain U V where UV: "openin Y U" "openin Y V" "f x \<in> U" "y \<in> V" "disjnt U V" |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
497 |
using Y Hausdorff_space_def by metis |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
498 |
show ?thesis |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
499 |
proof (intro exI conjI) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
500 |
show "openin X {x \<in> topspace X. f x \<in> U}" |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
501 |
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:
70019
diff
changeset
|
502 |
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:
70019
diff
changeset
|
503 |
using UV by (auto simp: disjnt_iff dest: openin_subset) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
504 |
qed (use UV \<open>x \<in> topspace X\<close> in auto) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
505 |
qed |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
506 |
qed |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
507 |
qed |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
508 |
|
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
509 |
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:
70178
diff
changeset
|
510 |
"\<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:
69986
diff
changeset
|
511 |
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:
69986
diff
changeset
|
512 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
513 |
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:
70178
diff
changeset
|
514 |
"\<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:
69986
diff
changeset
|
515 |
\<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:
69986
diff
changeset
|
516 |
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:
69986
diff
changeset
|
517 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
518 |
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:
70178
diff
changeset
|
519 |
"\<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:
69986
diff
changeset
|
520 |
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:
69986
diff
changeset
|
521 |
\<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:
69986
diff
changeset
|
522 |
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:
69986
diff
changeset
|
523 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
524 |
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:
70178
diff
changeset
|
525 |
"\<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:
69986
diff
changeset
|
526 |
\<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:
69986
diff
changeset
|
527 |
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:
69986
diff
changeset
|
528 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
529 |
lemma continuous_inverse_map: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
530 |
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:
69986
diff
changeset
|
531 |
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:
69986
diff
changeset
|
532 |
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:
69986
diff
changeset
|
533 |
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:
69986
diff
changeset
|
534 |
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:
69986
diff
changeset
|
535 |
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:
69986
diff
changeset
|
536 |
unfolding continuous_map_closedin |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
537 |
proof (intro conjI ballI allI impI) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
538 |
fix x |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
539 |
assume "x \<in> topspace (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:
69986
diff
changeset
|
540 |
then show "g x \<in> topspace X" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
541 |
by (auto simp: gf) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
542 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
543 |
fix C |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
544 |
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:
69986
diff
changeset
|
545 |
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:
69986
diff
changeset
|
546 |
{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:
69986
diff
changeset
|
547 |
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:
69986
diff
changeset
|
548 |
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:
69986
diff
changeset
|
549 |
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:
69986
diff
changeset
|
550 |
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:
69986
diff
changeset
|
551 |
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:
69986
diff
changeset
|
552 |
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:
69986
diff
changeset
|
553 |
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:
69986
diff
changeset
|
554 |
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:
69986
diff
changeset
|
555 |
by simp |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
556 |
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:
69986
diff
changeset
|
557 |
{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:
69986
diff
changeset
|
558 |
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:
69986
diff
changeset
|
559 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
560 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
561 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
562 |
|
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
563 |
lemma closed_map_paired_continuous_map_right: |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
564 |
"\<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:
70019
diff
changeset
|
565 |
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:
69986
diff
changeset
|
566 |
|
70178
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
567 |
lemma closed_map_paired_continuous_map_left: |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
568 |
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:
70019
diff
changeset
|
569 |
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:
70019
diff
changeset
|
570 |
proof - |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
571 |
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:
70019
diff
changeset
|
572 |
by auto |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
573 |
show ?thesis |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
574 |
unfolding eq |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
575 |
proof (rule closed_map_compose) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
576 |
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:
70019
diff
changeset
|
577 |
using Y closed_map_paired_continuous_map_right f by blast |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
578 |
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:
70019
diff
changeset
|
579 |
by (metis homeomorphic_map_swap homeomorphic_imp_closed_map) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
580 |
qed |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
581 |
qed |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
582 |
|
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
583 |
lemma proper_map_paired_continuous_map_right: |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
584 |
"\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk> |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
585 |
\<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:
70019
diff
changeset
|
586 |
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:
70019
diff
changeset
|
587 |
by (metis (mono_tags, lifting) Pair_inject inj_onI) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
588 |
|
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
589 |
lemma proper_map_paired_continuous_map_left: |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
590 |
"\<lbrakk>continuous_map X Y f; Hausdorff_space Y\<rbrakk> |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
591 |
\<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:
70019
diff
changeset
|
592 |
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:
70019
diff
changeset
|
593 |
by (metis (mono_tags, lifting) Pair_inject inj_onI) |
4900351361b0
Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
594 |
|
69994
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
595 |
lemma Hausdorff_space_prod_topology: |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
596 |
"Hausdorff_space(prod_topology X Y) \<longleftrightarrow> topspace(prod_topology X Y) = {} \<or> Hausdorff_space X \<and> Hausdorff_space Y" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
597 |
(is "?lhs = ?rhs") |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
598 |
proof |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
599 |
assume ?lhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
600 |
then show ?rhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
601 |
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:
69986
diff
changeset
|
602 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
603 |
assume R: ?rhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
604 |
show ?lhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
605 |
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:
69986
diff
changeset
|
606 |
case False |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
607 |
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:
69986
diff
changeset
|
608 |
by auto |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
609 |
show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
610 |
unfolding Hausdorff_space_def |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
611 |
proof clarify |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
612 |
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:
69986
diff
changeset
|
613 |
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:
69986
diff
changeset
|
614 |
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:
69986
diff
changeset
|
615 |
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:
69986
diff
changeset
|
616 |
\<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:
69986
diff
changeset
|
617 |
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:
69986
diff
changeset
|
618 |
using that |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
619 |
proof |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
620 |
assume "x \<noteq> x'" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
621 |
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:
69986
diff
changeset
|
622 |
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:
69986
diff
changeset
|
623 |
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:
69986
diff
changeset
|
624 |
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:
69986
diff
changeset
|
625 |
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:
69986
diff
changeset
|
626 |
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:
69986
diff
changeset
|
627 |
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:
69986
diff
changeset
|
628 |
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:
69986
diff
changeset
|
629 |
ultimately show False |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
630 |
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:
69986
diff
changeset
|
631 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
632 |
assume "y \<noteq> y'" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
633 |
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:
69986
diff
changeset
|
634 |
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:
69986
diff
changeset
|
635 |
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:
69986
diff
changeset
|
636 |
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:
69986
diff
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:
69986
diff
changeset
|
638 |
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:
69986
diff
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:
69986
diff
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:
69986
diff
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:
69986
diff
changeset
|
642 |
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:
69986
diff
changeset
|
643 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
644 |
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:
69986
diff
changeset
|
645 |
by blast |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
646 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
647 |
qed (simp add: Hausdorff_space_topspace_empty) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
648 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
649 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
650 |
|
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
651 |
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:
69986
diff
changeset
|
652 |
"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:
69986
diff
changeset
|
653 |
(is "?lhs = ?rhs") |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
654 |
proof |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
655 |
assume ?lhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
656 |
then show ?rhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
657 |
apply (rule topological_property_of_product_component) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
658 |
apply (blast dest: 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:
69986
diff
changeset
|
659 |
done |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
660 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
661 |
assume R: ?rhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
662 |
show ?lhs |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
663 |
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:
69986
diff
changeset
|
664 |
case True |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
665 |
then show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
666 |
by (simp add: Hausdorff_space_topspace_empty) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
667 |
next |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
668 |
case False |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
669 |
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:
69986
diff
changeset
|
670 |
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:
69986
diff
changeset
|
671 |
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:
69986
diff
changeset
|
672 |
proof - |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
673 |
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:
69986
diff
changeset
|
674 |
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:
69986
diff
changeset
|
675 |
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:
69986
diff
changeset
|
676 |
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:
69986
diff
changeset
|
677 |
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:
69986
diff
changeset
|
678 |
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:
69986
diff
changeset
|
679 |
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:
69986
diff
changeset
|
680 |
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:
69986
diff
changeset
|
681 |
show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
682 |
proof (intro exI conjI) |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
683 |
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:
69986
diff
changeset
|
684 |
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:
69986
diff
changeset
|
685 |
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:
69986
diff
changeset
|
686 |
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:
69986
diff
changeset
|
687 |
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:
69986
diff
changeset
|
688 |
show "f \<in> ?U" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
689 |
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:
69986
diff
changeset
|
690 |
show "g \<in> ?V" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
691 |
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:
69986
diff
changeset
|
692 |
show "disjnt ?U ?V" |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
693 |
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:
69986
diff
changeset
|
694 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
695 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
696 |
then show ?thesis |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
697 |
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:
69986
diff
changeset
|
698 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
699 |
qed |
cf7150ab1075
more stuff from HOL Light: Euclidean spaces and n-spheres, Hausdorff spaces, etc.
paulson <lp15@cam.ac.uk>
parents:
69986
diff
changeset
|
700 |
|
69874 | 701 |
end |