| author | wenzelm | 
| Sun, 29 Sep 2024 21:16:17 +0200 | |
| changeset 81008 | d0cd220d8e8b | 
| parent 79492 | c1b0f64eb865 | 
| child 82501 | 26f9f484f266 | 
| permissions | -rw-r--r-- | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | (* Author: L C Paulson, University of Cambridge [ported from HOL Light] *) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | section \<open>Various Forms of Topological Spaces\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | theory Abstract_Topological_Spaces | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 6 | imports Lindelof_Spaces Locally Abstract_Euclidean_Space Sum_Topology FSigma | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | begin | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | subsection\<open>Connected topological spaces\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | lemma connected_space_eq_frontier_eq_empty: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 |    "connected_space X \<longleftrightarrow> (\<forall>S. S \<subseteq> topspace X \<and> X frontier_of S = {} \<longrightarrow> S = {} \<or> S = topspace X)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | by (meson clopenin_eq_frontier_of connected_space_clopen_in) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | lemma connected_space_frontier_eq_empty: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | "connected_space X \<and> S \<subseteq> topspace X | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 |         \<Longrightarrow> (X frontier_of S = {} \<longleftrightarrow> S = {} \<or> S = topspace X)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | by (meson connected_space_eq_frontier_eq_empty frontier_of_empty frontier_of_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | lemma connectedin_eq_subset_separated_union: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | "connectedin X C \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | C \<subseteq> topspace X \<and> (\<forall>S T. separatedin X S T \<and> C \<subseteq> S \<union> T \<longrightarrow> C \<subseteq> S \<or> C \<subseteq> T)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | using connectedin_subset_topspace connectedin_subset_separated_union by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 | assume ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 | by (metis closure_of_subset connectedin_separation dual_order.eq_iff inf.orderE separatedin_def sup.boundedE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | lemma connectedin_clopen_cases: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | "\<lbrakk>connectedin X C; closedin X T; openin X T\<rbrakk> \<Longrightarrow> C \<subseteq> T \<or> disjnt C T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | by (metis Diff_eq_empty_iff Int_empty_right clopenin_eq_frontier_of connectedin_Int_frontier_of disjnt_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | lemma connected_space_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 | "\<lbrakk>retraction_map X X' r; connected_space X\<rbrakk> \<Longrightarrow> connected_space X'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | using connected_space_quotient_map_image retraction_imp_quotient_map by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | lemma connectedin_imp_perfect_gen: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 |   assumes X: "t1_space X" and S: "connectedin X S" and nontriv: "\<nexists>a. S = {a}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | shows "S \<subseteq> X derived_set_of S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | unfolding derived_set_of_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | proof (intro subsetI CollectI conjI strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | show XS: "x \<in> topspace X" if "x \<in> S" for x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | using that S connectedin by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 | show "\<exists>y. y \<noteq> x \<and> y \<in> S \<and> y \<in> T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | if "x \<in> S" and "x \<in> T \<and> openin X T" for x T | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 |     have opeXx: "openin X (topspace X - {x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | by (meson X openin_topspace t1_space_openin_delete_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | moreover | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 |     have "S \<subseteq> T \<union> (topspace X - {x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | using XS that(2) by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 |     moreover have "(topspace X - {x}) \<inter> S \<noteq> {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | by (metis Diff_triv S connectedin double_diff empty_subsetI inf_commute insert_subsetI nontriv that(1)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 |       using that connectedinD [OF S, of T "topspace X - {x}"]
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | lemma connectedin_imp_perfect: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 |   "\<lbrakk>Hausdorff_space X; connectedin X S; \<nexists>a. S = {a}\<rbrakk> \<Longrightarrow> S \<subseteq> X derived_set_of S"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | by (simp add: Hausdorff_imp_t1_space connectedin_imp_perfect_gen) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | subsection\<open>The notion of "separated between" (complement of "connected between)"\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | definition separated_between | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | where "separated_between X S T \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | \<exists>U V. openin X U \<and> openin X V \<and> U \<union> V = topspace X \<and> disjnt U V \<and> S \<subseteq> U \<and> T \<subseteq> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | lemma separated_between_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | "separated_between X S T \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | (\<exists>U V. closedin X U \<and> closedin X V \<and> U \<union> V = topspace X \<and> disjnt U V \<and> S \<subseteq> U \<and> T \<subseteq> V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | unfolding separated_between_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | by (metis separatedin_open_sets separation_closedin_Un_gen subtopology_topspace | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | separatedin_closed_sets separation_openin_Un_gen) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | lemma separated_between: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | "separated_between X S T \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | (\<exists>U. closedin X U \<and> openin X U \<and> S \<subseteq> U \<and> T \<subseteq> topspace X - U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | unfolding separated_between_def closedin_def disjnt_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | by (smt (verit, del_insts) Diff_cancel Diff_disjoint Diff_partition Un_Diff Un_Diff_Int openin_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | lemma separated_between_mono: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | "\<lbrakk>separated_between X S T; S' \<subseteq> S; T' \<subseteq> T\<rbrakk> \<Longrightarrow> separated_between X S' T'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | by (meson order.trans separated_between) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | lemma separated_between_refl: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 |    "separated_between X S S \<longleftrightarrow> S = {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | unfolding separated_between_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | by (metis Un_empty_right disjnt_def disjnt_empty2 disjnt_subset2 disjnt_sym le_iff_inf openin_empty openin_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | lemma separated_between_sym: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | "separated_between X S T \<longleftrightarrow> separated_between X T S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | by (metis disjnt_sym separated_between_alt sup_commute) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | lemma separated_between_imp_subset: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | "separated_between X S T \<Longrightarrow> S \<subseteq> topspace X \<and> T \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | by (metis le_supI1 le_supI2 separated_between_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | lemma separated_between_empty: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 |   "(separated_between X {} S \<longleftrightarrow> S \<subseteq> topspace X) \<and> (separated_between X S {} \<longleftrightarrow> S \<subseteq> topspace X)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | by (metis Diff_empty bot.extremum closedin_empty openin_empty separated_between separated_between_imp_subset separated_between_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | lemma separated_between_Un: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | "separated_between X S (T \<union> U) \<longleftrightarrow> separated_between X S T \<and> separated_between X S U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | by (auto simp: separated_between) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | lemma separated_between_Un': | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | "separated_between X (S \<union> T) U \<longleftrightarrow> separated_between X S U \<and> separated_between X T U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | by (simp add: separated_between_Un separated_between_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | lemma separated_between_imp_disjoint: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | "separated_between X S T \<Longrightarrow> disjnt S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | by (meson disjnt_iff separated_between_def subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | lemma separated_between_imp_separatedin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | "separated_between X S T \<Longrightarrow> separatedin X S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | by (meson separated_between_def separatedin_mono separatedin_open_sets) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | lemma separated_between_full: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | assumes "S \<union> T = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | shows "separated_between X S T \<longleftrightarrow> disjnt S T \<and> closedin X S \<and> openin X S \<and> closedin X T \<and> openin X T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | have "separated_between X S T \<longrightarrow> separatedin X S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | by (simp add: separated_between_imp_separatedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | unfolding separated_between_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | by (metis assms separation_closedin_Un_gen separation_openin_Un_gen subset_refl subtopology_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | lemma separated_between_eq_separatedin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | "S \<union> T = topspace X \<Longrightarrow> (separated_between X S T \<longleftrightarrow> separatedin X S T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | by (simp add: separated_between_full separatedin_full) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | lemma separated_between_pointwise_left: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | assumes "compactin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | shows "separated_between X S T \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 |          (S = {} \<longrightarrow> T \<subseteq> topspace X) \<and> (\<forall>x \<in> S. separated_between X {x} T)"  (is "?lhs=?rhs")
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | using separated_between_imp_subset separated_between_mono by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | then have "T \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | by (meson equals0I separated_between_imp_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | obtain U where U: "\<forall>x \<in> S. openin X (U x)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 |       "\<forall>x \<in> S. \<exists>V. openin X V \<and> U x \<union> V = topspace X \<and> disjnt (U x) V \<and> {x} \<subseteq> U x \<and> T \<subseteq> V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | using R unfolding separated_between_def by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 | then have "S \<subseteq> \<Union>(U ` S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 | then obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> (\<Union>i\<in>K. U i)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | using assms U unfolding compactin_def by (smt (verit) finite_subset_image imageE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | unfolding separated_between | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | have "\<And>x. x \<in> K \<Longrightarrow> closedin X (U x)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | by (smt (verit) \<open>K \<subseteq> S\<close> Diff_cancel U(2) Un_Diff Un_Diff_Int disjnt_def openin_closedin_eq subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | then show "closedin X (\<Union> (U ` K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | by (metis (mono_tags, lifting) \<open>finite K\<close> closedin_Union finite_imageI image_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | show "openin X (\<Union> (U ` K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | using U(1) \<open>K \<subseteq> S\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | show "S \<subseteq> \<Union> (U ` K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | by (simp add: K) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | have "\<And>x i. \<lbrakk>x \<in> T; i \<in> K; x \<in> U i\<rbrakk> \<Longrightarrow> False" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | by (meson U(2) \<open>K \<subseteq> S\<close> disjnt_iff subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | then show "T \<subseteq> topspace X - \<Union> (U ` K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | using \<open>T \<subseteq> topspace X\<close> by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | lemma separated_between_pointwise_right: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | "compactin X T | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 |         \<Longrightarrow> separated_between X S T \<longleftrightarrow> (T = {} \<longrightarrow> S \<subseteq> topspace X) \<and> (\<forall>y \<in> T. separated_between X S {y})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | by (meson separated_between_pointwise_left separated_between_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | lemma separated_between_closure_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | "S \<subseteq> topspace X \<Longrightarrow> separated_between X (X closure_of S) T \<longleftrightarrow> separated_between X S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | by (meson closure_of_minimal_eq separated_between_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | lemma separated_between_closure_of': | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 | "T \<subseteq> topspace X \<Longrightarrow> separated_between X S (X closure_of T) \<longleftrightarrow> separated_between X S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | by (meson separated_between_closure_of separated_between_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | lemma separated_between_closure_of_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | "separated_between X S T \<longleftrightarrow> S \<subseteq> topspace X \<and> separated_between X (X closure_of S) T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | by (metis separated_between_closure_of separated_between_imp_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | lemma separated_between_closure_of_eq': | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | "separated_between X S T \<longleftrightarrow> T \<subseteq> topspace X \<and> separated_between X S (X closure_of T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | by (metis separated_between_closure_of' separated_between_imp_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | lemma separated_between_frontier_of_eq': | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | "separated_between X S T \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | T \<subseteq> topspace X \<and> disjnt S T \<and> separated_between X S (X frontier_of T)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | by (metis interior_of_union_frontier_of separated_between_Un separated_between_closure_of_eq' | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | separated_between_imp_disjoint) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | then obtain U where U: "closedin X U" "openin X U" "S \<subseteq> U" "X closure_of T - X interior_of T \<subseteq> topspace X - U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | by (metis frontier_of_def separated_between) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | proof (rule separated_between_mono [of _ S "X closure_of T"]) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | have "separated_between X S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | unfolding separated_between | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | show "S \<subseteq> U - T" "T \<subseteq> topspace X - (U - T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | using R U(3) by (force simp: disjnt_iff)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | have "T \<subseteq> X closure_of T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | by (simp add: R closure_of_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | then have *: "U - T = U - X interior_of T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | using U(4) interior_of_subset by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | then show "closedin X (U - T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | by (simp add: U(1) closedin_diff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 |       have "U \<inter> X frontier_of T = {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | using U(4) frontier_of_def by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 | then show "openin X (U - T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | by (metis * Diff_Un U(2) Un_Diff_Int Un_Int_eq(1) closedin_closure_of interior_of_union_frontier_of openin_diff sup_bot_right) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | then show "separated_between X S (X closure_of T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | by (simp add: R separated_between_closure_of') | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | qed (auto simp: R closure_of_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | lemma separated_between_frontier_of_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | "separated_between X S T \<longleftrightarrow> S \<subseteq> topspace X \<and> disjnt S T \<and> separated_between X (X frontier_of S) T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | by (metis disjnt_sym separated_between_frontier_of_eq' separated_between_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | lemma separated_between_frontier_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | "\<lbrakk>S \<subseteq> topspace X; disjnt S T\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | \<Longrightarrow> (separated_between X (X frontier_of S) T \<longleftrightarrow> separated_between X S T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | using separated_between_frontier_of_eq by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | lemma separated_between_frontier_of': | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | "\<lbrakk>T \<subseteq> topspace X; disjnt S T\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | \<Longrightarrow> (separated_between X S (X frontier_of T) \<longleftrightarrow> separated_between X S T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | using separated_between_frontier_of_eq' by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | lemma connected_space_separated_between: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 |   "connected_space X \<longleftrightarrow> (\<forall>S T. separated_between X S T \<longrightarrow> S = {} \<or> T = {})" (is "?lhs=?rhs")
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | by (metis Diff_cancel connected_space_clopen_in separated_between subset_empty) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | by (meson connected_space_eq_not_separated separated_between_eq_separatedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | lemma connected_space_imp_separated_between_trivial: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | "connected_space X | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 |         \<Longrightarrow> (separated_between X S T \<longleftrightarrow> S = {} \<and> T \<subseteq> topspace X \<or> S \<subseteq> topspace X \<and> T = {})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | by (metis connected_space_separated_between separated_between_empty) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | subsection\<open>Connected components\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | lemma connected_component_of_subtopology_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | "connected_component_of (subtopology X U) a = connected_component_of X a \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | connected_component_of_set X a \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | by (force simp: connected_component_of_set connectedin_subtopology connected_component_of_def fun_eq_iff subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | lemma connected_components_of_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | assumes "C \<in> connected_components_of X" "C \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | shows "C \<in> connected_components_of (subtopology X U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | obtain a where a: "connected_component_of_set X a \<subseteq> U" and "a \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | and Ceq: "C = connected_component_of_set X a" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | using assms by (force simp: connected_components_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | then have "a \<in> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | by (simp add: connected_component_of_refl in_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | then have "connected_component_of_set X a = connected_component_of_set (subtopology X U) a" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | by (metis a connected_component_of_subtopology_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | by (simp add: Ceq \<open>a \<in> U\<close> \<open>a \<in> topspace X\<close> connected_component_in_connected_components_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | lemma open_in_finite_connected_components: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | assumes "finite(connected_components_of X)" "C \<in> connected_components_of X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | shows "openin X C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | have "closedin X (topspace X - C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | by (metis DiffD1 assms closedin_Union closedin_connected_components_of complement_connected_components_of_Union finite_Diff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | by (simp add: assms connected_components_of_subset openin_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | thm connected_component_of_eq_overlap | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | lemma connected_components_of_disjoint: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | assumes "C \<in> connected_components_of X" "C' \<in> connected_components_of X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | shows "(disjnt C C' \<longleftrightarrow> (C \<noteq> C'))" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 304 | using assms nonempty_connected_components_of pairwiseD pairwise_disjoint_connected_components_of by fastforce | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | lemma connected_components_of_overlap: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 |    "\<lbrakk>C \<in> connected_components_of X; C' \<in> connected_components_of X\<rbrakk> \<Longrightarrow> C \<inter> C' \<noteq> {} \<longleftrightarrow> C = C'"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | by (meson connected_components_of_disjoint disjnt_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | lemma pairwise_separated_connected_components_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | "pairwise (separatedin X) (connected_components_of X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | by (simp add: closedin_connected_components_of connected_components_of_disjoint pairwiseI separatedin_closed_sets) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | lemma finite_connected_components_of_finite: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | "finite(topspace X) \<Longrightarrow> finite(connected_components_of X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | by (simp add: Union_connected_components_of finite_UnionD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | lemma connected_component_of_unique: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | "\<lbrakk>x \<in> C; connectedin X C; \<And>C'. x \<in> C' \<and> connectedin X C' \<Longrightarrow> C' \<subseteq> C\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | \<Longrightarrow> connected_component_of_set X x = C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | by (meson connected_component_of_maximal connectedin_connected_component_of subsetD subset_antisym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 | lemma closedin_connected_component_of_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | "\<lbrakk>C \<in> connected_components_of (subtopology X s); X closure_of C \<subseteq> s\<rbrakk> \<Longrightarrow> closedin X C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | by (metis closedin_Int_closure_of closedin_connected_components_of closure_of_eq inf.absorb_iff2) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | lemma connected_component_of_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 |    "connected_component_of_set (discrete_topology U) x = (if x \<in> U then {x} else {})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | by (simp add: locally_path_connected_space_discrete_topology flip: path_component_eq_connected_component_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | lemma connected_components_of_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 |    "connected_components_of (discrete_topology U) = (\<lambda>x. {x}) ` U"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | by (simp add: connected_component_of_discrete_topology connected_components_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | lemma connected_component_of_continuous_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | "\<lbrakk>continuous_map X Y f; connected_component_of X x y\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | \<Longrightarrow> connected_component_of Y (f x) (f y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | by (meson connected_component_of_def connectedin_continuous_map_image image_eqI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | lemma homeomorphic_map_connected_component_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | assumes "homeomorphic_map X Y f" and x: "x \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | shows "connected_component_of_set Y (f x) = f ` (connected_component_of_set X x)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | obtain g where g: "continuous_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | "continuous_map Y X g " "\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | "\<And>y. y \<in> topspace Y \<Longrightarrow> f (g y) = y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | using assms(1) homeomorphic_map_maps homeomorphic_maps_def by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | using connected_component_in_topspace [of Y] x g | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | connected_component_of_continuous_image [of X Y f] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | connected_component_of_continuous_image [of Y X g] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | lemma homeomorphic_map_connected_components_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | assumes "homeomorphic_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | shows "connected_components_of Y = (image f) ` (connected_components_of X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | have "topspace Y = f ` topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | by (metis assms homeomorphic_imp_surjective_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | with homeomorphic_map_connected_component_of [OF assms] show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | by (auto simp: connected_components_of_def image_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | lemma connected_component_of_pair: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | "connected_component_of_set (prod_topology X Y) (x,y) = | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | connected_component_of_set X x \<times> connected_component_of_set Y y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | proof (cases "x \<in> topspace X \<and> y \<in> topspace Y") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | proof (rule connected_component_of_unique) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | show "(x, y) \<in> connected_component_of_set X x \<times> connected_component_of_set Y y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | using True by (simp add: connected_component_of_refl) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | show "connectedin (prod_topology X Y) (connected_component_of_set X x \<times> connected_component_of_set Y y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | by (metis connectedin_Times connectedin_connected_component_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | show "C \<subseteq> connected_component_of_set X x \<times> connected_component_of_set Y y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | if "(x, y) \<in> C \<and> connectedin (prod_topology X Y) C" for C | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | using that unfolding connected_component_of_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | apply clarsimp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | by (metis (no_types) connectedin_continuous_map_image continuous_map_fst continuous_map_snd fst_conv imageI snd_conv) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | case False then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | by (metis Sigma_empty1 Sigma_empty2 connected_component_of_eq_empty mem_Sigma_iff topspace_prod_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | lemma connected_components_of_prod_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | "connected_components_of (prod_topology X Y) = | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 |     {C \<times> D |C D. C \<in> connected_components_of X \<and> D \<in> connected_components_of Y}" (is "?lhs=?rhs")
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | show "?lhs \<subseteq> ?rhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | apply (clarsimp simp: connected_components_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | by (metis (no_types) connected_component_of_pair imageI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | show "?rhs \<subseteq> ?lhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | using connected_component_of_pair | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | by (fastforce simp: connected_components_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | lemma connected_component_of_product_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | "connected_component_of_set (product_topology X I) x = | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 |     (if x \<in> extensional I then PiE I (\<lambda>i. connected_component_of_set (X i) (x i)) else {})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | (is "?lhs = If _ ?R _") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | proof (cases "x \<in> topspace(product_topology X I)") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | have "?lhs = (\<Pi>\<^sub>E i\<in>I. connected_component_of_set (X i) (x i))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | if xX: "\<And>i. i\<in>I \<Longrightarrow> x i \<in> topspace (X i)" and ext: "x \<in> extensional I" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | proof (rule connected_component_of_unique) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | show "x \<in> ?R" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | by (simp add: PiE_iff connected_component_of_refl local.ext xX) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | show "connectedin (product_topology X I) ?R" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | by (simp add: connectedin_PiE connectedin_connected_component_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | show "C \<subseteq> ?R" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | if "x \<in> C \<and> connectedin (product_topology X I) C" for C | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | have "C \<subseteq> extensional I" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | using PiE_def connectedin_subset_topspace that by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | have "\<And>y. y \<in> C \<Longrightarrow> y \<in> (\<Pi> i\<in>I. connected_component_of_set (X i) (x i))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | apply (simp add: connected_component_of_def Pi_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | by (metis connectedin_continuous_map_image continuous_map_product_projection imageI that) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | using PiE_def \<open>C \<subseteq> extensional I\<close> by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | with True show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | by (simp add: PiE_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | then show ?thesis | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 431 | by (smt (verit, best) PiE_eq_empty_iff PiE_iff connected_component_of_eq_empty topspace_product_topology) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | lemma connected_components_of_product_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | "connected_components_of (product_topology X I) = | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 |     {PiE I B |B. \<forall>i \<in> I. B i \<in> connected_components_of(X i)}"  (is "?lhs=?rhs")
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | show "?lhs \<subseteq> ?rhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | by (auto simp: connected_components_of_def connected_component_of_product_topology PiE_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | show "?rhs \<subseteq> ?lhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | fix F | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | assume "F \<in> ?rhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 | then obtain B where Feq: "F = Pi\<^sub>E I B" and | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | "\<forall>i\<in>I. \<exists>x\<in>topspace (X i). B i = connected_component_of_set (X i) x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 | by (force simp: connected_components_of_def connected_component_of_product_topology image_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | then obtain f where | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 449 | f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> topspace (X i) \<and> B i = connected_component_of_set (X i) (f i)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | then have "(\<lambda>i\<in>I. f i) \<in> ((\<Pi>\<^sub>E i\<in>I. topspace (X i)) \<inter> extensional I)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 452 | by simp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | with f show "F \<in> ?lhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | unfolding Feq connected_components_of_def connected_component_of_product_topology image_iff | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | by (smt (verit, del_insts) PiE_cong restrict_PiE_iff restrict_apply' restrict_extensional topspace_product_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | subsection \<open>Monotone maps (in the general topological sense)\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | definition monotone_map | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 | where "monotone_map X Y f == | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | f ` (topspace X) \<subseteq> topspace Y \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 |         (\<forall>y \<in> topspace Y. connectedin X {x \<in> topspace X. f x = y})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | lemma monotone_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | "monotone_map X Y f \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 |    f ` (topspace X) \<subseteq> topspace Y \<and> (\<forall>y. connectedin X {x \<in> topspace X. f x = y})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | apply (simp add: monotone_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 | by (metis (mono_tags, lifting) connectedin_empty [of X] Collect_empty_eq image_subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | lemma monotone_map_in_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | "monotone_map X (subtopology Y S) f \<longleftrightarrow> monotone_map X Y f \<and> f ` (topspace X) \<subseteq> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 477 | by (smt (verit, del_insts) le_inf_iff monotone_map topspace_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 | lemma monotone_map_from_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 480 | assumes "monotone_map X Y f" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 481 | "\<And>x y. \<lbrakk>x \<in> topspace X; y \<in> topspace X; x \<in> S; f x = f y\<rbrakk> \<Longrightarrow> y \<in> S" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | shows "monotone_map (subtopology X S) Y f" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 483 | proof - | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 484 |   have "\<And>y. y \<in> topspace Y \<Longrightarrow> connectedin X {x \<in> topspace X. x \<in> S \<and> f x = y}"
 | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 485 | by (smt (verit) Collect_cong assms connectedin_empty empty_def monotone_map_def) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 486 | then show ?thesis | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 487 | using assms by (auto simp: monotone_map_def connectedin_subtopology) | 
| 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 488 | qed | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | lemma monotone_map_restriction: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 |   "monotone_map X Y f \<and> {x \<in> topspace X. f x \<in> v} = u
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 492 | \<Longrightarrow> monotone_map (subtopology X u) (subtopology Y v) f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 | by (smt (verit, best) IntI Int_Collect image_subset_iff mem_Collect_eq monotone_map monotone_map_from_subtopology topspace_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | lemma injective_imp_monotone_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 496 | assumes "f ` topspace X \<subseteq> topspace Y" "inj_on f (topspace X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | shows "monotone_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | unfolding monotone_map_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | proof (intro conjI assms strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 | fix y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 501 | assume "y \<in> topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 |   then have "{x \<in> topspace X. f x = y} = {} \<or> (\<exists>a \<in> topspace X. {x \<in> topspace X. f x = y} = {a})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 503 | using assms(2) unfolding inj_on_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 |   then show "connectedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 505 | by (metis (no_types, lifting) connectedin_empty connectedin_sing) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 506 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 507 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 | lemma embedding_imp_monotone_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 509 | "embedding_map X Y f \<Longrightarrow> monotone_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 510 | by (metis (no_types) embedding_map_def homeomorphic_eq_everything_map inf.absorb_iff2 injective_imp_monotone_map topspace_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | lemma section_imp_monotone_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 | "section_map X Y f \<Longrightarrow> monotone_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 514 | by (simp add: embedding_imp_monotone_map section_imp_embedding_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 515 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 516 | lemma homeomorphic_imp_monotone_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 517 | "homeomorphic_map X Y f \<Longrightarrow> monotone_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | by (meson section_and_retraction_eq_homeomorphic_map section_imp_monotone_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 519 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 520 | proposition connected_space_monotone_quotient_map_preimage: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | assumes f: "monotone_map X Y f" "quotient_map X Y f" and "connected_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 | shows "connected_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 523 | proof (rule ccontr) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 524 | assume "\<not> connected_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 525 |   then obtain U V where "openin X U" "openin X V" "U \<inter> V = {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 |     "U \<noteq> {}" "V \<noteq> {}" and topUV: "topspace X \<subseteq> U \<union> V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 527 | by (auto simp: connected_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 528 | then have UVsub: "U \<subseteq> topspace X" "V \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 529 | by (auto simp: openin_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 | have "\<not> connected_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 531 | unfolding connected_space_def not_not | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 532 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 533 | show "topspace Y \<subseteq> f`U \<union> f`V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | by (metis f(2) image_Un quotient_imp_surjective_map subset_Un_eq topUV) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 |     show "f`U \<noteq> {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 |       by (simp add: \<open>U \<noteq> {}\<close>)
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 537 |     show "(f`V) \<noteq> {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 538 |       by (simp add: \<open>V \<noteq> {}\<close>)
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 539 | have *: "y \<notin> f ` V" if "y \<in> f ` U" for y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 540 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 541 |       have \<section>: "connectedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | using f(1) monotone_map by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 543 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 544 |         using connectedinD [OF \<section> \<open>openin X U\<close> \<open>openin X V\<close>] UVsub topUV \<open>U \<inter> V = {}\<close> that
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 545 | by (force simp: disjoint_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 546 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 547 |     then show "f`U \<inter> f`V = {}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 548 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 549 | show "openin Y (f`U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 550 | using f \<open>openin X U\<close> topUV * unfolding quotient_map_saturated_open by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | show "openin Y (f`V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 552 | using f \<open>openin X V\<close> topUV * unfolding quotient_map_saturated_open by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 554 | then show False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 555 | by (simp add: assms) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 | lemma connectedin_monotone_quotient_map_preimage: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | assumes "monotone_map X Y f" "quotient_map X Y f" "connectedin Y C" "openin Y C \<or> closedin Y C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 |   shows "connectedin X {x \<in> topspace X. f x \<in> C}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 561 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 562 |   have "connected_space (subtopology X {x \<in> topspace X. f x \<in> C})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 563 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | have "connected_space (subtopology Y C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | using \<open>connectedin Y C\<close> connectedin_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 |     moreover have "quotient_map (subtopology X {a \<in> topspace X. f a \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 567 | by (simp add: assms quotient_map_restriction) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 568 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 569 | using \<open>monotone_map X Y f\<close> connected_space_monotone_quotient_map_preimage monotone_map_restriction by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 572 | by (simp add: connectedin_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 574 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 575 | lemma monotone_open_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | assumes "continuous_map X Y f" "open_map X Y f" and fim: "f ` (topspace X) = topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 |   shows "monotone_map X Y f \<longleftrightarrow> (\<forall>C. connectedin Y C \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> C})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 580 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 581 | show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 582 | unfolding connectedin_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | proof (intro strip conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | fix C | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | assume C: "C \<subseteq> topspace Y \<and> connected_space (subtopology Y C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 |     show "connected_space (subtopology X {x \<in> topspace X. f x \<in> C})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | proof (rule connected_space_monotone_quotient_map_preimage) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 |       show "monotone_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | by (simp add: L monotone_map_restriction) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 |       show "quotient_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 591 | proof (rule continuous_open_imp_quotient_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 |         show "continuous_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 593 | using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 594 | qed (use open_map_restriction assms in fastforce)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | qed (simp add: C) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | qed auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 598 | assume ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 599 |   then have "\<forall>y. connectedin Y {y} \<longrightarrow> connectedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | by (smt (verit) Collect_cong singletonD singletonI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 | by (simp add: fim monotone_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 | lemma monotone_closed_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | assumes "continuous_map X Y f" "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 |   shows "monotone_map X Y f \<longleftrightarrow> (\<forall>C. connectedin Y C \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> C})" 
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 612 | unfolding connectedin_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 613 | proof (intro strip conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | fix C | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | assume C: "C \<subseteq> topspace Y \<and> connected_space (subtopology Y C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 |     show "connected_space (subtopology X {x \<in> topspace X. f x \<in> C})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | proof (rule connected_space_monotone_quotient_map_preimage) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 |       show "monotone_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | by (simp add: L monotone_map_restriction) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 |       show "quotient_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | proof (rule continuous_closed_imp_quotient_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 |         show "continuous_map (subtopology X {x \<in> topspace X. f x \<in> C}) (subtopology Y C) f"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | using assms continuous_map_from_subtopology continuous_map_in_subtopology by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | qed (use closed_map_restriction assms in fastforce)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 625 | qed (simp add: C) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 626 | qed auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 627 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 628 | assume ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 |   then have "\<forall>y. connectedin Y {y} \<longrightarrow> connectedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 630 | by (smt (verit) Collect_cong singletonD singletonI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 631 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | by (simp add: fim monotone_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 633 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | subsection\<open>Other countability properties\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 | definition second_countable | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | where "second_countable X \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 639 | \<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 640 | (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 641 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 642 | definition first_countable | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | where "first_countable X \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 644 | \<forall>x \<in> topspace X. | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 645 | \<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 | (\<forall>U. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 647 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 648 | definition separable_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 649 | where "separable_space X \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 | \<exists>C. countable C \<and> C \<subseteq> topspace X \<and> X closure_of C = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 | lemma second_countable: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 653 | "second_countable X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | (\<exists>\<B>. countable \<B> \<and> openin X = arbitrary union_of (\<lambda>x. x \<in> \<B>))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 655 | by (smt (verit) openin_topology_base_unique second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | lemma second_countable_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 | assumes "second_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | shows "second_countable (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 660 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 661 | obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 | using assms by (auto simp: second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 664 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 665 | unfolding second_countable_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 666 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 667 | show "\<forall>V\<in>((\<inter>)S) ` \<B>. openin (subtopology X S) V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 | using openin_subtopology_Int2 \<B> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 669 | show "\<forall>U x. openin (subtopology X S) U \<and> x \<in> U \<longrightarrow> (\<exists>V\<in>((\<inter>)S) ` \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 | using \<B> unfolding openin_subtopology | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | by (smt (verit, del_insts) IntI image_iff inf_commute inf_le1 subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 672 | qed (use \<B> in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | lemma second_countable_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | "second_countable(discrete_topology U) \<longleftrightarrow> countable U" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | then | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 | obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> V \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | "\<And>W x. W \<subseteq> U \<and> x \<in> W \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> W)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | by (auto simp: second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 |   then have "{x} \<in> \<B>" if "x \<in> U" for x
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 | by (metis empty_subsetI insertCI insert_subset subset_antisym that) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 | by (smt (verit) countable_subset image_subsetI \<open>countable \<B>\<close> countable_image_inj_on [OF _ inj_singleton]) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | assume ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | unfolding second_countable_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 692 |     by (rule_tac x="(\<lambda>x. {x}) ` U" in exI) auto
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | lemma second_countable_open_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | assumes "continuous_map X Y f" "open_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 | and fim: "f ` (topspace X) = topspace Y" and "second_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 698 | shows "second_countable Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 699 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 700 | have openXYf: "\<And>U. openin X U \<longrightarrow> openin Y (f ` U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 | using assms by (auto simp: open_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 703 | and *: "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 704 | using assms by (auto simp: second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 705 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | unfolding second_countable_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 | proof (intro exI conjI strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 708 | fix V y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 | assume V: "openin Y V \<and> y \<in> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 | then obtain x where "x \<in> topspace X" and x: "f x = y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | by (metis fim image_iff openin_subset subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 |     then obtain W where "W\<in>\<B>" "x \<in> W" "W \<subseteq> {x \<in> topspace X. f x \<in> V}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 714 |       using * [of "{x \<in> topspace X. f x \<in> V}" x] V assms openin_continuous_map_preimage 
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 | by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 716 | then show "\<exists>W \<in> (image f) ` \<B>. y \<in> W \<and> W \<subseteq> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 | using x by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 | qed (use \<B> openXYf in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 719 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 720 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 721 | lemma homeomorphic_space_second_countability: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 | "X homeomorphic_space Y \<Longrightarrow> (second_countable X \<longleftrightarrow> second_countable Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 723 | by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym second_countable_open_map_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 725 | lemma second_countable_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 | "\<lbrakk>retraction_map X Y r; second_countable X\<rbrakk> \<Longrightarrow> second_countable Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 | using hereditary_imp_retractive_property homeomorphic_space_second_countability second_countable_subtopology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 728 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | lemma second_countable_imp_first_countable: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | "second_countable X \<Longrightarrow> first_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | by (metis first_countable_def second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 733 | lemma first_countable_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | assumes "first_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 | shows "first_countable (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | unfolding first_countable_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | fix x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 739 | assume "x \<in> topspace (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 740 | then obtain \<B> where "countable \<B>" and \<B>: "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 741 | "\<And>U. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | using assms first_countable_def by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 743 | show "\<exists>\<B>. countable \<B> \<and> (\<forall>V\<in>\<B>. openin (subtopology X S) V) \<and> (\<forall>U. openin (subtopology X S) U \<and> x \<in> U \<longrightarrow> (\<exists>V\<in>\<B>. x \<in> V \<and> V \<subseteq> U))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 744 | proof (intro exI conjI strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 | show "countable (((\<inter>)S) ` \<B>)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | using \<open>countable \<B>\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 | show "openin (subtopology X S) V" if "V \<in> ((\<inter>)S) ` \<B>" for V | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 | using \<B> openin_subtopology_Int2 that by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 | show "\<exists>V\<in>((\<inter>)S) ` \<B>. x \<in> V \<and> V \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 | if "openin (subtopology X S) U \<and> x \<in> U" for U | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 | using that \<B>(2) by (clarsimp simp: openin_subtopology) (meson le_infI2) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 753 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 754 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 755 | lemma first_countable_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 756 | "first_countable (discrete_topology U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 757 | unfolding first_countable_def topspace_discrete_topology openin_discrete_topology | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 758 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 759 | fix x assume "x \<in> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 760 | show "\<exists>\<B>. countable \<B> \<and> (\<forall>V\<in>\<B>. V \<subseteq> U) \<and> (\<forall>Ua. Ua \<subseteq> U \<and> x \<in> Ua \<longrightarrow> (\<exists>V\<in>\<B>. x \<in> V \<and> V \<subseteq> Ua))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 |     using \<open>x \<in> U\<close> by (rule_tac x="{{x}}" in exI) auto
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 762 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 763 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 764 | lemma first_countable_open_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 765 | assumes "continuous_map X Y f" "open_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 766 | and fim: "f ` (topspace X) = topspace Y" and "first_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 767 | shows "first_countable Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 768 | unfolding first_countable_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 769 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 770 | fix y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 771 | assume "y \<in> topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 772 | have openXYf: "\<And>U. openin X U \<longrightarrow> openin Y (f ` U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 773 | using assms by (auto simp: open_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 774 | then obtain x where x: "x \<in> topspace X" "f x = y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 775 | by (metis \<open>y \<in> topspace Y\<close> fim imageE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 776 | obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 777 | and *: "\<And>U. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 778 | using assms x first_countable_def by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 779 | show "\<exists>\<B>. countable \<B> \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | (\<forall>V\<in>\<B>. openin Y V) \<and> (\<forall>U. openin Y U \<and> y \<in> U \<longrightarrow> (\<exists>V\<in>\<B>. y \<in> V \<and> V \<subseteq> U))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 | proof (intro exI conjI strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 782 | fix V assume "openin Y V \<and> y \<in> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 783 |     then have "\<exists>W\<in>\<B>. x \<in> W \<and> W \<subseteq> {x \<in> topspace X. f x \<in> V}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 784 |       using * [of "{x \<in> topspace X. f x \<in> V}"] assms openin_continuous_map_preimage x 
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 785 | by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 786 | then show "\<exists>V' \<in> (image f) ` \<B>. y \<in> V' \<and> V' \<subseteq> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 787 | using image_mono x by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 788 | qed (use \<B> openXYf in force)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 789 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 790 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 791 | lemma homeomorphic_space_first_countability: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 792 | "X homeomorphic_space Y \<Longrightarrow> first_countable X \<longleftrightarrow> first_countable Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 793 | by (meson first_countable_open_map_image homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 794 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 795 | lemma first_countable_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 | "\<lbrakk>retraction_map X Y r; first_countable X\<rbrakk> \<Longrightarrow> first_countable Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 | using first_countable_subtopology hereditary_imp_retractive_property homeomorphic_space_first_countability by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 798 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 799 | lemma separable_space_open_subset: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | assumes "separable_space X" "openin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 801 | shows "separable_space (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 802 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 803 | obtain C where C: "countable C" "C \<subseteq> topspace X" "X closure_of C = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 804 | by (meson assms separable_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 805 | then have "\<And>x T. \<lbrakk>x \<in> topspace X; x \<in> T; openin (subtopology X S) T\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 806 | \<Longrightarrow> \<exists>y. y \<in> S \<and> y \<in> C \<and> y \<in> T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 807 | by (smt (verit) \<open>openin X S\<close> in_closure_of openin_open_subtopology subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 808 | with C \<open>openin X S\<close> show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 809 | unfolding separable_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 810 | by (rule_tac x="S \<inter> C" in exI) (force simp: in_closure_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 811 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 812 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 | lemma separable_space_continuous_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | assumes "separable_space X" "continuous_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 | and fim: "f ` (topspace X) = topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 816 | shows "separable_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 817 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 818 | have cont: "\<And>S. f ` (X closure_of S) \<subseteq> Y closure_of f ` S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 819 | by (simp add: assms continuous_map_image_closure_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 820 | obtain C where C: "countable C" "C \<subseteq> topspace X" "X closure_of C = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 821 | by (meson assms separable_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 822 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 823 | unfolding separable_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 824 | by (metis cont fim closure_of_subset_topspace countable_image image_mono subset_antisym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 825 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 827 | lemma separable_space_quotient_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 828 | "\<lbrakk>quotient_map X Y q; separable_space X\<rbrakk> \<Longrightarrow> separable_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 829 | by (meson quotient_imp_continuous_map quotient_imp_surjective_map separable_space_continuous_map_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 830 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 831 | lemma separable_space_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 832 | "\<lbrakk>retraction_map X Y r; separable_space X\<rbrakk> \<Longrightarrow> separable_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 833 | using retraction_imp_quotient_map separable_space_quotient_map_image by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 834 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 835 | lemma homeomorphic_separable_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 836 | "X homeomorphic_space Y \<Longrightarrow> (separable_space X \<longleftrightarrow> separable_space Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 | by (meson homeomorphic_eq_everything_map homeomorphic_maps_map homeomorphic_space_def separable_space_continuous_map_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 838 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 839 | lemma separable_space_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 840 | "separable_space(discrete_topology U) \<longleftrightarrow> countable U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 841 | by (metis countable_Int2 discrete_topology_closure_of dual_order.refl inf.orderE separable_space_def topspace_discrete_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 842 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 843 | lemma second_countable_imp_separable_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 844 | assumes "second_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 845 | shows "separable_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 846 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 | obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 848 | and *: "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 849 | using assms by (auto simp: second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 850 |   obtain c where c: "\<And>V. \<lbrakk>V \<in> \<B>; V \<noteq> {}\<rbrakk> \<Longrightarrow> c V \<in> V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 851 | by (metis all_not_in_conv) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 852 |   then have **: "\<And>x. x \<in> topspace X \<Longrightarrow> x \<in> X closure_of c ` (\<B> - {{}})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 853 | using * by (force simp: closure_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 854 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 855 | unfolding separable_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 856 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 857 |     show "countable (c ` (\<B>-{{}}))"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 858 | using \<B>(1) by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 859 |     show "(c ` (\<B>-{{}})) \<subseteq> topspace X"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 860 | using \<B>(2) c openin_subset by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 861 |     show "X closure_of (c ` (\<B>-{{}})) = topspace X"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 862 | by (meson ** closure_of_subset_topspace subsetI subset_antisym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 863 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 864 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 865 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 | lemma second_countable_imp_Lindelof_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 867 | assumes "second_countable X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 868 | shows "Lindelof_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 869 | unfolding Lindelof_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 870 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 871 | fix \<U> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 872 | assume "\<forall>U \<in> \<U>. openin X U" and UU: "\<Union>\<U> = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 873 | obtain \<B> where \<B>: "countable \<B>" "\<And>V. V \<in> \<B> \<Longrightarrow> openin X V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 874 | and *: "\<And>U x. openin X U \<and> x \<in> U \<longrightarrow> (\<exists>V \<in> \<B>. x \<in> V \<and> V \<subseteq> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 875 | using assms by (auto simp: second_countable_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 |   define \<B>' where "\<B>' = {B \<in> \<B>. \<exists>U. U \<in> \<U> \<and> B \<subseteq> U}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | have \<B>': "countable \<B>'" "\<Union>\<B>' = \<Union>\<U>" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 878 | using \<B> using "*" \<open>\<forall>U\<in>\<U>. openin X U\<close> by (fastforce simp: \<B>'_def)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 879 | have "\<And>b. \<exists>U. b \<in> \<B>' \<longrightarrow> U \<in> \<U> \<and> b \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 880 | by (simp add: \<B>'_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 881 | then obtain G where G: "\<And>b. b \<in> \<B>' \<longrightarrow> G b \<in> \<U> \<and> b \<subseteq> G b" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 882 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | with \<B>' UU show "\<exists>\<V>. countable \<V> \<and> \<V> \<subseteq> \<U> \<and> \<Union>\<V> = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 884 | by (rule_tac x="G ` \<B>'" in exI) fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 885 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 886 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 887 | subsection \<open>Neigbourhood bases EXTRAS\<close> | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 888 | |
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 889 | text \<open>Neigbourhood bases: useful for "local" properties of various kinds\<close> | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 890 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 891 | lemma openin_topology_neighbourhood_base_unique: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 892 | "openin X = arbitrary union_of P \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 893 | (\<forall>u. P u \<longrightarrow> openin X u) \<and> neighbourhood_base_of P X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 894 | by (smt (verit, best) open_neighbourhood_base_of openin_topology_base_unique) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 895 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 896 | lemma neighbourhood_base_at_topology_base: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 897 | " openin X = arbitrary union_of b | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 898 | \<Longrightarrow> (neighbourhood_base_at x P X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 899 | (\<forall>w. b w \<and> x \<in> w \<longrightarrow> (\<exists>u v. openin X u \<and> P v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w)))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 900 | apply (simp add: neighbourhood_base_at_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 901 | by (smt (verit, del_insts) openin_topology_base_unique subset_trans) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 902 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 903 | lemma neighbourhood_base_of_unlocalized: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 904 |   assumes "\<And>S t. P S \<and> openin X t \<and> (t \<noteq> {}) \<and> t \<subseteq> S \<Longrightarrow> P t"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 905 | shows "neighbourhood_base_of P X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | (\<forall>x \<in> topspace X. \<exists>u v. openin X u \<and> P v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> topspace X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 907 | apply (simp add: neighbourhood_base_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | by (smt (verit, ccfv_SIG) assms empty_iff neighbourhood_base_at_unlocalized) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 909 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 910 | lemma neighbourhood_base_at_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 911 |    "neighbourhood_base_at x P (discrete_topology u) \<longleftrightarrow> x \<in> u \<Longrightarrow> P {x}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | apply (simp add: neighbourhood_base_at_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 913 | by (smt (verit) empty_iff empty_subsetI insert_subset singletonI subsetD subset_singletonD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 914 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 915 | lemma neighbourhood_base_of_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 |    "neighbourhood_base_of P (discrete_topology u) \<longleftrightarrow> (\<forall>x \<in> u. P {x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 917 | apply (simp add: neighbourhood_base_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 918 | using neighbourhood_base_at_discrete_topology[of _ P u] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 919 | by (metis empty_subsetI insert_subset neighbourhood_base_at_def openin_discrete_topology singletonI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 920 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 921 | lemma second_countable_neighbourhood_base_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 922 | "second_countable X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 923 | (\<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and> neighbourhood_base_of (\<lambda>A. A\<in>\<B>) X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 924 | by (metis (full_types) openin_topology_neighbourhood_base_unique second_countable) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 925 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 926 | lemma first_countable_neighbourhood_base_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 | "first_countable X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 928 | (\<forall>x \<in> topspace X. \<exists>\<B>. countable \<B> \<and> (\<forall>V \<in> \<B>. openin X V) \<and> neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 929 | unfolding first_countable_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | apply (intro ball_cong refl ex_cong conj_cong) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 931 | by (metis (mono_tags, lifting) open_neighbourhood_base_at) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 932 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 933 | lemma second_countable_neighbourhood_base: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 934 | "second_countable X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 935 | (\<exists>\<B>. countable \<B> \<and> neighbourhood_base_of (\<lambda>V. V \<in> \<B>) X)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 936 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 937 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 938 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 939 | using second_countable_neighbourhood_base_alt by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 940 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 941 | assume ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 942 | then obtain \<B> where "countable \<B>" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 943 | and \<B>: "\<And>W x. openin X W \<and> x \<in> W \<longrightarrow> (\<exists>U. openin X U \<and> (\<exists>V. V \<in> \<B> \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 944 | by (metis neighbourhood_base_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 945 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 946 | unfolding second_countable_neighbourhood_base_alt neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 947 | apply (rule_tac x="(\<lambda>u. X interior_of u) ` \<B>" in exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 948 | by (smt (verit, best) interior_of_eq interior_of_mono countable_image image_iff openin_interior_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 949 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 950 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 951 | lemma first_countable_neighbourhood_base: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 952 | "first_countable X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 953 | (\<forall>x \<in> topspace X. \<exists>\<B>. countable \<B> \<and> neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) X)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 954 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 955 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 956 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 957 | by (metis first_countable_neighbourhood_base_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 958 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 959 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 960 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 961 | unfolding first_countable_neighbourhood_base_alt | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 962 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 963 | fix x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 964 | assume "x \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 965 | with R obtain \<B> where "countable \<B>" and \<B>: "neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 966 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 967 | then | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 968 | show "\<exists>\<B>. countable \<B> \<and> Ball \<B> (openin X) \<and> neighbourhood_base_at x (\<lambda>V. V \<in> \<B>) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 969 | unfolding neighbourhood_base_at_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 970 | apply (rule_tac x="(\<lambda>u. X interior_of u) ` \<B>" in exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 971 | by (smt (verit, best) countable_image image_iff interior_of_eq interior_of_mono openin_interior_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 972 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 973 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 974 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 975 | |
| 77942 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 976 | subsection\<open>$T_0$ spaces and the Kolmogorov quotient\<close> | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 978 | definition t0_space where | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 979 | "t0_space X \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 980 | \<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x \<noteq> y \<longrightarrow> (\<exists>U. openin X U \<and> (x \<notin> U \<longleftrightarrow> y \<in> U))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 981 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 982 | lemma t0_space_expansive: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 983 | "\<lbrakk>topspace Y = topspace X; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> \<Longrightarrow> t0_space X \<Longrightarrow> t0_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 984 | by (metis t0_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 985 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 986 | lemma t1_imp_t0_space: "t1_space X \<Longrightarrow> t0_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 987 | by (metis t0_space_def t1_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 988 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 989 | lemma t1_eq_symmetric_t0_space_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 990 | "t1_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 991 | t0_space X \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 992 |       (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x \<in> X closure_of {y} \<longleftrightarrow> y \<in> X closure_of {x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 993 | apply (simp add: t0_space_def t1_space_def closure_of_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 994 | by (smt (verit, best) openin_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 995 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 996 | lemma t1_eq_symmetric_t0_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 997 |   "t1_space X \<longleftrightarrow> t0_space X \<and> (\<forall>x y. x \<in> X closure_of {y} \<longleftrightarrow> y \<in> X closure_of {x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 998 | by (auto simp: t1_eq_symmetric_t0_space_alt in_closure_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 999 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1000 | lemma Hausdorff_imp_t0_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1001 | "Hausdorff_space X \<Longrightarrow> t0_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1002 | by (simp add: Hausdorff_imp_t1_space t1_imp_t0_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1003 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1004 | lemma t0_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1005 | "t0_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1006 | (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. x \<noteq> y \<longrightarrow> (\<exists>C. closedin X C \<and> (x \<notin> C \<longleftrightarrow> y \<in> C)))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | unfolding t0_space_def by (metis Diff_iff closedin_def openin_closedin_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1008 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1009 | lemma homeomorphic_t0_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1010 | assumes "X homeomorphic_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1011 | shows "t0_space X \<longleftrightarrow> t0_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1013 | obtain f where f: "homeomorphic_map X Y f" and F: "inj_on f (topspace X)" and "topspace Y = f ` topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1014 | by (metis assms homeomorphic_imp_injective_map homeomorphic_imp_surjective_map homeomorphic_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1015 | with inj_on_image_mem_iff [OF F] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1016 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1017 | apply (simp add: t0_space_def homeomorphic_eq_everything_map continuous_map_def open_map_def inj_on_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1018 | by (smt (verit) mem_Collect_eq openin_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1019 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1020 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1021 | lemma t0_space_closure_of_sing: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1022 | "t0_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1023 |     (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. X closure_of {x} = X closure_of {y} \<longrightarrow> x = y)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1024 | by (simp add: t0_space_def closure_of_def set_eq_iff) (smt (verit)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1025 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1026 | lemma t0_space_discrete_topology: "t0_space (discrete_topology S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1027 | by (simp add: Hausdorff_imp_t0_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1028 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1029 | lemma t0_space_subtopology: "t0_space X \<Longrightarrow> t0_space (subtopology X U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1030 | by (simp add: t0_space_def openin_subtopology) (metis Int_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1031 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1032 | lemma t0_space_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1033 | "\<lbrakk>retraction_map X Y r; t0_space X\<rbrakk> \<Longrightarrow> t0_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1034 | using hereditary_imp_retractive_property homeomorphic_t0_space t0_space_subtopology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1035 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1036 | lemma t0_space_prod_topologyI: "\<lbrakk>t0_space X; t0_space Y\<rbrakk> \<Longrightarrow> t0_space (prod_topology X Y)" | 
| 79492 
c1b0f64eb865
A few new results (mostly brought in from other developments)
 paulson <lp15@cam.ac.uk> parents: 
78517diff
changeset | 1037 | by (simp add: t0_space_closure_of_sing closure_of_Times closure_of_eq_empty_gen times_eq_iff flip: sing_Times_sing insert_Times_insert) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1038 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1039 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1040 | lemma t0_space_prod_topology_iff: | 
| 78336 | 1041 | "t0_space (prod_topology X Y) \<longleftrightarrow> prod_topology X Y = trivial_topology \<or> t0_space X \<and> t0_space Y" (is "?lhs=?rhs") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1042 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1043 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1044 | then show ?rhs | 
| 78336 | 1045 | by (metis prod_topology_trivial_iff retraction_map_fst retraction_map_snd t0_space_retraction_map_image) | 
| 1046 | qed (metis t0_space_discrete_topology t0_space_prod_topologyI) | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1047 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1048 | proposition t0_space_product_topology: | 
| 78336 | 1049 | "t0_space (product_topology X I) \<longleftrightarrow> product_topology X I = trivial_topology \<or> (\<forall>i \<in> I. t0_space (X i))" | 
| 1050 | (is "?lhs=?rhs") | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1051 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1052 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1053 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1054 | by (meson retraction_map_product_projection t0_space_retraction_map_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1055 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1056 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1057 | show ?lhs | 
| 78336 | 1058 | proof (cases "product_topology X I = trivial_topology") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1060 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1061 | by (simp add: t0_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1062 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1063 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1065 | unfolding t0_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1066 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1067 | fix x y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1068 | assume x: "x \<in> topspace (product_topology X I)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 | and y: "y \<in> topspace (product_topology X I)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 | and "x \<noteq> y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1071 | then obtain i where "i \<in> I" "x i \<noteq> y i" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 | by (metis PiE_ext topspace_product_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | then have "t0_space (X i)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1074 | using False R by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1075 | then obtain U where "closedin (X i) U" "(x i \<notin> U \<longleftrightarrow> y i \<in> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1076 | by (metis t0_space PiE_mem \<open>i \<in> I\<close> \<open>x i \<noteq> y i\<close> topspace_product_topology x y) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1077 | with \<open>i \<in> I\<close> x y show "\<exists>U. closedin (product_topology X I) U \<and> (x \<notin> U) = (y \<in> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1078 | by (rule_tac x="PiE I (\<lambda>j. if j = i then U else topspace(X j))" in exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | (simp add: closedin_product_topology PiE_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1080 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1081 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1082 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1083 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1085 | subsection \<open>Kolmogorov quotients\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1086 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1087 | definition Kolmogorov_quotient | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1088 | where "Kolmogorov_quotient X \<equiv> \<lambda>x. @y. \<forall>U. openin X U \<longrightarrow> (y \<in> U \<longleftrightarrow> x \<in> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1089 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1090 | lemma Kolmogorov_quotient_in_open: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1091 | "openin X U \<Longrightarrow> (Kolmogorov_quotient X x \<in> U \<longleftrightarrow> x \<in> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1092 | by (smt (verit, ccfv_SIG) Kolmogorov_quotient_def someI_ex) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1093 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1094 | lemma Kolmogorov_quotient_in_topspace: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1095 | "Kolmogorov_quotient X x \<in> topspace X \<longleftrightarrow> x \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | by (simp add: Kolmogorov_quotient_in_open) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1097 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1098 | lemma Kolmogorov_quotient_in_closed: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1099 | "closedin X C \<Longrightarrow> (Kolmogorov_quotient X x \<in> C \<longleftrightarrow> x \<in> C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1100 | unfolding closedin_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1101 | by (meson DiffD2 DiffI Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace in_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1102 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 | lemma continuous_map_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1104 | "continuous_map X X (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1105 | using Kolmogorov_quotient_in_open openin_subopen openin_subset | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1106 | by (fastforce simp: continuous_map_def Kolmogorov_quotient_in_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1107 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1108 | lemma open_map_Kolmogorov_quotient_explicit: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1109 | "openin X U \<Longrightarrow> Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \<inter> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1110 | using Kolmogorov_quotient_in_open openin_subset by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1111 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1112 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1113 | lemma open_map_Kolmogorov_quotient_gen: | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1114 | "open_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1115 | proof (clarsimp simp: open_map_def openin_subtopology_alt image_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1116 | fix U | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1117 | assume "openin X U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1118 | then have "Kolmogorov_quotient X ` (S \<inter> U) = Kolmogorov_quotient X ` S \<inter> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1119 | using Kolmogorov_quotient_in_open [of X U] by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1120 | then show "\<exists>V. openin X V \<and> Kolmogorov_quotient X ` (S \<inter> U) = Kolmogorov_quotient X ` S \<inter> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1121 | using \<open>openin X U\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1122 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1123 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1124 | lemma open_map_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1125 | "open_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1126 | (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1127 | by (metis open_map_Kolmogorov_quotient_gen subtopology_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1128 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1129 | lemma closed_map_Kolmogorov_quotient_explicit: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1130 | "closedin X U \<Longrightarrow> Kolmogorov_quotient X ` U = Kolmogorov_quotient X ` topspace X \<inter> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1131 | using closedin_subset by (fastforce simp: Kolmogorov_quotient_in_closed) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1132 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1133 | lemma closed_map_Kolmogorov_quotient_gen: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1134 | "closed_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1135 | (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1136 | using Kolmogorov_quotient_in_closed by (force simp: closed_map_def closedin_subtopology_alt image_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1137 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1138 | lemma closed_map_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1139 | "closed_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1140 | (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1141 | by (metis closed_map_Kolmogorov_quotient_gen subtopology_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1142 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1143 | lemma quotient_map_Kolmogorov_quotient_gen: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1144 | "quotient_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1145 | proof (intro continuous_open_imp_quotient_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1146 | show "continuous_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1147 | by (simp add: continuous_map_Kolmogorov_quotient continuous_map_from_subtopology continuous_map_in_subtopology image_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1148 | show "open_map (subtopology X S) (subtopology X (Kolmogorov_quotient X ` S)) (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1149 | using open_map_Kolmogorov_quotient_gen by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1150 | show "Kolmogorov_quotient X ` topspace (subtopology X S) = topspace (subtopology X (Kolmogorov_quotient X ` S))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1151 | by (force simp: Kolmogorov_quotient_in_open) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1152 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1153 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1154 | lemma quotient_map_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1155 | "quotient_map X (subtopology X (Kolmogorov_quotient X ` topspace X)) (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1156 | by (metis quotient_map_Kolmogorov_quotient_gen subtopology_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1157 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1158 | lemma Kolmogorov_quotient_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1159 | "Kolmogorov_quotient X x = Kolmogorov_quotient X y \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1160 | (\<forall>U. openin X U \<longrightarrow> (x \<in> U \<longleftrightarrow> y \<in> U))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1161 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1162 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1163 | by (metis Kolmogorov_quotient_in_open) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1164 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1165 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1166 | by (simp add: Kolmogorov_quotient_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1167 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1169 | lemma Kolmogorov_quotient_eq_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1170 | "Kolmogorov_quotient X x = Kolmogorov_quotient X y \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1171 | (\<forall>U. closedin X U \<longrightarrow> (x \<in> U \<longleftrightarrow> y \<in> U))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1172 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1173 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1174 | by (metis Kolmogorov_quotient_in_closed) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1175 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1176 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1177 | by (smt (verit) Diff_iff Kolmogorov_quotient_eq closedin_topspace in_mono openin_closedin_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1178 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1179 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1180 | lemma Kolmogorov_quotient_continuous_map: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1181 | assumes "continuous_map X Y f" "t0_space Y" and x: "x \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1182 | shows "f (Kolmogorov_quotient X x) = f x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1183 | using assms unfolding continuous_map_def t0_space_def | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1184 | by (smt (verit, ccfv_threshold) Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace PiE mem_Collect_eq) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1185 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1186 | lemma t0_space_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1187 | "t0_space (subtopology X (Kolmogorov_quotient X ` topspace X))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1188 | apply (clarsimp simp: t0_space_def ) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1189 | by (smt (verit, best) Kolmogorov_quotient_eq imageE image_eqI open_map_Kolmogorov_quotient open_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1190 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1191 | lemma Kolmogorov_quotient_id: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1192 | "t0_space X \<Longrightarrow> x \<in> topspace X \<Longrightarrow> Kolmogorov_quotient X x = x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1193 | by (metis Kolmogorov_quotient_in_open Kolmogorov_quotient_in_topspace t0_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1194 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1195 | lemma Kolmogorov_quotient_idemp: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1196 | "Kolmogorov_quotient X (Kolmogorov_quotient X x) = Kolmogorov_quotient X x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1197 | by (simp add: Kolmogorov_quotient_eq Kolmogorov_quotient_in_open) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1198 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1199 | lemma retraction_maps_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1200 | "retraction_maps X | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1201 | (subtopology X (Kolmogorov_quotient X ` topspace X)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1202 | (Kolmogorov_quotient X) id" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1203 | unfolding retraction_maps_def continuous_map_in_subtopology | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1204 | using Kolmogorov_quotient_idemp continuous_map_Kolmogorov_quotient by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1206 | lemma retraction_map_Kolmogorov_quotient: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1207 | "retraction_map X | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1208 | (subtopology X (Kolmogorov_quotient X ` topspace X)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1209 | (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1210 | using retraction_map_def retraction_maps_Kolmogorov_quotient by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1211 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1212 | lemma retract_of_space_Kolmogorov_quotient_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1213 | "Kolmogorov_quotient X ` topspace X retract_of_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1214 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1215 | have "continuous_map X X (Kolmogorov_quotient X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1216 | by (simp add: continuous_map_Kolmogorov_quotient) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1217 | then have "Kolmogorov_quotient X ` topspace X \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1218 | by (simp add: continuous_map_image_subset_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1219 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1220 | by (meson retract_of_space_retraction_maps retraction_maps_Kolmogorov_quotient) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1221 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1222 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1223 | lemma Kolmogorov_quotient_lift_exists: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1224 | assumes "S \<subseteq> topspace X" "t0_space Y" and f: "continuous_map (subtopology X S) Y f" | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1225 | obtains g where "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) Y g" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1226 | "\<And>x. x \<in> S \<Longrightarrow> g(Kolmogorov_quotient X x) = f x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1227 | proof - | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1228 | have "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; Kolmogorov_quotient X x = Kolmogorov_quotient X y\<rbrakk> \<Longrightarrow> f x = f y" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 | using assms | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1230 | apply (simp add: Kolmogorov_quotient_eq t0_space_def continuous_map_def Int_absorb1 openin_subtopology) | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1231 | by (smt (verit, del_insts) Int_iff mem_Collect_eq Pi_iff) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1232 | then obtain g where g: "continuous_map (subtopology X (Kolmogorov_quotient X ` S)) Y g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1233 | "g ` (topspace X \<inter> Kolmogorov_quotient X ` S) = f ` S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1234 | "\<And>x. x \<in> S \<Longrightarrow> g (Kolmogorov_quotient X x) = f x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1235 | using quotient_map_lift_exists [OF quotient_map_Kolmogorov_quotient_gen [of X S] f] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1236 | by (metis assms(1) topspace_subtopology topspace_subtopology_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1237 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1238 | proof qed (use g in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1239 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1240 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1241 | subsection\<open>Closed diagonals and graphs\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1242 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1243 | lemma Hausdorff_space_closedin_diagonal: | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1244 | "Hausdorff_space X \<longleftrightarrow> closedin (prod_topology X X) ((\<lambda>x. (x,x)) ` topspace X)" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1245 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1246 | have \<section>: "((\<lambda>x. (x, x)) ` topspace X) \<subseteq> topspace X \<times> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1247 | by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1248 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1249 | apply (simp add: closedin_def openin_prod_topology_alt Hausdorff_space_def disjnt_iff \<section>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1250 | apply (intro all_cong1 imp_cong ex_cong1 conj_cong refl) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1251 | by (force dest!: openin_subset)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1252 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1253 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1254 | lemma closed_map_diag_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1255 | "closed_map X (prod_topology X X) (\<lambda>x. (x,x)) \<longleftrightarrow> Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1256 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1257 | have "section_map X (prod_topology X X) (\<lambda>x. (x, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1258 | unfolding section_map_def retraction_maps_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1259 | by (smt (verit) continuous_map_fst continuous_map_of_fst continuous_map_on_empty continuous_map_pairwise fst_conv fst_diag_fst snd_diag_fst) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1260 | then have "embedding_map X (prod_topology X X) (\<lambda>x. (x, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1261 | by (rule section_imp_embedding_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1262 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1263 | using Hausdorff_space_closedin_diagonal embedding_imp_closed_map_eq by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1264 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1265 | |
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1266 | lemma proper_map_diag_eq [simp]: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1267 | "proper_map X (prod_topology X X) (\<lambda>x. (x,x)) \<longleftrightarrow> Hausdorff_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1268 | by (simp add: closed_map_diag_eq inj_on_convol_ident injective_imp_proper_eq_closed_map) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1269 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1270 | lemma closedin_continuous_maps_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1271 | assumes "Hausdorff_space Y" and f: "continuous_map X Y f" and g: "continuous_map X Y g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1272 |   shows "closedin X {x \<in> topspace X. f x = g x}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1273 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1274 |   have \<section>:"{x \<in> topspace X. f x = g x} = {x \<in> topspace X. (f x,g x) \<in> ((\<lambda>y.(y,y)) ` topspace Y)}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1275 | using f continuous_map_image_subset_topspace by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1276 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1277 | unfolding \<section> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1278 | proof (intro closedin_continuous_map_preimage) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1279 | show "continuous_map X (prod_topology Y Y) (\<lambda>x. (f x, g x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1280 | by (simp add: continuous_map_pairedI f g) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1281 | show "closedin (prod_topology Y Y) ((\<lambda>y. (y, y)) ` topspace Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1282 | using Hausdorff_space_closedin_diagonal assms by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1283 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1284 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1285 | |
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1286 | lemma forall_in_closure_of: | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1287 | assumes "x \<in> X closure_of S" "\<And>x. x \<in> S \<Longrightarrow> P x" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1288 |     and "closedin X {x \<in> topspace X. P x}"
 | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1289 | shows "P x" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1290 | by (smt (verit, ccfv_threshold) Diff_iff assms closedin_def in_closure_of mem_Collect_eq) | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1291 | |
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1292 | lemma forall_in_closure_of_eq: | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1293 | assumes x: "x \<in> X closure_of S" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1294 | and Y: "Hausdorff_space Y" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1295 | and f: "continuous_map X Y f" and g: "continuous_map X Y g" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1296 | and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = g x" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1297 | shows "f x = g x" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1298 | proof - | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1299 |   have "closedin X {x \<in> topspace X. f x = g x}"
 | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1300 | using Y closedin_continuous_maps_eq f g by blast | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1301 | then show ?thesis | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1302 | using forall_in_closure_of [OF x fg] | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1303 | by fastforce | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1304 | qed | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 1305 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1306 | lemma retract_of_space_imp_closedin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1307 | assumes "Hausdorff_space X" and S: "S retract_of_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1308 | shows "closedin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1309 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1310 | obtain r where r: "continuous_map X (subtopology X S) r" "\<forall>x\<in>S. r x = x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1311 | using assms by (meson retract_of_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1312 |   then have \<section>: "S = {x \<in> topspace X. r x = x}"
 | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1313 | using S retract_of_space_imp_subset by (force simp: continuous_map_def Pi_iff) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1314 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1315 | unfolding \<section> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1316 | using r continuous_map_into_fulltopology assms | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1317 | by (force intro: closedin_continuous_maps_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1318 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1319 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1320 | lemma homeomorphic_maps_graph: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1321 | "homeomorphic_maps X (subtopology (prod_topology X Y) ((\<lambda>x. (x, f x)) ` (topspace X))) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1322 | (\<lambda>x. (x, f x)) fst \<longleftrightarrow> continuous_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1323 | (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1324 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1325 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1326 | then | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1327 | have h: "homeomorphic_map X (subtopology (prod_topology X Y) ((\<lambda>x. (x, f x)) ` topspace X)) (\<lambda>x. (x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1328 | by (auto simp: homeomorphic_maps_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1329 | have "f = snd \<circ> (\<lambda>x. (x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1330 | by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1331 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1332 | by (metis (no_types, lifting) h continuous_map_in_subtopology continuous_map_snd_of homeomorphic_eq_everything_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1333 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1334 | assume ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1335 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1336 | unfolding homeomorphic_maps_def | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1337 | by (smt (verit, del_insts) continuous_map_eq continuous_map_subtopology_fst embedding_map_def | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1338 | embedding_map_graph homeomorphic_eq_everything_map image_cong image_iff prod.sel(1)) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1340 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1341 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1342 | subsection \<open> KC spaces, those where all compact sets are closed.\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1343 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1344 | definition kc_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 | where "kc_space X \<equiv> \<forall>S. compactin X S \<longrightarrow> closedin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1346 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 1347 | lemma kc_space_euclidean: "kc_space (euclidean :: 'a::metric_space topology)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 1348 | by (simp add: compact_imp_closed kc_space_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 1349 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1350 | lemma kc_space_expansive: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1351 | "\<lbrakk>kc_space X; topspace Y = topspace X; \<And>U. openin X U \<Longrightarrow> openin Y U\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1352 | \<Longrightarrow> kc_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 | by (meson compactin_contractive kc_space_def topology_finer_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1354 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1355 | lemma compactin_imp_closedin_gen: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 | "\<lbrakk>kc_space X; compactin X S\<rbrakk> \<Longrightarrow> closedin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1357 | using kc_space_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1359 | lemma Hausdorff_imp_kc_space: "Hausdorff_space X \<Longrightarrow> kc_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1360 | by (simp add: compactin_imp_closedin kc_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1361 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1362 | lemma kc_imp_t1_space: "kc_space X \<Longrightarrow> t1_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1363 | by (simp add: finite_imp_compactin kc_space_def t1_space_closedin_finite) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1364 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1365 | lemma kc_space_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | "kc_space X \<Longrightarrow> kc_space(subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1367 | by (metis closedin_Int_closure_of closure_of_eq compactin_subtopology inf.absorb2 kc_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1368 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1369 | lemma kc_space_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1370 | "kc_space(discrete_topology U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1371 | using Hausdorff_space_discrete_topology compactin_imp_closedin kc_space_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1372 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1373 | lemma kc_space_continuous_injective_map_preimage: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1374 | assumes "kc_space Y" "continuous_map X Y f" and injf: "inj_on f (topspace X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1375 | shows "kc_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1376 | unfolding kc_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1377 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1378 | fix S | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1379 | assume S: "compactin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 |   have "S = {x \<in> topspace X. f x \<in> f ` S}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1381 | using S compactin_subset_topspace inj_onD [OF injf] by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1382 | with assms S show "closedin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1383 | by (metis (no_types, lifting) Collect_cong closedin_continuous_map_preimage compactin_imp_closedin_gen image_compactin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1384 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1385 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1386 | lemma kc_space_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1387 | assumes "retraction_map X Y r" "kc_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1388 | shows "kc_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1389 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1390 | obtain s where s: "continuous_map X Y r" "continuous_map Y X s" "\<And>x. x \<in> topspace Y \<Longrightarrow> r (s x) = x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1391 | using assms by (force simp: retraction_map_def retraction_maps_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1392 | then have inj: "inj_on s (topspace Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1393 | by (metis inj_on_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1394 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1395 | unfolding kc_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1396 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1397 | fix S | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1398 | assume S: "compactin Y S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1399 |     have "S = {x \<in> topspace Y. s x \<in> s ` S}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1400 | using S compactin_subset_topspace inj_onD [OF inj] by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1401 | with assms S show "closedin Y S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1402 | by (meson compactin_imp_closedin_gen inj kc_space_continuous_injective_map_preimage s(2)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1403 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1404 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1405 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1406 | lemma homeomorphic_kc_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1407 | "X homeomorphic_space Y \<Longrightarrow> kc_space X \<longleftrightarrow> kc_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1408 | by (meson homeomorphic_eq_everything_map homeomorphic_space homeomorphic_space_sym kc_space_continuous_injective_map_preimage) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1409 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1410 | lemma compact_kc_eq_maximal_compact_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1411 | assumes "compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1412 | shows "kc_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1413 | (\<forall>Y. topspace Y = topspace X \<and> (\<forall>S. openin X S \<longrightarrow> openin Y S) \<and> compact_space Y \<longrightarrow> Y = X)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1414 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1415 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1416 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1417 | by (metis closedin_compact_space compactin_contractive kc_space_def topology_eq topology_finer_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1418 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1419 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1420 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1421 | unfolding kc_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1422 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1423 | fix S | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1424 | assume S: "compactin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1425 | define Y where | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1426 | "Y \<equiv> topology (arbitrary union_of (finite intersection_of (\<lambda>A. A = topspace X - S \<or> openin X A) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1427 | relative_to (topspace X)))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1428 | have "topspace Y = topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1429 | by (auto simp: Y_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1430 | have "openin X T \<longrightarrow> openin Y T" for T | 
| 78250 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1431 | by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase openin_subset relative_to_subset_inc) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1432 | have "compact_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1433 | proof (rule Alexander_subbase_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1434 | show "\<exists>\<F>'. finite \<F>' \<and> \<F>' \<subseteq> \<C> \<and> topspace X \<subseteq> \<Union> \<F>'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1435 | if \<C>: "\<C> \<subseteq> insert (topspace X - S) (Collect (openin X))" and sub: "topspace X \<subseteq> \<Union>\<C>" for \<C> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1436 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1437 | consider "\<C> \<subseteq> Collect (openin X)" | \<V> where "\<C> = insert (topspace X - S) \<V>" "\<V> \<subseteq> Collect (openin X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1438 | using \<C> by (metis insert_Diff subset_insert_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1439 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1440 | proof cases | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1441 | case 1 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1442 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1443 | by (metis assms compact_space_alt mem_Collect_eq subsetD that(2)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1444 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1445 | case 2 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1446 | then have "S \<subseteq> \<Union>\<V>" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1447 | using S sub compactin_subset_topspace by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1448 | with 2 obtain \<F> where "finite \<F> \<and> \<F> \<subseteq> \<V> \<and> S \<subseteq> \<Union>\<F>" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1449 | using S unfolding compactin_def by (metis Ball_Collect) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1450 | with 2 show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1451 | by (rule_tac x="insert (topspace X - S) \<F>" in exI) blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1452 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1453 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1454 | qed (auto simp: Y_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1455 | have "Y = X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1456 | using R \<open>\<And>S. openin X S \<longrightarrow> openin Y S\<close> \<open>compact_space Y\<close> \<open>topspace Y = topspace X\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1457 | moreover have "openin Y (topspace X - S)" | 
| 78250 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 1458 | by (simp add: Y_def arbitrary_union_of_inc finite_intersection_of_inc openin_subbase relative_to_subset_inc) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1459 | ultimately show "closedin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1460 | unfolding closedin_def using S compactin_subset_topspace by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1461 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1462 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1463 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1464 | lemma continuous_imp_closed_map_gen: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1465 | "\<lbrakk>compact_space X; kc_space Y; continuous_map X Y f\<rbrakk> \<Longrightarrow> closed_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1466 | by (meson closed_map_def closedin_compact_space compactin_imp_closedin_gen image_compactin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1467 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1468 | lemma kc_space_compact_subtopologies: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1469 | "kc_space X \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> kc_space(subtopology X K))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1470 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1471 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1472 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1473 | by (auto simp: kc_space_def closedin_closed_subtopology compactin_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1474 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1475 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1476 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1477 | unfolding kc_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1478 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1479 | fix K | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1480 | assume K: "compactin X K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1481 | then have "K \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1482 | by (simp add: compactin_subset_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1483 | moreover have "X closure_of K \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1484 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1485 | fix x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1486 | assume x: "x \<in> X closure_of K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1487 | have "kc_space (subtopology X K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1488 | by (simp add: R \<open>compactin X K\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1489 | have "compactin X (insert x K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1490 | by (metis K x compactin_Un compactin_sing in_closure_of insert_is_Un) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1491 | then show "x \<in> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1492 | by (metis R x K Int_insert_left_if1 closedin_Int_closure_of compact_imp_compactin_subtopology | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1493 | insertCI kc_space_def subset_insertI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1494 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1495 | ultimately show "closedin X K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1496 | using closure_of_subset_eq by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1497 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1498 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1499 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1500 | lemma kc_space_compact_prod_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1501 | assumes "compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1502 | shows "kc_space(prod_topology X X) \<longleftrightarrow> Hausdorff_space X" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1503 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1504 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1505 | show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1506 | unfolding closed_map_diag_eq [symmetric] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1507 | proof (intro continuous_imp_closed_map_gen) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1508 | show "continuous_map X (prod_topology X X) (\<lambda>x. (x, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1509 | by (intro continuous_intros) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1510 | qed (use L assms in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1511 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1512 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1513 | by (simp add: Hausdorff_imp_kc_space Hausdorff_space_prod_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1514 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1515 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1516 | lemma kc_space_prod_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1517 | "kc_space(prod_topology X X) \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> Hausdorff_space(subtopology X K))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1518 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1519 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1520 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1521 | by (metis compactin_subspace kc_space_compact_prod_topology kc_space_subtopology subtopology_Times) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1522 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1523 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1524 | have "kc_space (subtopology (prod_topology X X) L)" if "compactin (prod_topology X X) L" for L | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1525 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1526 | define K where "K \<equiv> fst ` L \<union> snd ` L" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1527 | have "L \<subseteq> K \<times> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1528 | by (force simp: K_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1529 | have "compactin X K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1530 | by (metis K_def compactin_Un continuous_map_fst continuous_map_snd image_compactin that) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1531 | then have "Hausdorff_space (subtopology X K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1532 | by (simp add: R) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1533 | then have "kc_space (prod_topology (subtopology X K) (subtopology X K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1534 | by (simp add: \<open>compactin X K\<close> compact_space_subtopology kc_space_compact_prod_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1535 | then have "kc_space (subtopology (prod_topology (subtopology X K) (subtopology X K)) L)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1536 | using kc_space_subtopology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1537 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1538 | using \<open>L \<subseteq> K \<times> K\<close> subtopology_Times subtopology_subtopology | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1539 | by (metis (no_types, lifting) Sigma_cong inf.absorb_iff2) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1540 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1541 | then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1542 | using kc_space_compact_subtopologies by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1543 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1544 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1545 | lemma kc_space_prod_topology_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1546 | "kc_space(prod_topology X X) \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1547 | kc_space X \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1548 | (\<forall>K. compactin X K \<longrightarrow> Hausdorff_space(subtopology X K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1549 | using Hausdorff_imp_kc_space kc_space_compact_subtopologies kc_space_prod_topology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1550 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1551 | proposition kc_space_prod_topology_left: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1552 | assumes X: "kc_space X" and Y: "Hausdorff_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1553 | shows "kc_space (prod_topology X Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1554 | unfolding kc_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1555 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1556 | fix K | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1557 | assume K: "compactin (prod_topology X Y) K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1558 | then have "K \<subseteq> topspace X \<times> topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1559 | using compactin_subset_topspace topspace_prod_topology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1560 | moreover have "\<exists>T. openin (prod_topology X Y) T \<and> (a,b) \<in> T \<and> T \<subseteq> (topspace X \<times> topspace Y) - K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1561 | if ab: "(a,b) \<in> (topspace X \<times> topspace Y) - K" for a b | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1562 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1563 |     have "compactin Y {b}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1564 | using that by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1565 | moreover | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1566 |     have "compactin Y {y \<in> topspace Y. (a,y) \<in> K}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1567 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1568 |       have "compactin (prod_topology X Y) (K \<inter> {a} \<times> topspace Y)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1569 | using that compact_Int_closedin [OF K] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1570 | by (simp add: X closedin_prod_Times_iff compactin_imp_closedin_gen) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1571 |       moreover have "subtopology (prod_topology X Y) (K \<inter> {a} \<times> topspace Y)  homeomorphic_space 
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1572 |                      subtopology Y {y \<in> topspace Y. (a, y) \<in> K}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1573 | unfolding homeomorphic_space_def homeomorphic_maps_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1574 | using that | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1575 | apply (rule_tac x="snd" in exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1576 | apply (rule_tac x="Pair a" in exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1577 | by (force simp: continuous_map_in_subtopology continuous_map_from_subtopology continuous_map_subtopology_snd continuous_map_paired) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1578 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1579 | by (simp add: compactin_subspace homeomorphic_compact_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1580 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1581 |     moreover have "disjnt {b} {y \<in> topspace Y. (a,y) \<in> K}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1582 | using ab by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1583 |     ultimately obtain V U where VU: "openin Y V" "openin Y U" "{b} \<subseteq> V" "{y \<in> topspace Y. (a,y) \<in> K} \<subseteq> U" "disjnt V U"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1584 | using Hausdorff_space_compact_separation [OF Y] by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1585 | define V' where "V' \<equiv> topspace Y - U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1586 |     have W: "closedin Y V'" "{y \<in> topspace Y. (a,y) \<in> K} \<subseteq> topspace Y - V'" "disjnt V (topspace Y - V')"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1587 | using VU by (auto simp: V'_def disjnt_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1588 | with VU obtain "V \<subseteq> topspace Y" "V' \<subseteq> topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1589 | by (meson closedin_subset openin_closedin_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1590 |     then obtain "b \<in> V" "disjnt {y \<in> topspace Y. (a,y) \<in> K} V'" "V \<subseteq> V'"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1591 | using VU unfolding disjnt_iff V'_def by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1592 |     define C where "C \<equiv> image fst (K \<inter> {z \<in> topspace(prod_topology X Y). snd z \<in> V'})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1593 |     have "closedin (prod_topology X Y) {z \<in> topspace (prod_topology X Y). snd z \<in> V'}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1594 | using closedin_continuous_map_preimage \<open>closedin Y V'\<close> continuous_map_snd by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1595 | then have "compactin X C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1596 | unfolding C_def by (meson K compact_Int_closedin continuous_map_fst image_compactin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1597 | then have "closedin X C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1598 | using assms by (auto simp: kc_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1599 | show ?thesis | 
| 77942 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1600 | proof (intro exI conjI) | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1601 | show "openin (prod_topology X Y) ((topspace X - C) \<times> V)" | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1602 | by (simp add: VU \<open>closedin X C\<close> openin_diff openin_prod_Times_iff) | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1603 | have "a \<notin> C" | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1604 | using VU by (auto simp: C_def V'_def) | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1605 | then show "(a, b) \<in> (topspace X - C) \<times> V" | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1606 | using \<open>a \<notin> C\<close> \<open>b \<in> V\<close> that by blast | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1607 | show "(topspace X - C) \<times> V \<subseteq> topspace X \<times> topspace Y - K" | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1608 | using \<open>V \<subseteq> V'\<close> \<open>V \<subseteq> topspace Y\<close> | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1609 | apply (simp add: C_def ) | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1610 | by (smt (verit, ccfv_threshold) DiffE DiffI IntI SigmaE SigmaI image_eqI mem_Collect_eq prod.sel(1) snd_conv subset_iff) | 
| 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1611 | qed | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1612 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1613 | ultimately show "closedin (prod_topology X Y) K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1614 | by (metis surj_pair closedin_def openin_subopen topspace_prod_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1615 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1616 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1617 | lemma kc_space_prod_topology_right: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1618 | "\<lbrakk>Hausdorff_space X; kc_space Y\<rbrakk> \<Longrightarrow> kc_space (prod_topology X Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1619 | using kc_space_prod_topology_left homeomorphic_kc_space homeomorphic_space_prod_topology_swap by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1620 | |
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1621 | subsection \<open>Technical results about proper maps, perfect maps, etc\<close> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1622 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1623 | lemma compact_imp_proper_map_gen: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1624 | assumes Y: "\<And>S. \<lbrakk>S \<subseteq> topspace Y; \<And>K. compactin Y K \<Longrightarrow> compactin Y (S \<inter> K)\<rbrakk> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1625 | \<Longrightarrow> closedin Y S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1626 | and fim: "f ` (topspace X) \<subseteq> topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1627 | and f: "continuous_map X Y f \<or> kc_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1628 |     and YX: "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1629 | shows "proper_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1630 | unfolding proper_map_alt closed_map_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1631 | proof (intro conjI strip) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1632 | fix C | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1633 | assume C: "closedin X C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1634 | show "closedin Y (f ` C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1635 | proof (intro Y) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1636 | show "f ` C \<subseteq> topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1637 | using C closedin_subset fim by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1638 | fix K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1639 | assume K: "compactin Y K" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1640 |     define A where "A \<equiv> {x \<in> topspace X. f x \<in> K}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1641 |     have eq: "f ` C \<inter> K = f ` ({x \<in> topspace X. f x \<in> K} \<inter> C)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1642 | using C closedin_subset by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1643 | show "compactin Y (f ` C \<inter> K)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1644 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1645 | proof (rule image_compactin) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1646 |       show "compactin (subtopology X A) ({x \<in> topspace X. f x \<in> K} \<inter> C)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1647 | proof (rule closedin_compact_space) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1648 | show "compact_space (subtopology X A)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1649 | by (simp add: A_def K YX compact_space_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1650 |         show "closedin (subtopology X A) ({x \<in> topspace X. f x \<in> K} \<inter> C)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1651 | using A_def C closedin_subtopology by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1652 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1653 | have "continuous_map (subtopology X A) (subtopology Y K) f" if "kc_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1654 | unfolding continuous_map_closedin | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1655 | proof (intro conjI strip) | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1656 | show "f \<in> topspace (subtopology X A) \<rightarrow> topspace (subtopology Y K)" | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1657 | using A_def K compactin_subset_topspace by fastforce | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1658 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1659 | fix C | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1660 | assume C: "closedin (subtopology Y K) C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1661 |         show "closedin (subtopology X A) {x \<in> topspace (subtopology X A). f x \<in> C}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1662 | proof (rule compactin_imp_closedin_gen) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1663 | show "kc_space (subtopology X A)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1664 | by (simp add: kc_space_subtopology that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1665 |           have [simp]: "{x \<in> topspace X. f x \<in> K \<and> f x \<in> C} = {x \<in> topspace X. f x \<in> C}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1666 | using C closedin_imp_subset by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1667 | have "compactin (subtopology Y K) C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1668 | by (simp add: C K closedin_compact_space compact_space_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1669 |           then have "compactin X {x \<in> topspace X. x \<in> A \<and> f x \<in> C}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1670 | by (auto simp: A_def compactin_subtopology dest: YX) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1671 |           then show "compactin (subtopology X A) {x \<in> topspace (subtopology X A). f x \<in> C}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1672 | by (auto simp add: compactin_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1673 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1674 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1675 | with f show "continuous_map (subtopology X A) Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1676 | using continuous_map_from_subtopology continuous_map_in_subtopology by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1677 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1678 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1679 | qed (simp add: YX) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1680 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1681 | lemma tube_lemma_left: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1682 | assumes W: "openin (prod_topology X Y) W" and C: "compactin X C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1683 |     and y: "y \<in> topspace Y" and subW: "C \<times> {y} \<subseteq> W"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1684 | shows "\<exists>U V. openin X U \<and> openin Y V \<and> C \<subseteq> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1685 | proof (cases "C = {}")
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1686 | case True | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1687 | with y show ?thesis by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1688 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1689 | case False | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1690 | have "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1691 | if "x \<in> C" for x | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1692 | using W openin_prod_topology_alt subW subsetD that by fastforce | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1693 | then obtain U V where UV: "\<And>x. x \<in> C \<Longrightarrow> openin X (U x) \<and> openin Y (V x) \<and> x \<in> U x \<and> y \<in> V x \<and> U x \<times> V x \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1694 | by metis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1695 | then obtain D where D: "finite D" "D \<subseteq> C" "C \<subseteq> \<Union> (U ` D)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1696 | using compactinD [OF C, of "U`C"] | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1697 | by (smt (verit) UN_I finite_subset_image imageE subsetI) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1698 | show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1699 | proof (intro exI conjI) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1700 | show "openin X (\<Union> (U ` D))" "openin Y (\<Inter> (V ` D))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1701 | using D False UV by blast+ | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1702 | show "y \<in> \<Inter> (V ` D)" "C \<subseteq> \<Union> (U ` D)" "\<Union>(U ` D) \<times> \<Inter>(V ` D) \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1703 | using D UV by force+ | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1704 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1705 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1706 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1707 | lemma Wallace_theorem_prod_topology: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1708 | assumes "compactin X K" "compactin Y L" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1709 | and W: "openin (prod_topology X Y) W" and subW: "K \<times> L \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1710 | obtains U V where "openin X U" "openin Y V" "K \<subseteq> U" "L \<subseteq> V" "U \<times> V \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1711 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1712 | have "\<And>y. y \<in> L \<Longrightarrow> \<exists>U V. openin X U \<and> openin Y V \<and> K \<subseteq> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1713 | proof (intro tube_lemma_left assms) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1714 | fix y assume "y \<in> L" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1715 | show "y \<in> topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1716 | using assms \<open>y \<in> L\<close> compactin_subset_topspace by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1717 |     show "K \<times> {y} \<subseteq> W"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1718 | using \<open>y \<in> L\<close> subW by force | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1719 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1720 | then obtain U V where UV: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1721 | "\<And>y. y \<in> L \<Longrightarrow> openin X (U y) \<and> openin Y (V y) \<and> K \<subseteq> U y \<and> y \<in> V y \<and> U y \<times> V y \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1722 | by metis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1723 | then obtain M where "finite M" "M \<subseteq> L" and M: "L \<subseteq> \<Union> (V ` M)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1724 | using \<open>compactin Y L\<close> unfolding compactin_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1725 | by (smt (verit) UN_iff finite_subset_image imageE subset_iff) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1726 | show thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1727 |   proof (cases "M={}")
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1728 | case True | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1729 |     with M have "L={}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1730 | by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1731 | then show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1732 | using \<open>compactin X K\<close> compactin_subset_topspace that by fastforce | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1733 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1734 | case False | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1735 | show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1736 | proof | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1737 | show "openin X (\<Inter>(U`M))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1738 | using False UV \<open>M \<subseteq> L\<close> \<open>finite M\<close> by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1739 | show "openin Y (\<Union>(V`M))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1740 | using UV \<open>M \<subseteq> L\<close> by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1741 | show "K \<subseteq> \<Inter>(U`M)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1742 | by (meson INF_greatest UV \<open>M \<subseteq> L\<close> subsetD) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1743 | show "L \<subseteq> \<Union>(V`M)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1744 | by (simp add: M) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1745 | show "\<Inter>(U`M) \<times> \<Union>(V`M) \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1746 | using UV \<open>M \<subseteq> L\<close> by fastforce | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1747 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1748 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1749 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1750 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1751 | lemma proper_map_prod: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1752 | "proper_map (prod_topology X Y) (prod_topology X' Y') (\<lambda>(x,y). (f x, g y)) \<longleftrightarrow> | 
| 78336 | 1753 | (prod_topology X Y) = trivial_topology \<or> proper_map X X' f \<and> proper_map Y Y' g" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1754 | (is "?lhs \<longleftrightarrow> _ \<or> ?rhs") | 
| 78336 | 1755 | proof (cases "(prod_topology X Y) = trivial_topology") | 
| 1756 | case True then show ?thesis by auto | |
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1757 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1758 | case False | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1759 |   then have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1760 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1761 | define h where "h \<equiv> \<lambda>(x,y). (f x, g y)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1762 | have "proper_map X X' f" "proper_map Y Y' g" if ?lhs | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1763 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1764 | have cm: "closed_map X X' f" "closed_map Y Y' g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1765 | using that False closed_map_prod proper_imp_closed_map by blast+ | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1766 | show "proper_map X X' f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1767 | proof (clarsimp simp add: proper_map_def cm) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1768 | fix y | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1769 | assume y: "y \<in> topspace X'" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1770 | obtain z where z: "z \<in> topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1771 | using ne by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1772 |       then have eq: "{x \<in> topspace X. f x = y} =
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1773 |                      fst ` {u \<in> topspace X \<times> topspace Y. h u = (y,g z)}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1774 | by (force simp: h_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1775 |       show "compactin X {x \<in> topspace X. f x = y}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1776 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1777 | proof (intro image_compactin) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1778 | have "g z \<in> topspace Y'" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1779 | by (meson closed_map_def closedin_subset closedin_topspace cm image_subset_iff z) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1780 |         with y show "compactin (prod_topology X Y) {u \<in> topspace X \<times> topspace Y. (h u) = (y, g z)}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1781 | using that by (simp add: h_def proper_map_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1782 | show "continuous_map (prod_topology X Y) X fst" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1783 | by (simp add: continuous_map_fst) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1784 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1785 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1786 | show "proper_map Y Y' g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1787 | proof (clarsimp simp add: proper_map_def cm) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1788 | fix y | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1789 | assume y: "y \<in> topspace Y'" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1790 | obtain z where z: "z \<in> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1791 | using ne by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1792 |       then have eq: "{x \<in> topspace Y. g x = y} =
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1793 |                      snd ` {u \<in> topspace X \<times> topspace Y. h u = (f z,y)}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1794 | by (force simp: h_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1795 |       show "compactin Y {x \<in> topspace Y. g x = y}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1796 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1797 | proof (intro image_compactin) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1798 | have "f z \<in> topspace X'" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1799 | by (meson closed_map_def closedin_subset closedin_topspace cm image_subset_iff z) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1800 |         with y show "compactin (prod_topology X Y) {u \<in> topspace X \<times> topspace Y. (h u) = (f z, y)}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1801 | using that by (simp add: proper_map_def h_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1802 | show "continuous_map (prod_topology X Y) Y snd" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1803 | by (simp add: continuous_map_snd) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1804 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1805 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1806 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1807 | moreover | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1808 |   { assume R: ?rhs
 | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1809 | then have fgim: "f \<in> topspace X \<rightarrow> topspace X'" "g \<in> topspace Y \<rightarrow> topspace Y'" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1810 | and cm: "closed_map X X' f" "closed_map Y Y' g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1811 | by (auto simp: proper_map_def closed_map_imp_subset_topspace) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1812 | have "closed_map (prod_topology X Y) (prod_topology X' Y') h" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1813 | unfolding closed_map_fibre_neighbourhood imp_conjL | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1814 | proof (intro conjI strip) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1815 | show "h \<in> topspace (prod_topology X Y) \<rightarrow> topspace (prod_topology X' Y')" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1816 | unfolding h_def using fgim by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1817 | fix W w | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1818 | assume W: "openin (prod_topology X Y) W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1819 | and w: "w \<in> topspace (prod_topology X' Y')" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1820 |         and subW: "{x \<in> topspace (prod_topology X Y). h x = w} \<subseteq> W"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1821 | then obtain x' y' where weq: "w = (x',y')" "x' \<in> topspace X'" "y' \<in> topspace Y'" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1822 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1823 |       have eq: "{u \<in> topspace X \<times> topspace Y. h u = (x',y')} = {x \<in> topspace X. f x = x'} \<times> {y \<in> topspace Y. g y = y'}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1824 | by (auto simp: h_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1825 | obtain U V where "openin X U" "openin Y V" "U \<times> V \<subseteq> W" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1826 |         and U: "{x \<in> topspace X. f x = x'} \<subseteq> U" 
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1827 |         and V: "{x \<in> topspace Y. g x = y'} \<subseteq> V" 
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1828 | proof (rule Wallace_theorem_prod_topology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1829 |         show "compactin X {x \<in> topspace X. f x = x'}" "compactin Y {x \<in> topspace Y. g x = y'}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1830 | using R weq unfolding proper_map_def closed_map_fibre_neighbourhood by fastforce+ | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1831 |         show "{x \<in> topspace X. f x = x'} \<times> {x \<in> topspace Y. g x = y'} \<subseteq> W"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1832 | using weq subW by (auto simp: h_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1833 | qed (use W in auto) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1834 |       obtain U' where "openin X' U'" "x' \<in> U'" and U': "{x \<in> topspace X. f x \<in> U'} \<subseteq> U"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1835 | using cm U \<open>openin X U\<close> weq unfolding closed_map_fibre_neighbourhood by meson | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1836 |       obtain V' where "openin Y' V'" "y' \<in> V'" and V': "{x \<in> topspace Y. g x \<in> V'} \<subseteq> V"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1837 | using cm V \<open>openin Y V\<close> weq unfolding closed_map_fibre_neighbourhood by meson | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1838 |       show "\<exists>V. openin (prod_topology X' Y') V \<and> w \<in> V \<and> {x \<in> topspace (prod_topology X Y). h x \<in> V} \<subseteq> W"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1839 | proof (intro conjI exI) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1840 | show "openin (prod_topology X' Y') (U' \<times> V')" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1841 | by (simp add: \<open>openin X' U'\<close> \<open>openin Y' V'\<close> openin_prod_Times_iff) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1842 | show "w \<in> U' \<times> V'" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1843 | using \<open>x' \<in> U'\<close> \<open>y' \<in> V'\<close> weq by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1844 |         show "{x \<in> topspace (prod_topology X Y). h x \<in> U' \<times> V'} \<subseteq> W"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1845 | using \<open>U \<times> V \<subseteq> W\<close> U' V' h_def by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1846 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1847 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1848 | moreover | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1849 |     have "compactin (prod_topology X Y) {u \<in> topspace X \<times> topspace Y. h u = (w, z)}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1850 | if "w \<in> topspace X'" and "z \<in> topspace Y'" for w z | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1851 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1852 |       have eq: "{u \<in> topspace X \<times> topspace Y. h u = (w,z)} =
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1853 |                 {u \<in> topspace X. f u = w} \<times> {y. y \<in> topspace Y \<and> g y = z}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1854 | by (auto simp: h_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1855 | show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1856 | using R that by (simp add: eq compactin_Times proper_map_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1857 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1858 | ultimately have ?lhs | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1859 | by (auto simp: h_def proper_map_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1860 | } | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1861 | ultimately show ?thesis using False by metis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1862 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1863 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1864 | lemma proper_map_paired: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1865 | assumes "Hausdorff_space X \<and> proper_map X Y f \<and> proper_map X Z g \<or> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1866 | Hausdorff_space Y \<and> continuous_map X Y f \<and> proper_map X Z g \<or> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1867 | Hausdorff_space Z \<and> proper_map X Y f \<and> continuous_map X Z g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1868 | shows "proper_map X (prod_topology Y Z) (\<lambda>x. (f x,g x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1869 | using assms | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1870 | proof (elim disjE conjE) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1871 | assume \<section>: "Hausdorff_space X" "proper_map X Y f" "proper_map X Z g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1872 | have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x, y). (f x, g y)) \<circ> (\<lambda>x. (x, x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1873 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1874 | show "proper_map X (prod_topology Y Z) (\<lambda>x. (f x, g x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1875 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1876 | proof (rule proper_map_compose) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1877 | show "proper_map X (prod_topology X X) (\<lambda>x. (x,x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1878 | by (simp add: \<section>) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1879 | show "proper_map (prod_topology X X) (prod_topology Y Z) (\<lambda>(x,y). (f x, g y))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1880 | by (simp add: \<section> proper_map_prod) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1881 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1882 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1883 | assume \<section>: "Hausdorff_space Y" "continuous_map X Y f" "proper_map X Z g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1884 | have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x,y). (x,g y)) \<circ> (\<lambda>x. (f x,x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1885 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1886 | show "proper_map X (prod_topology Y Z) (\<lambda>x. (f x, g x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1887 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1888 | proof (rule proper_map_compose) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1889 | show "proper_map X (prod_topology Y X) (\<lambda>x. (f x,x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1890 | by (simp add: \<section> proper_map_paired_continuous_map_left) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1891 | show "proper_map (prod_topology Y X) (prod_topology Y Z) (\<lambda>(x,y). (x,g y))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1892 | by (simp add: \<section> proper_map_prod proper_map_id [unfolded id_def]) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1893 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1894 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1895 | assume \<section>: "Hausdorff_space Z" "proper_map X Y f" "continuous_map X Z g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1896 | have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x,y). (f x,y)) \<circ> (\<lambda>x. (x,g x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1897 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1898 | show "proper_map X (prod_topology Y Z) (\<lambda>x. (f x, g x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1899 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1900 | proof (rule proper_map_compose) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1901 | show "proper_map X (prod_topology X Z) (\<lambda>x. (x, g x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1902 | using \<section> proper_map_paired_continuous_map_right by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1903 | show "proper_map (prod_topology X Z) (prod_topology Y Z) (\<lambda>(x,y). (f x,y))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1904 | by (simp add: \<section> proper_map_prod proper_map_id [unfolded id_def]) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1905 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1906 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1907 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1908 | lemma proper_map_pairwise: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1909 | assumes | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1910 | "Hausdorff_space X \<and> proper_map X Y (fst \<circ> f) \<and> proper_map X Z (snd \<circ> f) \<or> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1911 | Hausdorff_space Y \<and> continuous_map X Y (fst \<circ> f) \<and> proper_map X Z (snd \<circ> f) \<or> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1912 | Hausdorff_space Z \<and> proper_map X Y (fst \<circ> f) \<and> continuous_map X Z (snd \<circ> f)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1913 | shows "proper_map X (prod_topology Y Z) f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1914 | using proper_map_paired [OF assms] by (simp add: o_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1915 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1916 | lemma proper_map_from_composition_right: | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1917 | assumes "Hausdorff_space Y" "proper_map X Z (g \<circ> f)" and contf: "continuous_map X Y f" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1918 | and contg: "continuous_map Y Z g" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1919 | shows "proper_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1920 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1921 | define YZ where "YZ \<equiv> subtopology (prod_topology Y Z) ((\<lambda>x. (x, g x)) ` topspace Y)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1922 | have "proper_map X Y (fst \<circ> (\<lambda>x. (f x, (g \<circ> f) x)))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1923 | proof (rule proper_map_compose) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1924 | have [simp]: "x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y" for x | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 1925 | using contf continuous_map_preimage_topspace by auto | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1926 | show "proper_map X YZ (\<lambda>x. (f x, (g \<circ> f) x))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1927 | unfolding YZ_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1928 | using assms | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1929 | by (force intro!: proper_map_into_subtopology proper_map_paired simp: o_def image_iff)+ | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1930 | show "proper_map YZ Y fst" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1931 | using contg | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1932 | by (simp flip: homeomorphic_maps_graph add: YZ_def homeomorphic_maps_map homeomorphic_imp_proper_map) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1933 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1934 | moreover have "fst \<circ> (\<lambda>x. (f x, (g \<circ> f) x)) = f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1935 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1936 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1937 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1938 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1939 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1940 | lemma perfect_map_from_composition_right: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1941 | "\<lbrakk>Hausdorff_space Y; perfect_map X Z (g \<circ> f); | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1942 | continuous_map X Y f; continuous_map Y Z g; f ` topspace X = topspace Y\<rbrakk> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1943 | \<Longrightarrow> perfect_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1944 | by (meson perfect_map_def proper_map_from_composition_right) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1945 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1946 | lemma perfect_map_from_composition_right_inj: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1947 | "\<lbrakk>perfect_map X Z (g \<circ> f); f ` topspace X = topspace Y; | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1948 | continuous_map X Y f; continuous_map Y Z g; inj_on g (topspace Y)\<rbrakk> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1949 | \<Longrightarrow> perfect_map X Y f" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 1950 | by (meson continuous_map_openin_preimage_eq perfect_map_def proper_map_from_composition_right_inj) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 1951 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1952 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1953 | subsection \<open>Regular spaces\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1954 | |
| 77942 
30f69046f120
fixes esp to theory presentation
 paulson <lp15@cam.ac.uk> parents: 
77941diff
changeset | 1955 | text \<open>Regular spaces are *not* a priori assumed to be Hausdorff or $T_1$\<close> | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1956 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1957 | definition regular_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1958 | where "regular_space X \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1959 | \<forall>C a. closedin X C \<and> a \<in> topspace X - C | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1960 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1961 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1962 | lemma homeomorphic_regular_space_aux: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1963 | assumes hom: "X homeomorphic_space Y" and X: "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1964 | shows "regular_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1965 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1966 | obtain f g where hmf: "homeomorphic_map X Y f" and hmg: "homeomorphic_map Y X g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1967 | and fg: "(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> topspace Y. f(g y) = y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1968 | using assms X homeomorphic_maps_map homeomorphic_space_def by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1969 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1970 | unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1971 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1972 | fix C a | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1973 | assume Y: "closedin Y C" "a \<in> topspace Y" and "a \<notin> C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1974 | then obtain "closedin X (g ` C)" "g a \<in> topspace X" "g a \<notin> g ` C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1975 | using \<open>closedin Y C\<close> hmg homeomorphic_map_closedness_eq | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1976 | by (smt (verit, ccfv_SIG) fg homeomorphic_imp_surjective_map image_iff in_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1977 | then obtain S T where ST: "openin X S" "openin X T" "g a \<in> S" "g`C \<subseteq> T" "disjnt S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1978 | using X unfolding regular_space_def by (metis DiffI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1979 | then have "openin Y (f`S)" "openin Y (f`T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1980 | by (meson hmf homeomorphic_map_openness_eq)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1981 | moreover have "a \<in> f`S \<and> C \<subseteq> f`T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1982 | using ST by (smt (verit, best) Y closedin_subset fg image_eqI subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1983 | moreover have "disjnt (f`S) (f`T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1984 | using ST by (smt (verit, ccfv_SIG) disjnt_iff fg image_iff openin_subset subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1985 | ultimately show "\<exists>U V. openin Y U \<and> openin Y V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1986 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1987 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1988 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1989 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1990 | lemma homeomorphic_regular_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1991 | "X homeomorphic_space Y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1992 | \<Longrightarrow> (regular_space X \<longleftrightarrow> regular_space Y)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1993 | by (meson homeomorphic_regular_space_aux homeomorphic_space_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1994 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1995 | lemma regular_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1996 | "regular_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1997 | (\<forall>C a. closedin X C \<and> a \<in> topspace X - C | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1998 | \<longrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> disjnt C (X closure_of U)))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1999 | unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2000 | proof (intro all_cong1 imp_cong refl ex_cong1) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2001 | fix C a U | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2002 | assume C: "closedin X C \<and> a \<in> topspace X - C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2003 | show "(\<exists>V. openin X U \<and> openin X V \<and> a \<in> U \<and> C \<subseteq> V \<and> disjnt U V) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2004 | \<longleftrightarrow> (openin X U \<and> a \<in> U \<and> disjnt C (X closure_of U))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2005 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2006 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2007 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2008 | by (smt (verit, best) disjnt_iff in_closure_of subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2009 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2010 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2011 | then have "disjnt U (topspace X - X closure_of U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2012 | by (meson DiffD2 closure_of_subset disjnt_iff openin_subset subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2013 | moreover have "C \<subseteq> topspace X - X closure_of U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2014 | by (meson C DiffI R closedin_subset disjnt_iff subset_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2015 | ultimately show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2016 | using R by (rule_tac x="topspace X - X closure_of U" in exI) auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2017 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2018 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2019 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2020 | lemma neighbourhood_base_of_closedin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2021 | "neighbourhood_base_of (closedin X) X \<longleftrightarrow> regular_space X" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2022 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2023 | have "?lhs \<longleftrightarrow> (\<forall>W x. openin X W \<and> x \<in> W \<longrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2024 | (\<exists>U. openin X U \<and> (\<exists>V. closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W)))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2025 | by (simp add: neighbourhood_base_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2026 | also have "\<dots> \<longleftrightarrow> (\<forall>W x. closedin X W \<and> x \<in> topspace X - W \<longrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2027 | (\<exists>U. openin X U \<and> (\<exists>V. closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> topspace X - W)))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2028 | by (smt (verit) Diff_Diff_Int closedin_def inf.absorb_iff2 openin_closedin_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2029 | also have "\<dots> \<longleftrightarrow> ?rhs" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2030 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2031 | have \<section>: "(\<exists>V. closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> topspace X - W) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2032 | \<longleftrightarrow> (\<exists>V. openin X V \<and> x \<in> U \<and> W \<subseteq> V \<and> disjnt U V)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2033 | if "openin X U" "closedin X W" "x \<in> topspace X" "x \<notin> W" for W U x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2034 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2035 | assume ?lhs with \<open>closedin X W\<close> show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2036 | unfolding closedin_def by (smt (verit) Diff_mono disjnt_Diff1 double_diff subset_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2037 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2038 | assume ?rhs with \<open>openin X U\<close> show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2039 | unfolding openin_closedin_eq disjnt_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2040 | by (smt (verit) Diff_Diff_Int Diff_disjoint Diff_eq_empty_iff Int_Diff inf.orderE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2041 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2042 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2043 | unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2044 | by (intro all_cong1 ex_cong1 imp_cong refl) (metis \<section> DiffE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2045 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2046 | finally show ?thesis . | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2047 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2048 | |
| 78336 | 2049 | lemma regular_space_discrete_topology [simp]: | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2050 | "regular_space (discrete_topology S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2051 | using neighbourhood_base_of_closedin neighbourhood_base_of_discrete_topology by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2052 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2053 | lemma regular_space_subtopology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2054 | "regular_space X \<Longrightarrow> regular_space (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2055 | unfolding regular_space_def openin_subtopology_alt closedin_subtopology_alt disjnt_iff | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2056 | by clarsimp (smt (verit, best) inf.orderE inf_le1 le_inf_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2057 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2058 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2059 | lemma regular_space_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2060 | "\<lbrakk>retraction_map X Y r; regular_space X\<rbrakk> \<Longrightarrow> regular_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2061 | using hereditary_imp_retractive_property homeomorphic_regular_space regular_space_subtopology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2062 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2063 | lemma regular_t0_imp_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2064 | "\<lbrakk>regular_space X; t0_space X\<rbrakk> \<Longrightarrow> Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2065 | apply (clarsimp simp: regular_space_def t0_space Hausdorff_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2066 | by (metis disjnt_sym subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2067 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2068 | lemma regular_t0_eq_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2069 | "regular_space X \<Longrightarrow> (t0_space X \<longleftrightarrow> Hausdorff_space X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2070 | using Hausdorff_imp_t0_space regular_t0_imp_Hausdorff_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2071 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2072 | lemma regular_t1_imp_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2073 | "\<lbrakk>regular_space X; t1_space X\<rbrakk> \<Longrightarrow> Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2074 | by (simp add: regular_t0_imp_Hausdorff_space t1_imp_t0_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2075 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2076 | lemma regular_t1_eq_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2077 | "regular_space X \<Longrightarrow> t1_space X \<longleftrightarrow> Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2078 | using regular_t0_imp_Hausdorff_space t1_imp_t0_space t1_or_Hausdorff_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2079 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2080 | lemma compact_Hausdorff_imp_regular_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2081 | assumes "compact_space X" "Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2082 | shows "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2083 | unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2084 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2085 | fix S a | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2086 | assume "closedin X S" and "a \<in> topspace X" and "a \<notin> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2087 | then show "\<exists>U V. openin X U \<and> openin X V \<and> a \<in> U \<and> S \<subseteq> V \<and> disjnt U V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2088 | using assms unfolding Hausdorff_space_compact_sets | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2089 | by (metis closedin_compact_space compactin_sing disjnt_empty1 insert_subset disjnt_insert1) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2090 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2091 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2092 | lemma neighbourhood_base_of_closed_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2093 | "regular_space X \<and> Hausdorff_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2094 | neighbourhood_base_of (\<lambda>C. closedin X C \<and> Hausdorff_space(subtopology X C)) X" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2095 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2096 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2097 | by (simp add: Hausdorff_space_subtopology neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2098 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2099 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2100 | by (metis (mono_tags, lifting) Hausdorff_space_closed_neighbourhood neighbourhood_base_of neighbourhood_base_of_closedin openin_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2101 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2102 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2103 | lemma locally_compact_imp_kc_eq_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2104 | "neighbourhood_base_of (compactin X) X \<Longrightarrow> kc_space X \<longleftrightarrow> Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2105 | by (metis Hausdorff_imp_kc_space kc_imp_t1_space kc_space_def neighbourhood_base_of_closedin neighbourhood_base_of_mono regular_t1_imp_Hausdorff_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2106 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2107 | lemma regular_space_compact_closed_separation: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2108 | assumes X: "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2109 | and S: "compactin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2110 | and T: "closedin X T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2111 | and "disjnt S T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2112 | shows "\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2113 | proof (cases "S={}")
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2114 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2115 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2116 | by (meson T closedin_def disjnt_empty1 empty_subsetI openin_empty openin_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2117 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2118 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2119 | then have "\<exists>U V. x \<in> S \<longrightarrow> openin X U \<and> openin X V \<and> x \<in> U \<and> T \<subseteq> V \<and> disjnt U V" for x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2120 | using assms unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2121 | by (smt (verit) Diff_iff compactin_subset_topspace disjnt_iff subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2122 | then obtain U V where UV: "\<And>x. x \<in> S \<Longrightarrow> 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)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2123 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2124 | then obtain \<F> where "finite \<F>" "\<F> \<subseteq> U ` S" "S \<subseteq> \<Union> \<F>" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2125 | using S unfolding compactin_def by (smt (verit) UN_iff image_iff subsetI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2126 | then obtain K where "finite K" "K \<subseteq> S" and K: "S \<subseteq> \<Union>(U ` K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2127 | by (metis finite_subset_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2128 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2129 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2130 | show "openin X (\<Union>(U ` K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2131 | using \<open>K \<subseteq> S\<close> UV by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2132 | show "openin X (\<Inter>(V ` K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2133 | using False K UV \<open>K \<subseteq> S\<close> \<open>finite K\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2134 | show "S \<subseteq> \<Union>(U ` K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2135 | by (simp add: K) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2136 | show "T \<subseteq> \<Inter>(V ` K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2137 | using UV \<open>K \<subseteq> S\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2138 | show "disjnt (\<Union>(U ` K)) (\<Inter>(V ` K))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2139 | by (smt (verit) Inter_iff UN_E UV \<open>K \<subseteq> S\<close> disjnt_iff image_eqI subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2140 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2141 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2142 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2143 | lemma regular_space_compact_closed_sets: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2144 | "regular_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2145 | (\<forall>S T. compactin X S \<and> closedin X T \<and> disjnt S T | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2146 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2147 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2148 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2149 | using regular_space_compact_closed_separation by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2150 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2151 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2152 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2153 | unfolding regular_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2154 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2155 | fix S x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2156 | assume "closedin X S" and "x \<in> topspace X" and "x \<notin> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2157 |     then obtain U V where "openin X U \<and> openin X V \<and> {x} \<subseteq> U \<and> S \<subseteq> V \<and> disjnt U V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2158 | by (metis R compactin_sing disjnt_empty1 disjnt_insert1) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2159 | then show "\<exists>U. openin X U \<and> x \<in> U \<and> disjnt S (X closure_of U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2160 | by (smt (verit, best) disjnt_iff in_closure_of insert_subset subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2161 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2162 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2163 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2164 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2165 | lemma regular_space_prod_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2166 | "regular_space (prod_topology X Y) \<longleftrightarrow> | 
| 78336 | 2167 | X = trivial_topology \<or> Y = trivial_topology \<or> regular_space X \<and> regular_space Y" (is "?lhs=?rhs") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2168 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2169 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2170 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2171 | by (metis regular_space_retraction_map_image retraction_map_fst retraction_map_snd) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2172 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2173 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2174 | show ?lhs | 
| 78336 | 2175 | proof (cases "X = trivial_topology \<or> Y = trivial_topology") | 
| 2176 | case True then show ?thesis by auto | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2177 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2178 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2179 | then have "regular_space X" "regular_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2180 | using R by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2181 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2182 | unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2183 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2184 | fix W x y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2185 | assume W: "openin (prod_topology X Y) W" and "(x,y) \<in> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2186 | then obtain U V where U: "openin X U" "x \<in> U" and V: "openin Y V" "y \<in> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2187 | and "U \<times> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2188 | by (metis openin_prod_topology_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2189 | obtain D1 C1 where 1: "openin X D1" "closedin X C1" "x \<in> D1" "D1 \<subseteq> C1" "C1 \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2190 | by (metis \<open>regular_space X\<close> U neighbourhood_base_of neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2191 | obtain D2 C2 where 2: "openin Y D2" "closedin Y C2" "y \<in> D2" "D2 \<subseteq> C2" "C2 \<subseteq> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2192 | by (metis \<open>regular_space Y\<close> V neighbourhood_base_of neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2193 | show "\<exists>U V. openin (prod_topology X Y) U \<and> closedin (prod_topology X Y) V \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2194 | (x,y) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2195 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2196 | show "openin (prod_topology X Y) (D1 \<times> D2)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2197 | by (simp add: 1 2 openin_prod_Times_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2198 | show "closedin (prod_topology X Y) (C1 \<times> C2)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2199 | by (simp add: 1 2 closedin_prod_Times_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2200 | qed (use 1 2 \<open>U \<times> V \<subseteq> W\<close> in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2201 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2202 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2203 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2204 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2205 | lemma regular_space_product_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2206 | "regular_space (product_topology X I) \<longleftrightarrow> | 
| 78336 | 2207 | (product_topology X I) = trivial_topology \<or> (\<forall>i \<in> I. regular_space (X i))" (is "?lhs=?rhs") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2208 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2209 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2210 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2211 | by (meson regular_space_retraction_map_image retraction_map_product_projection) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2212 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2213 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2214 | show ?lhs | 
| 78336 | 2215 | proof (cases "product_topology X I = trivial_topology") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2216 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2217 | then show ?thesis | 
| 78336 | 2218 | by auto | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2219 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2220 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2221 | then obtain x where x: "x \<in> topspace (product_topology X I)" | 
| 78336 | 2222 | by (meson ex_in_conv null_topspace_iff_trivial) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2223 |     define \<F> where "\<F> \<equiv> {Pi\<^sub>E I U |U. finite {i \<in> I. U i \<noteq> topspace (X i)}
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2224 | \<and> (\<forall>i\<in>I. openin (X i) (U i))}" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2225 | have oo: "openin (product_topology X I) = arbitrary union_of (\<lambda>W. W \<in> \<F>)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2226 | by (simp add: \<F>_def openin_product_topology product_topology_base_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2227 | have "\<exists>U V. openin (product_topology X I) U \<and> closedin (product_topology X I) V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> Pi\<^sub>E I W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2228 |       if fin: "finite {i \<in> I. W i \<noteq> topspace (X i)}" 
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2229 | and opeW: "\<And>k. k \<in> I \<Longrightarrow> openin (X k) (W k)" and x: "x \<in> PiE I W" for W x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2230 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2231 | have "\<And>i. i \<in> I \<Longrightarrow> \<exists>U V. openin (X i) U \<and> closedin (X i) V \<and> x i \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W i" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2232 | by (metis False PiE_iff R neighbourhood_base_of neighbourhood_base_of_closedin opeW x) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2233 | then obtain U C where UC: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2234 | "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (U i) \<and> closedin (X i) (C i) \<and> x i \<in> U i \<and> U i \<subseteq> C i \<and> C i \<subseteq> W i" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2235 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2236 | define PI where "PI \<equiv> \<lambda>V. PiE I (\<lambda>i. if W i = topspace(X i) then topspace(X i) else V i)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2237 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2238 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2239 | have "\<forall>i\<in>I. W i \<noteq> topspace (X i) \<longrightarrow> openin (X i) (U i)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2240 | using UC by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2241 | with fin show "openin (product_topology X I) (PI U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2242 | by (simp add: Collect_mono_iff PI_def openin_PiE_gen rev_finite_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2243 | show "closedin (product_topology X I) (PI C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2244 | by (simp add: UC closedin_product_topology PI_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2245 | show "x \<in> PI U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2246 | using UC x by (fastforce simp: PI_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2247 | show "PI U \<subseteq> PI C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2248 | by (smt (verit) UC Orderings.order_eq_iff PiE_mono PI_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2249 | show "PI C \<subseteq> Pi\<^sub>E I W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2250 | by (simp add: UC PI_def subset_PiE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2251 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2252 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2253 | then have "neighbourhood_base_of (closedin (product_topology X I)) (product_topology X I)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2254 | unfolding neighbourhood_base_of_topology_base [OF oo] by (force simp: \<F>_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2255 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2256 | by (simp add: neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2257 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2258 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2259 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2260 | lemma closed_map_paired_gen_aux: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2261 | assumes "regular_space Y" and f: "closed_map Z X f" and g: "closed_map Z Y g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2262 |     and clo: "\<And>y. y \<in> topspace X \<Longrightarrow> closedin Z {x \<in> topspace Z. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2263 | and contg: "continuous_map Z Y g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2264 | shows "closed_map Z (prod_topology X Y) (\<lambda>x. (f x, g x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2265 | unfolding closed_map_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2266 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2267 | fix C assume "closedin Z C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2268 | then have "C \<subseteq> topspace Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2269 | by (simp add: closedin_subset) | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 2270 | have "f \<in> topspace Z \<rightarrow> topspace X" "g \<in> topspace Z \<rightarrow> topspace Y" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2271 | by (simp_all add: assms closed_map_imp_subset_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2272 | show "closedin (prod_topology X Y) ((\<lambda>x. (f x, g x)) ` C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2273 | unfolding closedin_def topspace_prod_topology | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2274 | proof (intro conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2275 | have "closedin Y (g ` C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2276 | using \<open>closedin Z C\<close> assms(3) closed_map_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2277 | with assms show "(\<lambda>x. (f x, g x)) ` C \<subseteq> topspace X \<times> topspace Y" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 2278 | by (smt (verit) SigmaI \<open>closedin Z C\<close> closed_map_def closedin_subset image_subset_iff) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2279 | have *: "\<exists>T. openin (prod_topology X Y) T \<and> (y1,y2) \<in> T \<and> T \<subseteq> topspace X \<times> topspace Y - (\<lambda>x. (f x, g x)) ` C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2280 | if "(y1,y2) \<notin> (\<lambda>x. (f x, g x)) ` C" and y1: "y1 \<in> topspace X" and y2: "y2 \<in> topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2281 | for y1 y2 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2282 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2283 |       define A where "A \<equiv> topspace Z - (C \<inter> {x \<in> topspace Z. f x = y1})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2284 |       have A: "openin Z A" "{x \<in> topspace Z. g x = y2} \<subseteq> A"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2285 | using that \<open>closedin Z C\<close> clo that(2) by (auto simp: A_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2286 |       obtain V0 where "openin Y V0 \<and> y2 \<in> V0" and UA: "{x \<in> topspace Z. g x \<in> V0} \<subseteq> A"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2287 | using g A y2 unfolding closed_map_fibre_neighbourhood by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2288 | then obtain V V' where VV: "openin Y V \<and> closedin Y V' \<and> y2 \<in> V \<and> V \<subseteq> V'" and "V' \<subseteq> V0" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2289 | by (metis (no_types, lifting) \<open>regular_space Y\<close> neighbourhood_base_of neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2290 |       with UA have subA: "{x \<in> topspace Z. g x \<in> V'} \<subseteq> A"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2291 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2292 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2293 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2294 |         define B where "B \<equiv> topspace Z - (C \<inter> {x \<in> topspace Z. g x \<in> V'})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2295 | have "openin Z B" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2296 | using VV \<open>closedin Z C\<close> contg by (fastforce simp: B_def continuous_map_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2297 |         have "{x \<in> topspace Z. f x = y1} \<subseteq> B"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2298 | using A_def subA by (auto simp: A_def B_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2299 |         then obtain U where "openin X U" "y1 \<in> U" and subB: "{x \<in> topspace Z. f x \<in> U} \<subseteq> B"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2300 | using \<open>openin Z B\<close> y1 f unfolding closed_map_fibre_neighbourhood by meson | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2301 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2302 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2303 | show "openin (prod_topology X Y) (U \<times> V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2304 | by (metis VV \<open>openin X U\<close> openin_prod_Times_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2305 | show "(y1, y2) \<in> U \<times> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2306 | by (simp add: VV \<open>y1 \<in> U\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2307 | show "U \<times> V \<subseteq> topspace X \<times> topspace Y - (\<lambda>x. (f x, g x)) ` C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2308 | using VV \<open>C \<subseteq> topspace Z\<close> \<open>openin X U\<close> subB | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2309 | by (force simp: image_iff B_def subset_iff dest: openin_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2310 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2311 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2312 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2313 | then show "openin (prod_topology X Y) (topspace X \<times> topspace Y - (\<lambda>x. (f x, g x)) ` C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2314 | by (smt (verit, ccfv_threshold) Diff_iff SigmaE openin_subopen) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2315 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2316 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2317 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2318 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2319 | lemma closed_map_paired_gen: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2320 | assumes f: "closed_map Z X f" and g: "closed_map Z Y g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2321 |   and D: "(regular_space X \<and> continuous_map Z X f \<and> (\<forall>z \<in> topspace Y. closedin Z {x \<in> topspace Z. g x = z})
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2322 |          \<or> regular_space Y \<and> continuous_map Z Y g \<and> (\<forall>y \<in> topspace X. closedin Z {x \<in> topspace Z. f x = y}))"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2323 | (is "?RSX \<or> ?RSY") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2324 | shows "closed_map Z (prod_topology X Y) (\<lambda>x. (f x, g x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2325 | using D | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2326 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2327 | assume RSX: ?RSX | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2328 | have eq: "(\<lambda>x. (f x, g x)) = (\<lambda>(x,y). (y,x)) \<circ> (\<lambda>x. (g x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2329 | by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2330 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2331 | unfolding eq | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2332 | proof (rule closed_map_compose) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2333 | show "closed_map Z (prod_topology Y X) (\<lambda>x. (g x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2334 | using RSX closed_map_paired_gen_aux f g by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2335 | show "closed_map (prod_topology Y X) (prod_topology X Y) (\<lambda>(x, y). (y, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2336 | using homeomorphic_imp_closed_map homeomorphic_map_swap by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2337 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2338 | qed (blast intro: assms closed_map_paired_gen_aux) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2339 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2340 | lemma closed_map_paired: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2341 | assumes "closed_map Z X f" and contf: "continuous_map Z X f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2342 | "closed_map Z Y g" and contg: "continuous_map Z Y g" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2343 | and D: "t1_space X \<and> regular_space Y \<or> regular_space X \<and> t1_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2344 | shows "closed_map Z (prod_topology X Y) (\<lambda>x. (f x, g x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2345 | proof (rule closed_map_paired_gen) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2346 |   show "regular_space X \<and> continuous_map Z X f \<and> (\<forall>z\<in>topspace Y. closedin Z {x \<in> topspace Z. g x = z}) \<or> regular_space Y \<and> continuous_map Z Y g \<and> (\<forall>y\<in>topspace X. closedin Z {x \<in> topspace Z. f x = y})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2347 | using D contf contg | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2348 | by (smt (verit, del_insts) Collect_cong closedin_continuous_map_preimage t1_space_closedin_singleton singleton_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2349 | qed (use assms in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2350 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2351 | lemma closed_map_pairwise: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2352 | assumes "closed_map Z X (fst \<circ> f)" "continuous_map Z X (fst \<circ> f)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2353 | "closed_map Z Y (snd \<circ> f)" "continuous_map Z Y (snd \<circ> f)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2354 | "t1_space X \<and> regular_space Y \<or> regular_space X \<and> t1_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2355 | shows "closed_map Z (prod_topology X Y) f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2356 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2357 | have "closed_map Z (prod_topology X Y) (\<lambda>a. ((fst \<circ> f) a, (snd \<circ> f) a))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2358 | using assms closed_map_paired by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2359 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2360 | by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2361 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2362 | |
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2363 | lemma continuous_imp_proper_map: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2364 | "\<lbrakk>compact_space X; kc_space Y; continuous_map X Y f\<rbrakk> \<Longrightarrow> proper_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2365 | by (simp add: continuous_closed_imp_proper_map continuous_imp_closed_map_gen kc_imp_t1_space) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2366 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2367 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2368 | lemma tube_lemma_right: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2369 | assumes W: "openin (prod_topology X Y) W" and C: "compactin Y C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2370 |     and x: "x \<in> topspace X" and subW: "{x} \<times> C \<subseteq> W"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2371 | shows "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> C \<subseteq> V \<and> U \<times> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2372 | proof (cases "C = {}")
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2373 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2374 | with x show ?thesis by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2375 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2376 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2377 | have "\<exists>U V. openin X U \<and> openin Y V \<and> x \<in> U \<and> y \<in> V \<and> U \<times> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2378 | if "y \<in> C" for y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2379 | using W openin_prod_topology_alt subW subsetD that by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2380 | then obtain U V where UV: "\<And>y. y \<in> C \<Longrightarrow> openin X (U y) \<and> openin Y (V y) \<and> x \<in> U y \<and> y \<in> V y \<and> U y \<times> V y \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2381 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2382 | then obtain D where D: "finite D" "D \<subseteq> C" "C \<subseteq> \<Union> (V ` D)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2383 | using compactinD [OF C, of "V`C"] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2384 | by (smt (verit) UN_I finite_subset_image imageE subsetI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2385 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2386 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2387 | show "openin X (\<Inter> (U ` D))" "openin Y (\<Union> (V ` D))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2388 | using D False UV by blast+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2389 | show "x \<in> \<Inter> (U ` D)" "C \<subseteq> \<Union> (V ` D)" "\<Inter> (U ` D) \<times> \<Union> (V ` D) \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2390 | using D UV by force+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2391 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2392 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2393 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2394 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2395 | lemma closed_map_fst: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2396 | assumes "compact_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2397 | shows "closed_map (prod_topology X Y) X fst" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2398 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2399 |   have *: "{x \<in> topspace X \<times> topspace Y. fst x \<in> U} = U \<times> topspace Y"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2400 | if "U \<subseteq> topspace X" for U | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2401 | using that by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2402 | have **: "\<And>U y. \<lbrakk>openin (prod_topology X Y) U; y \<in> topspace X; | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2403 |             {x \<in> topspace X \<times> topspace Y. fst x = y} \<subseteq> U\<rbrakk>
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2404 | \<Longrightarrow> \<exists>V. openin X V \<and> y \<in> V \<and> V \<times> topspace Y \<subseteq> U" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2405 | using tube_lemma_right[of X Y _ "topspace Y"] assms by (fastforce simp: compact_space_def) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2406 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2407 | unfolding closed_map_fibre_neighbourhood | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2408 | by (force simp: * openin_subset cong: conj_cong intro: **) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2409 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2410 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2411 | lemma closed_map_snd: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2412 | assumes "compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2413 | shows "closed_map (prod_topology X Y) Y snd" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2414 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2415 | have "snd = fst o prod.swap" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2416 | by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2417 | moreover have "closed_map (prod_topology X Y) Y (fst o prod.swap)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2418 | proof (rule closed_map_compose) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2419 | show "closed_map (prod_topology X Y) (prod_topology Y X) prod.swap" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2420 | by (metis (no_types, lifting) homeomorphic_imp_closed_map homeomorphic_map_eq homeomorphic_map_swap prod.swap_def split_beta) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2421 | show "closed_map (prod_topology Y X) Y fst" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2422 | by (simp add: closed_map_fst assms) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2423 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2424 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2425 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2426 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2427 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2428 | lemma closed_map_paired_closed_map_right: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2429 | "\<lbrakk>closed_map X Y f; regular_space X; | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2430 |      \<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}\<rbrakk>
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2431 | \<Longrightarrow> closed_map X (prod_topology X Y) (\<lambda>x. (x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2432 | by (rule closed_map_paired_gen [OF closed_map_id, unfolded id_def]) auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2433 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2434 | lemma closed_map_paired_closed_map_left: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2435 | assumes "closed_map X Y f" "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2436 |     "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2437 | shows "closed_map X (prod_topology Y X) (\<lambda>x. (f x, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2438 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2439 | have eq: "(\<lambda>x. (f x, x)) = (\<lambda>(x,y). (y,x)) \<circ> (\<lambda>x. (x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2440 | by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2441 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2442 | unfolding eq | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2443 | proof (rule closed_map_compose) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2444 | show "closed_map X (prod_topology X Y) (\<lambda>x. (x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2445 | by (simp add: assms closed_map_paired_closed_map_right) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2446 | show "closed_map (prod_topology X Y) (prod_topology Y X) (\<lambda>(x, y). (y, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2447 | using homeomorphic_imp_closed_map homeomorphic_map_swap by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2448 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2449 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2450 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2451 | lemma closed_map_imp_closed_graph: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2452 | assumes "closed_map X Y f" "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2453 |           "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2454 | shows "closedin (prod_topology X Y) ((\<lambda>x. (x, f x)) ` topspace X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2455 | using assms closed_map_def closed_map_paired_closed_map_right by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2456 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2457 | lemma proper_map_paired_closed_map_right: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2458 | assumes "closed_map X Y f" "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2459 |     "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2460 | shows "proper_map X (prod_topology X Y) (\<lambda>x. (x, f x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2461 | by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_right) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2462 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2463 | lemma proper_map_paired_closed_map_left: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2464 | assumes "closed_map X Y f" "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2465 |     "\<And>y. y \<in> topspace Y \<Longrightarrow> closedin X {x \<in> topspace X. f x = y}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2466 | shows "proper_map X (prod_topology Y X) (\<lambda>x. (f x, x))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2467 | by (simp add: assms closed_injective_imp_proper_map inj_on_def closed_map_paired_closed_map_left) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2468 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2469 | proposition regular_space_continuous_proper_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2470 | assumes "regular_space X" and contf: "continuous_map X Y f" and pmapf: "proper_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2471 | and fim: "f ` (topspace X) = topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2472 | shows "regular_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2473 | unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2474 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2475 | fix C y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2476 | assume "closedin Y C" and "y \<in> topspace Y" and "y \<notin> C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2477 |   have "closed_map X Y f" "(\<forall>y \<in> topspace Y. compactin X {x \<in> topspace X. f x = y})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2478 | using pmapf proper_map_def by force+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2479 |   moreover have "closedin X {z \<in> topspace X. f z \<in> C}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2480 | using \<open>closedin Y C\<close> contf continuous_map_closedin by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2481 |   moreover have "disjnt {z \<in> topspace X. f z = y} {z \<in> topspace X. f z \<in> C}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2482 | using \<open>y \<notin> C\<close> disjnt_iff by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2483 | ultimately | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2484 |   obtain U V where UV: "openin X U" "openin X V" "{z \<in> topspace X. f z = y} \<subseteq> U" "{z \<in> topspace X. f z \<in> C} \<subseteq> V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2485 | and dUV: "disjnt U V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2486 | using \<open>y \<in> topspace Y\<close> \<open>regular_space X\<close> unfolding regular_space_compact_closed_sets | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2487 | by meson | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2488 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2489 |   have *: "\<And>U T. openin X U \<and> T \<subseteq> topspace Y \<and> {x \<in> topspace X. f x \<in> T} \<subseteq> U \<longrightarrow>
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2490 |          (\<exists>V. openin Y V \<and> T \<subseteq> V \<and> {x \<in> topspace X. f x \<in> V} \<subseteq> U)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2491 | using \<open>closed_map X Y f\<close> unfolding closed_map_preimage_neighbourhood by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2492 |   obtain V1 where V1: "openin Y V1 \<and> y \<in> V1" and sub1: "{x \<in> topspace X. f x \<in> V1} \<subseteq> U"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2493 |     using * [of U "{y}"] UV \<open>y \<in> topspace Y\<close> by auto
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2494 | moreover | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2495 |   obtain V2 where "openin Y V2 \<and> C \<subseteq> V2" and sub2: "{x \<in> topspace X. f x \<in> V2} \<subseteq> V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2496 | by (smt (verit, ccfv_SIG) * UV \<open>closedin Y C\<close> closedin_subset mem_Collect_eq subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2497 | moreover have "disjnt V1 V2" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2498 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2499 | have "\<And>x. \<lbrakk>\<forall>x. x \<in> U \<longrightarrow> x \<notin> V; x \<in> V1; x \<in> V2\<rbrakk> \<Longrightarrow> False" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2500 | by (smt (verit) V1 fim image_iff mem_Collect_eq openin_subset sub1 sub2 subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2501 | with dUV show ?thesis by (auto simp: disjnt_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2502 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2503 | ultimately show "\<exists>U V. openin Y U \<and> openin Y V \<and> y \<in> U \<and> C \<subseteq> V \<and> disjnt U V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2504 | by meson | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2505 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2506 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2507 | lemma regular_space_perfect_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2508 | "\<lbrakk>regular_space X; perfect_map X Y f\<rbrakk> \<Longrightarrow> regular_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2509 | by (meson perfect_map_def regular_space_continuous_proper_map_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2510 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2511 | proposition regular_space_perfect_map_image_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2512 | assumes "Hausdorff_space X" and perf: "perfect_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2513 | shows "regular_space X \<longleftrightarrow> regular_space Y" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2514 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2515 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2516 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2517 | using perf regular_space_perfect_map_image by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2518 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2519 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2520 | have "continuous_map X Y f" "proper_map X Y f" and fim: "f ` (topspace X) = topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2521 | using perf by (auto simp: perfect_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2522 |   then have "closed_map X Y f" and preYf: "(\<forall>y \<in> topspace Y. compactin X {x \<in> topspace X. f x = y})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2523 | by (simp_all add: proper_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2524 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2525 | unfolding regular_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2526 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2527 | fix C x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2528 | assume "closedin X C" and "x \<in> topspace X" and "x \<notin> C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2529 |     obtain U1 U2 where "openin X U1" "openin X U2" "{x} \<subseteq> U1" and "disjnt U1 U2"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2530 |       and subV: "C \<inter> {z \<in> topspace X. f z = f x} \<subseteq> U2"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2531 |     proof (rule Hausdorff_space_compact_separation [of X "{x}" "C \<inter> {z \<in> topspace X. f z = f x}", OF \<open>Hausdorff_space X\<close>])
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2532 |       show "compactin X {x}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2533 | by (simp add: \<open>x \<in> topspace X\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2534 |       show "compactin X (C \<inter> {z \<in> topspace X. f z = f x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2535 | using \<open>closedin X C\<close> fim \<open>x \<in> topspace X\<close> closed_Int_compactin preYf by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2536 |       show "disjnt {x} (C \<inter> {z \<in> topspace X. f z = f x})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2537 | using \<open>x \<notin> C\<close> by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2538 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2539 | have "closedin Y (f ` (C - U2))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2540 | using \<open>closed_map X Y f\<close> \<open>closedin X C\<close> \<open>openin X U2\<close> closed_map_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2541 | moreover | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2542 | have "f x \<in> topspace Y - f ` (C - U2)" | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 2543 | using \<open>closedin X C\<close> \<open>continuous_map X Y f\<close> \<open>x \<in> topspace X\<close> closedin_subset continuous_map_def subV | 
| 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 2544 | by (fastforce simp: Pi_iff) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2545 | ultimately | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2546 | obtain V1 V2 where VV: "openin Y V1" "openin Y V2" "f x \<in> V1" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2547 | and subV2: "f ` (C - U2) \<subseteq> V2" and "disjnt V1 V2" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2548 | by (meson R regular_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2549 | show "\<exists>U U'. openin X U \<and> openin X U' \<and> x \<in> U \<and> C \<subseteq> U' \<and> disjnt U U'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2550 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2551 |       show "openin X (U1 \<inter> {x \<in> topspace X. f x \<in> V1})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2552 | using VV(1) \<open>continuous_map X Y f\<close> \<open>openin X U1\<close> continuous_map by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2553 |       show "openin X (U2 \<union> {x \<in> topspace X. f x \<in> V2})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2554 | using VV(2) \<open>continuous_map X Y f\<close> \<open>openin X U2\<close> continuous_map by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2555 |       show "x \<in> U1 \<inter> {x \<in> topspace X. f x \<in> V1}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2556 |         using VV(3) \<open>x \<in> topspace X\<close> \<open>{x} \<subseteq> U1\<close> by auto
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2557 |       show "C \<subseteq> U2 \<union> {x \<in> topspace X. f x \<in> V2}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2558 | using \<open>closedin X C\<close> closedin_subset subV2 by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2559 |       show "disjnt (U1 \<inter> {x \<in> topspace X. f x \<in> V1}) (U2 \<union> {x \<in> topspace X. f x \<in> V2})"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2560 | using \<open>disjnt U1 U2\<close> \<open>disjnt V1 V2\<close> by (auto simp: disjnt_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2561 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2562 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2563 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2564 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2565 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2566 | subsection\<open>Locally compact spaces\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2567 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2568 | definition locally_compact_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2569 | where "locally_compact_space X \<equiv> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2570 | \<forall>x \<in> topspace X. \<exists>U K. openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2571 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2572 | lemma homeomorphic_locally_compact_spaceD: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2573 | assumes X: "locally_compact_space X" and "X homeomorphic_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2574 | shows "locally_compact_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2575 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2576 | obtain f where hmf: "homeomorphic_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2577 | using assms homeomorphic_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2578 | then have eq: "topspace Y = f ` (topspace X)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2579 | by (simp add: homeomorphic_imp_surjective_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2580 | have "\<exists>V K. openin Y V \<and> compactin Y K \<and> f x \<in> V \<and> V \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2581 | if "x \<in> topspace X" "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K" for x U K | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2582 | using that | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2583 | by (meson hmf homeomorphic_map_compactness_eq homeomorphic_map_openness_eq image_mono image_eqI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2584 | with X show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2585 | by (smt (verit) eq image_iff locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2586 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2587 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2588 | lemma homeomorphic_locally_compact_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2589 | assumes "X homeomorphic_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2590 | shows "locally_compact_space X \<longleftrightarrow> locally_compact_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2591 | by (meson assms homeomorphic_locally_compact_spaceD homeomorphic_space_sym) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2592 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2593 | lemma locally_compact_space_retraction_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2594 | assumes "retraction_map X Y r" and X: "locally_compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2595 | shows "locally_compact_space Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2596 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2597 | obtain s where s: "retraction_maps X Y r s" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2598 | using assms retraction_map_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2599 | obtain T where "T retract_of_space X" and Teq: "T = s ` topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2600 | using retraction_maps_section_image1 s by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2601 | then obtain r where r: "continuous_map X (subtopology X T) r" "\<forall>x\<in>T. r x = x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2602 | by (meson retract_of_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2603 | have "locally_compact_space (subtopology X T)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2604 | unfolding locally_compact_space_def openin_subtopology_alt | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2605 | proof clarsimp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2606 | fix x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2607 | assume "x \<in> topspace X" "x \<in> T" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2608 | obtain U K where UK: "openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2609 | by (meson X \<open>x \<in> topspace X\<close> locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2610 | then have "compactin (subtopology X T) (r ` K) \<and> T \<inter> U \<subseteq> r ` K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2611 | by (smt (verit) IntD1 image_compactin image_iff inf_le2 r subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2612 | then show "\<exists>U. openin X U \<and> (\<exists>K. compactin (subtopology X T) K \<and> x \<in> U \<and> T \<inter> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2613 | using UK by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2614 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2615 | with Teq show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2616 | using homeomorphic_locally_compact_space retraction_maps_section_image2 s by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2617 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2618 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2619 | lemma compact_imp_locally_compact_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2620 | "compact_space X \<Longrightarrow> locally_compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2621 | using compact_space_def locally_compact_space_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2622 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2623 | lemma neighbourhood_base_imp_locally_compact_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2624 | "neighbourhood_base_of (compactin X) X \<Longrightarrow> locally_compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2625 | by (metis locally_compact_space_def neighbourhood_base_of openin_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2626 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2627 | lemma locally_compact_imp_neighbourhood_base: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2628 | assumes loc: "locally_compact_space X" and reg: "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2629 | shows "neighbourhood_base_of (compactin X) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2630 | unfolding neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2631 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2632 | fix W x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2633 | assume "openin X W" and "x \<in> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2634 | then obtain U K where "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2635 | by (metis loc locally_compact_space_def openin_subset subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2636 | moreover have "openin X (U \<inter> W) \<and> x \<in> U \<inter> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2637 | using \<open>openin X W\<close> \<open>x \<in> W\<close> \<open>openin X U\<close> \<open>x \<in> U\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2638 | then have "\<exists>u' v. openin X u' \<and> closedin X v \<and> x \<in> u' \<and> u' \<subseteq> v \<and> v \<subseteq> U \<and> v \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2639 | using reg | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2640 | by (metis le_infE neighbourhood_base_of neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2641 | then show "\<exists>U V. openin X U \<and> compactin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2642 | by (meson \<open>U \<subseteq> K\<close> \<open>compactin X K\<close> closed_compactin subset_trans) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2643 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2644 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2645 | lemma Hausdorff_regular: "\<lbrakk>Hausdorff_space X; neighbourhood_base_of (compactin X) X\<rbrakk> \<Longrightarrow> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2646 | by (metis compactin_imp_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2647 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2648 | lemma locally_compact_Hausdorff_imp_regular_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2649 | assumes loc: "locally_compact_space X" and "Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2650 | shows "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2651 | unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2652 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2653 | fix W x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2654 | assume "openin X W" and "x \<in> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2655 | then have "x \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2656 | using openin_subset by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2657 | then obtain U K where "openin X U" "compactin X K" and UK: "x \<in> U" "U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2658 | by (meson loc locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2659 | with \<open>Hausdorff_space X\<close> have "regular_space (subtopology X K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2660 | using Hausdorff_space_subtopology compact_Hausdorff_imp_regular_space compact_space_subtopology by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2661 | then have "\<exists>U' V'. openin (subtopology X K) U' \<and> closedin (subtopology X K) V' \<and> x \<in> U' \<and> U' \<subseteq> V' \<and> V' \<subseteq> K \<inter> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2662 | unfolding neighbourhood_base_of_closedin [symmetric] neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2663 | by (meson IntI \<open>U \<subseteq> K\<close> \<open>openin X W\<close> \<open>x \<in> U\<close> \<open>x \<in> W\<close> openin_subtopology_Int2 subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2664 | then obtain V C where "openin X V" "closedin X C" and VC: "x \<in> K \<inter> V" "K \<inter> V \<subseteq> K \<inter> C" "K \<inter> C \<subseteq> K \<inter> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2665 | by (metis Int_commute closedin_subtopology openin_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2666 | show "\<exists>U V. openin X U \<and> closedin X V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2667 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2668 | show "openin X (U \<inter> V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2669 | using \<open>openin X U\<close> \<open>openin X V\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2670 | show "closedin X (K \<inter> C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2671 | using \<open>closedin X C\<close> \<open>compactin X K\<close> compactin_imp_closedin \<open>Hausdorff_space X\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2672 | qed (use UK VC in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2673 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2674 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2675 | lemma locally_compact_space_neighbourhood_base: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2676 | "Hausdorff_space X \<or> regular_space X | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2677 | \<Longrightarrow> locally_compact_space X \<longleftrightarrow> neighbourhood_base_of (compactin X) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2678 | by (metis locally_compact_imp_neighbourhood_base locally_compact_Hausdorff_imp_regular_space | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2679 | neighbourhood_base_imp_locally_compact_space) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2680 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2681 | lemma locally_compact_Hausdorff_or_regular: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2682 | "locally_compact_space X \<and> (Hausdorff_space X \<or> regular_space X) \<longleftrightarrow> locally_compact_space X \<and> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2683 | using locally_compact_Hausdorff_imp_regular_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2684 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2685 | lemma locally_compact_space_compact_closedin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2686 | assumes "Hausdorff_space X \<or> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2687 | shows "locally_compact_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2688 | (\<forall>x \<in> topspace X. \<exists>U K. openin X U \<and> compactin X K \<and> closedin X K \<and> x \<in> U \<and> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2689 | using locally_compact_Hausdorff_or_regular unfolding locally_compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2690 | by (metis assms closed_compactin inf.absorb_iff2 le_infE neighbourhood_base_of neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2691 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2692 | lemma locally_compact_space_compact_closure_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2693 | assumes "Hausdorff_space X \<or> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2694 | shows "locally_compact_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2695 | (\<forall>x \<in> topspace X. \<exists>U. openin X U \<and> compactin X (X closure_of U) \<and> x \<in> U)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2696 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2697 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2698 | by (metis assms closed_compactin closedin_closure_of closure_of_eq closure_of_mono locally_compact_space_compact_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2699 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2700 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2701 | by (meson closure_of_subset locally_compact_space_def openin_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2702 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2703 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2704 | lemma locally_compact_space_neighbourhood_base_closedin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2705 | assumes "Hausdorff_space X \<or> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2706 | shows "locally_compact_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>C. compactin X C \<and> closedin X C) X" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2707 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2708 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2709 | then have "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2710 | using assms locally_compact_Hausdorff_imp_regular_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2711 | with L have "neighbourhood_base_of (compactin X) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2712 | by (simp add: locally_compact_imp_neighbourhood_base) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2713 | with \<open>regular_space X\<close> show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2714 | by (smt (verit, ccfv_threshold) closed_compactin neighbourhood_base_of subset_trans neighbourhood_base_of_closedin) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2715 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2716 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2717 | using neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_mono by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2718 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2719 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2720 | lemma locally_compact_space_neighbourhood_base_closure_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2721 | assumes "Hausdorff_space X \<or> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2722 | shows "locally_compact_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>T. compactin X (X closure_of T)) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2723 | (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2724 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2725 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2726 | then have "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2727 | using assms locally_compact_Hausdorff_imp_regular_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2728 | with L have "neighbourhood_base_of (\<lambda>A. compactin X A \<and> closedin X A) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2729 | using locally_compact_space_neighbourhood_base_closedin by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2730 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2731 | by (simp add: closure_of_closedin neighbourhood_base_of_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2732 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2733 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2734 | unfolding locally_compact_space_def neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2735 | by (meson closure_of_subset openin_topspace subset_trans) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2736 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2737 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2738 | lemma locally_compact_space_neighbourhood_base_open_closure_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2739 | assumes "Hausdorff_space X \<or> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2740 | shows "locally_compact_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2741 | neighbourhood_base_of (\<lambda>U. openin X U \<and> compactin X (X closure_of U)) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2742 | (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2743 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2744 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2745 | then have "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2746 | using assms locally_compact_Hausdorff_imp_regular_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2747 | then have "neighbourhood_base_of (\<lambda>T. compactin X (X closure_of T)) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2748 | using L locally_compact_space_neighbourhood_base_closure_of by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2749 | with L show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2750 | unfolding neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2751 | by (meson closed_compactin closure_of_closure_of closure_of_eq closure_of_mono subset_trans) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2752 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2753 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2754 | unfolding locally_compact_space_def neighbourhood_base_of | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2755 | by (meson closure_of_subset openin_topspace subset_trans) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2756 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2757 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2758 | lemma locally_compact_space_compact_closed_compact: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2759 | assumes "Hausdorff_space X \<or> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2760 | shows "locally_compact_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2761 | (\<forall>K. compactin X K | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2762 | \<longrightarrow> (\<exists>U L. openin X U \<and> compactin X L \<and> closedin X L \<and> K \<subseteq> U \<and> U \<subseteq> L))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2763 | (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2764 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2765 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2766 | then obtain U L where UL: "\<forall>x \<in> topspace X. openin X (U x) \<and> compactin X (L x) \<and> closedin X (L x) \<and> x \<in> U x \<and> U x \<subseteq> L x" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2767 | unfolding locally_compact_space_compact_closedin [OF assms] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2768 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2769 | show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2770 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2771 | fix K | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2772 | assume "compactin X K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2773 | then have "K \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2774 | by (simp add: compactin_subset_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2775 | then have *: "(\<forall>U\<in>U ` K. openin X U) \<and> K \<subseteq> \<Union> (U ` K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2776 | using UL by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2777 | with \<open>compactin X K\<close> obtain KF where KF: "finite KF" "KF \<subseteq> K" "K \<subseteq> \<Union>(U ` KF)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2778 | by (metis compactinD finite_subset_image) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2779 | show "\<exists>U L. openin X U \<and> compactin X L \<and> closedin X L \<and> K \<subseteq> U \<and> U \<subseteq> L" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2780 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2781 | show "openin X (\<Union> (U ` KF))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2782 | using "*" \<open>KF \<subseteq> K\<close> by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2783 | show "compactin X (\<Union> (L ` KF))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2784 | by (smt (verit) UL \<open>K \<subseteq> topspace X\<close> KF compactin_Union finite_imageI imageE subset_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2785 | show "closedin X (\<Union> (L ` KF))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2786 | by (smt (verit) UL \<open>K \<subseteq> topspace X\<close> KF closedin_Union finite_imageI imageE subsetD) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2787 | qed (use UL \<open>K \<subseteq> topspace X\<close> KF in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2788 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2789 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2790 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2791 | by (metis compactin_sing insert_subset locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2792 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2793 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2794 | lemma locally_compact_regular_space_neighbourhood_base: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2795 | "locally_compact_space X \<and> regular_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2796 | neighbourhood_base_of (\<lambda>C. compactin X C \<and> closedin X C) X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2797 | using locally_compact_space_neighbourhood_base_closedin neighbourhood_base_of_closedin neighbourhood_base_of_mono by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2798 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2799 | lemma locally_compact_kc_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2800 | "neighbourhood_base_of (compactin X) X \<and> kc_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2801 | locally_compact_space X \<and> Hausdorff_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2802 | using Hausdorff_imp_kc_space locally_compact_imp_kc_eq_Hausdorff_space locally_compact_space_neighbourhood_base by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2803 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2804 | lemma locally_compact_kc_space_alt: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2805 | "neighbourhood_base_of (compactin X) X \<and> kc_space X \<longleftrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2806 | locally_compact_space X \<and> Hausdorff_space X \<and> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2807 | using Hausdorff_regular locally_compact_kc_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2808 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2809 | lemma locally_compact_kc_imp_regular_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2810 | "\<lbrakk>neighbourhood_base_of (compactin X) X; kc_space X\<rbrakk> \<Longrightarrow> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2811 | using Hausdorff_regular locally_compact_imp_kc_eq_Hausdorff_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2812 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2813 | lemma kc_locally_compact_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2814 | "kc_space X | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2815 | \<Longrightarrow> neighbourhood_base_of (compactin X) X \<longleftrightarrow> locally_compact_space X \<and> Hausdorff_space X \<and> regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2816 | using Hausdorff_regular locally_compact_kc_space by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2817 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2818 | lemma locally_compact_space_closed_subset: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2819 | assumes loc: "locally_compact_space X" and "closedin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2820 | shows "locally_compact_space (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2821 | proof (clarsimp simp: locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2822 | fix x assume x: "x \<in> topspace X" "x \<in> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2823 | then obtain U K where UK: "openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2824 | by (meson loc locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2825 | show "\<exists>U. openin (subtopology X S) U \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2826 | (\<exists>K. compactin (subtopology X S) K \<and> x \<in> U \<and> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2827 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2828 | show "openin (subtopology X S) (S \<inter> U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2829 | by (simp add: UK openin_subtopology_Int2) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2830 | show "compactin (subtopology X S) (S \<inter> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2831 | by (simp add: UK assms(2) closed_Int_compactin compactin_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2832 | qed (use UK x in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2833 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2834 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2835 | lemma locally_compact_space_open_subset: | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2836 | assumes X: "Hausdorff_space X \<or> regular_space X" and loc: "locally_compact_space X" and "openin X S" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2837 | shows "locally_compact_space (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2838 | proof (clarsimp simp: locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2839 | fix x assume x: "x \<in> topspace X" "x \<in> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2840 | then obtain U K where UK: "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2841 | by (meson loc locally_compact_space_def) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2842 | moreover have reg: "regular_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2843 | using X loc locally_compact_Hausdorff_imp_regular_space by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2844 | moreover have "openin X (U \<inter> S)" | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2845 | by (simp add: UK \<open>openin X S\<close> openin_Int) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2846 | ultimately obtain V C | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2847 | where VC: "openin X V" "closedin X C" "x \<in> V" "V \<subseteq> C" "C \<subseteq> U" "C \<subseteq> S" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2848 | by (metis \<open>x \<in> S\<close> IntI le_inf_iff neighbourhood_base_of neighbourhood_base_of_closedin) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2849 | show "\<exists>U. openin (subtopology X S) U \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2850 | (\<exists>K. compactin (subtopology X S) K \<and> x \<in> U \<and> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2851 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2852 | show "openin (subtopology X S) V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2853 | using VC by (meson \<open>openin X S\<close> openin_open_subtopology order_trans) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2854 | show "compactin (subtopology X S) (C \<inter> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2855 | using UK VC closed_Int_compactin compactin_subtopology by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2856 | qed (use UK VC x in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2857 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2858 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2859 | lemma locally_compact_space_discrete_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2860 | "locally_compact_space (discrete_topology U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2861 | by (simp add: neighbourhood_base_imp_locally_compact_space neighbourhood_base_of_discrete_topology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2862 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2863 | lemma locally_compact_space_continuous_open_map_image: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2864 | "\<lbrakk>continuous_map X X' f; open_map X X' f; | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2865 | f ` topspace X = topspace X'; locally_compact_space X\<rbrakk> \<Longrightarrow> locally_compact_space X'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2866 | unfolding locally_compact_space_def open_map_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2867 | by (smt (verit, ccfv_SIG) image_compactin image_iff image_mono) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2868 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2869 | lemma locally_compact_subspace_openin_closure_of: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2870 | assumes "Hausdorff_space X" and S: "S \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2871 | and loc: "locally_compact_space (subtopology X S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2872 | shows "openin (subtopology X (X closure_of S)) S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2873 | unfolding openin_subopen [where S=S] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2874 | proof clarify | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2875 | fix a assume "a \<in> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2876 | then obtain T K where *: "openin X T" "compactin X K" "K \<subseteq> S" "a \<in> S" "a \<in> T" "S \<inter> T \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2877 | using loc unfolding locally_compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2878 | by (metis IntE S compactin_subtopology inf_commute openin_subtopology topspace_subtopology_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2879 | have "T \<inter> X closure_of S \<subseteq> X closure_of (T \<inter> S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2880 | by (simp add: "*"(1) openin_Int_closure_of_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2881 | also have "... \<subseteq> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2882 | using * \<open>Hausdorff_space X\<close> by (metis closure_of_minimal compactin_imp_closedin order.trans inf_commute) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2883 | finally have "T \<inter> X closure_of S \<subseteq> T \<inter> S" by simp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2884 | then have "openin (subtopology X (X closure_of S)) (T \<inter> S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2885 | unfolding openin_subtopology using \<open>openin X T\<close> S closure_of_subset by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2886 | with * show "\<exists>T. openin (subtopology X (X closure_of S)) T \<and> a \<in> T \<and> T \<subseteq> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2887 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2888 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2889 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2890 | lemma locally_compact_subspace_closed_Int_openin: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2891 | "\<lbrakk>Hausdorff_space X \<and> S \<subseteq> topspace X \<and> locally_compact_space(subtopology X S)\<rbrakk> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2892 | \<Longrightarrow> \<exists>C U. closedin X C \<and> openin X U \<and> C \<inter> U = S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2893 | by (metis closedin_closure_of inf_commute locally_compact_subspace_openin_closure_of openin_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2894 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2895 | lemma locally_compact_subspace_open_in_closure_of_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2896 | assumes "Hausdorff_space X" and loc: "locally_compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2897 | shows "openin (subtopology X (X closure_of S)) S \<longleftrightarrow> S \<subseteq> topspace X \<and> locally_compact_space(subtopology X S)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2898 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2899 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2900 | then obtain "S \<subseteq> topspace X" "regular_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2901 | using assms locally_compact_Hausdorff_imp_regular_space openin_subset by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2902 | then have "locally_compact_space (subtopology (subtopology X (X closure_of S)) S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2903 | by (simp add: L loc locally_compact_space_closed_subset locally_compact_space_open_subset regular_space_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2904 | then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2905 | by (metis L inf.orderE inf_commute le_inf_iff openin_subset subtopology_subtopology topspace_subtopology) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2906 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2907 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2908 | using assms locally_compact_subspace_openin_closure_of by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2909 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2910 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2911 | lemma locally_compact_subspace_closed_Int_openin_eq: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2912 | assumes "Hausdorff_space X" and loc: "locally_compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2913 | shows "(\<exists>C U. closedin X C \<and> openin X U \<and> C \<inter> U = S) \<longleftrightarrow> S \<subseteq> topspace X \<and> locally_compact_space(subtopology X S)" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2914 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2915 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2916 | then obtain C U where "closedin X C" "openin X U" and Seq: "S = C \<inter> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2917 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2918 | then have "C \<subseteq> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2919 | by (simp add: closedin_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2920 | have "locally_compact_space (subtopology (subtopology X C) (topspace (subtopology X C) \<inter> U))" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2921 | proof (rule locally_compact_space_open_subset) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2922 | show "locally_compact_space (subtopology X C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2923 | by (simp add: \<open>closedin X C\<close> loc locally_compact_space_closed_subset) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2924 | show "openin (subtopology X C) (topspace (subtopology X C) \<inter> U)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2925 | by (simp add: \<open>openin X U\<close> Int_left_commute inf_commute openin_Int openin_subtopology_Int2) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2926 | qed (simp add: Hausdorff_space_subtopology \<open>Hausdorff_space X\<close>) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2927 | then show ?rhs | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2928 | by (metis Seq \<open>C \<subseteq> topspace X\<close> inf.coboundedI1 subtopology_subtopology subtopology_topspace) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2929 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2930 | assume ?rhs then show ?lhs | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 2931 | using assms locally_compact_subspace_closed_Int_openin by blast | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2932 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2933 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2934 | lemma dense_locally_compact_openin_Hausdorff_space: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2935 | "\<lbrakk>Hausdorff_space X; S \<subseteq> topspace X; X closure_of S = topspace X; | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2936 | locally_compact_space (subtopology X S)\<rbrakk> \<Longrightarrow> openin X S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2937 | by (metis locally_compact_subspace_openin_closure_of subtopology_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2938 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2939 | lemma locally_compact_space_prod_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2940 | "locally_compact_space (prod_topology X Y) \<longleftrightarrow> | 
| 78336 | 2941 | (prod_topology X Y) = trivial_topology \<or> | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2942 | locally_compact_space X \<and> locally_compact_space Y" (is "?lhs=?rhs") | 
| 78336 | 2943 | proof (cases "(prod_topology X Y) = trivial_topology") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2944 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2945 | then show ?thesis | 
| 78336 | 2946 | using locally_compact_space_discrete_topology by force | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2947 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2948 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2949 | then obtain w z where wz: "w \<in> topspace X" "z \<in> topspace Y" | 
| 78336 | 2950 | by fastforce | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2951 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2952 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2953 | assume L: ?lhs then show ?rhs | 
| 78336 | 2954 | by (metis locally_compact_space_retraction_map_image prod_topology_trivial_iff retraction_map_fst retraction_map_snd) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2955 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2956 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2957 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2958 | unfolding locally_compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2959 | proof clarsimp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2960 | fix x y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2961 | assume "x \<in> topspace X" and "y \<in> topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2962 | obtain U C where "openin X U" "compactin X C" "x \<in> U" "U \<subseteq> C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2963 | by (meson False R \<open>x \<in> topspace X\<close> locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2964 | obtain V D where "openin Y V" "compactin Y D" "y \<in> V" "V \<subseteq> D" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2965 | by (meson False R \<open>y \<in> topspace Y\<close> locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2966 | show "\<exists>U. openin (prod_topology X Y) U \<and> (\<exists>K. compactin (prod_topology X Y) K \<and> (x, y) \<in> U \<and> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2967 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2968 | show "openin (prod_topology X Y) (U \<times> V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2969 | by (simp add: \<open>openin X U\<close> \<open>openin Y V\<close> openin_prod_Times_iff) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2970 | show "compactin (prod_topology X Y) (C \<times> D)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2971 | by (simp add: \<open>compactin X C\<close> \<open>compactin Y D\<close> compactin_Times) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2972 | show "(x, y) \<in> U \<times> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2973 | by (simp add: \<open>x \<in> U\<close> \<open>y \<in> V\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2974 | show "U \<times> V \<subseteq> C \<times> D" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2975 | by (simp add: Sigma_mono \<open>U \<subseteq> C\<close> \<open>V \<subseteq> D\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2976 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2977 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2978 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2979 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2980 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2981 | lemma locally_compact_space_product_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2982 | "locally_compact_space(product_topology X I) \<longleftrightarrow> | 
| 78336 | 2983 | product_topology X I = trivial_topology \<or> | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2984 |         finite {i \<in> I. \<not> compact_space(X i)} \<and> (\<forall>i \<in> I. locally_compact_space(X i))" (is "?lhs=?rhs")
 | 
| 78336 | 2985 | proof (cases "(product_topology X I) = trivial_topology") | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2986 | case True | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2987 | then show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2988 | by (simp add: locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2989 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2990 | case False | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2991 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2992 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2993 | assume L: ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2994 | obtain z where z: "z \<in> topspace (product_topology X I)" | 
| 78336 | 2995 | using False | 
| 2996 | by (meson ex_in_conv null_topspace_iff_trivial) | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2997 | with L z obtain U C where "openin (product_topology X I) U" "compactin (product_topology X I) C" "z \<in> U" "U \<subseteq> C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2998 | by (meson locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2999 |     then obtain V where finV: "finite {i \<in> I. V i \<noteq> topspace (X i)}" and "\<forall>i \<in> I. openin (X i) (V i)" 
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3000 | and "z \<in> PiE I V" "PiE I V \<subseteq> U" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3001 | by (auto simp: openin_product_topology_alt) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3002 | have "compact_space (X i)" if "i \<in> I" "V i = topspace (X i)" for i | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3003 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3004 | have "compactin (X i) ((\<lambda>x. x i) ` C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3005 | using \<open>compactin (product_topology X I) C\<close> image_compactin | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3006 | by (metis continuous_map_product_projection \<open>i \<in> I\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3007 | moreover have "V i \<subseteq> (\<lambda>x. x i) ` C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3008 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3009 | have "V i \<subseteq> (\<lambda>x. x i) ` Pi\<^sub>E I V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3010 | by (metis \<open>z \<in> Pi\<^sub>E I V\<close> empty_iff image_projection_PiE order_refl \<open>i \<in> I\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3011 | also have "\<dots> \<subseteq> (\<lambda>x. x i) ` C" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3012 | using \<open>U \<subseteq> C\<close> \<open>Pi\<^sub>E I V \<subseteq> U\<close> by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3013 | finally show ?thesis . | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3014 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3015 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3016 | by (metis closed_compactin closedin_topspace compact_space_def that(2)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3017 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3018 |     with finV have "finite {i \<in> I. \<not> compact_space (X i)}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3019 | by (metis (mono_tags, lifting) mem_Collect_eq finite_subset subsetI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3020 | moreover have "locally_compact_space (X i)" if "i \<in> I" for i | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3021 | by (meson False L locally_compact_space_retraction_map_image retraction_map_product_projection that) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3022 | ultimately show ?rhs by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3023 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3024 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3025 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3026 | unfolding locally_compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3027 | proof clarsimp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3028 | fix z | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3029 | assume z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3030 | have "\<exists>U C. openin (X i) U \<and> compactin (X i) C \<and> z i \<in> U \<and> U \<subseteq> C \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3031 | (compact_space(X i) \<longrightarrow> U = topspace(X i) \<and> C = topspace(X i))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3032 | if "i \<in> I" for i | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3033 | using that R z unfolding locally_compact_space_def compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3034 | by (metis (no_types, lifting) False PiE_mem openin_topspace set_eq_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3035 | then obtain U C where UC: "\<And>i. i \<in> I \<Longrightarrow> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3036 | openin (X i) (U i) \<and> compactin (X i) (C i) \<and> z i \<in> U i \<and> U i \<subseteq> C i \<and> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3037 | (compact_space(X i) \<longrightarrow> U i = topspace(X i) \<and> C i = topspace(X i))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3038 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3039 | show "\<exists>U. openin (product_topology X I) U \<and> (\<exists>K. compactin (product_topology X I) K \<and> z \<in> U \<and> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3040 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3041 | show "openin (product_topology X I) (Pi\<^sub>E I U)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3042 | by (smt (verit) Collect_cong False R UC compactin_subspace openin_PiE_gen subset_antisym subtopology_topspace) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3043 | show "compactin (product_topology X I) (Pi\<^sub>E I C)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3044 | by (simp add: UC compactin_PiE) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3045 | qed (use UC z in blast)+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3046 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3047 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3048 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3049 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3050 | lemma locally_compact_space_sum_topology: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3051 | "locally_compact_space (sum_topology X I) \<longleftrightarrow> (\<forall>i \<in> I. locally_compact_space (X i))" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3052 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3053 | assume ?lhs then show ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3054 | by (metis closed_map_component_injection embedding_map_imp_homeomorphic_space embedding_map_component_injection | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3055 | embedding_imp_closed_map_eq homeomorphic_locally_compact_space locally_compact_space_closed_subset) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3056 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3057 | assume R: ?rhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3058 | show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3059 | unfolding locally_compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3060 | proof clarsimp | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3061 | fix i y | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3062 | assume "i \<in> I" and y: "y \<in> topspace (X i)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3063 | then obtain U K where UK: "openin (X i) U" "compactin (X i) K" "y \<in> U" "U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3064 | using R by (fastforce simp: locally_compact_space_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3065 | then show "\<exists>U. openin (sum_topology X I) U \<and> (\<exists>K. compactin (sum_topology X I) K \<and> (i, y) \<in> U \<and> U \<subseteq> K)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3066 | by (metis \<open>i \<in> I\<close> continuous_map_component_injection image_compactin image_mono | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3067 | imageI open_map_component_injection open_map_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3068 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3069 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3070 | |
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3071 | lemma locally_compact_space_euclidean: | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3072 | "locally_compact_space (euclidean::'a::heine_borel topology)" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3073 | unfolding locally_compact_space_def | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3074 | proof (intro strip) | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3075 | fix x::'a | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3076 | assume "x \<in> topspace euclidean" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3077 | have "ball x 1 \<subseteq> cball x 1" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3078 | by auto | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3079 | then show "\<exists>U K. openin euclidean U \<and> compactin euclidean K \<and> x \<in> U \<and> U \<subseteq> K" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3080 | by (metis Elementary_Metric_Spaces.open_ball centre_in_ball compact_cball compactin_euclidean_iff open_openin zero_less_one) | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3081 | qed | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3082 | |
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3083 | lemma locally_compact_Euclidean_space: | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3084 | "locally_compact_space(Euclidean_space n)" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3085 | using homeomorphic_locally_compact_space [OF homeomorphic_Euclidean_space_product_topology] | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3086 | using locally_compact_space_product_topology locally_compact_space_euclidean by fastforce | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
78093diff
changeset | 3087 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3088 | proposition quotient_map_prod_right: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3089 | assumes loc: "locally_compact_space Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3090 | and reg: "Hausdorff_space Z \<or> regular_space Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3091 | and f: "quotient_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3092 | shows "quotient_map (prod_topology Z X) (prod_topology Z Y) (\<lambda>(x,y). (x,f y))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3093 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3094 | define h where "h \<equiv> (\<lambda>(x::'a,y). (x,f y))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3095 | have "continuous_map (prod_topology Z X) Y (f o snd)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3096 | by (simp add: continuous_map_of_snd f quotient_imp_continuous_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3097 | then have cmh: "continuous_map (prod_topology Z X) (prod_topology Z Y) h" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3098 | by (simp add: h_def continuous_map_paired split_def continuous_map_fst o_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3099 | have fim: "f ` topspace X = topspace Y" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3100 | by (simp add: f quotient_imp_surjective_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3101 | moreover | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3102 |   have "openin (prod_topology Z X) {u \<in> topspace Z \<times> topspace X. h u \<in> W}
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3103 | \<longleftrightarrow> openin (prod_topology Z Y) W" (is "?lhs=?rhs") | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3104 | if W: "W \<subseteq> topspace Z \<times> topspace Y" for W | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3105 | proof | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3106 |     define S where "S \<equiv> {u \<in> topspace Z \<times> topspace X. h u \<in> W}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3107 | assume ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3108 | then have L: "openin (prod_topology Z X) S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3109 | using S_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3110 | have "\<exists>T. openin (prod_topology Z Y) T \<and> (x0, z0) \<in> T \<and> T \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3111 | if \<section>: "(x0,z0) \<in> W" for x0 z0 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3112 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3113 | have x0: "x0 \<in> topspace Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3114 | using W that by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3115 | obtain y0 where y0: "y0 \<in> topspace X" "f y0 = z0" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3116 | by (metis W fim imageE insert_absorb insert_subset mem_Sigma_iff \<section>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3117 | then have "(x0, y0) \<in> S" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3118 | by (simp add: S_def h_def that x0) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3119 | have "continuous_map Z (prod_topology Z X) (\<lambda>x. (x, y0))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3120 | by (simp add: continuous_map_paired y0) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3121 | with openin_continuous_map_preimage [OF _ L] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3122 |       have ope_ZS: "openin Z {x \<in> topspace Z. (x,y0) \<in> S}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3123 | by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3124 | obtain U U' where "openin Z U" "compactin Z U'" "closedin Z U'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3125 |         "x0 \<in> U"  "U \<subseteq> U'" "U' \<subseteq> {x \<in> topspace Z. (x,y0) \<in> S}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3126 | using loc ope_ZS x0 \<open>(x0, y0) \<in> S\<close> | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3127 | by (force simp: locally_compact_space_neighbourhood_base_closedin [OF reg] | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3128 | neighbourhood_base_of) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3129 |       then have D: "U' \<times> {y0} \<subseteq> S"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3130 | by (auto simp: ) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3131 |       define V where "V \<equiv> {z \<in> topspace Y. U' \<times> {y \<in> topspace X. f y = z} \<subseteq> S}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3132 | have "z0 \<in> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3133 | using D y0 Int_Collect fim by (fastforce simp: h_def V_def S_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3134 |       have "openin X {x \<in> topspace X. f x \<in> V} \<Longrightarrow> openin Y V"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3135 | using f unfolding V_def quotient_map_def subset_iff | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3136 | by (smt (verit, del_insts) Collect_cong mem_Collect_eq) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3137 |       moreover have "openin X {x \<in> topspace X. f x \<in> V}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3138 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3139 | let ?Z = "subtopology Z U'" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3140 |         have *: "{x \<in> topspace X. f x \<in> V} = topspace X - snd ` (U' \<times> topspace X - S)"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3141 | by (force simp: V_def S_def h_def simp flip: fim) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3142 | have "compact_space ?Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3143 | using \<open>compactin Z U'\<close> compactin_subspace by auto | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3144 | moreover have "closedin (prod_topology ?Z X) (U' \<times> topspace X - S)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3145 | by (simp add: L \<open>closedin Z U'\<close> closedin_closed_subtopology closedin_diff closedin_prod_Times_iff | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3146 | prod_topology_subtopology(1)) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3147 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3148 | using "*" closed_map_snd closed_map_def by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3149 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3150 | ultimately have "openin Y V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3151 | by metis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3152 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3153 | proof (intro conjI exI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3154 | show "openin (prod_topology Z Y) (U \<times> V)" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3155 | by (simp add: openin_prod_Times_iff \<open>openin Z U\<close> \<open>openin Y V\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3156 | show "(x0, z0) \<in> U \<times> V" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3157 | by (simp add: \<open>x0 \<in> U\<close> \<open>z0 \<in> V\<close>) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3158 | show "U \<times> V \<subseteq> W" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3159 | using \<open>U \<subseteq> U'\<close> by (force simp: V_def S_def h_def simp flip: fim) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3160 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3161 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3162 | with openin_subopen show ?rhs by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3163 | next | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3164 | assume ?rhs then show ?lhs | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3165 | using openin_continuous_map_preimage cmh by fastforce | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3166 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3167 | ultimately show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3168 | by (fastforce simp: image_iff quotient_map_def h_def) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3169 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3170 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3171 | lemma quotient_map_prod_left: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3172 | assumes loc: "locally_compact_space Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3173 | and reg: "Hausdorff_space Z \<or> regular_space Z" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3174 | and f: "quotient_map X Y f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3175 | shows "quotient_map (prod_topology X Z) (prod_topology Y Z) (\<lambda>(x,y). (f x,y))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3176 | proof - | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3177 | have "(\<lambda>(x,y). (f x,y)) = prod.swap \<circ> (\<lambda>(x,y). (x,f y)) \<circ> prod.swap" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3178 | by force | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3179 | then | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3180 | show ?thesis | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3181 | apply (rule ssubst) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3182 | proof (intro quotient_map_compose) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3183 | show "quotient_map (prod_topology X Z) (prod_topology Z X) prod.swap" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3184 | "quotient_map (prod_topology Z Y) (prod_topology Y Z) prod.swap" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3185 | using homeomorphic_map_def homeomorphic_map_swap quotient_map_eq by fastforce+ | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3186 | show "quotient_map (prod_topology Z X) (prod_topology Z Y) (\<lambda>(x, y). (x, f y))" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3187 | by (simp add: f loc quotient_map_prod_right reg) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3188 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3189 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3190 | |
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3191 | lemma locally_compact_space_perfect_map_preimage: | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3192 | assumes "locally_compact_space X'" and f: "perfect_map X X' f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3193 | shows "locally_compact_space X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3194 | unfolding locally_compact_space_def | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3195 | proof (intro strip) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3196 | fix x | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3197 | assume x: "x \<in> topspace X" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3198 | then obtain U K where "openin X' U" "compactin X' K" "f x \<in> U" "U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3199 | using assms unfolding locally_compact_space_def perfect_map_def | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 3200 | by (metis (no_types, lifting) continuous_map_closedin Pi_iff) | 
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3201 | show "\<exists>U K. openin X U \<and> compactin X K \<and> x \<in> U \<and> U \<subseteq> K" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3202 | proof (intro exI conjI) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3203 | have "continuous_map X X' f" | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3204 | using f perfect_map_def by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3205 |     then show "openin X {x \<in> topspace X. f x \<in> U}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3206 | by (simp add: \<open>openin X' U\<close> continuous_map) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3207 |     show "compactin X {x \<in> topspace X. f x \<in> K}"
 | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3208 | using \<open>compactin X' K\<close> f perfect_imp_proper_map proper_map_alt by blast | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3209 | qed (use x \<open>f x \<in> U\<close> \<open>U \<subseteq> K\<close> in auto) | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3210 | qed | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3211 | |
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3212 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3213 | subsection\<open>Special characterizations of classes of functions into and out of R\<close> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3214 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3215 | lemma monotone_map_into_euclideanreal_alt: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3216 | assumes "continuous_map X euclideanreal f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3217 |   shows "(\<forall>k. is_interval k \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> k}) \<longleftrightarrow>
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3218 | connected_space X \<and> monotone_map X euclideanreal f" (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3219 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3220 | assume L: ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3221 | show ?rhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3222 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3223 | show "connected_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3224 | using L connected_space_subconnected by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3225 |     have "connectedin X {x \<in> topspace X. f x \<in> {y}}" for y
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3226 | by (metis L is_interval_1 nle_le singletonD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3227 | then show "monotone_map X euclideanreal f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3228 | by (simp add: monotone_map) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3229 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3230 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3231 | assume R: ?rhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3232 | then | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3233 | have *: False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3234 |       if "a < b" "closedin X U" "closedin X V" "U \<noteq> {}" "V \<noteq> {}" "disjnt U V"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3235 |          and UV: "{x \<in> topspace X. f x \<in> {a..b}} = U \<union> V" 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3236 |          and dis: "disjnt U {x \<in> topspace X. f x = b}" "disjnt V {x \<in> topspace X. f x = a}" 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3237 | for a b U V | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3238 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3239 |     define E1 where "E1 \<equiv> U \<union> {x \<in> topspace X. f x \<in> {c. c \<le> a}}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3240 |     define E2 where "E2 \<equiv> V \<union> {x \<in> topspace X. f x \<in> {c. b \<le> c}}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3241 |     have "closedin X {x \<in> topspace X. f x \<le> a}" "closedin X {x \<in> topspace X. b \<le> f x}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3242 | using assms continuous_map_upper_lower_semicontinuous_le by blast+ | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3243 | then have "closedin X E1" "closedin X E2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3244 | unfolding E1_def E2_def using that by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3245 | moreover | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3246 |     have "E1 \<inter> E2 = {}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3247 | unfolding E1_def E2_def using \<open>a<b\<close> \<open>disjnt U V\<close> dis UV | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3248 | by (simp add: disjnt_def set_eq_iff) (smt (verit)) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3249 | have "topspace X \<subseteq> E1 \<union> E2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3250 | unfolding E1_def E2_def using UV by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3251 |     have "E1 = {} \<or> E2 = {}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3252 | using R connected_space_closedin | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3253 |       using \<open>E1 \<inter> E2 = {}\<close> \<open>closedin X E1\<close> \<open>closedin X E2\<close> \<open>topspace X \<subseteq> E1 \<union> E2\<close> by blast
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3254 | then show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3255 |       using E1_def E2_def \<open>U \<noteq> {}\<close> \<open>V \<noteq> {}\<close> by fastforce
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3256 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3257 | show ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3258 | proof (intro strip) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3259 | fix K :: "real set" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3260 | assume "is_interval K" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3261 | have False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3262 | if "a \<in> K" "b \<in> K" and clo: "closedin X U" "closedin X V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3263 |          and UV: "{x. x \<in> topspace X \<and> f x \<in> K} \<subseteq> U \<union> V"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3264 |                  "U \<inter> V \<inter> {x. x \<in> topspace X \<and> f x \<in> K} = {}" 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3265 |          and nondis: "\<not> disjnt U {x. x \<in> topspace X \<and> f x = a}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3266 |                      "\<not> disjnt V {x. x \<in> topspace X \<and> f x = b}" 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3267 | for a b U V | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3268 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3269 |       have "\<forall>y. connectedin X {x. x \<in> topspace X \<and> f x = y}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3270 | using R monotone_map by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3271 | then have **: False if "p \<in> U \<and> q \<in> V \<and> f p = f q \<and> f q \<in> K" for p q | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3272 | unfolding connectedin_closedin | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3273 | using \<open>a \<in> K\<close> \<open>b \<in> K\<close> UV clo that | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3274 | by (smt (verit, ccfv_threshold) closedin_subset disjoint_iff mem_Collect_eq subset_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3275 | consider "a < b" | "a = b" | "b < a" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3276 | by linarith | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3277 | then show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3278 | proof cases | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3279 | case 1 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3280 |         define W where "W \<equiv> {x \<in> topspace X. f x \<in> {a..b}}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3281 | have "closedin X W" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3282 | unfolding W_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3283 | by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3284 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3285 | proof (rule * [OF 1 , of "U \<inter> W" "V \<inter> W"]) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3286 | show "closedin X (U \<inter> W)" "closedin X (V \<inter> W)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3287 | using \<open>closedin X W\<close> clo by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3288 |           show "U \<inter> W \<noteq> {}" "V \<inter> W \<noteq> {}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3289 | using nondis 1 by (auto simp: disjnt_iff W_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3290 | show "disjnt (U \<inter> W) (V \<inter> W)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3291 | using \<open>is_interval K\<close> unfolding is_interval_1 disjnt_iff W_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3292 | by (metis (mono_tags, lifting) \<open>a \<in> K\<close> \<open>b \<in> K\<close> ** Int_Collect atLeastAtMost_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3293 | have "\<And>x. \<lbrakk>x \<in> topspace X; a \<le> f x; f x \<le> b\<rbrakk> \<Longrightarrow> x \<in> U \<or> x \<in> V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3294 | using \<open>a \<in> K\<close> \<open>b \<in> K\<close> \<open>is_interval K\<close> UV unfolding is_interval_1 disjnt_iff | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3295 | by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3296 |           then show "{x \<in> topspace X. f x \<in> {a..b}} = U \<inter> W \<union> V \<inter> W"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3297 | by (auto simp: W_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3298 |           show "disjnt (U \<inter> W) {x \<in> topspace X. f x = b}" "disjnt (V \<inter> W) {x \<in> topspace X. f x = a}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3299 | using ** \<open>a \<in> K\<close> \<open>b \<in> K\<close> nondis by (force simp: disjnt_iff)+ | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3300 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3301 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3302 | case 2 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3303 | then show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3304 | using ** nondis \<open>b \<in> K\<close> by (force simp add: disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3305 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3306 | case 3 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3307 |         define W where "W \<equiv> {x \<in> topspace X. f x \<in> {b..a}}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3308 | have "closedin X W" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3309 | unfolding W_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3310 | by (metis (no_types) assms closed_real_atLeastAtMost closed_closedin continuous_map_closedin) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3311 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3312 | proof (rule * [OF 3, of "V \<inter> W" "U \<inter> W"]) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3313 | show "closedin X (U \<inter> W)" "closedin X (V \<inter> W)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3314 | using \<open>closedin X W\<close> clo by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3315 |           show "U \<inter> W \<noteq> {}" "V \<inter> W \<noteq> {}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3316 | using nondis 3 by (auto simp: disjnt_iff W_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3317 | show "disjnt (V \<inter> W) (U \<inter> W)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3318 | using \<open>is_interval K\<close> unfolding is_interval_1 disjnt_iff W_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3319 | by (metis (mono_tags, lifting) \<open>a \<in> K\<close> \<open>b \<in> K\<close> ** Int_Collect atLeastAtMost_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3320 | have "\<And>x. \<lbrakk>x \<in> topspace X; b \<le> f x; f x \<le> a\<rbrakk> \<Longrightarrow> x \<in> U \<or> x \<in> V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3321 | using \<open>a \<in> K\<close> \<open>b \<in> K\<close> \<open>is_interval K\<close> UV unfolding is_interval_1 disjnt_iff | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3322 | by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3323 |           then show "{x \<in> topspace X. f x \<in> {b..a}} = V \<inter> W \<union> U \<inter> W"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3324 | by (auto simp: W_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3325 |           show "disjnt (V \<inter> W) {x \<in> topspace X. f x = a}" "disjnt (U \<inter> W) {x \<in> topspace X. f x = b}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3326 | using ** \<open>a \<in> K\<close> \<open>b \<in> K\<close> nondis by (force simp: disjnt_iff)+ | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3327 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3328 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3329 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3330 |     then show "connectedin X {x \<in> topspace X. f x \<in> K}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3331 | unfolding connectedin_closedin disjnt_iff by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3332 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3333 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3334 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3335 | lemma monotone_map_into_euclideanreal: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3336 | "\<lbrakk>connected_space X; continuous_map X euclideanreal f\<rbrakk> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3337 | \<Longrightarrow> monotone_map X euclideanreal f \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3338 |         (\<forall>k. is_interval k \<longrightarrow> connectedin X {x \<in> topspace X. f x \<in> k})"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3339 | by (simp add: monotone_map_into_euclideanreal_alt) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3340 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3341 | lemma monotone_map_euclideanreal_alt: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3342 |    "(\<forall>I::real set. is_interval I \<longrightarrow> is_interval {x::real. x \<in> S \<and> f x \<in> I}) \<longleftrightarrow>
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3343 | is_interval S \<and> (mono_on S f \<or> antimono_on S f)" (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3344 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3345 | assume L [rule_format]: ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3346 | show ?rhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3347 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3348 | show "is_interval S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3349 | using L is_interval_1 by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3350 | have False if "a \<in> S" "b \<in> S" "c \<in> S" "a<b" "b<c" and d: "f a < f b \<and> f c < f b \<or> f a > f b \<and> f c > f b" for a b c | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3351 | using d | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3352 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3353 | assume "f a < f b \<and> f c < f b" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3354 | then show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3355 |         using L [of "{y.  y < f b}"] unfolding is_interval_1
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3356 | by (smt (verit, best) mem_Collect_eq that) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3357 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3358 | assume "f b < f a \<and> f b < f c" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3359 | then show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3360 |         using L [of "{y.  y > f b}"] unfolding is_interval_1
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3361 | by (smt (verit, best) mem_Collect_eq that) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3362 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3363 | then show "mono_on S f \<or> monotone_on S (\<le>) (\<ge>) f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3364 | unfolding monotone_on_def by (smt (verit)) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3365 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3366 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3367 | assume ?rhs then show ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3368 | unfolding is_interval_1 monotone_on_def by simp meson | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3369 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3370 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3371 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3372 | lemma monotone_map_euclideanreal: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3373 | fixes S :: "real set" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3374 | shows | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3375 | "\<lbrakk>is_interval S; continuous_on S f\<rbrakk> \<Longrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3376 | monotone_map (top_of_set S) euclideanreal f \<longleftrightarrow> (mono_on S f \<or> monotone_on S (\<le>) (\<ge>) f)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3377 | using monotone_map_euclideanreal_alt | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3378 | by (simp add: monotone_map_into_euclideanreal connectedin_subtopology is_interval_connected_1) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3379 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3380 | lemma injective_eq_monotone_map: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3381 | fixes f :: "real \<Rightarrow> real" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3382 | assumes "is_interval S" "continuous_on S f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3383 | shows "inj_on f S \<longleftrightarrow> strict_mono_on S f \<or> strict_antimono_on S f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3384 | by (metis assms injective_imp_monotone_map monotone_map_euclideanreal strict_antimono_iff_antimono | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3385 | strict_mono_iff_mono top_greatest topspace_euclidean topspace_euclidean_subtopology) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3386 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3387 | |
| 78277 | 3388 | subsection\<open>Normal spaces\<close> | 
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3389 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3390 | definition normal_space | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3391 | where "normal_space X \<equiv> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3392 | \<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3393 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3394 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3395 | lemma normal_space_retraction_map_image: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3396 | assumes r: "retraction_map X Y r" and X: "normal_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3397 | shows "normal_space Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3398 | unfolding normal_space_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3399 | proof clarify | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3400 | fix S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3401 | assume "closedin Y S" and "closedin Y T" and "disjnt S T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3402 | obtain r' where r': "retraction_maps X Y r r'" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3403 | using r retraction_map_def by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3404 |   have "closedin X {x \<in> topspace X. r x \<in> S}" "closedin X {x \<in> topspace X. r x \<in> T}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3405 | using closedin_continuous_map_preimage \<open>closedin Y S\<close> \<open>closedin Y T\<close> r' | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3406 | by (auto simp: retraction_maps_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3407 | moreover | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3408 |   have "disjnt {x \<in> topspace X. r x \<in> S} {x \<in> topspace X. r x \<in> T}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3409 | using \<open>disjnt S T\<close> by (auto simp: disjnt_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3410 | ultimately | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3411 |   obtain U V where UV: "openin X U \<and> openin X V \<and> {x \<in> topspace X. r x \<in> S} \<subseteq> U \<and> {x \<in> topspace X. r x \<in> T} \<subseteq> V" "disjnt U V"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3412 | by (meson X normal_space_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3413 | show "\<exists>U V. openin Y U \<and> openin Y V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3414 | proof (intro exI conjI) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3415 |     show "openin Y {x \<in> topspace Y. r' x \<in> U}" "openin Y {x \<in> topspace Y. r' x \<in> V}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3416 | using openin_continuous_map_preimage UV r' | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3417 | by (auto simp: retraction_maps_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3418 |     show "S \<subseteq> {x \<in> topspace Y. r' x \<in> U}" "T \<subseteq> {x \<in> topspace Y. r' x \<in> V}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3419 | using openin_continuous_map_preimage UV r' \<open>closedin Y S\<close> \<open>closedin Y T\<close> | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 3420 | by (auto simp add: closedin_def continuous_map_closedin retraction_maps_def subset_iff Pi_iff) | 
| 77943 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3421 |     show "disjnt {x \<in> topspace Y. r' x \<in> U} {x \<in> topspace Y. r' x \<in> V}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3422 | using \<open>disjnt U V\<close> by (auto simp: disjnt_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3423 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3424 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3425 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3426 | lemma homeomorphic_normal_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3427 | "X homeomorphic_space Y \<Longrightarrow> normal_space X \<longleftrightarrow> normal_space Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3428 | unfolding homeomorphic_space_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3429 | by (meson homeomorphic_imp_retraction_maps homeomorphic_maps_sym normal_space_retraction_map_image retraction_map_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3430 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3431 | lemma normal_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3432 | "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3433 | (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3434 | \<longrightarrow> (\<exists>U. openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U)))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3435 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3436 | have "(\<exists>V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V) \<longleftrightarrow> openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3437 | (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3438 | if "closedin X S" "closedin X T" "disjnt S T" for S T U | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3439 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3440 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3441 | by (smt (verit, best) disjnt_iff in_closure_of subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3442 | assume R: ?rhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3443 |     then have "(U \<union> S) \<inter> (topspace X - X closure_of U) = {}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3444 | by (metis Diff_eq_empty_iff Int_Diff Int_Un_eq(4) closure_of_subset inf.orderE openin_subset) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3445 | moreover have "T \<subseteq> topspace X - X closure_of U" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3446 | by (meson DiffI R closedin_subset disjnt_iff subsetD subsetI that(2)) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3447 | ultimately show ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3448 | by (metis R closedin_closure_of closedin_def disjnt_def sup.orderE) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3449 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3450 | then show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3451 | unfolding normal_space_def by meson | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3452 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3453 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3454 | lemma normal_space_alt: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3455 | "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3456 | (\<forall>S U. closedin X S \<and> openin X U \<and> S \<subseteq> U \<longrightarrow> (\<exists>V. openin X V \<and> S \<subseteq> V \<and> X closure_of V \<subseteq> U))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3457 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3458 | have "\<exists>V. openin X V \<and> S \<subseteq> V \<and> X closure_of V \<subseteq> U" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3459 | if "\<And>T. closedin X T \<longrightarrow> disjnt S T \<longrightarrow> (\<exists>U. openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3460 | "closedin X S" "openin X U" "S \<subseteq> U" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3461 | for S U | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3462 | using that | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3463 | by (smt (verit) Diff_eq_empty_iff Int_Diff closure_of_subset_topspace disjnt_def inf.orderE inf_commute openin_closedin_eq) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3464 | moreover have "\<exists>U. openin X U \<and> S \<subseteq> U \<and> disjnt T (X closure_of U)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3465 | if "\<And>U. openin X U \<and> S \<subseteq> U \<longrightarrow> (\<exists>V. openin X V \<and> S \<subseteq> V \<and> X closure_of V \<subseteq> U)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3466 | and "closedin X S" "closedin X T" "disjnt S T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3467 | for S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3468 | using that | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3469 | by (smt (verit) Diff_Diff_Int Diff_eq_empty_iff Int_Diff closedin_def disjnt_def inf.absorb_iff2 inf.orderE) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3470 | ultimately show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3471 | by (fastforce simp: normal_space) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3472 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3473 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3474 | lemma normal_space_closures: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3475 | "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3476 | (\<forall>S T. S \<subseteq> topspace X \<and> T \<subseteq> topspace X \<and> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3477 | disjnt (X closure_of S) (X closure_of T) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3478 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3479 | (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3480 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3481 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3482 | by (meson closedin_closure_of closure_of_subset normal_space_def order.trans) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3483 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3484 | by (metis closedin_subset closure_of_eq normal_space_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3485 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3486 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3487 | lemma normal_space_disjoint_closures: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3488 | "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3489 | (\<forall>S T. closedin X S \<and> closedin X T \<and> disjnt S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3490 | \<longrightarrow> (\<exists>U V. openin X U \<and> openin X V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3491 | disjnt (X closure_of U) (X closure_of V)))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3492 | (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3493 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3494 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3495 | by (metis closedin_closure_of normal_space) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3496 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3497 | by (smt (verit) closure_of_subset disjnt_iff normal_space openin_subset subset_eq) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3498 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3499 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3500 | lemma normal_space_dual: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3501 | "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3502 | (\<forall>U V. openin X U \<longrightarrow> openin X V \<and> U \<union> V = topspace X | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3503 | \<longrightarrow> (\<exists>S T. closedin X S \<and> closedin X T \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> S \<union> T = topspace X))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3504 | (is "_ = ?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3505 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3506 | have "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3507 | (\<forall>U V. closedin X U \<longrightarrow> closedin X V \<longrightarrow> disjnt U V \<longrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3508 | (\<exists>S T. \<not> (openin X S \<and> openin X T \<longrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3509 | \<not> (U \<subseteq> S \<and> V \<subseteq> T \<and> disjnt S T))))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3510 | unfolding normal_space_def by meson | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3511 | also have "... \<longleftrightarrow> (\<forall>U V. openin X U \<longrightarrow> openin X V \<and> disjnt (topspace X - U) (topspace X - V) \<longrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3512 | (\<exists>S T. \<not> (openin X S \<and> openin X T \<longrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3513 | \<not> (topspace X - U \<subseteq> S \<and> topspace X - V \<subseteq> T \<and> disjnt S T))))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3514 | by (auto simp: all_closedin) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3515 | also have "... \<longleftrightarrow> ?rhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3516 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3517 | have *: "disjnt (topspace X - U) (topspace X - V) \<longleftrightarrow> U \<union> V = topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3518 | if "U \<subseteq> topspace X" "V \<subseteq> topspace X" for U V | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3519 | using that by (auto simp: disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3520 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3521 | using ex_closedin * | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3522 | apply (simp add: ex_closedin * [OF openin_subset openin_subset] cong: conj_cong) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3523 | apply (intro all_cong1 ex_cong1 imp_cong refl) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3524 | by (smt (verit, best) "*" Diff_Diff_Int Diff_subset Diff_subset_conv inf.orderE inf_commute openin_subset sup_commute) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3525 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3526 | finally show ?thesis . | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3527 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3528 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3529 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3530 | lemma normal_t1_imp_Hausdorff_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3531 | assumes "normal_space X" "t1_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3532 | shows "Hausdorff_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3533 | unfolding Hausdorff_space_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3534 | proof clarify | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3535 | fix x y | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3536 | assume xy: "x \<in> topspace X" "y \<in> topspace X" "x \<noteq> y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3537 |   then have "disjnt {x} {y}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3538 | by (auto simp: disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3539 | then show "\<exists>U V. openin X U \<and> openin X V \<and> x \<in> U \<and> y \<in> V \<and> disjnt U V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3540 | using assms xy closedin_t1_singleton normal_space_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3541 | by (metis singletonI subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3542 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3543 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3544 | lemma normal_t1_eq_Hausdorff_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3545 | "normal_space X \<Longrightarrow> t1_space X \<longleftrightarrow> Hausdorff_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3546 | using normal_t1_imp_Hausdorff_space t1_or_Hausdorff_space by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3547 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3548 | lemma normal_t1_imp_regular_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3549 | "\<lbrakk>normal_space X; t1_space X\<rbrakk> \<Longrightarrow> regular_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3550 | by (metis compactin_imp_closedin normal_space_def normal_t1_eq_Hausdorff_space regular_space_compact_closed_sets) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3551 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3552 | lemma compact_Hausdorff_or_regular_imp_normal_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3553 | "\<lbrakk>compact_space X; Hausdorff_space X \<or> regular_space X\<rbrakk> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3554 | \<Longrightarrow> normal_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3555 | by (metis Hausdorff_space_compact_sets closedin_compact_space normal_space_def regular_space_compact_closed_sets) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3556 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3557 | lemma normal_space_discrete_topology: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3558 | "normal_space(discrete_topology U)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3559 | by (metis discrete_topology_closure_of inf_le2 normal_space_alt) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3560 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3561 | lemma normal_space_fsigmas: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3562 | "normal_space X \<longleftrightarrow> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3563 | (\<forall>S T. fsigma_in X S \<and> fsigma_in X T \<and> separatedin X S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3564 | \<longrightarrow> (\<exists>U B. openin X U \<and> openin X B \<and> S \<subseteq> U \<and> T \<subseteq> B \<and> disjnt U B))" (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3565 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3566 | assume L: ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3567 | show ?rhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3568 | proof clarify | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3569 | fix S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3570 | assume "fsigma_in X S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3571 | then obtain C where C: "\<And>n. closedin X (C n)" "\<And>n. C n \<subseteq> C (Suc n)" "\<Union> (range C) = S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3572 | by (meson fsigma_in_ascending) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3573 | assume "fsigma_in X T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3574 | then obtain D where D: "\<And>n. closedin X (D n)" "\<And>n. D n \<subseteq> D (Suc n)" "\<Union> (range D) = T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3575 | by (meson fsigma_in_ascending) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3576 | assume "separatedin X S T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3577 | have "\<And>n. disjnt (D n) (X closure_of S)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3578 | by (metis D(3) \<open>separatedin X S T\<close> disjnt_Union1 disjnt_def rangeI separatedin_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3579 | then have "\<And>n. \<exists>V V'. openin X V \<and> openin X V' \<and> D n \<subseteq> V \<and> X closure_of S \<subseteq> V' \<and> disjnt V V'" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3580 | by (metis D(1) L closedin_closure_of normal_space_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3581 | then obtain V V' where V: "\<And>n. openin X (V n)" and "\<And>n. openin X (V' n)" "\<And>n. disjnt (V n) (V' n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3582 | and DV: "\<And>n. D n \<subseteq> V n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3583 | and subV': "\<And>n. X closure_of S \<subseteq> V' n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3584 | by metis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3585 |     then have VV: "V' n \<inter> X closure_of V n = {}" for n
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3586 | using openin_Int_closure_of_eq_empty [of X "V' n" "V n"] by (simp add: Int_commute disjnt_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3587 | have "\<And>n. disjnt (C n) (X closure_of T)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3588 | by (metis C(3) \<open>separatedin X S T\<close> disjnt_Union1 disjnt_def rangeI separatedin_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3589 | then have "\<And>n. \<exists>U U'. openin X U \<and> openin X U' \<and> C n \<subseteq> U \<and> X closure_of T \<subseteq> U' \<and> disjnt U U'" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3590 | by (metis C(1) L closedin_closure_of normal_space_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3591 | then obtain U U' where U: "\<And>n. openin X (U n)" and "\<And>n. openin X (U' n)" "\<And>n. disjnt (U n) (U' n)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3592 | and CU: "\<And>n. C n \<subseteq> U n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3593 | and subU': "\<And>n. X closure_of T \<subseteq> U' n" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3594 | by metis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3595 |     then have UU: "U' n \<inter> X closure_of U n = {}" for n
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3596 | using openin_Int_closure_of_eq_empty [of X "U' n" "U n"] by (simp add: Int_commute disjnt_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3597 | show "\<exists>U B. openin X U \<and> openin X B \<and> S \<subseteq> U \<and> T \<subseteq> B \<and> disjnt U B" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3598 | proof (intro conjI exI) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3599 | have "\<And>S n. closedin X (\<Union>m\<le>n. X closure_of V m)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3600 | by (force intro: closedin_Union) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3601 | then show "openin X (\<Union>n. U n - (\<Union>m\<le>n. X closure_of V m))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3602 | using U by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3603 | have "\<And>S n. closedin X (\<Union>m\<le>n. X closure_of U m)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3604 | by (force intro: closedin_Union) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3605 | then show "openin X (\<Union>n. V n - (\<Union>m\<le>n. X closure_of U m))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3606 | using V by blast | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3607 | have "S \<subseteq> topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3608 | by (simp add: \<open>fsigma_in X S\<close> fsigma_in_subset) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3609 | then show "S \<subseteq> (\<Union>n. U n - (\<Union>m\<le>n. X closure_of V m))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3610 | apply (clarsimp simp: Ball_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3611 | by (metis VV C(3) CU IntI UN_E closure_of_subset empty_iff subV' subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3612 | have "T \<subseteq> topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3613 | by (simp add: \<open>fsigma_in X T\<close> fsigma_in_subset) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3614 | then show "T \<subseteq> (\<Union>n. V n - (\<Union>m\<le>n. X closure_of U m))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3615 | apply (clarsimp simp: Ball_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3616 | by (metis UU D(3) DV IntI UN_E closure_of_subset empty_iff subU' subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3617 | have "\<And>x m n. \<lbrakk>x \<in> U n; x \<in> V m; \<forall>k\<le>m. x \<notin> X closure_of U k\<rbrakk> \<Longrightarrow> \<exists>k\<le>n. x \<in> X closure_of V k" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3618 | by (meson U V closure_of_subset nat_le_linear openin_subset subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3619 | then show "disjnt (\<Union>n. U n - (\<Union>m\<le>n. X closure_of V m)) (\<Union>n. V n - (\<Union>m\<le>n. X closure_of U m))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3620 | by (force simp: disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3621 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3622 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3623 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3624 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3625 | by (simp add: closed_imp_fsigma_in normal_space_def separatedin_closed_sets) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3626 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3627 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3628 | lemma normal_space_fsigma_subtopology: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3629 | assumes "normal_space X" "fsigma_in X S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3630 | shows "normal_space (subtopology X S)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3631 | unfolding normal_space_fsigmas | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3632 | proof clarify | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3633 | fix T U | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3634 | assume "fsigma_in (subtopology X S) T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3635 | and "fsigma_in (subtopology X S) U" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3636 | and TU: "separatedin (subtopology X S) T U" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3637 | then obtain A B where "openin X A \<and> openin X B \<and> T \<subseteq> A \<and> U \<subseteq> B \<and> disjnt A B" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3638 | by (metis assms fsigma_in_fsigma_subtopology normal_space_fsigmas separatedin_subtopology) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3639 | then | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3640 | show "\<exists>A B. openin (subtopology X S) A \<and> openin (subtopology X S) B \<and> T \<subseteq> A \<and> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3641 | U \<subseteq> B \<and> disjnt A B" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3642 | using TU | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3643 | by (force simp add: separatedin_subtopology openin_subtopology_alt disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3644 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3645 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3646 | lemma normal_space_closed_subtopology: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3647 | assumes "normal_space X" "closedin X S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3648 | shows "normal_space (subtopology X S)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3649 | by (simp add: assms closed_imp_fsigma_in normal_space_fsigma_subtopology) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3650 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3651 | lemma normal_space_continuous_closed_map_image: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3652 | assumes "normal_space X" and contf: "continuous_map X Y f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3653 | and clof: "closed_map X Y f" and fim: "f ` topspace X = topspace Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3654 | shows "normal_space Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3655 | unfolding normal_space_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3656 | proof clarify | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3657 | fix S T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3658 | assume "closedin Y S" and "closedin Y T" and "disjnt S T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3659 |   have "closedin X {x \<in> topspace X. f x \<in> S}" "closedin X {x \<in> topspace X. f x \<in> T}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3660 | using \<open>closedin Y S\<close> \<open>closedin Y T\<close> closedin_continuous_map_preimage contf by auto | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3661 | moreover | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3662 |   have "disjnt {x \<in> topspace X. f x \<in> S} {x \<in> topspace X. f x \<in> T}"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3663 | using \<open>disjnt S T\<close> by (auto simp: disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3664 | ultimately | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3665 | obtain U V where "closedin X U" "closedin X V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3666 |     and subXU: "{x \<in> topspace X. f x \<in> S} \<subseteq> topspace X - U" 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3667 |     and subXV: "{x \<in> topspace X. f x \<in> T} \<subseteq> topspace X - V" 
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3668 | and dis: "disjnt (topspace X - U) (topspace X -V)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3669 | using \<open>normal_space X\<close> by (force simp add: normal_space_def ex_openin) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3670 | have "closedin Y (f ` U)" "closedin Y (f ` V)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3671 | using \<open>closedin X U\<close> \<open>closedin X V\<close> clof closed_map_def by blast+ | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3672 | moreover have "S \<subseteq> topspace Y - f ` U" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3673 | using \<open>closedin Y S\<close> \<open>closedin X U\<close> subXU by (force dest: closedin_subset) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3674 | moreover have "T \<subseteq> topspace Y - f ` V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3675 | using \<open>closedin Y T\<close> \<open>closedin X V\<close> subXV by (force dest: closedin_subset) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3676 | moreover have "disjnt (topspace Y - f ` U) (topspace Y - f ` V)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3677 | using fim dis by (force simp add: disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3678 | ultimately show "\<exists>U V. openin Y U \<and> openin Y V \<and> S \<subseteq> U \<and> T \<subseteq> V \<and> disjnt U V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3679 | by (force simp add: ex_openin) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3680 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3681 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3682 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3683 | subsection \<open>Hereditary topological properties\<close> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3684 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3685 | definition hereditarily | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3686 | where "hereditarily P X \<equiv> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3687 | \<forall>S. S \<subseteq> topspace X \<longrightarrow> P(subtopology X S)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3688 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3689 | lemma hereditarily: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3690 | "hereditarily P X \<longleftrightarrow> (\<forall>S. P(subtopology X S))" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3691 | by (metis Int_lower1 hereditarily_def subtopology_restrict) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3692 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3693 | lemma hereditarily_mono: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3694 | "\<lbrakk>hereditarily P X; \<And>x. P x \<Longrightarrow> Q x\<rbrakk> \<Longrightarrow> hereditarily Q X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3695 | by (simp add: hereditarily) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3696 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3697 | lemma hereditarily_inc: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3698 | "hereditarily P X \<Longrightarrow> P X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3699 | by (metis hereditarily subtopology_topspace) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3700 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3701 | lemma hereditarily_subtopology: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3702 | "hereditarily P X \<Longrightarrow> hereditarily P (subtopology X S)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3703 | by (simp add: hereditarily subtopology_subtopology) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3704 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3705 | lemma hereditarily_normal_space_continuous_closed_map_image: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3706 | assumes X: "hereditarily normal_space X" and contf: "continuous_map X Y f" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3707 | and clof: "closed_map X Y f" and fim: "f ` (topspace X) = topspace Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3708 | shows "hereditarily normal_space Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3709 | unfolding hereditarily_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3710 | proof (intro strip) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3711 | fix T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3712 | assume "T \<subseteq> topspace Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3713 |   then have nx: "normal_space (subtopology X {x \<in> topspace X. f x \<in> T})"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3714 | by (meson X hereditarily) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3715 |   moreover have "continuous_map (subtopology X {x \<in> topspace X. f x \<in> T}) (subtopology Y T) f"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3716 | by (simp add: contf continuous_map_from_subtopology continuous_map_in_subtopology image_subset_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3717 |   moreover have "closed_map (subtopology X {x \<in> topspace X. f x \<in> T}) (subtopology Y T) f"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3718 | by (simp add: clof closed_map_restriction) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3719 | ultimately show "normal_space (subtopology Y T)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3720 | using fim normal_space_continuous_closed_map_image by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3721 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3722 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3723 | lemma homeomorphic_hereditarily_normal_space: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3724 | "X homeomorphic_space Y | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3725 | \<Longrightarrow> (hereditarily normal_space X \<longleftrightarrow> hereditarily normal_space Y)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3726 | by (meson hereditarily_normal_space_continuous_closed_map_image homeomorphic_eq_everything_map | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3727 | homeomorphic_space homeomorphic_space_sym) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3728 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3729 | lemma hereditarily_normal_space_retraction_map_image: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3730 | "\<lbrakk>retraction_map X Y r; hereditarily normal_space X\<rbrakk> \<Longrightarrow> hereditarily normal_space Y" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3731 | by (smt (verit) hereditarily_subtopology hereditary_imp_retractive_property homeomorphic_hereditarily_normal_space) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3732 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3733 | subsection\<open>Limits in a topological space\<close> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3734 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3735 | lemma limitin_const_iff: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3736 | assumes "t1_space X" "\<not> trivial_limit F" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3737 | shows "limitin X (\<lambda>k. a) l F \<longleftrightarrow> l \<in> topspace X \<and> a = l" (is "?lhs=?rhs") | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3738 | proof | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3739 | assume ?lhs then show ?rhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3740 | using assms unfolding limitin_def t1_space_def by (metis eventually_const openin_topspace) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3741 | next | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3742 | assume ?rhs then show ?lhs | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3743 | using assms by (auto simp: limitin_def t1_space_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3744 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3745 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3746 | lemma compactin_sequence_with_limit: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3747 | assumes lim: "limitin X \<sigma> l sequentially" and "S \<subseteq> range \<sigma>" and SX: "S \<subseteq> topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3748 | shows "compactin X (insert l S)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3749 | unfolding compactin_def | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3750 | proof (intro conjI strip) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3751 | show "insert l S \<subseteq> topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3752 | by (meson SX insert_subset lim limitin_topspace) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3753 | fix \<U> | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3754 | assume \<section>: "Ball \<U> (openin X) \<and> insert l S \<subseteq> \<Union> \<U>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3755 | have "\<exists>V. finite V \<and> V \<subseteq> \<U> \<and> (\<exists>t \<in> V. l \<in> t) \<and> S \<subseteq> \<Union> V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3756 | if *: "\<forall>x \<in> S. \<exists>T \<in> \<U>. x \<in> T" and "T \<in> \<U>" "l \<in> T" for T | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3757 | proof - | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3758 | obtain V where V: "\<And>x. x \<in> S \<Longrightarrow> V x \<in> \<U> \<and> x \<in> V x" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3759 | using * by metis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3760 | obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> \<sigma> n \<in> T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3761 | by (meson "\<section>" \<open>T \<in> \<U>\<close> \<open>l \<in> T\<close> lim limitin_sequentially) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3762 | show ?thesis | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3763 | proof (intro conjI exI) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3764 | have "x \<in> T" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3765 | if "x \<in> S" and "\<forall>A. (\<forall>x \<in> S. (\<forall>n\<le>N. x \<noteq> \<sigma> n) \<or> A \<noteq> V x) \<or> x \<notin> A" for x | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3766 | by (metis (no_types) N V that assms(2) imageE nle_le subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3767 |       then show "S \<subseteq> \<Union> (insert T (V ` (S \<inter> \<sigma> ` {0..N})))"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3768 | by force | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3769 | qed (use V that in auto) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3770 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3771 | then show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> insert l S \<subseteq> \<Union> \<F>" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3772 | by (smt (verit, best) Union_iff \<section> insert_subset subsetD) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3773 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3774 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3775 | lemma limitin_Hausdorff_unique: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3776 | assumes "limitin X f l1 F" "limitin X f l2 F" "\<not> trivial_limit F" "Hausdorff_space X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3777 | shows "l1 = l2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3778 | proof (rule ccontr) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3779 | assume "l1 \<noteq> l2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3780 | with assms obtain U V where "openin X U" "openin X V" "l1 \<in> U" "l2 \<in> V" "disjnt U V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3781 | by (metis Hausdorff_space_def limitin_topspace) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3782 | then have "eventually (\<lambda>x. f x \<in> U) F" "eventually (\<lambda>x. f x \<in> V) F" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3783 | using assms by (fastforce simp: limitin_def)+ | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3784 | then have "\<exists>x. f x \<in> U \<and> f x \<in> V" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3785 | using assms eventually_elim2 filter_eq_iff by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3786 | with assms \<open>disjnt U V\<close> show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3787 | by (meson disjnt_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3788 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3789 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3790 | lemma limitin_kc_unique: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3791 | assumes "kc_space X" and lim1: "limitin X f l1 sequentially" and lim2: "limitin X f l2 sequentially" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3792 | shows "l1 = l2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3793 | proof (rule ccontr) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3794 | assume "l1 \<noteq> l2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3795 |   define A where "A \<equiv> insert l1 (range f - {l2})"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3796 | have "l1 \<in> topspace X" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3797 | using lim1 limitin_def by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3798 |   moreover have "compactin X (insert l1 (topspace X \<inter> (range f - {l2})))"
 | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3799 | by (meson Diff_subset compactin_sequence_with_limit inf_le1 inf_le2 lim1 subset_trans) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3800 | ultimately have "compactin X (topspace X \<inter> A)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3801 | by (simp add: A_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3802 | then have OXA: "openin X (topspace X - A)" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3803 | by (metis Diff_Diff_Int Diff_subset \<open>kc_space X\<close> kc_space_def openin_closedin_eq) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3804 | have "l2 \<in> topspace X - A" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3805 | using \<open>l1 \<noteq> l2\<close> A_def lim2 limitin_topspace by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3806 | then have "\<forall>\<^sub>F x in sequentially. f x = l2" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3807 | using limitinD [OF lim2 OXA] by (auto simp: A_def eventually_conj_iff) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3808 | then show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3809 | using limitin_transform_eventually [OF _ lim1] | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3810 | limitin_const_iff [OF kc_imp_t1_space trivial_limit_sequentially] | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3811 | using \<open>l1 \<noteq> l2\<close> \<open>kc_space X\<close> by fastforce | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3812 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3813 | |
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3814 | lemma limitin_closedin: | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3815 | assumes lim: "limitin X f l F" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3816 | and "closedin X S" and ev: "eventually (\<lambda>x. f x \<in> S) F" "\<not> trivial_limit F" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3817 | shows "l \<in> S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3818 | proof (rule ccontr) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3819 | assume "l \<notin> S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3820 | have "\<forall>\<^sub>F x in F. f x \<in> topspace X - S" | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3821 | by (metis Diff_iff \<open>l \<notin> S\<close> \<open>closedin X S\<close> closedin_def lim limitin_def) | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3822 | with ev eventually_elim2 trivial_limit_def show False | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3823 | by force | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3824 | qed | 
| 
ffdad62bc235
Importation of additional lemmas from metric.ml
 paulson <lp15@cam.ac.uk> parents: 
77942diff
changeset | 3825 | |
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3826 | subsection\<open>Quasi-components\<close> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3827 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3828 | definition quasi_component_of :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3829 | where | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3830 | "quasi_component_of X x y \<equiv> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3831 | x \<in> topspace X \<and> y \<in> topspace X \<and> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3832 | (\<forall>T. closedin X T \<and> openin X T \<longrightarrow> (x \<in> T \<longleftrightarrow> y \<in> T))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3833 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3834 | abbreviation "quasi_component_of_set S x \<equiv> Collect (quasi_component_of S x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3835 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3836 | definition quasi_components_of :: "'a topology \<Rightarrow> ('a set) set"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3837 | where | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3838 | "quasi_components_of X = quasi_component_of_set X ` topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3839 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3840 | lemma quasi_component_in_topspace: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3841 | "quasi_component_of X x y \<Longrightarrow> x \<in> topspace X \<and> y \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3842 | by (simp add: quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3843 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3844 | lemma quasi_component_of_refl [simp]: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3845 | "quasi_component_of X x x \<longleftrightarrow> x \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3846 | by (simp add: quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3847 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3848 | lemma quasi_component_of_sym: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3849 | "quasi_component_of X x y \<longleftrightarrow> quasi_component_of X y x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3850 | by (meson quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3851 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3852 | lemma quasi_component_of_trans: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3853 | "\<lbrakk>quasi_component_of X x y; quasi_component_of X y z\<rbrakk> \<Longrightarrow> quasi_component_of X x z" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3854 | by (simp add: quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3855 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3856 | lemma quasi_component_of_subset_topspace: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3857 | "quasi_component_of_set X x \<subseteq> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3858 | using quasi_component_of_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3859 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3860 | lemma quasi_component_of_eq_empty: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3861 |    "quasi_component_of_set X x = {} \<longleftrightarrow> (x \<notin> topspace X)"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3862 | using quasi_component_of_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3863 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3864 | lemma quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3865 | "quasi_component_of X x y \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3866 | x \<in> topspace X \<and> y \<in> topspace X \<and> (\<forall>T. x \<in> T \<and> closedin X T \<and> openin X T \<longrightarrow> y \<in> T)" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 3867 | unfolding quasi_component_of_def by (metis Diff_iff closedin_def openin_closedin_eq) | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3868 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3869 | lemma quasi_component_of_alt: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3870 | "quasi_component_of X x y \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3871 | x \<in> topspace X \<and> y \<in> topspace X \<and> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3872 | \<not> (\<exists>U V. openin X U \<and> openin X V \<and> U \<union> V = topspace X \<and> disjnt U V \<and> x \<in> U \<and> y \<in> V)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3873 | (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3874 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3875 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3876 | unfolding quasi_component_of_def | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3877 | by (metis disjnt_iff separatedin_full separatedin_open_sets) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3878 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3879 | unfolding quasi_component_of_def | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3880 | by (metis Diff_disjoint Diff_iff Un_Diff_cancel closedin_def disjnt_def inf_commute sup.orderE sup_commute) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3881 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3882 | |
| 78250 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 3883 | lemma quasi_components_lepoll_topspace: "quasi_components_of X \<lesssim> topspace X" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 3884 | by (simp add: image_lepoll quasi_components_of_def) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78248diff
changeset | 3885 | |
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3886 | lemma quasi_component_of_separated: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3887 | "quasi_component_of X x y \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3888 | x \<in> topspace X \<and> y \<in> topspace X \<and> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3889 | \<not> (\<exists>U V. separatedin X U V \<and> U \<union> V = topspace X \<and> x \<in> U \<and> y \<in> V)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3890 | by (meson quasi_component_of_alt separatedin_full separatedin_open_sets) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3891 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3892 | lemma quasi_component_of_subtopology: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3893 | "quasi_component_of (subtopology X s) x y \<Longrightarrow> quasi_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3894 | unfolding quasi_component_of_def | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3895 | by (simp add: closedin_subtopology) (metis Int_iff inf_commute openin_subtopology_Int2) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3896 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3897 | lemma quasi_component_of_mono: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3898 | "quasi_component_of (subtopology X S) x y \<and> S \<subseteq> T | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3899 | \<Longrightarrow> quasi_component_of (subtopology X T) x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3900 | by (metis inf.absorb_iff2 quasi_component_of_subtopology subtopology_subtopology) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3901 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3902 | lemma quasi_component_of_equiv: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3903 | "quasi_component_of X x y \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3904 | x \<in> topspace X \<and> y \<in> topspace X \<and> quasi_component_of X x = quasi_component_of X y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3905 | using quasi_component_of_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3906 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3907 | lemma quasi_component_of_disjoint [simp]: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3908 | "disjnt (quasi_component_of_set X x) (quasi_component_of_set X y) \<longleftrightarrow> \<not> (quasi_component_of X x y)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3909 | by (metis disjnt_iff quasi_component_of_equiv mem_Collect_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3910 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3911 | lemma quasi_component_of_eq: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3912 | "quasi_component_of X x = quasi_component_of X y \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3913 | (x \<notin> topspace X \<and> y \<notin> topspace X) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3914 | \<or> x \<in> topspace X \<and> y \<in> topspace X \<and> quasi_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3915 | by (metis Collect_empty_eq_bot quasi_component_of_eq_empty quasi_component_of_equiv) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3916 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3917 | lemma topspace_imp_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3918 | assumes "x \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3919 | obtains C where "C \<in> quasi_components_of X" "x \<in> C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3920 | by (metis assms imageI mem_Collect_eq quasi_component_of_refl quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3921 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3922 | lemma Union_quasi_components_of: "\<Union> (quasi_components_of X) = topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3923 | by (auto simp: quasi_components_of_def quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3924 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3925 | lemma pairwise_disjoint_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3926 | "pairwise disjnt (quasi_components_of X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3927 | by (auto simp: quasi_components_of_def quasi_component_of_def disjoint_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3928 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3929 | lemma complement_quasi_components_of_Union: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3930 | assumes "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3931 |   shows "topspace X - C = \<Union> (quasi_components_of X - {C})"  (is "?lhs = ?rhs")
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3932 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3933 | show "?lhs \<subseteq> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3934 | using Union_quasi_components_of by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3935 | show "?rhs \<subseteq> ?lhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3936 | using assms | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3937 | using quasi_component_of_equiv by (fastforce simp add: quasi_components_of_def image_iff subset_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3938 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3939 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3940 | lemma nonempty_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3941 |    "C \<in> quasi_components_of X \<Longrightarrow> C \<noteq> {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3942 | by (metis imageE quasi_component_of_eq_empty quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3943 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3944 | lemma quasi_components_of_subset: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3945 | "C \<in> quasi_components_of X \<Longrightarrow> C \<subseteq> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3946 | using Union_quasi_components_of by force | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3947 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3948 | lemma quasi_component_in_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3949 | "quasi_component_of_set X a \<in> quasi_components_of X \<longleftrightarrow> a \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3950 | by (metis (no_types, lifting) image_iff quasi_component_of_eq_empty quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3951 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3952 | lemma quasi_components_of_eq_empty [simp]: | 
| 78336 | 3953 |    "quasi_components_of X = {} \<longleftrightarrow> X = trivial_topology"
 | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3954 | by (simp add: quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3955 | |
| 78336 | 3956 | lemma quasi_components_of_empty_space [simp]: | 
| 3957 |    "quasi_components_of trivial_topology = {}"
 | |
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3958 | by simp | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3959 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3960 | lemma quasi_component_of_set: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3961 | "quasi_component_of_set X x = | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3962 | (if x \<in> topspace X | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3963 |         then \<Inter> {t. closedin X t \<and> openin X t \<and> x \<in> t}
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3964 |         else {})"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3965 | by (auto simp: quasi_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3966 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3967 | lemma closedin_quasi_component_of: "closedin X (quasi_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3968 | by (auto simp: quasi_component_of_set) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3969 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3970 | lemma closedin_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3971 | "C \<in> quasi_components_of X \<Longrightarrow> closedin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3972 | by (auto simp: quasi_components_of_def closedin_quasi_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3973 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3974 | lemma openin_finite_quasi_components: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3975 | "\<lbrakk>finite(quasi_components_of X); C \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> openin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3976 | apply (simp add:openin_closedin_eq quasi_components_of_subset complement_quasi_components_of_Union) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3977 | by (meson DiffD1 closedin_Union closedin_quasi_components_of finite_Diff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3978 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3979 | lemma quasi_component_of_eq_overlap: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3980 | "quasi_component_of X x = quasi_component_of X y \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3981 | (x \<notin> topspace X \<and> y \<notin> topspace X) \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3982 |       \<not> (quasi_component_of_set X x \<inter> quasi_component_of_set X y = {})"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3983 | using quasi_component_of_equiv by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3984 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3985 | lemma quasi_component_of_nonoverlap: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3986 |    "quasi_component_of_set X x \<inter> quasi_component_of_set X y = {} \<longleftrightarrow>
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3987 | (x \<notin> topspace X) \<or> (y \<notin> topspace X) \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3988 | \<not> (quasi_component_of X x = quasi_component_of X y)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3989 | by (metis inf.idem quasi_component_of_eq_empty quasi_component_of_eq_overlap) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3990 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3991 | lemma quasi_component_of_overlap: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3992 |    "\<not> (quasi_component_of_set X x \<inter> quasi_component_of_set X y = {}) \<longleftrightarrow>
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3993 | x \<in> topspace X \<and> y \<in> topspace X \<and> quasi_component_of X x = quasi_component_of X y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3994 | by (meson quasi_component_of_nonoverlap) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3995 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3996 | lemma quasi_components_of_disjoint: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3997 | "\<lbrakk>C \<in> quasi_components_of X; D \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> disjnt C D \<longleftrightarrow> C \<noteq> D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3998 | by (metis disjnt_self_iff_empty nonempty_quasi_components_of pairwiseD pairwise_disjoint_quasi_components_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 3999 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4000 | lemma quasi_components_of_overlap: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4001 |    "\<lbrakk>C \<in> quasi_components_of X; D \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> \<not> (C \<inter> D = {}) \<longleftrightarrow> C = D"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4002 | by (metis disjnt_def quasi_components_of_disjoint) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4003 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4004 | lemma pairwise_separated_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4005 | "pairwise (separatedin X) (quasi_components_of X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4006 | by (metis closedin_quasi_components_of pairwise_def pairwise_disjoint_quasi_components_of separatedin_closed_sets) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4007 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4008 | lemma finite_quasi_components_of_finite: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4009 | "finite(topspace X) \<Longrightarrow> finite(quasi_components_of X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4010 | by (simp add: Union_quasi_components_of finite_UnionD) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4011 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4012 | lemma connected_imp_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4013 | assumes "connected_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4014 | shows "quasi_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4015 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4016 | have "x \<in> topspace X" "y \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4017 | by (meson assms connected_component_of_equiv)+ | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4018 | with assms show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4019 | apply (clarsimp simp add: quasi_component_of connected_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4020 | by (meson connectedin_clopen_cases disjnt_iff subsetD) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4021 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4022 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4023 | lemma connected_component_subset_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4024 | "connected_component_of_set X x \<subseteq> quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4025 | using connected_imp_quasi_component_of by force | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4026 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4027 | lemma quasi_component_as_connected_component_Union: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4028 | "quasi_component_of_set X x = | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4029 | \<Union> (connected_component_of_set X ` quasi_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4030 | (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4031 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4032 | show "?lhs \<subseteq> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4033 | using connected_component_of_refl quasi_component_of by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4034 | show "?rhs \<subseteq> ?lhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4035 | apply (rule SUP_least) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4036 | by (simp add: connected_component_subset_quasi_component_of quasi_component_of_equiv) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4037 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4038 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4039 | lemma quasi_components_as_connected_components_Union: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4040 | assumes "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4041 | obtains \<T> where "\<T> \<subseteq> connected_components_of X" "\<Union>\<T> = C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4042 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4043 | obtain x where "x \<in> topspace X" and Ceq: "C = quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4044 | by (metis assms imageE quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4045 | define \<T> where "\<T> \<equiv> connected_component_of_set X ` quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4046 | show thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4047 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4048 | show "\<T> \<subseteq> connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4049 | by (simp add: \<T>_def connected_components_of_def image_mono quasi_component_of_subset_topspace) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4050 | show "\<Union>\<T> = C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4051 | by (metis \<T>_def Ceq quasi_component_as_connected_component_Union) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4052 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4053 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4054 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4055 | lemma path_imp_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4056 | "path_component_of X x y \<Longrightarrow> quasi_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4057 | by (simp add: connected_imp_quasi_component_of path_imp_connected_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4058 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4059 | lemma path_component_subset_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4060 | "path_component_of_set X x \<subseteq> quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4061 | by (simp add: Collect_mono path_imp_quasi_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4062 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4063 | lemma connected_space_iff_quasi_component: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4064 | "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. \<forall>y \<in> topspace X. quasi_component_of X x y)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4065 | unfolding connected_space_clopen_in closedin_def quasi_component_of | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4066 | by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4067 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4068 | lemma connected_space_imp_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4069 | " \<lbrakk>connected_space X; a \<in> topspace X; b \<in> topspace X\<rbrakk> \<Longrightarrow> quasi_component_of X a b" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4070 | by (simp add: connected_space_iff_quasi_component) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4071 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4072 | lemma connected_space_quasi_component_set: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4073 | "connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. quasi_component_of_set X x = topspace X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4074 | by (metis Ball_Collect connected_space_iff_quasi_component quasi_component_of_subset_topspace subset_antisym) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4075 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4076 | lemma connected_space_iff_quasi_components_eq: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4077 | "connected_space X \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4078 | (\<forall>C \<in> quasi_components_of X. \<forall>D \<in> quasi_components_of X. C = D)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4079 | apply (simp add: quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4080 | by (metis connected_space_iff_quasi_component mem_Collect_eq quasi_component_of_equiv) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4081 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4082 | lemma quasi_components_of_subset_sing: | 
| 78336 | 4083 |    "quasi_components_of X \<subseteq> {S} \<longleftrightarrow> connected_space X \<and> (X = trivial_topology \<or> topspace X = S)"
 | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4084 | proof (cases "quasi_components_of X = {}")
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4085 | case True | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4086 | then show ?thesis | 
| 78336 | 4087 | by (simp add: subset_singleton_iff) | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4088 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4089 | case False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4090 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4091 | apply (simp add: connected_space_iff_quasi_components_eq subset_iff Ball_def) | 
| 78336 | 4092 | by (metis False Union_quasi_components_of ccpo_Sup_singleton insert_iff is_singletonE is_singletonI') | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4093 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4094 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4095 | lemma connected_space_iff_quasi_components_subset_sing: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4096 |    "connected_space X \<longleftrightarrow> (\<exists>a. quasi_components_of X \<subseteq> {a})"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4097 | by (simp add: quasi_components_of_subset_sing) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4098 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4099 | lemma quasi_components_of_eq_singleton: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4100 |    "quasi_components_of X = {S} \<longleftrightarrow>
 | 
| 78336 | 4101 | connected_space X \<and> \<not> (X = trivial_topology) \<and> S = topspace X" | 
| 4102 | by (metis empty_not_insert quasi_components_of_eq_empty quasi_components_of_subset_sing subset_singleton_iff) | |
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4103 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4104 | lemma quasi_components_of_connected_space: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4105 | "connected_space X | 
| 78336 | 4106 |         \<Longrightarrow> quasi_components_of X = (if X = trivial_topology then {} else {topspace X})"
 | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4107 | by (simp add: quasi_components_of_eq_singleton) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4108 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4109 | lemma separated_between_singletons: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4110 |    "separated_between X {x} {y} \<longleftrightarrow>
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4111 | x \<in> topspace X \<and> y \<in> topspace X \<and> \<not> (quasi_component_of X x y)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4112 | proof (cases "x \<in> topspace X \<and> y \<in> topspace X") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4113 | case True | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4114 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4115 | by (auto simp add: separated_between_def quasi_component_of_alt) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4116 | qed (use separated_between_imp_subset in blast) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4117 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4118 | lemma quasi_component_nonseparated: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4119 |    "quasi_component_of X x y \<longleftrightarrow> x \<in> topspace X \<and> y \<in> topspace X \<and> \<not> (separated_between X {x} {y})"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4120 | by (metis quasi_component_of_equiv separated_between_singletons) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4121 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4122 | lemma separated_between_quasi_component_pointwise_left: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4123 | assumes "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4124 |   shows "separated_between X C S \<longleftrightarrow> (\<exists>x \<in> C. separated_between X {x} S)"  (is "?lhs = ?rhs")
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4125 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4126 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4127 | using assms quasi_components_of_disjoint separated_between_mono by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4128 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4129 | assume ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4130 |   then obtain y where "separated_between X {y} S" and "y \<in> C"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4131 | by metis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4132 | with assms show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4133 | by (force simp add: separated_between quasi_components_of_def quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4134 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4135 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4136 | lemma separated_between_quasi_component_pointwise_right: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4137 |    "C \<in> quasi_components_of X \<Longrightarrow> separated_between X S C \<longleftrightarrow> (\<exists>x \<in> C. separated_between X S {x})"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4138 | by (simp add: separated_between_quasi_component_pointwise_left separated_between_sym) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4139 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4140 | lemma separated_between_quasi_component_point: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4141 | assumes "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4142 |   shows "separated_between X C {x} \<longleftrightarrow> x \<in> topspace X - C" (is "?lhs = ?rhs")
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4143 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4144 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4145 | by (meson DiffI disjnt_insert2 insert_subset separated_between_imp_disjoint separated_between_imp_subset) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4146 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4147 | assume ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4148 | with assms show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4149 | unfolding quasi_components_of_def image_iff Diff_iff separated_between_quasi_component_pointwise_left [OF assms] | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4150 | by (metis mem_Collect_eq quasi_component_of_refl separated_between_singletons) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4151 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4152 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4153 | lemma separated_between_point_quasi_component: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4154 |    "C \<in> quasi_components_of X \<Longrightarrow> separated_between X {x} C \<longleftrightarrow> x \<in> topspace X - C"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4155 | by (simp add: separated_between_quasi_component_point separated_between_sym) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4156 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4157 | lemma separated_between_quasi_component_compact: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4158 | "\<lbrakk>C \<in> quasi_components_of X; compactin X K\<rbrakk> \<Longrightarrow> (separated_between X C K \<longleftrightarrow> disjnt C K)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4159 | unfolding disjnt_iff | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4160 | using compactin_subset_topspace quasi_components_of_subset separated_between_pointwise_right separated_between_quasi_component_point by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4161 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4162 | lemma separated_between_compact_quasi_component: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4163 | "\<lbrakk>compactin X K; C \<in> quasi_components_of X\<rbrakk> \<Longrightarrow> separated_between X K C \<longleftrightarrow> disjnt K C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4164 | using disjnt_sym separated_between_quasi_component_compact separated_between_sym by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4165 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4166 | lemma separated_between_quasi_components: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4167 | assumes C: "C \<in> quasi_components_of X" and D: "D \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4168 | shows "separated_between X C D \<longleftrightarrow> disjnt C D" (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4169 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4170 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4171 | by (simp add: separated_between_imp_disjoint) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4172 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4173 | assume ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4174 | obtain x y where x: "C = quasi_component_of_set X x" and "x \<in> C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4175 | and y: "D = quasi_component_of_set X y" and "y \<in> D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4176 | using assms by (auto simp: quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4177 |   then have "separated_between X {x} {y}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4178 | using \<open>disjnt C D\<close> separated_between_singletons by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4179 | with \<open>x \<in> C\<close> \<open>y \<in> D\<close> show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4180 | by (auto simp: assms separated_between_quasi_component_pointwise_left separated_between_quasi_component_pointwise_right) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4181 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4182 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4183 | lemma quasi_eq_connected_component_of_eq: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4184 | "quasi_component_of X x = connected_component_of X x \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4185 | connectedin X (quasi_component_of_set X x)" (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4186 | proof (cases "x \<in> topspace X") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4187 | case True | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4188 | show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4189 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4190 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4191 | by (simp add: connectedin_connected_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4192 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4193 | assume ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4194 | then have "\<And>y. quasi_component_of X x y = connected_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4195 | by (metis connected_component_of_def connected_imp_quasi_component_of mem_Collect_eq quasi_component_of_equiv) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4196 | then show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4197 | by force | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4198 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4199 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4200 | case False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4201 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4202 | by (metis Collect_empty_eq_bot connected_component_of_eq_empty connectedin_empty quasi_component_of_eq_empty) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4203 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4204 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4205 | lemma connected_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4206 | assumes "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4207 | shows "C \<in> connected_components_of X \<longleftrightarrow> connectedin X C" (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4208 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4209 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4210 | using assms | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4211 | by (simp add: connectedin_connected_components_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4212 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4213 | assume ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4214 | with assms show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4215 | unfolding quasi_components_of_def connected_components_of_def image_iff | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4216 | by (metis quasi_eq_connected_component_of_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4217 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4218 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4219 | lemma quasi_component_of_clopen_cases: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4220 | "\<lbrakk>C \<in> quasi_components_of X; closedin X T; openin X T\<rbrakk> \<Longrightarrow> C \<subseteq> T \<or> disjnt C T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4221 | by (smt (verit) disjnt_iff image_iff mem_Collect_eq quasi_component_of_def quasi_components_of_def subset_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4222 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4223 | lemma quasi_components_of_set: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4224 | assumes "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4225 |   shows "\<Inter> {T. closedin X T \<and> openin X T \<and> C \<subseteq> T} = C"  (is "?lhs = ?rhs")
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4226 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4227 |   have "x \<in> C" if "x \<in> \<Inter> {T. closedin X T \<and> openin X T \<and> C \<subseteq> T}" for x
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4228 | proof (rule ccontr) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4229 | assume "x \<notin> C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4230 | have "x \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4231 | using assms quasi_components_of_subset that by force | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4232 |     then have "separated_between X C {x}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4233 | by (simp add: \<open>x \<notin> C\<close> assms separated_between_quasi_component_point) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4234 | with that show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4235 | by (auto simp: separated_between) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4236 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4237 | then show "?lhs \<subseteq> ?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4238 | by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4239 | qed blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4240 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4241 | lemma open_quasi_eq_connected_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4242 | assumes "openin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4243 | shows "C \<in> quasi_components_of X \<longleftrightarrow> C \<in> connected_components_of X" (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4244 | proof (cases "closedin X C") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4245 | case True | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4246 | show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4247 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4248 | assume L: ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4249 |     have "T = {} \<or> T = topspace X \<inter> C"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4250 | if "openin (subtopology X C) T" "closedin (subtopology X C) T" for T | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4251 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4252 | have "C \<subseteq> T \<or> disjnt C T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4253 | by (meson L True assms closedin_trans_full openin_trans_full quasi_component_of_clopen_cases that) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4254 | with that show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4255 | by (metis Int_absorb2 True closedin_imp_subset closure_of_subset_eq disjnt_def inf_absorb2) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4256 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4257 | with L assms show "?rhs" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4258 | by (simp add: connected_quasi_component_of connected_space_clopen_in connectedin_def openin_subset) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4259 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4260 | assume ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4261 | then obtain x where "x \<in> topspace X" and x: "C = connected_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4262 | by (metis connected_components_of_def imageE) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4263 | have "C = quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4264 | using True assms connected_component_of_refl connected_imp_quasi_component_of quasi_component_of_def x by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4265 | then show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4266 | using \<open>x \<in> topspace X\<close> quasi_components_of_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4267 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4268 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4269 | case False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4270 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4271 | using closedin_connected_components_of closedin_quasi_components_of by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4272 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4273 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4274 | lemma quasi_component_of_continuous_image: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4275 | assumes f: "continuous_map X Y f" and qc: "quasi_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4276 | shows "quasi_component_of Y (f x) (f y)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4277 | unfolding quasi_component_of_def | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4278 | proof (intro strip conjI) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4279 | show "f x \<in> topspace Y" "f y \<in> topspace Y" | 
| 78320 
eb9a9690b8f5
cosmetic improvements, new lemmas, especially more uses of function space
 paulson <lp15@cam.ac.uk> parents: 
78277diff
changeset | 4280 | using assms by (simp_all add: continuous_map_def quasi_component_of_def Pi_iff) | 
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4281 | fix T | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4282 | assume "closedin Y T \<and> openin Y T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4283 | with assms show "(f x \<in> T) = (f y \<in> T)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4284 | by (smt (verit) continuous_map_closedin continuous_map_def mem_Collect_eq quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4285 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4286 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4287 | lemma quasi_component_of_discrete_topology: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4288 |    "quasi_component_of_set (discrete_topology U) x = (if x \<in> U then {x} else {})"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4289 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4290 |   have "quasi_component_of_set (discrete_topology U) y = {y}" if "y \<in> U" for y
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4291 | using that | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4292 | apply (simp add: set_eq_iff quasi_component_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4293 | by (metis Set.set_insert insertE subset_insertI) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4294 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4295 | by (simp add: quasi_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4296 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4297 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4298 | lemma quasi_components_of_discrete_topology: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4299 |    "quasi_components_of (discrete_topology U) = (\<lambda>x. {x}) ` U"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4300 | by (auto simp add: quasi_components_of_def quasi_component_of_discrete_topology) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4301 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4302 | lemma homeomorphic_map_quasi_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4303 | assumes hmf: "homeomorphic_map X Y f" and "x \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4304 | shows "quasi_component_of_set Y (f x) = f ` (quasi_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4305 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4306 | obtain g where hmg: "homeomorphic_map Y X g" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4307 | and contf: "continuous_map X Y f" and contg: "continuous_map Y X g" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4308 | and fg: "(\<forall>x \<in> topspace X. g(f x) = x) \<and> (\<forall>y \<in> topspace Y. f(g y) = y)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4309 | by (smt (verit, best) hmf homeomorphic_map_maps homeomorphic_maps_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4310 | show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4311 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4312 | show "quasi_component_of_set Y (f x) \<subseteq> f ` quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4313 | using quasi_component_of_continuous_image [OF contg] | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4314 | \<open>x \<in> topspace X\<close> fg image_iff quasi_component_of_subset_topspace by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4315 | show "f ` quasi_component_of_set X x \<subseteq> quasi_component_of_set Y (f x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4316 | using quasi_component_of_continuous_image [OF contf] by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4317 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4318 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4319 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4320 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4321 | lemma homeomorphic_map_quasi_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4322 | assumes "homeomorphic_map X Y f" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4323 | shows "quasi_components_of Y = image (image f) (quasi_components_of X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4324 | using assms | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4325 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4326 | have "\<exists>x\<in>topspace X. quasi_component_of_set Y y = f ` quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4327 | if "y \<in> topspace Y" for y | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4328 | by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of image_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4329 | moreover have "\<exists>x\<in>topspace Y. f ` quasi_component_of_set X u = quasi_component_of_set Y x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4330 | if "u \<in> topspace X" for u | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4331 | by (metis that assms homeomorphic_imp_surjective_map homeomorphic_map_quasi_component_of imageI) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4332 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4333 | by (auto simp: quasi_components_of_def image_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4334 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4335 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4336 | lemma openin_quasi_component_of_locally_connected_space: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4337 | assumes "locally_connected_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4338 | shows "openin X (quasi_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4339 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4340 | have *: "openin X (connected_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4341 | by (simp add: assms openin_connected_component_of_locally_connected_space) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4342 | moreover have "connected_component_of_set X x = quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4343 | using * closedin_connected_component_of connected_component_of_refl connected_imp_quasi_component_of | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4344 | quasi_component_of_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4345 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4346 | by simp | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4347 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4348 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4349 | lemma openin_quasi_components_of_locally_connected_space: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4350 | "locally_connected_space X \<and> c \<in> quasi_components_of X | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4351 | \<Longrightarrow> openin X c" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4352 | by (smt (verit, best) image_iff openin_quasi_component_of_locally_connected_space quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4353 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4354 | lemma quasi_eq_connected_components_of_alt: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4355 | "quasi_components_of X = connected_components_of X \<longleftrightarrow> (\<forall>C \<in> quasi_components_of X. connectedin X C)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4356 | (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4357 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4358 | assume R: ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4359 | moreover have "connected_components_of X \<subseteq> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4360 | using R unfolding quasi_components_of_def connected_components_of_def | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4361 | by (force simp flip: quasi_eq_connected_component_of_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4362 | ultimately show ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4363 | using connected_quasi_component_of by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4364 | qed (use connected_quasi_component_of in blast) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4365 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4366 | lemma connected_subset_quasi_components_of_pointwise: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4367 | "connected_components_of X \<subseteq> quasi_components_of X \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4368 | (\<forall>x \<in> topspace X. quasi_component_of X x = connected_component_of X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4369 | (is "?lhs = ?rhs") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4370 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4371 | assume L: ?lhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4372 | have "connectedin X (quasi_component_of_set X x)" if "x \<in> topspace X" for x | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4373 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4374 | have "\<exists>y\<in>topspace X. connected_component_of_set X x = quasi_component_of_set X y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4375 | using L that by (force simp: quasi_components_of_def connected_components_of_def image_subset_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4376 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4377 | by (metis connected_component_of_equiv connectedin_connected_component_of mem_Collect_eq quasi_component_of_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4378 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4379 | then show ?rhs | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4380 | by (simp add: quasi_eq_connected_component_of_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4381 | qed (simp add: connected_components_of_def quasi_components_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4382 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4383 | lemma quasi_subset_connected_components_of_pointwise: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4384 | "quasi_components_of X \<subseteq> connected_components_of X \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4385 | (\<forall>x \<in> topspace X. quasi_component_of X x = connected_component_of X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4386 | by (simp add: connected_quasi_component_of image_subset_iff quasi_components_of_def quasi_eq_connected_component_of_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4387 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4388 | lemma quasi_eq_connected_components_of_pointwise: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4389 | "quasi_components_of X = connected_components_of X \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4390 | (\<forall>x \<in> topspace X. quasi_component_of X x = connected_component_of X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4391 | using connected_subset_quasi_components_of_pointwise quasi_subset_connected_components_of_pointwise by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4392 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4393 | lemma quasi_eq_connected_components_of_pointwise_alt: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4394 | "quasi_components_of X = connected_components_of X \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4395 | (\<forall>x. quasi_component_of X x = connected_component_of X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4396 | unfolding quasi_eq_connected_components_of_pointwise | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4397 | by (metis connectedin_empty quasi_component_of_eq_empty quasi_eq_connected_component_of_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4398 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4399 | lemma quasi_eq_connected_components_of_inclusion: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4400 | "quasi_components_of X = connected_components_of X \<longleftrightarrow> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4401 | connected_components_of X \<subseteq> quasi_components_of X \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4402 | quasi_components_of X \<subseteq> connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4403 | by (simp add: connected_subset_quasi_components_of_pointwise dual_order.eq_iff quasi_subset_connected_components_of_pointwise) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4404 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4405 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4406 | lemma quasi_eq_connected_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4407 | "finite(connected_components_of X) \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4408 | finite(quasi_components_of X) \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4409 | locally_connected_space X \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4410 | compact_space X \<and> (Hausdorff_space X \<or> regular_space X \<or> normal_space X) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4411 | \<Longrightarrow> quasi_components_of X = connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4412 | proof (elim disjE) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4413 | show "quasi_components_of X = connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4414 | if "finite (connected_components_of X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4415 | unfolding quasi_eq_connected_components_of_inclusion | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4416 | using that open_in_finite_connected_components open_quasi_eq_connected_components_of by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4417 | show "quasi_components_of X = connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4418 | if "finite (quasi_components_of X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4419 | unfolding quasi_eq_connected_components_of_inclusion | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4420 | using that open_quasi_eq_connected_components_of openin_finite_quasi_components by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4421 | show "quasi_components_of X = connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4422 | if "locally_connected_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4423 | unfolding quasi_eq_connected_components_of_inclusion | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4424 | using that open_quasi_eq_connected_components_of openin_quasi_components_of_locally_connected_space by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4425 | show "quasi_components_of X = connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4426 | if "compact_space X \<and> (Hausdorff_space X \<or> regular_space X \<or> normal_space X)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4427 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4428 | show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4429 | unfolding quasi_eq_connected_components_of_alt | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4430 | proof (intro strip) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4431 | fix C | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4432 | assume C: "C \<in> quasi_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4433 | then have cloC: "closedin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4434 | by (simp add: closedin_quasi_components_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4435 | have "normal_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4436 | using that compact_Hausdorff_or_regular_imp_normal_space by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4437 | show "connectedin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4438 | proof (clarsimp simp add: connectedin_def connected_space_closedin_eq closedin_closed_subtopology cloC closedin_subset [OF cloC]) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4439 | fix S T | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4440 |         assume "S \<subseteq> C" and "closedin X S" and "S \<inter> T = {}" and SUT: "S \<union> T = topspace X \<inter> C"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4441 |           and T: "T \<subseteq> C" "T \<noteq> {}" and "closedin X T" 
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4442 | with \<open>normal_space X\<close> obtain U V where UV: "openin X U" "openin X V" "S \<subseteq> U" "T \<subseteq> V" "disjnt U V" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4443 | by (meson disjnt_def normal_space_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4444 | moreover have "compactin X (topspace X - (U \<union> V))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4445 | using UV that by (intro closedin_compact_space closedin_diff openin_Un) auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4446 | ultimately have "separated_between X C (topspace X - (U \<union> V)) \<longleftrightarrow> disjnt C (topspace X - (U \<union> V))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4447 | by (simp add: \<open>C \<in> quasi_components_of X\<close> separated_between_quasi_component_compact) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4448 | moreover have "disjnt C (topspace X - (U \<union> V))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4449 | using UV SUT disjnt_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4450 | ultimately have "separated_between X C (topspace X - (U \<union> V))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4451 | by simp | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4452 | then obtain A B where "openin X A" "openin X B" "A \<union> B = topspace X" "disjnt A B" "C \<subseteq> A" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4453 | and subB: "topspace X - (U \<union> V) \<subseteq> B" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4454 | by (meson separated_between_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4455 | have "B \<union> U = topspace X - (A \<inter> V)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4456 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4457 | show "B \<union> U \<subseteq> topspace X - A \<inter> V" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4458 | using \<open>openin X U\<close> \<open>disjnt U V\<close> \<open>disjnt A B\<close> \<open>openin X B\<close> disjnt_iff openin_closedin_eq by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4459 | show "topspace X - A \<inter> V \<subseteq> B \<union> U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4460 | using \<open>A \<union> B = topspace X\<close> subB by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4461 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4462 | then have "closedin X (B \<union> U)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4463 | using \<open>openin X V\<close> \<open>openin X A\<close> by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4464 | then have "C \<subseteq> B \<union> U \<or> disjnt C (B \<union> U)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4465 | using quasi_component_of_clopen_cases [OF C] \<open>openin X U\<close> \<open>openin X B\<close> by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4466 |         with UV show "S = {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4467 | by (metis UnE \<open>C \<subseteq> A\<close> \<open>S \<subseteq> C\<close> T \<open>disjnt A B\<close> all_not_in_conv disjnt_Un2 disjnt_iff subset_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4468 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4469 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4470 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4471 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4472 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4473 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4474 | lemma quasi_eq_connected_component_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4475 | "finite(connected_components_of X) \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4476 | finite(quasi_components_of X) \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4477 | locally_connected_space X \<or> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4478 | compact_space X \<and> (Hausdorff_space X \<or> regular_space X \<or> normal_space X) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4479 | \<Longrightarrow> quasi_component_of X x = connected_component_of X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4480 | by (metis quasi_eq_connected_components_of quasi_eq_connected_components_of_pointwise_alt) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4481 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4482 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4483 | subsection\<open>Additional quasicomponent and continuum properties like Boundary Bumping\<close> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4484 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4485 | lemma cut_wire_fence_theorem_gen: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4486 | assumes "compact_space X" and X: "Hausdorff_space X \<or> regular_space X \<or> normal_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4487 | and S: "compactin X S" and T: "closedin X T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4488 | and dis: "\<And>C. connectedin X C \<Longrightarrow> disjnt C S \<or> disjnt C T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4489 | shows "separated_between X S T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4490 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4491 |   have "x \<in> topspace X" if "x \<in> S" and "T = {}" for x
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4492 | using that S compactin_subset_topspace by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4493 |   moreover have "separated_between X {x} {y}" if "x \<in> S" and "y \<in> T" for x y
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4494 | proof (cases "x \<in> topspace X \<and> y \<in> topspace X") | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4495 | case True | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4496 | then have "\<not> connected_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4497 | by (meson dis connected_component_of_def disjnt_iff that) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4498 | with True X \<open>compact_space X\<close> show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4499 | by (metis quasi_component_nonseparated quasi_eq_connected_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4500 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4501 | case False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4502 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4503 | using S T compactin_subset_topspace closedin_subset that by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4504 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4505 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4506 | using assms | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4507 | by (simp add: separated_between_pointwise_left separated_between_pointwise_right | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4508 | closedin_compact_space closedin_subset) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4509 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4510 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4511 | lemma cut_wire_fence_theorem: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4512 | "\<lbrakk>compact_space X; Hausdorff_space X; closedin X S; closedin X T; | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4513 | \<And>C. connectedin X C \<Longrightarrow> disjnt C S \<or> disjnt C T\<rbrakk> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4514 | \<Longrightarrow> separated_between X S T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4515 | by (simp add: closedin_compact_space cut_wire_fence_theorem_gen) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4516 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4517 | lemma separated_between_from_closed_subtopology: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4518 | assumes XC: "separated_between (subtopology X C) S (X frontier_of C)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4519 | and ST: "separated_between (subtopology X C) S T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4520 | shows "separated_between X S T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4521 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4522 | obtain U where clo: "closedin (subtopology X C) U" and ope: "openin (subtopology X C) U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4523 | and "S \<subseteq> U" and sub: "X frontier_of C \<union> T \<subseteq> topspace (subtopology X C) - U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4524 | by (meson assms separated_between separated_between_Un) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4525 | then have "X frontier_of C \<union> T \<subseteq> topspace X \<inter> C - U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4526 | by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4527 | have "closedin X (topspace X \<inter> C)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4528 | by (metis XC frontier_of_restrict frontier_of_subset_eq inf_le1 separated_between_imp_subset topspace_subtopology) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4529 | then have "closedin X U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4530 | by (metis clo closedin_closed_subtopology subtopology_restrict) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4531 | moreover have "openin (subtopology X C) U \<longleftrightarrow> openin X U \<and> U \<subseteq> C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4532 | using disjnt_iff sub by (force intro!: openin_subset_topspace_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4533 | with ope have "openin X U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4534 | by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4535 | moreover have "T \<subseteq> topspace X - U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4536 | using ope openin_closedin_eq sub by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4537 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4538 | using \<open>S \<subseteq> U\<close> separated_between by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4539 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4540 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4541 | lemma separated_between_from_closed_subtopology_frontier: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4542 | "separated_between (subtopology X T) S (X frontier_of T) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4543 | \<Longrightarrow> separated_between X S (X frontier_of T)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4544 | using separated_between_from_closed_subtopology by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4545 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4546 | lemma separated_between_from_frontier_of_closed_subtopology: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4547 | assumes "separated_between (subtopology X T) S (X frontier_of T)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4548 | shows "separated_between X S (topspace X - T)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4549 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4550 | have "disjnt S (topspace X - T)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4551 | using assms disjnt_iff separated_between_imp_subset by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4552 | then show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4553 | by (metis Diff_subset assms frontier_of_complement separated_between_from_closed_subtopology separated_between_frontier_of_eq') | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4554 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4555 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4556 | lemma separated_between_compact_connected_component: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4557 | assumes "locally_compact_space X" "Hausdorff_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4558 | and C: "C \<in> connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4559 | and "compactin X C" "closedin X T" "disjnt C T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4560 | shows "separated_between X C T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4561 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4562 | have Csub: "C \<subseteq> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4563 | by (simp add: assms(4) compactin_subset_topspace) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4564 | have "Hausdorff_space (subtopology X (topspace X - T))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4565 | using Hausdorff_space_subtopology assms(2) by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4566 | moreover have "compactin (subtopology X (topspace X - T)) C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4567 | using assms Csub by (metis Diff_Int_distrib Diff_empty compact_imp_compactin_subtopology disjnt_def le_iff_inf) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4568 | moreover have "locally_compact_space (subtopology X (topspace X - T))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4569 | by (meson assms closedin_def locally_compact_Hausdorff_imp_regular_space locally_compact_space_open_subset) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4570 | ultimately | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4571 | obtain N L where "openin X N" "compactin X L" "closedin X L" "C \<subseteq> N" "N \<subseteq> L" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4572 | and Lsub: "L \<subseteq> topspace X - T" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4573 | using \<open>Hausdorff_space X\<close> \<open>closedin X T\<close> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4574 | apply (simp add: locally_compact_space_compact_closed_compact compactin_subtopology) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4575 | by (meson closedin_def compactin_imp_closedin openin_trans_full) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4576 | then have disC: "disjnt C (topspace X - L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4577 | by (meson DiffD2 disjnt_iff subset_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4578 | have "separated_between (subtopology X L) C (X frontier_of L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4579 | proof (rule cut_wire_fence_theorem) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4580 | show "compact_space (subtopology X L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4581 | by (simp add: \<open>compactin X L\<close> compact_space_subtopology) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4582 | show "Hausdorff_space (subtopology X L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4583 | by (simp add: Hausdorff_space_subtopology \<open>Hausdorff_space X\<close>) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4584 | show "closedin (subtopology X L) C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4585 | by (meson \<open>C \<subseteq> N\<close> \<open>N \<subseteq> L\<close> \<open>Hausdorff_space X\<close> \<open>compactin X C\<close> closedin_subset_topspace compactin_imp_closedin subset_trans) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4586 | show "closedin (subtopology X L) (X frontier_of L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4587 | by (simp add: \<open>closedin X L\<close> closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4588 | show "disjnt D C \<or> disjnt D (X frontier_of L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4589 | if "connectedin (subtopology X L) D" for D | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4590 | proof (rule ccontr) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4591 | assume "\<not> (disjnt D C \<or> disjnt D (X frontier_of L))" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4592 | moreover have "connectedin X D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4593 | using connectedin_subtopology that by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4594 | ultimately show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4595 | using that connected_components_of_maximal [of C X D] C | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4596 | apply (simp add: disjnt_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4597 | by (metis Diff_eq_empty_iff \<open>C \<subseteq> N\<close> \<open>N \<subseteq> L\<close> \<open>openin X N\<close> disjoint_iff frontier_of_openin_straddle_Int(2) subsetD) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4598 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4599 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4600 | then have "separated_between X (X frontier_of C) (topspace X - L)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4601 | using separated_between_from_frontier_of_closed_subtopology separated_between_frontier_of_eq by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4602 | with \<open>closedin X T\<close> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4603 | separated_between_frontier_of [OF Csub disC] | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4604 | show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4605 | unfolding separated_between by (smt (verit) Diff_iff Lsub closedin_subset subset_iff) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4606 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4607 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4608 | lemma wilder_locally_compact_component_thm: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4609 | assumes "locally_compact_space X" "Hausdorff_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4610 | and "C \<in> connected_components_of X" "compactin X C" "openin X W" "C \<subseteq> W" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4611 | obtains U V where "openin X U" "openin X V" "disjnt U V" "U \<union> V = topspace X" "C \<subseteq> U" "U \<subseteq> W" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4612 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4613 | have "closedin X (topspace X - W)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4614 | using \<open>openin X W\<close> by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4615 | moreover have "disjnt C (topspace X - W)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4616 | using \<open>C \<subseteq> W\<close> disjnt_def by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4617 | ultimately have "separated_between X C (topspace X - W)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4618 | using separated_between_compact_connected_component assms by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4619 | then show thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4620 | by (smt (verit, del_insts) DiffI disjnt_iff openin_subset separated_between_def subset_iff that) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4621 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4622 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4623 | lemma compact_quasi_eq_connected_components_of: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4624 | assumes "locally_compact_space X" "Hausdorff_space X" "compactin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4625 | shows "C \<in> quasi_components_of X \<longleftrightarrow> C \<in> connected_components_of X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4626 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4627 | have "compactin X (connected_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4628 | if "x \<in> topspace X" "compactin X (quasi_component_of_set X x)" for x | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4629 | proof (rule closed_compactin) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4630 | show "compactin X (quasi_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4631 | by (simp add: that) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4632 | show "connected_component_of_set X x \<subseteq> quasi_component_of_set X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4633 | by (simp add: connected_component_subset_quasi_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4634 | show "closedin X (connected_component_of_set X x)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4635 | by (simp add: closedin_connected_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4636 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4637 | moreover have "connected_component_of X x = quasi_component_of X x" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4638 | if \<section>: "x \<in> topspace X" "compactin X (connected_component_of_set X x)" for x | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4639 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4640 | have "\<And>y. connected_component_of X x y \<Longrightarrow> quasi_component_of X x y" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4641 | by (simp add: connected_imp_quasi_component_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4642 | moreover have False if non: "\<not> connected_component_of X x y" and quasi: "quasi_component_of X x y" for y | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4643 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4644 | have "y \<in> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4645 | by (meson quasi_component_of_equiv that) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4646 |       then have "closedin X {y}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4647 | by (simp add: \<open>Hausdorff_space X\<close> compactin_imp_closedin) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4648 |       moreover have "disjnt (connected_component_of_set X x) {y}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4649 | by (simp add: non) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4650 |       moreover have "\<not> separated_between X (connected_component_of_set X x) {y}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4651 | using \<section> quasi separated_between_pointwise_left | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4652 | by (fastforce simp: quasi_component_nonseparated connected_component_of_refl) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4653 | ultimately show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4654 | using assms by (metis \<section> connected_component_in_connected_components_of separated_between_compact_connected_component) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4655 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4656 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4657 | by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4658 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4659 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4660 | using \<open>compactin X C\<close> unfolding connected_components_of_def image_iff quasi_components_of_def by metis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4661 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4662 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4663 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4664 | lemma boundary_bumping_theorem_closed_gen: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4665 | assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" "closedin X S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4666 | "S \<noteq> topspace X" and C: "compactin X C" "C \<in> connected_components_of (subtopology X S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4667 |   shows "C \<inter> X frontier_of S \<noteq> {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4668 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4669 |   assume \<section>: "C \<inter> X frontier_of S = {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4670 |   consider "C \<noteq> {}" "X frontier_of S \<subseteq> topspace X" | "C \<subseteq> topspace X" "S = {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4671 | using C by (metis frontier_of_subset_topspace nonempty_connected_components_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4672 | then show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4673 | proof cases | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4674 | case 1 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4675 | have "separated_between (subtopology X S) C (X frontier_of S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4676 | proof (rule separated_between_compact_connected_component) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4677 | show "compactin (subtopology X S) C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4678 | using C compact_imp_compactin_subtopology connected_components_of_subset by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4679 | show "closedin (subtopology X S) (X frontier_of S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4680 | by (simp add: \<open>closedin X S\<close> closedin_frontier_of closedin_subset_topspace frontier_of_subset_closedin) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4681 | show "disjnt C (X frontier_of S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4682 | using \<section> by (simp add: disjnt_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4683 | qed (use assms Hausdorff_space_subtopology locally_compact_space_closed_subset in auto) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4684 | then have "separated_between X C (X frontier_of S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4685 | using separated_between_from_closed_subtopology by auto | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4686 |     then have "X frontier_of S = {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4687 |       using \<open>C \<noteq> {}\<close> \<open>connected_space X\<close> connected_space_separated_between by blast
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4688 | moreover have "C \<subseteq> S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4689 | using C connected_components_of_subset by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4690 | ultimately show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4691 | using 1 assms by (metis closedin_subset connected_space_eq_frontier_eq_empty subset_empty) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4692 | next | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4693 | case 2 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4694 | then show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4695 | using C connected_components_of_eq_empty by fastforce | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4696 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4697 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4698 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4699 | lemma boundary_bumping_theorem_closed: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4700 | assumes "connected_space X" "compact_space X" "Hausdorff_space X" "closedin X S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4701 | "S \<noteq> topspace X" "C \<in> connected_components_of(subtopology X S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4702 |   shows "C \<inter> X frontier_of S \<noteq> {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4703 | by (meson assms boundary_bumping_theorem_closed_gen closedin_compact_space closedin_connected_components_of | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4704 | closedin_trans_full compact_imp_locally_compact_space) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4705 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4706 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4707 | lemma intermediate_continuum_exists: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4708 | assumes "connected_space X" "locally_compact_space X" "Hausdorff_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4709 |     and C: "compactin X C" "connectedin X C" "C \<noteq> {}" "C \<noteq> topspace X"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4710 | and U: "openin X U" "C \<subseteq> U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4711 | obtains D where "compactin X D" "connectedin X D" "C \<subset> D" "D \<subset> U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4712 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4713 | have "C \<subseteq> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4714 | by (simp add: C compactin_subset_topspace) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4715 | with C obtain a where a: "a \<in> topspace X" "a \<notin> C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4716 | by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4717 |   moreover have "compactin (subtopology X (U - {a})) C"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4718 | by (simp add: C U a compact_imp_compactin_subtopology subset_Diff_insert) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4719 |   moreover have "Hausdorff_space (subtopology X (U - {a}))"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4720 | using Hausdorff_space_subtopology assms(3) by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4721 | moreover | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4722 |   have "locally_compact_space (subtopology X (U - {a}))"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4723 | by (rule locally_compact_space_open_subset) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4724 | (auto simp: locally_compact_Hausdorff_imp_regular_space open_in_Hausdorff_delete assms) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4725 | ultimately obtain V K where V: "openin X V" "a \<notin> V" "V \<subseteq> U" and K: "compactin X K" "a \<notin> K" "K \<subseteq> U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4726 |     and cloK: "closedin (subtopology X (U - {a})) K" and "C \<subseteq> V" "V \<subseteq> K"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4727 |     using locally_compact_space_compact_closed_compact [of "subtopology X (U - {a})"] assms
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4728 | by (smt (verit, del_insts) Diff_empty compactin_subtopology open_in_Hausdorff_delete openin_open_subtopology subset_Diff_insert) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4729 | then obtain D where D: "D \<in> connected_components_of (subtopology X K)" and "C \<subseteq> D" | 
| 78336 | 4730 | using C | 
| 4731 | by (metis compactin_subset_topspace connected_component_in_connected_components_of | |
| 4732 | connected_component_of_maximal connectedin_subtopology subset_empty subset_eq topspace_subtopology_subset) | |
| 78038 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4733 | show thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4734 | proof | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4735 | have cloD: "closedin (subtopology X K) D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4736 | by (simp add: D closedin_connected_components_of) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4737 | then have XKD: "compactin (subtopology X K) D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4738 | by (simp add: K closedin_compact_space compact_space_subtopology) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4739 | then show "compactin X D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4740 | using compactin_subtopology_imp_compact by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4741 | show "connectedin X D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4742 | using D connectedin_connected_components_of connectedin_subtopology by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4743 | have "K \<noteq> topspace X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4744 | using K a by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4745 | moreover have "V \<subseteq> X interior_of K" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4746 | by (simp add: \<open>openin X V\<close> \<open>V \<subseteq> K\<close> interior_of_maximal) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4747 | ultimately have "C \<noteq> D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4748 | using boundary_bumping_theorem_closed_gen [of X K C] D \<open>C \<subseteq> V\<close> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4749 | by (auto simp add: assms K compactin_imp_closedin frontier_of_def) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4750 | then show "C \<subset> D" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4751 | using \<open>C \<subseteq> D\<close> by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4752 | have "D \<subseteq> U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4753 | using K(3) \<open>closedin (subtopology X K) D\<close> closedin_imp_subset by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4754 | moreover have "D \<noteq> U" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4755 | using K XKD \<open>C \<subset> D\<close> assms | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4756 | by (metis \<open>K \<noteq> topspace X\<close> cloD closedin_imp_subset compactin_imp_closedin connected_space_clopen_in | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4757 | inf_bot_left inf_le2 subset_antisym) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4758 | ultimately | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4759 | show "D \<subset> U" by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4760 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4761 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4762 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4763 | lemma boundary_bumping_theorem_gen: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4764 | assumes X: "connected_space X" "locally_compact_space X" "Hausdorff_space X" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4765 | and "S \<subset> topspace X" and C: "C \<in> connected_components_of(subtopology X S)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4766 | and compC: "compactin X (X closure_of C)" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4767 |  shows "X frontier_of C \<inter> X frontier_of S \<noteq> {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4768 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4769 | have Csub: "C \<subseteq> topspace X" "C \<subseteq> S" and "connectedin X C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4770 | using C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4771 | by fastforce+ | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4772 |   have "C \<noteq> {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4773 | using C nonempty_connected_components_of by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4774 | obtain "X interior_of C \<subseteq> X interior_of S" "X closure_of C \<subseteq> X closure_of S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4775 | by (simp add: Csub closure_of_mono interior_of_mono) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4776 | moreover have False if "X closure_of C \<subseteq> X interior_of S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4777 | proof - | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4778 | have "X closure_of C = C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4779 | by (meson C closedin_connected_component_of_subtopology closure_of_eq interior_of_subset order_trans that) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4780 | with that have "C \<subseteq> X interior_of S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4781 | by simp | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4782 | then obtain D where "compactin X D" and "connectedin X D" and "C \<subset> D" and "D \<subset> X interior_of S" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4783 | using intermediate_continuum_exists assms \<open>X closure_of C = C\<close> compC Csub | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4784 |       by (metis \<open>C \<noteq> {}\<close> \<open>connectedin X C\<close> openin_interior_of psubsetE)
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4785 | then have "D \<subseteq> C" | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4786 |       by (metis C \<open>C \<noteq> {}\<close> connected_components_of_maximal connectedin_subtopology disjnt_def inf.orderE interior_of_subset order_trans psubsetE)
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4787 | then show False | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4788 | using \<open>C \<subset> D\<close> by blast | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4789 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4790 | ultimately show ?thesis | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4791 | by (smt (verit, ccfv_SIG) DiffI disjoint_iff_not_equal frontier_of_def subset_eq) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4792 | qed | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4793 | |
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4794 | lemma boundary_bumping_theorem: | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4795 | "\<lbrakk>connected_space X; compact_space X; Hausdorff_space X; S \<subset> topspace X; | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4796 | C \<in> connected_components_of(subtopology X S)\<rbrakk> | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4797 |     \<Longrightarrow> X frontier_of C \<inter> X frontier_of S \<noteq> {}"
 | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4798 | by (simp add: boundary_bumping_theorem_gen closedin_compact_space compact_imp_locally_compact_space) | 
| 
2c1b01563163
New material from the HOL Light metric space library, mostly about quasi components
 paulson <lp15@cam.ac.uk> parents: 
78037diff
changeset | 4799 | |
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4800 | subsection \<open>Compactly generated spaces (k-spaces)\<close> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4801 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4802 | text \<open>These don't have to be Hausdorff\<close> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4803 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4804 | definition k_space where | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4805 | "k_space X \<equiv> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4806 | \<forall>S. S \<subseteq> topspace X \<longrightarrow> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4807 | (closedin X S \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> S)))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4808 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4809 | lemma k_space: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4810 | "k_space X \<longleftrightarrow> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4811 | (\<forall>S. S \<subseteq> topspace X \<and> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4812 | (\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> S)) \<longrightarrow> closedin X S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4813 | by (metis closedin_subtopology inf_commute k_space_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4814 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4815 | lemma k_space_open: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4816 | "k_space X \<longleftrightarrow> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4817 | (\<forall>S. S \<subseteq> topspace X \<and> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4818 | (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S)) \<longrightarrow> openin X S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4819 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4820 | have "openin X S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4821 | if "k_space X" "S \<subseteq> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4822 | and "\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S)" for S | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4823 | using that unfolding k_space openin_closedin_eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4824 | by (metis Diff_Int_distrib2 Diff_subset inf_commute topspace_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4825 | moreover have "k_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4826 | if "\<forall>S. S \<subseteq> topspace X \<and> (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S)) \<longrightarrow> openin X S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4827 | unfolding k_space openin_closedin_eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4828 | by (simp add: Diff_Int_distrib closedin_def inf_commute that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4829 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4830 | by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4831 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4832 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4833 | lemma k_space_alt: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4834 | "k_space X \<longleftrightarrow> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4835 | (\<forall>S. S \<subseteq> topspace X | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4836 | \<longrightarrow> (openin X S \<longleftrightarrow> (\<forall>K. compactin X K \<longrightarrow> openin (subtopology X K) (K \<inter> S))))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4837 | by (meson k_space_open openin_subtopology_Int2) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4838 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4839 | lemma k_space_quotient_map_image: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4840 | assumes q: "quotient_map X Y q" and X: "k_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4841 | shows "k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4842 | unfolding k_space | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4843 | proof clarify | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4844 | fix S | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4845 | assume "S \<subseteq> topspace Y" and S: "\<forall>K. compactin Y K \<longrightarrow> closedin (subtopology Y K) (K \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4846 |   then have iff: "closedin X {x \<in> topspace X. q x \<in> S} \<longleftrightarrow> closedin Y S"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4847 | using q quotient_map_closedin by fastforce | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4848 |   have "closedin (subtopology X K) (K \<inter> {x \<in> topspace X. q x \<in> S})" if "compactin X K" for K
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4849 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4850 |     have "{x \<in> topspace X. q x \<in> q ` K} \<inter> K = K"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4851 | using compactin_subset_topspace that by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4852 |     then have *: "subtopology X K = subtopology (subtopology X {x \<in> topspace X. q x \<in> q ` K}) K"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4853 | by (simp add: subtopology_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4854 |     have **: "K \<inter> {x \<in> topspace X. q x \<in> S} =
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4855 |               K \<inter> {x \<in> topspace (subtopology X {x \<in> topspace X. q x \<in> q ` K}). q x \<in> q ` K \<inter> S}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4856 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4857 | have "K \<subseteq> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4858 | by (simp add: compactin_subset_topspace that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4859 | show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4860 | unfolding * ** | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4861 | proof (intro closedin_continuous_map_preimage closedin_subtopology_Int_closed) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4862 |       show "continuous_map (subtopology X {x \<in> topspace X. q x \<in> q ` K}) (subtopology Y (q ` K)) q"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4863 | by (auto simp add: continuous_map_in_subtopology continuous_map_from_subtopology q quotient_imp_continuous_map) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4864 | show "closedin (subtopology Y (q ` K)) (q ` K \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4865 | by (meson S image_compactin q quotient_imp_continuous_map that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4866 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4867 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4868 |   then have "closedin X {x \<in> topspace X. q x \<in> S}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4869 | by (metis (no_types, lifting) X k_space mem_Collect_eq subsetI) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4870 | with iff show "closedin Y S" by simp | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4871 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4872 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4873 | lemma k_space_retraction_map_image: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4874 | "\<lbrakk>retraction_map X Y r; k_space X\<rbrakk> \<Longrightarrow> k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4875 | using k_space_quotient_map_image retraction_imp_quotient_map by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4876 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4877 | lemma homeomorphic_k_space: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4878 | "X homeomorphic_space Y \<Longrightarrow> k_space X \<longleftrightarrow> k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4879 | by (meson homeomorphic_map_def homeomorphic_space homeomorphic_space_sym k_space_quotient_map_image) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4880 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4881 | lemma k_space_perfect_map_image: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4882 | "\<lbrakk>k_space X; perfect_map X Y f\<rbrakk> \<Longrightarrow> k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4883 | using k_space_quotient_map_image perfect_imp_quotient_map by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4884 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4885 | lemma locally_compact_imp_k_space: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4886 | assumes "locally_compact_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4887 | shows "k_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4888 | unfolding k_space | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4889 | proof clarify | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4890 | fix S | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4891 | assume "S \<subseteq> topspace X" and S: "\<forall>K. compactin X K \<longrightarrow> closedin (subtopology X K) (K \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4892 | have False if non: "\<not> (X closure_of S \<subseteq> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4893 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4894 | obtain x where "x \<in> X closure_of S" "x \<notin> S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4895 | using non by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4896 | then have "x \<in> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4897 | by (simp add: in_closure_of) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4898 | then obtain K U where "openin X U" "compactin X K" "x \<in> U" "U \<subseteq> K" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4899 | by (meson assms locally_compact_space_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4900 | then show False | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4901 | using \<open>x \<in> X closure_of S\<close> openin_Int_closure_of_eq [OF \<open>openin X U\<close>] | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4902 | by (smt (verit, ccfv_threshold) Int_iff S \<open>x \<notin> S\<close> closedin_Int_closure_of inf.orderE inf_assoc) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4903 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4904 | then show "closedin X S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4905 | using S \<open>S \<subseteq> topspace X\<close> closure_of_subset_eq by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4906 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4907 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4908 | lemma compact_imp_k_space: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4909 | "compact_space X \<Longrightarrow> k_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4910 | by (simp add: compact_imp_locally_compact_space locally_compact_imp_k_space) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4911 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4912 | lemma k_space_discrete_topology: "k_space(discrete_topology U)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4913 | by (simp add: k_space_open) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4914 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4915 | lemma k_space_closed_subtopology: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4916 | assumes "k_space X" "closedin X C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4917 | shows "k_space (subtopology X C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4918 | unfolding k_space compactin_subtopology | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4919 | proof clarsimp | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4920 | fix S | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4921 | assume Ssub: "S \<subseteq> topspace X" "S \<subseteq> C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4922 | and S: "\<forall>K. compactin X K \<and> K \<subseteq> C \<longrightarrow> closedin (subtopology (subtopology X C) K) (K \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4923 | have "closedin (subtopology X K) (K \<inter> S)" if "compactin X K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4924 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4925 | have "closedin (subtopology (subtopology X C) (K \<inter> C)) ((K \<inter> C) \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4926 | by (simp add: S \<open>closedin X C\<close> compact_Int_closedin that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4927 | then show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4928 | using \<open>closedin X C\<close> Ssub by (auto simp add: closedin_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4929 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4930 | then show "closedin (subtopology X C) S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4931 | by (metis Ssub \<open>k_space X\<close> closedin_subset_topspace k_space_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4932 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4933 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4934 | lemma k_space_subtopology: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4935 | assumes 1: "\<And>T. \<lbrakk>T \<subseteq> topspace X; T \<subseteq> S; | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4936 | \<And>K. compactin X K \<Longrightarrow> closedin (subtopology X (K \<inter> S)) (K \<inter> T)\<rbrakk> \<Longrightarrow> closedin (subtopology X S) T" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4937 | assumes 2: "\<And>K. compactin X K \<Longrightarrow> k_space(subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4938 | shows "k_space (subtopology X S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4939 | unfolding k_space | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4940 | proof (intro conjI strip) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4941 | fix U | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4942 | assume \<section>: "U \<subseteq> topspace (subtopology X S) \<and> (\<forall>K. compactin (subtopology X S) K \<longrightarrow> closedin (subtopology (subtopology X S) K) (K \<inter> U))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4943 | have "closedin (subtopology X (K \<inter> S)) (K \<inter> U)" if "compactin X K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4944 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4945 | have "K \<inter> U \<subseteq> topspace (subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4946 | using "\<section>" by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4947 | moreover | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4948 | have "\<And>K'. compactin (subtopology X (K \<inter> S)) K' \<Longrightarrow> closedin (subtopology (subtopology X (K \<inter> S)) K') (K' \<inter> K \<inter> U)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4949 | by (metis "\<section>" compactin_subtopology inf.orderE inf_commute subtopology_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4950 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4951 | by (metis (no_types, opaque_lifting) "2" inf.assoc k_space_def that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4952 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4953 | then show "closedin (subtopology X S) U" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4954 | using "1" \<section> by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4955 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4956 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4957 | lemma k_space_subtopology_open: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4958 | assumes 1: "\<And>T. \<lbrakk>T \<subseteq> topspace X; T \<subseteq> S; | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4959 | \<And>K. compactin X K \<Longrightarrow> openin (subtopology X (K \<inter> S)) (K \<inter> T)\<rbrakk> \<Longrightarrow> openin (subtopology X S) T" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4960 | assumes 2: "\<And>K. compactin X K \<Longrightarrow> k_space(subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4961 | shows "k_space (subtopology X S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4962 | unfolding k_space_open | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4963 | proof (intro conjI strip) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4964 | fix U | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4965 | assume \<section>: "U \<subseteq> topspace (subtopology X S) \<and> (\<forall>K. compactin (subtopology X S) K \<longrightarrow> openin (subtopology (subtopology X S) K) (K \<inter> U))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4966 | have "openin (subtopology X (K \<inter> S)) (K \<inter> U)" if "compactin X K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4967 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4968 | have "K \<inter> U \<subseteq> topspace (subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4969 | using "\<section>" by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4970 | moreover | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4971 | have "\<And>K'. compactin (subtopology X (K \<inter> S)) K' \<Longrightarrow> openin (subtopology (subtopology X (K \<inter> S)) K') (K' \<inter> K \<inter> U)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4972 | by (metis "\<section>" compactin_subtopology inf.orderE inf_commute subtopology_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4973 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4974 | by (metis (no_types, opaque_lifting) "2" inf.assoc k_space_open that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4975 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4976 | then show "openin (subtopology X S) U" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4977 | using "1" \<section> by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4978 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4979 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4980 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4981 | lemma k_space_open_subtopology_aux: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4982 | assumes "kc_space X" "compact_space X" "openin X V" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4983 | shows "k_space (subtopology X V)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4984 | proof (clarsimp simp: k_space subtopology_subtopology compactin_subtopology Int_absorb1) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4985 | fix S | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4986 | assume "S \<subseteq> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4987 | and "S \<subseteq> V" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4988 | and S: "\<forall>K. compactin X K \<and> K \<subseteq> V \<longrightarrow> closedin (subtopology X K) (K \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4989 | then have "V \<subseteq> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4990 | using assms openin_subset by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4991 | have "S = V \<inter> ((topspace X - V) \<union> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4992 | using \<open>S \<subseteq> V\<close> by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4993 | moreover have "closedin (subtopology X V) (V \<inter> ((topspace X - V) \<union> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4994 | proof (intro closedin_subtopology_Int_closed compactin_imp_closedin_gen \<open>kc_space X\<close>) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4995 | show "compactin X (topspace X - V \<union> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4996 | unfolding compactin_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4997 | proof (intro conjI strip) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4998 | show "topspace X - V \<union> S \<subseteq> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 4999 | by (simp add: \<open>S \<subseteq> topspace X\<close>) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5000 | fix \<U> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5001 | assume \<U>: "Ball \<U> (openin X) \<and> topspace X - V \<union> S \<subseteq> \<Union>\<U>" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5002 | moreover | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5003 | have "compactin X (topspace X - V)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5004 | using assms closedin_compact_space by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5005 | ultimately obtain \<G> where "finite \<G>" "\<G> \<subseteq> \<U>" and \<G>: "topspace X - V \<subseteq> \<Union>\<G>" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5006 | unfolding compactin_def using \<open>V \<subseteq> topspace X\<close> by (metis le_sup_iff) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5007 | then have "topspace X - \<Union>\<G> \<subseteq> V" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5008 | by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5009 | then have "closedin (subtopology X (topspace X - \<Union>\<G>)) ((topspace X - \<Union>\<G>) \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5010 | by (meson S \<U> \<open>\<G> \<subseteq> \<U>\<close> \<open>compact_space X\<close> closedin_compact_space openin_Union openin_closedin_eq subset_iff) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5011 | then have "compactin X ((topspace X - \<Union>\<G>) \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5012 | by (meson \<U> \<open>\<G> \<subseteq> \<U>\<close>\<open>compact_space X\<close> closedin_compact_space closedin_trans_full openin_Union openin_closedin_eq subset_iff) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5013 | then obtain \<H> where "finite \<H>" "\<H> \<subseteq> \<U>" "(topspace X - \<Union>\<G>) \<inter> S \<subseteq> \<Union>\<H>" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5014 | unfolding compactin_def by (smt (verit, best) \<U> inf_le2 subset_trans sup.boundedE) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5015 | with \<G> have "topspace X - V \<union> S \<subseteq> \<Union>(\<G> \<union> \<H>)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5016 | using \<open>S \<subseteq> topspace X\<close> by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5017 | then show "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> \<U> \<and> topspace X - V \<union> S \<subseteq> \<Union>\<F>" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5018 | by (metis \<open>\<G> \<subseteq> \<U>\<close> \<open>\<H> \<subseteq> \<U>\<close> \<open>finite \<G>\<close> \<open>finite \<H>\<close> finite_Un le_sup_iff) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5019 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5020 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5021 | ultimately show "closedin (subtopology X V) S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5022 | by metis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5023 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5024 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5025 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5026 | lemma k_space_open_subtopology: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5027 | assumes X: "kc_space X \<or> Hausdorff_space X \<or> regular_space X" and "k_space X" "openin X S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5028 | shows "k_space(subtopology X S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5029 | proof (rule k_space_subtopology_open) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5030 | fix T | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5031 | assume "T \<subseteq> topspace X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5032 | and "T \<subseteq> S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5033 | and T: "\<And>K. compactin X K \<Longrightarrow> openin (subtopology X (K \<inter> S)) (K \<inter> T)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5034 | have "openin (subtopology X K) (K \<inter> T)" if "compactin X K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5035 | by (smt (verit, ccfv_threshold) T assms(3) inf_assoc inf_commute openin_Int openin_subtopology that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5036 | then show "openin (subtopology X S) T" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5037 | by (metis \<open>T \<subseteq> S\<close> \<open>T \<subseteq> topspace X\<close> assms(2) k_space_alt subset_openin_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5038 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5039 | fix K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5040 | assume "compactin X K" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5041 | then have KS: "openin (subtopology X K) (K \<inter> S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5042 | by (simp add: \<open>openin X S\<close> openin_subtopology_Int2) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5043 | have XK: "compact_space (subtopology X K)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5044 | by (simp add: \<open>compactin X K\<close> compact_space_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5045 | show "k_space (subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5046 | using X | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5047 | proof (rule disjE) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5048 | assume "kc_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5049 | then show "k_space (subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5050 | using k_space_open_subtopology_aux [of "subtopology X K" "K \<inter> S"] | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5051 | by (simp add: KS XK kc_space_subtopology subtopology_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5052 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5053 | assume "Hausdorff_space X \<or> regular_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5054 | then have "locally_compact_space (subtopology (subtopology X K) (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5055 | using locally_compact_space_open_subset Hausdorff_space_subtopology KS XK | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5056 | compact_imp_locally_compact_space regular_space_subtopology by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5057 | then show "k_space (subtopology X (K \<inter> S))" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5058 | by (simp add: locally_compact_imp_k_space subtopology_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5059 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5060 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5061 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5062 | lemma k_kc_space_subtopology: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5063 | "\<lbrakk>k_space X; kc_space X; openin X S \<or> closedin X S\<rbrakk> \<Longrightarrow> k_space(subtopology X S) \<and> kc_space(subtopology X S)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5064 | by (metis k_space_closed_subtopology k_space_open_subtopology kc_space_subtopology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5065 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5066 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5067 | lemma k_space_as_quotient_explicit: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5068 |   "k_space X \<longleftrightarrow> quotient_map (sum_topology (subtopology X) {K. compactin X K}) X snd"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5069 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5070 |   have [simp]: "{x \<in> topspace X. x \<in> K \<and> x \<in> U} = K \<inter> U" if "U \<subseteq> topspace X" for K U
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5071 | using that by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5072 | show "?thesis" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5073 | apply (simp add: quotient_map_def openin_sum_topology snd_image_Sigma k_space_alt) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5074 | by (smt (verit, del_insts) Union_iff compactin_sing inf.orderE mem_Collect_eq singletonI subsetI) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5075 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5076 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5077 | lemma k_space_as_quotient: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5078 | fixes X :: "'a topology" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5079 |   shows "k_space X \<longleftrightarrow> (\<exists>q. \<exists>Y:: ('a set * 'a) topology. locally_compact_space Y \<and> quotient_map Y X q)" 
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5080 | (is "?lhs=?rhs") | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5081 | proof | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5082 | show "k_space X" if ?rhs | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5083 | using that k_space_quotient_map_image locally_compact_imp_k_space by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5084 | next | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5085 | assume "k_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5086 | show ?rhs | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5087 | proof (intro exI conjI) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5088 |     show "locally_compact_space (sum_topology (subtopology X) {K. compactin X K})"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5089 | by (simp add: compact_imp_locally_compact_space compact_space_subtopology locally_compact_space_sum_topology) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5090 |     show "quotient_map (sum_topology (subtopology X) {K. compactin X K}) X snd"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5091 | using \<open>k_space X\<close> k_space_as_quotient_explicit by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5092 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5093 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5094 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5095 | lemma k_space_prod_topology_left: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5096 | assumes X: "locally_compact_space X" "Hausdorff_space X \<or> regular_space X" and "k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5097 | shows "k_space (prod_topology X Y)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5098 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5099 |   obtain q and Z :: "('b set * 'b) topology" where "locally_compact_space Z" and q: "quotient_map Z Y q"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5100 | using \<open>k_space Y\<close> k_space_as_quotient by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5101 | then show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5102 | using quotient_map_prod_right [OF X q] X k_space_quotient_map_image locally_compact_imp_k_space | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5103 | locally_compact_space_prod_topology by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5104 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5105 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5106 | lemma k_space_prod_topology_right: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5107 | assumes "k_space X" and Y: "locally_compact_space Y" "Hausdorff_space Y \<or> regular_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5108 | shows "k_space (prod_topology X Y)" | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 5109 | using assms homeomorphic_k_space homeomorphic_space_prod_topology_swap k_space_prod_topology_left by blast | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5110 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5111 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5112 | lemma continuous_map_from_k_space: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5113 | assumes "k_space X" and f: "\<And>K. compactin X K \<Longrightarrow> continuous_map(subtopology X K) Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5114 | shows "continuous_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5115 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5116 | have "\<And>x. x \<in> topspace X \<Longrightarrow> f x \<in> topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5117 | by (metis compactin_absolute compactin_sing f image_compactin image_empty image_insert) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5118 |   moreover have "closedin X {x \<in> topspace X. f x \<in> C}" if "closedin Y C" for C
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5119 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5120 |     have "{x \<in> topspace X. f x \<in> C} \<subseteq> topspace X"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5121 | by fastforce | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5122 | moreover | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5123 |     have eq: "K \<inter> {x \<in> topspace X. f x \<in> C} = {x. x \<in> topspace(subtopology X K) \<and> f x \<in> (f ` K \<inter> C)}" for K
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5124 | by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5125 |     have "closedin (subtopology X K) (K \<inter> {x \<in> topspace X. f x \<in> C})" if "compactin X K" for K
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5126 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5127 | proof (rule closedin_continuous_map_preimage) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5128 | show "continuous_map (subtopology X K) (subtopology Y (f`K)) f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5129 | by (simp add: continuous_map_in_subtopology f image_mono that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5130 | show "closedin (subtopology Y (f`K)) (f ` K \<inter> C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5131 | using \<open>closedin Y C\<close> closedin_subtopology by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5132 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5133 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5134 | using \<open>k_space X\<close> k_space by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5135 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5136 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5137 | by (simp add: continuous_map_closedin) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5138 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5139 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5140 | lemma closed_map_into_k_space: | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5141 | assumes "k_space Y" and fim: "f \<in> (topspace X) \<rightarrow> topspace Y" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5142 | and f: "\<And>K. compactin Y K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5143 |                 \<Longrightarrow> closed_map(subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5144 | shows "closed_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5145 | unfolding closed_map_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5146 | proof (intro strip) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5147 | fix C | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5148 | assume "closedin X C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5149 | have "closedin (subtopology Y K) (K \<inter> f ` C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5150 | if "compactin Y K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5151 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5152 |     have eq: "K \<inter> f ` C = f ` ({x \<in> topspace X. f x \<in> K} \<inter> C)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5153 | using \<open>closedin X C\<close> closedin_subset by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5154 | show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5155 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5156 | by (metis (no_types, lifting) \<open>closedin X C\<close> closed_map_def closedin_subtopology f inf_commute that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5157 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5158 | then show "closedin Y (f ` C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5159 | using \<open>k_space Y\<close> unfolding k_space | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5160 | by (meson \<open>closedin X C\<close> closedin_subset order.trans fim funcset_image subset_image_iff) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5161 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5162 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5163 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5164 | text \<open>Essentially the same proof\<close> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5165 | lemma open_map_into_k_space: | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5166 | assumes "k_space Y" and fim: "f \<in> (topspace X) \<rightarrow> topspace Y" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5167 | and f: "\<And>K. compactin Y K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5168 |                  \<Longrightarrow> open_map (subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5169 | shows "open_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5170 | unfolding open_map_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5171 | proof (intro strip) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5172 | fix C | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5173 | assume "openin X C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5174 | have "openin (subtopology Y K) (K \<inter> f ` C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5175 | if "compactin Y K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5176 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5177 |     have eq: "K \<inter> f ` C = f ` ({x \<in> topspace X. f x \<in> K} \<inter> C)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5178 | using \<open>openin X C\<close> openin_subset by auto | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5179 | show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5180 | unfolding eq | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5181 | by (metis (no_types, lifting) \<open>openin X C\<close> open_map_def openin_subtopology f inf_commute that) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5182 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5183 | then show "openin Y (f ` C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5184 | using \<open>k_space Y\<close> unfolding k_space_open | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5185 | by (meson \<open>openin X C\<close> openin_subset order.trans fim funcset_image subset_image_iff) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5186 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5187 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5188 | lemma quotient_map_into_k_space: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5189 | fixes f :: "'a \<Rightarrow> 'b" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5190 | assumes "k_space Y" and cmf: "continuous_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5191 | and fim: "f ` (topspace X) = topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5192 | and f: "\<And>k. compactin Y k | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5193 |                  \<Longrightarrow> quotient_map (subtopology X {x \<in> topspace X. f x \<in> k})
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5194 | (subtopology Y k) f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5195 | shows "quotient_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5196 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5197 | have "closedin Y C" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5198 |     if "C \<subseteq> topspace Y" and K: "closedin X {x \<in> topspace X. f x \<in> C}" for C
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5199 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5200 | have "closedin (subtopology Y K) (K \<inter> C)" if "compactin Y K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5201 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5202 |       define Kf where "Kf \<equiv> {x \<in> topspace X. f x \<in> K}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5203 | have *: "K \<inter> C \<subseteq> topspace Y \<and> K \<inter> C \<subseteq> K" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5204 | using \<open>C \<subseteq> topspace Y\<close> by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5205 |       then have eq: "closedin (subtopology X Kf) (Kf \<inter> {x \<in> topspace X. f x \<in> C}) =
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5206 | closedin (subtopology Y K) (K \<inter> C)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5207 | using f [OF that] * unfolding quotient_map_closedin Kf_def | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5208 | by (smt (verit, ccfv_SIG) Collect_cong Int_def compactin_subset_topspace mem_Collect_eq that topspace_subtopology topspace_subtopology_subset) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5209 |       have dd: "{x \<in> topspace X \<inter> Kf. f x \<in> K \<inter> C} = Kf \<inter> {x \<in> topspace X. f x \<in> C}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5210 | by (auto simp add: Kf_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5211 |       have "closedin (subtopology X Kf) {x \<in> topspace X. x \<in> Kf \<and> f x \<in> K \<and> f x \<in> C}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5212 | using K closedin_subtopology by (fastforce simp add: Kf_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5213 | with K closedin_subtopology_Int_closed eq show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5214 | by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5215 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5216 | then show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5217 | using \<open>k_space Y\<close> that unfolding k_space by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5218 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5219 |   moreover have "closedin X {x \<in> topspace X. f x \<in> K}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5220 | if "K \<subseteq> topspace Y" "closedin Y K" for K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5221 | using that cmf continuous_map_closedin by fastforce | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5222 | ultimately show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5223 | unfolding quotient_map_closedin using fim by blast | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5224 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5225 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5226 | lemma quotient_map_into_k_space_eq: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5227 | assumes "k_space Y" "kc_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5228 | shows "quotient_map X Y f \<longleftrightarrow> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5229 | continuous_map X Y f \<and> f ` (topspace X) = topspace Y \<and> | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5230 | (\<forall>K. compactin Y K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5231 |               \<longrightarrow> quotient_map (subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5232 | using assms | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5233 | by (auto simp: kc_space_def intro: quotient_map_into_k_space quotient_map_restriction | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5234 | dest: quotient_imp_continuous_map quotient_imp_surjective_map) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5235 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5236 | lemma open_map_into_k_space_eq: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5237 | assumes "k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5238 | shows "open_map X Y f \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5239 | f \<in> (topspace X) \<rightarrow> topspace Y \<and> | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5240 | (\<forall>k. compactin Y k | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5241 |               \<longrightarrow> open_map (subtopology X {x \<in> topspace X. f x \<in> k}) (subtopology Y k) f)"
 | 
| 78517 
28c1f4f5335f
Numerous minor tweaks and simplifications
 paulson <lp15@cam.ac.uk> parents: 
78336diff
changeset | 5242 | using assms open_map_imp_subset_topspace open_map_into_k_space open_map_restriction by fastforce | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5243 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5244 | lemma closed_map_into_k_space_eq: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5245 | assumes "k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5246 | shows "closed_map X Y f \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5247 | f \<in> (topspace X) \<rightarrow> topspace Y \<and> | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5248 | (\<forall>k. compactin Y k | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5249 |               \<longrightarrow> closed_map (subtopology X {x \<in> topspace X. f x \<in> k}) (subtopology Y k) f)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5250 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5251 | proof | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5252 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5253 | by (simp add: closed_map_imp_subset_topspace closed_map_restriction) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5254 | show "?rhs \<Longrightarrow> ?lhs" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5255 | by (simp add: assms closed_map_into_k_space) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5256 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5257 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5258 | lemma proper_map_into_k_space: | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5259 | assumes "k_space Y" and fim: "f \<in> (topspace X) \<rightarrow> topspace Y" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5260 | and f: "\<And>K. compactin Y K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5261 |                  \<Longrightarrow> proper_map (subtopology X {x \<in> topspace X. f x \<in> K})
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5262 | (subtopology Y K) f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5263 | shows "proper_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5264 | proof - | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5265 | have "closed_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5266 | by (meson assms closed_map_into_k_space fim proper_map_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5267 | with f topspace_subtopology_subset show ?thesis | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5268 | apply (simp add: proper_map_alt) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5269 | by (smt (verit, best) Collect_cong compactin_absolute) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5270 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5271 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5272 | lemma proper_map_into_k_space_eq: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5273 | assumes "k_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5274 | shows "proper_map X Y f \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5275 | f \<in> (topspace X) \<rightarrow> topspace Y \<and> | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5276 | (\<forall>K. compactin Y K | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5277 |               \<longrightarrow> proper_map (subtopology X {x \<in> topspace X. f x \<in> K}) (subtopology Y K) f)"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5278 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5279 | proof | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5280 | show "?lhs \<Longrightarrow> ?rhs" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5281 | by (simp add: proper_map_imp_subset_topspace proper_map_restriction) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5282 | show "?rhs \<Longrightarrow> ?lhs" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5283 | by (simp add: assms funcset_image proper_map_into_k_space) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5284 | qed | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5285 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5286 | lemma compact_imp_proper_map: | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5287 | assumes "k_space Y" "kc_space Y" and fim: "f \<in> (topspace X) \<rightarrow> topspace Y" | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5288 | and f: "continuous_map X Y f \<or> kc_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5289 |     and comp: "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5290 | shows "proper_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5291 | proof (rule compact_imp_proper_map_gen) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5292 | fix S | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5293 | assume "S \<subseteq> topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5294 | and "\<And>K. compactin Y K \<Longrightarrow> compactin Y (S \<inter> K)" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5295 | with assms show "closedin Y S" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5296 | by (simp add: closedin_subset_topspace inf_commute k_space kc_space_def) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5297 | qed (use assms in auto) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5298 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5299 | lemma proper_eq_compact_map: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5300 | assumes "k_space Y" "kc_space Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5301 | and f: "continuous_map X Y f \<or> kc_space X" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5302 | shows "proper_map X Y f \<longleftrightarrow> | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5303 | f \<in> (topspace X) \<rightarrow> topspace Y \<and> | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5304 |              (\<forall>K. compactin Y K \<longrightarrow> compactin X {x \<in> topspace X. f x \<in> K})"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5305 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5306 | proof | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5307 | show "?lhs \<Longrightarrow> ?rhs" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5308 | using \<open>k_space Y\<close> compactin_proper_map_preimage proper_map_into_k_space_eq by blast | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5309 | qed (use assms compact_imp_proper_map in auto) | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5310 | |
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5311 | lemma compact_imp_perfect_map: | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5312 | assumes "k_space Y" "kc_space Y" and "f ` (topspace X) = topspace Y" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5313 | and "continuous_map X Y f" | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5314 |     and "\<And>K. compactin Y K \<Longrightarrow> compactin X {x \<in> topspace X. f x \<in> K}"
 | 
| 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5315 | shows "perfect_map X Y f" | 
| 78248 
740b23f1138a
EXPERIMENTAL replacement of f ` A <= B by f : A -> B in Analysis
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 5316 | by (simp add: assms compact_imp_proper_map perfect_map_def flip: image_subset_iff_funcset) | 
| 78093 
cec875dcc59e
Finally, the abstract metric space development
 paulson <lp15@cam.ac.uk> parents: 
78038diff
changeset | 5317 | |
| 77941 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5318 | end | 
| 
c35f06b0931e
new material ported from HOL Light's metric.ml
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5319 |