| 71137 |      1 | section \<open>Neighbourhood bases and Locally path-connected spaces\<close>
 | 
| 69945 |      2 | 
 | 
|  |      3 | theory Locally
 | 
|  |      4 |   imports
 | 
|  |      5 |     Path_Connected Function_Topology
 | 
|  |      6 | begin
 | 
|  |      7 | 
 | 
| 71137 |      8 | subsection\<open>Neighbourhood Bases\<close>
 | 
|  |      9 | 
 | 
|  |     10 | text \<open>Useful for "local" properties of various kinds\<close>
 | 
| 69945 |     11 | 
 | 
|  |     12 | definition neighbourhood_base_at where
 | 
|  |     13 |  "neighbourhood_base_at x P X \<equiv>
 | 
|  |     14 |         \<forall>W. openin X W \<and> x \<in> W
 | 
|  |     15 |             \<longrightarrow> (\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W)"
 | 
|  |     16 | 
 | 
|  |     17 | definition neighbourhood_base_of where
 | 
|  |     18 |  "neighbourhood_base_of P X \<equiv>
 | 
|  |     19 |         \<forall>x \<in> topspace X. neighbourhood_base_at x P X"
 | 
|  |     20 | 
 | 
|  |     21 | lemma neighbourhood_base_of:
 | 
|  |     22 |    "neighbourhood_base_of P X \<longleftrightarrow>
 | 
|  |     23 |         (\<forall>W x. openin X W \<and> x \<in> W
 | 
|  |     24 |           \<longrightarrow> (\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W))"
 | 
|  |     25 |   unfolding neighbourhood_base_at_def neighbourhood_base_of_def
 | 
|  |     26 |   using openin_subset by blast
 | 
|  |     27 | 
 | 
|  |     28 | lemma neighbourhood_base_at_mono:
 | 
|  |     29 |    "\<lbrakk>neighbourhood_base_at x P X; \<And>S. \<lbrakk>P S; x \<in> S\<rbrakk> \<Longrightarrow> Q S\<rbrakk> \<Longrightarrow> neighbourhood_base_at x Q X"
 | 
|  |     30 |   unfolding neighbourhood_base_at_def
 | 
|  |     31 |   by (meson subset_eq)
 | 
|  |     32 | 
 | 
|  |     33 | lemma neighbourhood_base_of_mono:
 | 
|  |     34 |    "\<lbrakk>neighbourhood_base_of P X; \<And>S. P S \<Longrightarrow> Q S\<rbrakk> \<Longrightarrow> neighbourhood_base_of Q X"
 | 
|  |     35 |   unfolding neighbourhood_base_of_def
 | 
|  |     36 |   using neighbourhood_base_at_mono by force
 | 
|  |     37 | 
 | 
|  |     38 | lemma open_neighbourhood_base_at:
 | 
|  |     39 |    "(\<And>S. \<lbrakk>P S; x \<in> S\<rbrakk> \<Longrightarrow> openin X S)
 | 
|  |     40 |         \<Longrightarrow> neighbourhood_base_at x P X \<longleftrightarrow> (\<forall>W. openin X W \<and> x \<in> W \<longrightarrow> (\<exists>U. P U \<and> x \<in> U \<and> U \<subseteq> W))"
 | 
|  |     41 |   unfolding neighbourhood_base_at_def
 | 
|  |     42 |   by (meson subsetD)
 | 
|  |     43 | 
 | 
|  |     44 | lemma open_neighbourhood_base_of:
 | 
|  |     45 |   "(\<forall>S. P S \<longrightarrow> openin X S)
 | 
|  |     46 |         \<Longrightarrow> neighbourhood_base_of P X \<longleftrightarrow> (\<forall>W x. openin X W \<and> x \<in> W \<longrightarrow> (\<exists>U. P U \<and> x \<in> U \<and> U \<subseteq> W))"
 | 
|  |     47 |   apply (simp add: neighbourhood_base_of, safe, blast)
 | 
|  |     48 |   by meson
 | 
|  |     49 | 
 | 
|  |     50 | lemma neighbourhood_base_of_open_subset:
 | 
|  |     51 |    "\<lbrakk>neighbourhood_base_of P X; openin X S\<rbrakk>
 | 
|  |     52 |         \<Longrightarrow> neighbourhood_base_of P (subtopology X S)"
 | 
|  |     53 |   apply (clarsimp simp add: neighbourhood_base_of openin_subtopology_alt image_def)
 | 
|  |     54 |   apply (rename_tac x V)
 | 
|  |     55 |   apply (drule_tac x="S \<inter> V" in spec)
 | 
|  |     56 |   apply (simp add: Int_ac)
 | 
|  |     57 |   by (metis IntI le_infI1 openin_Int)
 | 
|  |     58 | 
 | 
|  |     59 | lemma neighbourhood_base_of_topology_base:
 | 
|  |     60 |    "openin X = arbitrary union_of (\<lambda>W. W \<in> \<B>)
 | 
|  |     61 |         \<Longrightarrow> neighbourhood_base_of P X \<longleftrightarrow>
 | 
|  |     62 |              (\<forall>W x. W \<in> \<B> \<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))"
 | 
|  |     63 |   apply (auto simp: openin_topology_base_unique neighbourhood_base_of)
 | 
|  |     64 |   by (meson subset_trans)
 | 
|  |     65 | 
 | 
|  |     66 | lemma neighbourhood_base_at_unlocalized:
 | 
|  |     67 |   assumes "\<And>S T. \<lbrakk>P S; openin X T; x \<in> T; T \<subseteq> S\<rbrakk> \<Longrightarrow> P T"
 | 
|  |     68 |   shows "neighbourhood_base_at x P X
 | 
|  |     69 |      \<longleftrightarrow> (x \<in> topspace X \<longrightarrow> (\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> topspace X))"
 | 
|  |     70 |          (is "?lhs = ?rhs")
 | 
|  |     71 | proof
 | 
|  |     72 |   assume R: ?rhs show ?lhs
 | 
|  |     73 |     unfolding neighbourhood_base_at_def
 | 
|  |     74 |   proof clarify
 | 
|  |     75 |     fix W
 | 
|  |     76 |     assume "openin X W" "x \<in> W"
 | 
|  |     77 |     then have "x \<in> topspace X"
 | 
|  |     78 |       using openin_subset by blast
 | 
|  |     79 |     with R obtain U V where "openin X U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> topspace X"
 | 
|  |     80 |       by blast
 | 
|  |     81 |     then show "\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
 | 
|  |     82 |       by (metis IntI \<open>openin X W\<close> \<open>x \<in> W\<close> assms inf_le1 inf_le2 openin_Int)
 | 
|  |     83 |   qed
 | 
|  |     84 | qed (simp add: neighbourhood_base_at_def)
 | 
|  |     85 | 
 | 
|  |     86 | lemma neighbourhood_base_at_with_subset:
 | 
|  |     87 |    "\<lbrakk>openin X U; x \<in> U\<rbrakk>
 | 
|  |     88 |         \<Longrightarrow> (neighbourhood_base_at x P X \<longleftrightarrow> neighbourhood_base_at x (\<lambda>T. T \<subseteq> U \<and> P T) X)"
 | 
|  |     89 |   apply (simp add: neighbourhood_base_at_def)
 | 
|  |     90 |   apply (metis IntI Int_subset_iff openin_Int)
 | 
|  |     91 |   done
 | 
|  |     92 | 
 | 
|  |     93 | lemma neighbourhood_base_of_with_subset:
 | 
|  |     94 |    "neighbourhood_base_of P X \<longleftrightarrow> neighbourhood_base_of (\<lambda>t. t \<subseteq> topspace X \<and> P t) X"
 | 
|  |     95 |   using neighbourhood_base_at_with_subset
 | 
|  |     96 |   by (fastforce simp add: neighbourhood_base_of_def)
 | 
|  |     97 | 
 | 
|  |     98 | subsection\<open>Locally path-connected spaces\<close>
 | 
|  |     99 | 
 | 
|  |    100 | definition weakly_locally_path_connected_at
 | 
|  |    101 |   where "weakly_locally_path_connected_at x X \<equiv> neighbourhood_base_at x (path_connectedin X) X"
 | 
|  |    102 | 
 | 
|  |    103 | definition locally_path_connected_at
 | 
|  |    104 |   where "locally_path_connected_at x X \<equiv>
 | 
|  |    105 |     neighbourhood_base_at x (\<lambda>U. openin X U \<and> path_connectedin X U) X"
 | 
|  |    106 | 
 | 
|  |    107 | definition locally_path_connected_space
 | 
|  |    108 |   where "locally_path_connected_space X \<equiv> neighbourhood_base_of (path_connectedin X) X"
 | 
|  |    109 | 
 | 
|  |    110 | lemma locally_path_connected_space_alt:
 | 
|  |    111 |   "locally_path_connected_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>U. openin X U \<and> path_connectedin X U) X"
 | 
|  |    112 |   (is "?P = ?Q")
 | 
|  |    113 |   and locally_path_connected_space_eq_open_path_component_of:
 | 
|  |    114 |   "locally_path_connected_space X \<longleftrightarrow>
 | 
|  |    115 |         (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (Collect (path_component_of (subtopology X U) x)))"
 | 
|  |    116 |   (is "?P = ?R")
 | 
|  |    117 | proof -
 | 
|  |    118 |   have ?P if ?Q
 | 
|  |    119 |     using locally_path_connected_space_def neighbourhood_base_of_mono that by auto
 | 
|  |    120 |   moreover have ?R if P: ?P
 | 
|  |    121 |   proof -
 | 
|  |    122 |     show ?thesis
 | 
|  |    123 |     proof clarify
 | 
|  |    124 |       fix U y
 | 
|  |    125 |       assume "openin X U" "y \<in> U"
 | 
|  |    126 |       have "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> Collect (path_component_of (subtopology X U) y)"
 | 
|  |    127 |         if "path_component_of (subtopology X U) y x" for x
 | 
|  |    128 | 
 | 
|  |    129 |       proof -
 | 
|  |    130 |         have "x \<in> U"
 | 
|  |    131 |           using path_component_of_equiv that topspace_subtopology by fastforce
 | 
|  |    132 |         then have "\<exists>Ua. openin X Ua \<and> (\<exists>V. path_connectedin X V \<and> x \<in> Ua \<and> Ua \<subseteq> V \<and> V \<subseteq> U)"
 | 
|  |    133 |       using P
 | 
|  |    134 |       by (simp add: \<open>openin X U\<close> locally_path_connected_space_def neighbourhood_base_of)
 | 
|  |    135 |         then show ?thesis
 | 
|  |    136 |           by (metis dual_order.trans path_component_of_equiv path_component_of_maximal path_connectedin_subtopology subset_iff that)
 | 
|  |    137 |       qed
 | 
|  |    138 |       then show "openin X (Collect (path_component_of (subtopology X U) y))"
 | 
|  |    139 |         using openin_subopen by force
 | 
|  |    140 |     qed
 | 
|  |    141 |   qed
 | 
|  |    142 |   moreover have ?Q if ?R
 | 
|  |    143 |     using that
 | 
|  |    144 |     apply (simp add: open_neighbourhood_base_of)
 | 
|  |    145 |     by (metis mem_Collect_eq openin_subset path_component_of_refl path_connectedin_path_component_of path_connectedin_subtopology that topspace_subtopology_subset)
 | 
|  |    146 |   ultimately show "?P = ?Q" "?P = ?R"
 | 
|  |    147 |     by blast+
 | 
|  |    148 | qed
 | 
|  |    149 | 
 | 
|  |    150 | lemma locally_path_connected_space:
 | 
|  |    151 |    "locally_path_connected_space X
 | 
|  |    152 |    \<longleftrightarrow> (\<forall>V x. openin X V \<and> x \<in> V \<longrightarrow> (\<exists>U. openin X U \<and> path_connectedin X U \<and> x \<in> U \<and> U \<subseteq> V))"
 | 
|  |    153 |   by (simp add: locally_path_connected_space_alt open_neighbourhood_base_of)
 | 
|  |    154 | 
 | 
|  |    155 | lemma locally_path_connected_space_open_path_components:
 | 
|  |    156 |    "locally_path_connected_space X \<longleftrightarrow>
 | 
|  |    157 |         (\<forall>U c. openin X U \<and> c \<in> path_components_of(subtopology X U) \<longrightarrow> openin X c)"
 | 
| 71172 |    158 |   apply (auto simp: locally_path_connected_space_eq_open_path_component_of path_components_of_def)
 | 
| 69945 |    159 |   by (metis imageI inf.absorb_iff2 openin_closedin_eq)
 | 
|  |    160 | 
 | 
|  |    161 | lemma openin_path_component_of_locally_path_connected_space:
 | 
|  |    162 |    "locally_path_connected_space X \<Longrightarrow> openin X (Collect (path_component_of X x))"
 | 
|  |    163 |   apply (auto simp: locally_path_connected_space_eq_open_path_component_of)
 | 
|  |    164 |   by (metis openin_empty openin_topspace path_component_of_eq_empty subtopology_topspace)
 | 
|  |    165 | 
 | 
|  |    166 | lemma openin_path_components_of_locally_path_connected_space:
 | 
|  |    167 |    "\<lbrakk>locally_path_connected_space X; c \<in> path_components_of X\<rbrakk> \<Longrightarrow> openin X c"
 | 
|  |    168 |   apply (auto simp: locally_path_connected_space_eq_open_path_component_of)
 | 
|  |    169 |   by (metis (no_types, lifting) imageE openin_topspace path_components_of_def subtopology_topspace)
 | 
|  |    170 | 
 | 
|  |    171 | lemma closedin_path_components_of_locally_path_connected_space:
 | 
|  |    172 |    "\<lbrakk>locally_path_connected_space X; c \<in> path_components_of X\<rbrakk> \<Longrightarrow> closedin X c"
 | 
|  |    173 |   by (simp add: closedin_def complement_path_components_of_Union openin_path_components_of_locally_path_connected_space openin_clauses(3) path_components_of_subset)
 | 
|  |    174 | 
 | 
|  |    175 | lemma closedin_path_component_of_locally_path_connected_space:
 | 
|  |    176 |   assumes "locally_path_connected_space X"
 | 
|  |    177 |   shows "closedin X (Collect (path_component_of X x))"
 | 
|  |    178 | proof (cases "x \<in> topspace X")
 | 
|  |    179 |   case True
 | 
|  |    180 |   then show ?thesis
 | 
|  |    181 |     by (simp add: assms closedin_path_components_of_locally_path_connected_space path_component_in_path_components_of)
 | 
|  |    182 | next
 | 
|  |    183 |   case False
 | 
|  |    184 |   then show ?thesis
 | 
|  |    185 |     by (metis closedin_empty path_component_of_eq_empty)
 | 
|  |    186 | qed
 | 
|  |    187 | 
 | 
|  |    188 | lemma weakly_locally_path_connected_at:
 | 
|  |    189 |    "weakly_locally_path_connected_at x X \<longleftrightarrow>
 | 
|  |    190 |     (\<forall>V. openin X V \<and> x \<in> V
 | 
|  |    191 |           \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> U \<subseteq> V \<and>
 | 
|  |    192 |                   (\<forall>y \<in> U. \<exists>C. path_connectedin X C \<and> C \<subseteq> V \<and> x \<in> C \<and> y \<in> C)))"
 | 
|  |    193 |          (is "?lhs = ?rhs")
 | 
|  |    194 | proof
 | 
|  |    195 |   assume ?lhs then show ?rhs
 | 
|  |    196 |     apply (simp add: weakly_locally_path_connected_at_def neighbourhood_base_at_def)
 | 
|  |    197 |     by (meson order_trans subsetD)
 | 
|  |    198 | next
 | 
|  |    199 |   have *: "\<exists>V. path_connectedin X V \<and> U \<subseteq> V \<and> V \<subseteq> W"
 | 
|  |    200 |     if "(\<forall>y\<in>U. \<exists>C. path_connectedin X C \<and> C \<subseteq> W \<and> x \<in> C \<and> y \<in> C)"
 | 
|  |    201 |     for W U
 | 
|  |    202 |   proof (intro exI conjI)
 | 
|  |    203 |     let ?V = "(Collect (path_component_of (subtopology X W) x))"
 | 
|  |    204 |       show "path_connectedin X (Collect (path_component_of (subtopology X W) x))"
 | 
|  |    205 |         by (meson path_connectedin_path_component_of path_connectedin_subtopology)
 | 
|  |    206 |       show "U \<subseteq> ?V"
 | 
|  |    207 |         by (auto simp: path_component_of path_connectedin_subtopology that)
 | 
|  |    208 |       show "?V \<subseteq> W"
 | 
|  |    209 |         by (meson path_connectedin_path_component_of path_connectedin_subtopology)
 | 
|  |    210 |     qed
 | 
|  |    211 |   assume ?rhs
 | 
|  |    212 |   then show ?lhs
 | 
|  |    213 |     unfolding weakly_locally_path_connected_at_def neighbourhood_base_at_def by (metis "*")
 | 
|  |    214 | qed
 | 
|  |    215 | 
 | 
|  |    216 | lemma locally_path_connected_space_im_kleinen:
 | 
|  |    217 |    "locally_path_connected_space X \<longleftrightarrow>
 | 
|  |    218 |       (\<forall>V x. openin X V \<and> x \<in> V
 | 
|  |    219 |              \<longrightarrow> (\<exists>U. openin X U \<and>
 | 
|  |    220 |                     x \<in> U \<and> U \<subseteq> V \<and>
 | 
|  |    221 |                     (\<forall>y \<in> U. \<exists>c. path_connectedin X c \<and>
 | 
|  |    222 |                                 c \<subseteq> V \<and> x \<in> c \<and> y \<in> c)))"
 | 
|  |    223 |   apply (simp add: locally_path_connected_space_def neighbourhood_base_of_def)
 | 
|  |    224 |   apply (simp add: weakly_locally_path_connected_at flip: weakly_locally_path_connected_at_def)
 | 
|  |    225 |   using openin_subset apply force
 | 
|  |    226 |   done
 | 
|  |    227 | 
 | 
|  |    228 | lemma locally_path_connected_space_open_subset:
 | 
|  |    229 |    "\<lbrakk>locally_path_connected_space X; openin X s\<rbrakk>
 | 
|  |    230 |         \<Longrightarrow> locally_path_connected_space (subtopology X s)"
 | 
|  |    231 |   apply (simp add: locally_path_connected_space_def neighbourhood_base_of openin_open_subtopology path_connectedin_subtopology)
 | 
|  |    232 |   by (meson order_trans)
 | 
|  |    233 | 
 | 
|  |    234 | lemma locally_path_connected_space_quotient_map_image:
 | 
|  |    235 |   assumes f: "quotient_map X Y f" and X: "locally_path_connected_space X"
 | 
|  |    236 |   shows "locally_path_connected_space Y"
 | 
|  |    237 |   unfolding locally_path_connected_space_open_path_components
 | 
|  |    238 | proof clarify
 | 
|  |    239 |   fix V C
 | 
|  |    240 |   assume V: "openin Y V" and c: "C \<in> path_components_of (subtopology Y V)"
 | 
|  |    241 |   then have sub: "C \<subseteq> topspace Y"
 | 
|  |    242 |     using path_connectedin_path_components_of path_connectedin_subset_topspace path_connectedin_subtopology by blast
 | 
|  |    243 |   have "openin X {x \<in> topspace X. f x \<in> C}"
 | 
|  |    244 |   proof (subst openin_subopen, clarify)
 | 
|  |    245 |     fix x
 | 
|  |    246 |     assume x: "x \<in> topspace X" and "f x \<in> C"
 | 
|  |    247 |     let ?T = "Collect (path_component_of (subtopology X {z \<in> topspace X. f z \<in> V}) x)"
 | 
|  |    248 |     show "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace X. f x \<in> C}"
 | 
|  |    249 |     proof (intro exI conjI)
 | 
|  |    250 |       have "\<exists>U. openin X U \<and> ?T \<in> path_components_of (subtopology X U)"
 | 
|  |    251 |       proof (intro exI conjI)
 | 
|  |    252 |         show "openin X ({z \<in> topspace X. f z \<in> V})"
 | 
|  |    253 |           using V f openin_subset quotient_map_def by fastforce
 | 
|  |    254 |         show "Collect (path_component_of (subtopology X {z \<in> topspace X. f z \<in> V}) x)
 | 
|  |    255 |         \<in> path_components_of (subtopology X {z \<in> topspace X. f z \<in> V})"
 | 
|  |    256 |           by (metis (no_types, lifting) Int_iff \<open>f x \<in> C\<close> c mem_Collect_eq path_component_in_path_components_of path_components_of_subset topspace_subtopology topspace_subtopology_subset x)
 | 
|  |    257 |       qed
 | 
|  |    258 |       with X show "openin X ?T"
 | 
|  |    259 |         using locally_path_connected_space_open_path_components by blast
 | 
|  |    260 |       show x: "x \<in> ?T"
 | 
|  |    261 |         using V \<open>f x \<in> C\<close> c openin_subset path_component_of_equiv path_components_of_subset topspace_subtopology topspace_subtopology_subset x
 | 
|  |    262 |         by fastforce
 | 
|  |    263 |       have "f ` ?T \<subseteq> C"
 | 
|  |    264 |       proof (rule path_components_of_maximal)
 | 
|  |    265 |         show "C \<in> path_components_of (subtopology Y V)"
 | 
|  |    266 |           by (simp add: c)
 | 
|  |    267 |         show "path_connectedin (subtopology Y V) (f ` ?T)"
 | 
|  |    268 |         proof -
 | 
|  |    269 |           have "quotient_map (subtopology X {a \<in> topspace X. f a \<in> V}) (subtopology Y V) f"
 | 
|  |    270 |             by (simp add: V f quotient_map_restriction)
 | 
|  |    271 |           then show ?thesis
 | 
|  |    272 |             by (meson path_connectedin_continuous_map_image path_connectedin_path_component_of quotient_imp_continuous_map)
 | 
|  |    273 |         qed
 | 
|  |    274 |         show "\<not> disjnt C (f ` ?T)"
 | 
|  |    275 |           by (metis (no_types, lifting) \<open>f x \<in> C\<close> x disjnt_iff image_eqI)
 | 
|  |    276 |       qed
 | 
|  |    277 |       then show "?T \<subseteq> {x \<in> topspace X. f x \<in> C}"
 | 
| 71172 |    278 |         by (force simp: path_component_of_equiv)
 | 
| 69945 |    279 |     qed
 | 
|  |    280 |   qed
 | 
|  |    281 |   then show "openin Y C"
 | 
|  |    282 |     using f sub by (simp add: quotient_map_def)
 | 
|  |    283 | qed
 | 
|  |    284 | 
 | 
|  |    285 | lemma homeomorphic_locally_path_connected_space:
 | 
|  |    286 |   assumes "X homeomorphic_space Y"
 | 
|  |    287 |   shows "locally_path_connected_space X \<longleftrightarrow> locally_path_connected_space Y"
 | 
|  |    288 | proof -
 | 
|  |    289 |   obtain f g where "homeomorphic_maps X Y f g"
 | 
|  |    290 |     using assms homeomorphic_space_def by fastforce
 | 
|  |    291 |   then show ?thesis
 | 
|  |    292 |     by (metis (no_types) homeomorphic_map_def homeomorphic_maps_map locally_path_connected_space_quotient_map_image)
 | 
|  |    293 | qed
 | 
|  |    294 | 
 | 
|  |    295 | lemma locally_path_connected_space_retraction_map_image:
 | 
|  |    296 |    "\<lbrakk>retraction_map X Y r; locally_path_connected_space X\<rbrakk>
 | 
|  |    297 |         \<Longrightarrow> locally_path_connected_space Y"
 | 
|  |    298 |   using Abstract_Topology.retraction_imp_quotient_map locally_path_connected_space_quotient_map_image by blast
 | 
|  |    299 | 
 | 
|  |    300 | lemma locally_path_connected_space_euclideanreal: "locally_path_connected_space euclideanreal"
 | 
|  |    301 |   unfolding locally_path_connected_space_def neighbourhood_base_of
 | 
|  |    302 |   proof clarsimp
 | 
|  |    303 |   fix W and x :: "real"
 | 
|  |    304 |   assume "open W" "x \<in> W"
 | 
|  |    305 |   then obtain e where "e > 0" and e: "\<And>x'. \<bar>x' - x\<bar> < e \<longrightarrow> x' \<in> W"
 | 
|  |    306 |     by (auto simp: open_real)
 | 
|  |    307 |   then show "\<exists>U. open U \<and> (\<exists>V. path_connected V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W)"
 | 
|  |    308 |     by (force intro!: convex_imp_path_connected exI [where x = "{x-e<..<x+e}"])
 | 
|  |    309 | qed
 | 
|  |    310 | 
 | 
|  |    311 | lemma locally_path_connected_space_discrete_topology:
 | 
|  |    312 |    "locally_path_connected_space (discrete_topology U)"
 | 
|  |    313 |   using locally_path_connected_space_open_path_components by fastforce
 | 
|  |    314 | 
 | 
|  |    315 | lemma path_component_eq_connected_component_of:
 | 
|  |    316 |   assumes "locally_path_connected_space X"
 | 
|  |    317 |   shows "(path_component_of_set X x = connected_component_of_set X x)"
 | 
|  |    318 | proof (cases "x \<in> topspace X")
 | 
|  |    319 |   case True
 | 
|  |    320 |   then show ?thesis
 | 
|  |    321 |     using connectedin_connected_component_of [of X x]
 | 
|  |    322 |     apply (clarsimp simp add: connectedin_def connected_space_clopen_in topspace_subtopology_subset cong: conj_cong)
 | 
|  |    323 |     apply (drule_tac x="path_component_of_set X x" in spec)
 | 
|  |    324 |     by (metis assms closedin_closed_subtopology closedin_connected_component_of closedin_path_component_of_locally_path_connected_space inf.absorb_iff2 inf.orderE openin_path_component_of_locally_path_connected_space openin_subtopology path_component_of_eq_empty path_component_subset_connected_component_of)
 | 
|  |    325 | next
 | 
|  |    326 |   case False
 | 
|  |    327 |   then show ?thesis
 | 
|  |    328 |     using connected_component_of_eq_empty path_component_of_eq_empty by fastforce
 | 
|  |    329 | qed
 | 
|  |    330 | 
 | 
|  |    331 | lemma path_components_eq_connected_components_of:
 | 
|  |    332 |    "locally_path_connected_space X \<Longrightarrow> (path_components_of X = connected_components_of X)"
 | 
|  |    333 |   by (simp add: path_components_of_def connected_components_of_def image_def path_component_eq_connected_component_of)
 | 
|  |    334 | 
 | 
|  |    335 | lemma path_connected_eq_connected_space:
 | 
|  |    336 |    "locally_path_connected_space X
 | 
|  |    337 |          \<Longrightarrow> path_connected_space X \<longleftrightarrow> connected_space X"
 | 
|  |    338 |   by (metis connected_components_of_subset_sing path_components_eq_connected_components_of path_components_of_subset_singleton)
 | 
|  |    339 | 
 | 
|  |    340 | lemma locally_path_connected_space_product_topology:
 | 
|  |    341 |    "locally_path_connected_space(product_topology X I) \<longleftrightarrow>
 | 
|  |    342 |         topspace(product_topology X I) = {} \<or>
 | 
|  |    343 |         finite {i. i \<in> I \<and> ~path_connected_space(X i)} \<and>
 | 
|  |    344 |         (\<forall>i \<in> I. locally_path_connected_space(X i))"
 | 
|  |    345 |     (is "?lhs \<longleftrightarrow> ?empty \<or> ?rhs")
 | 
|  |    346 | proof (cases ?empty)
 | 
|  |    347 |   case True
 | 
|  |    348 |   then show ?thesis
 | 
|  |    349 |     by (simp add: locally_path_connected_space_def neighbourhood_base_of openin_closedin_eq)
 | 
|  |    350 | next
 | 
|  |    351 |   case False
 | 
|  |    352 |   then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
 | 
|  |    353 |     by auto
 | 
|  |    354 |   have ?rhs if L: ?lhs
 | 
|  |    355 |   proof -
 | 
|  |    356 |     obtain U C where U: "openin (product_topology X I) U"
 | 
|  |    357 |       and V: "path_connectedin (product_topology X I) C"
 | 
|  |    358 |       and "z \<in> U" "U \<subseteq> C" and Csub: "C \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
 | 
|  |    359 |       using L apply (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of)
 | 
|  |    360 |       by (metis openin_topspace topspace_product_topology z)
 | 
|  |    361 |     then obtain V where finV: "finite {i \<in> I. V i \<noteq> topspace (X i)}"
 | 
|  |    362 |       and XV: "\<And>i. i\<in>I \<Longrightarrow> openin (X i) (V i)" and "z \<in> Pi\<^sub>E I V" and subU: "Pi\<^sub>E I V \<subseteq> U"
 | 
|  |    363 |       by (force simp: openin_product_topology_alt)
 | 
|  |    364 |     show ?thesis
 | 
|  |    365 |     proof (intro conjI ballI)
 | 
|  |    366 |       have "path_connected_space (X i)" if "i \<in> I" "V i = topspace (X i)" for i
 | 
|  |    367 |       proof -
 | 
|  |    368 |         have pc: "path_connectedin (X i) ((\<lambda>x. x i) ` C)"
 | 
|  |    369 |           apply (rule path_connectedin_continuous_map_image [OF _ V])
 | 
|  |    370 |           by (simp add: continuous_map_product_projection \<open>i \<in> I\<close>)
 | 
|  |    371 |         moreover have "((\<lambda>x. x i) ` C) = topspace (X i)"
 | 
|  |    372 |         proof
 | 
|  |    373 |           show "(\<lambda>x. x i) ` C \<subseteq> topspace (X i)"
 | 
|  |    374 |             by (simp add: pc path_connectedin_subset_topspace)
 | 
|  |    375 |           have "V i \<subseteq> (\<lambda>x. x i) ` (\<Pi>\<^sub>E i\<in>I. V i)"
 | 
|  |    376 |             by (metis \<open>z \<in> Pi\<^sub>E I V\<close> empty_iff image_projection_PiE order_refl that(1))
 | 
|  |    377 |           also have "\<dots> \<subseteq> (\<lambda>x. x i) ` U"
 | 
|  |    378 |             using subU by blast
 | 
|  |    379 |           finally show "topspace (X i) \<subseteq> (\<lambda>x. x i) ` C"
 | 
|  |    380 |             using \<open>U \<subseteq> C\<close> that by blast
 | 
|  |    381 |         qed
 | 
|  |    382 |         ultimately show ?thesis
 | 
|  |    383 |           by (simp add: path_connectedin_topspace)
 | 
|  |    384 |       qed
 | 
|  |    385 |       then have "{i \<in> I. \<not> path_connected_space (X i)} \<subseteq> {i \<in> I. V i \<noteq> topspace (X i)}"
 | 
|  |    386 |         by blast
 | 
|  |    387 |       with finV show "finite {i \<in> I. \<not> path_connected_space (X i)}"
 | 
|  |    388 |         using finite_subset by blast
 | 
|  |    389 |     next
 | 
|  |    390 |       show "locally_path_connected_space (X i)" if "i \<in> I" for i
 | 
|  |    391 |         apply (rule locally_path_connected_space_quotient_map_image [OF _ L, where f = "\<lambda>x. x i"])
 | 
|  |    392 |         by (metis False Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection that)
 | 
|  |    393 |     qed
 | 
|  |    394 |   qed
 | 
|  |    395 |   moreover have ?lhs if R: ?rhs
 | 
|  |    396 |   proof (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of)
 | 
|  |    397 |     fix F z
 | 
|  |    398 |     assume "openin (product_topology X I) F" and "z \<in> F"
 | 
|  |    399 |     then obtain W where finW: "finite {i \<in> I. W i \<noteq> topspace (X i)}"
 | 
|  |    400 |             and opeW: "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (W i)" and "z \<in> Pi\<^sub>E I W" "Pi\<^sub>E I W \<subseteq> F"
 | 
|  |    401 |       by (auto simp: openin_product_topology_alt)
 | 
|  |    402 |     have "\<forall>i \<in> I. \<exists>U C. openin (X i) U \<and> path_connectedin (X i) C \<and> z i \<in> U \<and> U \<subseteq> C \<and> C \<subseteq> W i \<and>
 | 
|  |    403 |                         (W i = topspace (X i) \<and>
 | 
|  |    404 |                          path_connected_space (X i) \<longrightarrow> U = topspace (X i) \<and> C = topspace (X i))"
 | 
|  |    405 |           (is "\<forall>i \<in> I. ?\<Phi> i")
 | 
|  |    406 |     proof
 | 
|  |    407 |       fix i assume "i \<in> I"
 | 
|  |    408 |       have "locally_path_connected_space (X i)"
 | 
|  |    409 |         by (simp add: R \<open>i \<in> I\<close>)
 | 
|  |    410 |       moreover have "openin (X i) (W i) " "z i \<in> W i"
 | 
|  |    411 |         using \<open>z \<in> Pi\<^sub>E I W\<close> opeW \<open>i \<in> I\<close> by auto
 | 
|  |    412 |       ultimately obtain U C where UC: "openin (X i) U" "path_connectedin (X i) C" "z i \<in> U" "U \<subseteq> C" "C \<subseteq> W i"
 | 
|  |    413 |         using \<open>i \<in> I\<close> by (force simp: locally_path_connected_space_def neighbourhood_base_of)
 | 
|  |    414 |       show "?\<Phi> i"
 | 
|  |    415 |       proof (cases "W i = topspace (X i) \<and> path_connected_space(X i)")
 | 
|  |    416 |         case True
 | 
|  |    417 |         then show ?thesis
 | 
|  |    418 |           using \<open>z i \<in> W i\<close> path_connectedin_topspace by blast
 | 
|  |    419 |       next
 | 
|  |    420 |         case False
 | 
|  |    421 |         then show ?thesis
 | 
|  |    422 |           by (meson UC)
 | 
|  |    423 |       qed
 | 
|  |    424 |     qed
 | 
|  |    425 |     then obtain U C where
 | 
|  |    426 |       *: "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (U i) \<and> path_connectedin (X i) (C i) \<and> z i \<in> (U i) \<and> (U i) \<subseteq> (C i) \<and> (C i) \<subseteq> W i \<and>
 | 
|  |    427 |                         (W i = topspace (X i) \<and> path_connected_space (X i)
 | 
|  |    428 |                          \<longrightarrow> (U i) = topspace (X i) \<and> (C i) = topspace (X i))"
 | 
|  |    429 |       by metis
 | 
|  |    430 |     let ?A = "{i \<in> I. \<not> path_connected_space (X i)} \<union> {i \<in> I. W i \<noteq> topspace (X i)}"
 | 
|  |    431 |     have "{i \<in> I. U i \<noteq> topspace (X i)} \<subseteq> ?A"
 | 
|  |    432 |       by (clarsimp simp add: "*")
 | 
|  |    433 |     moreover have "finite ?A"
 | 
|  |    434 |       by (simp add: that finW)
 | 
|  |    435 |     ultimately have "finite {i \<in> I. U i \<noteq> topspace (X i)}"
 | 
|  |    436 |       using finite_subset by auto
 | 
|  |    437 |     then have "openin (product_topology X I) (Pi\<^sub>E I U)"
 | 
|  |    438 |       using * by (simp add: openin_PiE_gen)
 | 
|  |    439 |     then show "\<exists>U. openin (product_topology X I) U \<and>
 | 
|  |    440 |             (\<exists>V. path_connectedin (product_topology X I) V \<and> z \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> F)"
 | 
|  |    441 |       apply (rule_tac x="PiE I U" in exI, simp)
 | 
|  |    442 |       apply (rule_tac x="PiE I C" in exI)
 | 
|  |    443 |       using \<open>z \<in> Pi\<^sub>E I W\<close> \<open>Pi\<^sub>E I W \<subseteq> F\<close> *
 | 
|  |    444 |       apply (simp add: path_connectedin_PiE subset_PiE PiE_iff PiE_mono dual_order.trans)
 | 
|  |    445 |       done
 | 
|  |    446 |   qed
 | 
|  |    447 |   ultimately show ?thesis
 | 
|  |    448 |     using False by blast
 | 
|  |    449 | qed
 | 
|  |    450 | 
 | 
|  |    451 | end
 |