src/HOL/Analysis/Locally.thy
author wenzelm
Tue, 20 Jun 2023 22:57:34 +0200
changeset 78183 8d57ed9e27a7
parent 78037 37894dff0111
child 78250 400aecdfd71f
permissions -rw-r--r--
store heaps within database server;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71137
nipkow
parents: 69945
diff changeset
     1
section \<open>Neighbourhood bases and Locally path-connected spaces\<close>
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
theory Locally
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
  imports
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
     5
    Path_Connected Function_Topology Sum_Topology
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
begin
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
71137
nipkow
parents: 69945
diff changeset
     8
subsection\<open>Neighbourhood Bases\<close>
nipkow
parents: 69945
diff changeset
     9
nipkow
parents: 69945
diff changeset
    10
text \<open>Useful for "local" properties of various kinds\<close>
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
definition neighbourhood_base_at where
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
 "neighbourhood_base_at x P X \<equiv>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
        \<forall>W. openin X W \<and> x \<in> W
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
            \<longrightarrow> (\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
definition neighbourhood_base_of where
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
 "neighbourhood_base_of P X \<equiv>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
        \<forall>x \<in> topspace X. neighbourhood_base_at x P X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
lemma neighbourhood_base_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
   "neighbourhood_base_of P X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
        (\<forall>W x. openin X W \<and> x \<in> W
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
          \<longrightarrow> (\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  unfolding neighbourhood_base_at_def neighbourhood_base_of_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  using openin_subset by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
lemma neighbourhood_base_at_mono:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    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"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  unfolding neighbourhood_base_at_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
  by (meson subset_eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
lemma neighbourhood_base_of_mono:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
   "\<lbrakk>neighbourhood_base_of P X; \<And>S. P S \<Longrightarrow> Q S\<rbrakk> \<Longrightarrow> neighbourhood_base_of Q X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
  unfolding neighbourhood_base_of_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  using neighbourhood_base_at_mono by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
lemma open_neighbourhood_base_at:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
   "(\<And>S. \<lbrakk>P S; x \<in> S\<rbrakk> \<Longrightarrow> openin X S)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    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))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
  unfolding neighbourhood_base_at_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  by (meson subsetD)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
lemma open_neighbourhood_base_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
  "(\<forall>S. P S \<longrightarrow> openin X S)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    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))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
  apply (simp add: neighbourhood_base_of, safe, blast)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  by meson
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
lemma neighbourhood_base_of_open_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
   "\<lbrakk>neighbourhood_base_of P X; openin X S\<rbrakk>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
        \<Longrightarrow> neighbourhood_base_of P (subtopology X S)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
  apply (clarsimp simp add: neighbourhood_base_of openin_subtopology_alt image_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
  apply (rename_tac x V)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
  apply (drule_tac x="S \<inter> V" in spec)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
  apply (simp add: Int_ac)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
  by (metis IntI le_infI1 openin_Int)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
lemma neighbourhood_base_of_topology_base:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
   "openin X = arbitrary union_of (\<lambda>W. W \<in> \<B>)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
        \<Longrightarrow> neighbourhood_base_of P X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    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))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  apply (auto simp: openin_topology_base_unique neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  by (meson subset_trans)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
lemma neighbourhood_base_at_unlocalized:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  assumes "\<And>S T. \<lbrakk>P S; openin X T; x \<in> T; T \<subseteq> S\<rbrakk> \<Longrightarrow> P T"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  shows "neighbourhood_base_at x P X
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    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))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
         (is "?lhs = ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
  assume R: ?rhs show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
    unfolding neighbourhood_base_at_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
    fix W
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
    assume "openin X W" "x \<in> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    then have "x \<in> topspace X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
      using openin_subset by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
    with R obtain U V where "openin X U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> topspace X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
      by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
    then show "\<exists>U V. openin X U \<and> P V \<and> x \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
      by (metis IntI \<open>openin X W\<close> \<open>x \<in> W\<close> assms inf_le1 inf_le2 openin_Int)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
qed (simp add: neighbourhood_base_at_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
lemma neighbourhood_base_at_with_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
   "\<lbrakk>openin X U; x \<in> U\<rbrakk>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
        \<Longrightarrow> (neighbourhood_base_at x P X \<longleftrightarrow> neighbourhood_base_at x (\<lambda>T. T \<subseteq> U \<and> P T) X)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
  apply (simp add: neighbourhood_base_at_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  apply (metis IntI Int_subset_iff openin_Int)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
lemma neighbourhood_base_of_with_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
   "neighbourhood_base_of P X \<longleftrightarrow> neighbourhood_base_of (\<lambda>t. t \<subseteq> topspace X \<and> P t) X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  using neighbourhood_base_at_with_subset
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  by (fastforce simp add: neighbourhood_base_of_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
subsection\<open>Locally path-connected spaces\<close>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
definition weakly_locally_path_connected_at
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  where "weakly_locally_path_connected_at x X \<equiv> neighbourhood_base_at x (path_connectedin X) X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
definition locally_path_connected_at
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  where "locally_path_connected_at x X \<equiv>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
    neighbourhood_base_at x (\<lambda>U. openin X U \<and> path_connectedin X U) X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
definition locally_path_connected_space
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  where "locally_path_connected_space X \<equiv> neighbourhood_base_of (path_connectedin X) X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
lemma locally_path_connected_space_alt:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
  "locally_path_connected_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>U. openin X U \<and> path_connectedin X U) X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  (is "?P = ?Q")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
  and locally_path_connected_space_eq_open_path_component_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  "locally_path_connected_space X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
        (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (Collect (path_component_of (subtopology X U) x)))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  (is "?P = ?R")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
  have ?P if ?Q
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
    using locally_path_connected_space_def neighbourhood_base_of_mono that by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  moreover have ?R if P: ?P
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
  proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
    show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
    proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
      fix U y
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
      assume "openin X U" "y \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
      have "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> Collect (path_component_of (subtopology X U) y)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
        if "path_component_of (subtopology X U) y x" for x
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
      proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
        have "x \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
          using path_component_of_equiv that topspace_subtopology by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
      using P
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
      by (simp add: \<open>openin X U\<close> locally_path_connected_space_def neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
        then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
          by (metis dual_order.trans path_component_of_equiv path_component_of_maximal path_connectedin_subtopology subset_iff that)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
      then show "openin X (Collect (path_component_of (subtopology X U) y))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
        using openin_subopen by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  moreover have ?Q if ?R
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
    using that
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
    apply (simp add: open_neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
    by (metis mem_Collect_eq openin_subset path_component_of_refl path_connectedin_path_component_of path_connectedin_subtopology that topspace_subtopology_subset)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
  ultimately show "?P = ?Q" "?P = ?R"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
    by blast+
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
lemma locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
   "locally_path_connected_space X
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
  by (simp add: locally_path_connected_space_alt open_neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
lemma locally_path_connected_space_open_path_components:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
   "locally_path_connected_space X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
        (\<forall>U c. openin X U \<and> c \<in> path_components_of(subtopology X U) \<longrightarrow> openin X c)"
71172
nipkow
parents: 71137
diff changeset
   158
  apply (auto simp: locally_path_connected_space_eq_open_path_component_of path_components_of_def)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
  by (metis imageI inf.absorb_iff2 openin_closedin_eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
lemma openin_path_component_of_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
   "locally_path_connected_space X \<Longrightarrow> openin X (Collect (path_component_of X x))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  apply (auto simp: locally_path_connected_space_eq_open_path_component_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
  by (metis openin_empty openin_topspace path_component_of_eq_empty subtopology_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
lemma openin_path_components_of_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
   "\<lbrakk>locally_path_connected_space X; c \<in> path_components_of X\<rbrakk> \<Longrightarrow> openin X c"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  apply (auto simp: locally_path_connected_space_eq_open_path_component_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  by (metis (no_types, lifting) imageE openin_topspace path_components_of_def subtopology_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
lemma closedin_path_components_of_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
   "\<lbrakk>locally_path_connected_space X; c \<in> path_components_of X\<rbrakk> \<Longrightarrow> closedin X c"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
lemma closedin_path_component_of_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
  assumes "locally_path_connected_space X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  shows "closedin X (Collect (path_component_of X x))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
proof (cases "x \<in> topspace X")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
    by (simp add: assms closedin_path_components_of_locally_path_connected_space path_component_in_path_components_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
    by (metis closedin_empty path_component_of_eq_empty)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
lemma weakly_locally_path_connected_at:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
   "weakly_locally_path_connected_at x X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
    (\<forall>V. openin X V \<and> x \<in> V
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
          \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> U \<subseteq> V \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
                  (\<forall>y \<in> U. \<exists>C. path_connectedin X C \<and> C \<subseteq> V \<and> x \<in> C \<and> y \<in> C)))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
         (is "?lhs = ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
  assume ?lhs then show ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
    apply (simp add: weakly_locally_path_connected_at_def neighbourhood_base_at_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
    by (meson order_trans subsetD)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  have *: "\<exists>V. path_connectedin X V \<and> U \<subseteq> V \<and> V \<subseteq> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    if "(\<forall>y\<in>U. \<exists>C. path_connectedin X C \<and> C \<subseteq> W \<and> x \<in> C \<and> y \<in> C)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
    for W U
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
  proof (intro exI conjI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
    let ?V = "(Collect (path_component_of (subtopology X W) x))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
      show "path_connectedin X (Collect (path_component_of (subtopology X W) x))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
        by (meson path_connectedin_path_component_of path_connectedin_subtopology)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
      show "U \<subseteq> ?V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
        by (auto simp: path_component_of path_connectedin_subtopology that)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
      show "?V \<subseteq> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
        by (meson path_connectedin_path_component_of path_connectedin_subtopology)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
  assume ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
  then show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
    unfolding weakly_locally_path_connected_at_def neighbourhood_base_at_def by (metis "*")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
lemma locally_path_connected_space_im_kleinen:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
   "locally_path_connected_space X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
      (\<forall>V x. openin X V \<and> x \<in> V
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
             \<longrightarrow> (\<exists>U. openin X U \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
                    x \<in> U \<and> U \<subseteq> V \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
                    (\<forall>y \<in> U. \<exists>c. path_connectedin X c \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
                                c \<subseteq> V \<and> x \<in> c \<and> y \<in> c)))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  apply (simp add: locally_path_connected_space_def neighbourhood_base_of_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
  apply (simp add: weakly_locally_path_connected_at flip: weakly_locally_path_connected_at_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
  using openin_subset apply force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
lemma locally_path_connected_space_open_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
   "\<lbrakk>locally_path_connected_space X; openin X s\<rbrakk>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
        \<Longrightarrow> locally_path_connected_space (subtopology X s)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
  apply (simp add: locally_path_connected_space_def neighbourhood_base_of openin_open_subtopology path_connectedin_subtopology)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
  by (meson order_trans)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
lemma locally_path_connected_space_quotient_map_image:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
  assumes f: "quotient_map X Y f" and X: "locally_path_connected_space X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
  shows "locally_path_connected_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
  unfolding locally_path_connected_space_open_path_components
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
  fix V C
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
  assume V: "openin Y V" and c: "C \<in> path_components_of (subtopology Y V)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
  then have sub: "C \<subseteq> topspace Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
    using path_connectedin_path_components_of path_connectedin_subset_topspace path_connectedin_subtopology by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
  have "openin X {x \<in> topspace X. f x \<in> C}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
  proof (subst openin_subopen, clarify)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
    fix x
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
    assume x: "x \<in> topspace X" and "f x \<in> C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
    let ?T = "Collect (path_component_of (subtopology X {z \<in> topspace X. f z \<in> V}) x)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
    show "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace X. f x \<in> C}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
    proof (intro exI conjI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
      have "\<exists>U. openin X U \<and> ?T \<in> path_components_of (subtopology X U)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
      proof (intro exI conjI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
        show "openin X ({z \<in> topspace X. f z \<in> V})"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
          using V f openin_subset quotient_map_def by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
        show "Collect (path_component_of (subtopology X {z \<in> topspace X. f z \<in> V}) x)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
        \<in> path_components_of (subtopology X {z \<in> topspace X. f z \<in> V})"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
      with X show "openin X ?T"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
        using locally_path_connected_space_open_path_components by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
      show x: "x \<in> ?T"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
        by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
      have "f ` ?T \<subseteq> C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
      proof (rule path_components_of_maximal)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
        show "C \<in> path_components_of (subtopology Y V)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
          by (simp add: c)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
        show "path_connectedin (subtopology Y V) (f ` ?T)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
        proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
          have "quotient_map (subtopology X {a \<in> topspace X. f a \<in> V}) (subtopology Y V) f"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
            by (simp add: V f quotient_map_restriction)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
          then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
            by (meson path_connectedin_continuous_map_image path_connectedin_path_component_of quotient_imp_continuous_map)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
        show "\<not> disjnt C (f ` ?T)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
          by (metis (no_types, lifting) \<open>f x \<in> C\<close> x disjnt_iff image_eqI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
      then show "?T \<subseteq> {x \<in> topspace X. f x \<in> C}"
71172
nipkow
parents: 71137
diff changeset
   278
        by (force simp: path_component_of_equiv)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
  then show "openin Y C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
    using f sub by (simp add: quotient_map_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
lemma homeomorphic_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
  assumes "X homeomorphic_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
  shows "locally_path_connected_space X \<longleftrightarrow> locally_path_connected_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
  obtain f g where "homeomorphic_maps X Y f g"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    using assms homeomorphic_space_def by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
    by (metis (no_types) homeomorphic_map_def homeomorphic_maps_map locally_path_connected_space_quotient_map_image)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
lemma locally_path_connected_space_retraction_map_image:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
   "\<lbrakk>retraction_map X Y r; locally_path_connected_space X\<rbrakk>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
        \<Longrightarrow> locally_path_connected_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
  using Abstract_Topology.retraction_imp_quotient_map locally_path_connected_space_quotient_map_image by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
lemma locally_path_connected_space_euclideanreal: "locally_path_connected_space euclideanreal"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
  unfolding locally_path_connected_space_def neighbourhood_base_of
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
  proof clarsimp
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
  fix W and x :: "real"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
  assume "open W" "x \<in> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
  then obtain e where "e > 0" and e: "\<And>x'. \<bar>x' - x\<bar> < e \<longrightarrow> x' \<in> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
    by (auto simp: open_real)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
    by (force intro!: convex_imp_path_connected exI [where x = "{x-e<..<x+e}"])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
lemma locally_path_connected_space_discrete_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
   "locally_path_connected_space (discrete_topology U)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
  using locally_path_connected_space_open_path_components by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
lemma path_component_eq_connected_component_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
  assumes "locally_path_connected_space X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
  shows "(path_component_of_set X x = connected_component_of_set X x)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
proof (cases "x \<in> topspace X")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
  case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
    using connectedin_connected_component_of [of X x]
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
    apply (clarsimp simp add: connectedin_def connected_space_clopen_in topspace_subtopology_subset cong: conj_cong)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
    apply (drule_tac x="path_component_of_set X x" in spec)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
    using connected_component_of_eq_empty path_component_of_eq_empty by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
lemma path_components_eq_connected_components_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
   "locally_path_connected_space X \<Longrightarrow> (path_components_of X = connected_components_of X)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
  by (simp add: path_components_of_def connected_components_of_def image_def path_component_eq_connected_component_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
lemma path_connected_eq_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
   "locally_path_connected_space X
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
         \<Longrightarrow> path_connected_space X \<longleftrightarrow> connected_space X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
  by (metis connected_components_of_subset_sing path_components_eq_connected_components_of path_components_of_subset_singleton)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
lemma locally_path_connected_space_product_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
   "locally_path_connected_space(product_topology X I) \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
        topspace(product_topology X I) = {} \<or>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
        finite {i. i \<in> I \<and> ~path_connected_space(X i)} \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
        (\<forall>i \<in> I. locally_path_connected_space(X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
    (is "?lhs \<longleftrightarrow> ?empty \<or> ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
proof (cases ?empty)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
  case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
    by (simp add: locally_path_connected_space_def neighbourhood_base_of openin_closedin_eq)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
  then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
    by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
  have ?rhs if L: ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
    obtain U C where U: "openin (product_topology X I) U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
      and V: "path_connectedin (product_topology X I) C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
      and "z \<in> U" "U \<subseteq> C" and Csub: "C \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
      using L apply (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
      by (metis openin_topspace topspace_product_topology z)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
    then obtain V where finV: "finite {i \<in> I. V i \<noteq> topspace (X i)}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
      by (force simp: openin_product_topology_alt)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
    show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
    proof (intro conjI ballI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
      have "path_connected_space (X i)" if "i \<in> I" "V i = topspace (X i)" for i
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
      proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
        have pc: "path_connectedin (X i) ((\<lambda>x. x i) ` C)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
          apply (rule path_connectedin_continuous_map_image [OF _ V])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
          by (simp add: continuous_map_product_projection \<open>i \<in> I\<close>)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
        moreover have "((\<lambda>x. x i) ` C) = topspace (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
        proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
          show "(\<lambda>x. x i) ` C \<subseteq> topspace (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
            by (simp add: pc path_connectedin_subset_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
          have "V i \<subseteq> (\<lambda>x. x i) ` (\<Pi>\<^sub>E i\<in>I. V i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
            by (metis \<open>z \<in> Pi\<^sub>E I V\<close> empty_iff image_projection_PiE order_refl that(1))
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
          also have "\<dots> \<subseteq> (\<lambda>x. x i) ` U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
            using subU by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
          finally show "topspace (X i) \<subseteq> (\<lambda>x. x i) ` C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
            using \<open>U \<subseteq> C\<close> that by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
        ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
          by (simp add: path_connectedin_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
      then have "{i \<in> I. \<not> path_connected_space (X i)} \<subseteq> {i \<in> I. V i \<noteq> topspace (X i)}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
        by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
      with finV show "finite {i \<in> I. \<not> path_connected_space (X i)}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
        using finite_subset by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
    next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
      show "locally_path_connected_space (X i)" if "i \<in> I" for i
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
        apply (rule locally_path_connected_space_quotient_map_image [OF _ L, where f = "\<lambda>x. x i"])
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
        by (metis False Abstract_Topology.retraction_imp_quotient_map retraction_map_product_projection that)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
  moreover have ?lhs if R: ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
  proof (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
    fix F z
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
    assume "openin (product_topology X I) F" and "z \<in> F"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
    then obtain W where finW: "finite {i \<in> I. W i \<noteq> topspace (X i)}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
      by (auto simp: openin_product_topology_alt)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
                        (W i = topspace (X i) \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
                         path_connected_space (X i) \<longrightarrow> U = topspace (X i) \<and> C = topspace (X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
          (is "\<forall>i \<in> I. ?\<Phi> i")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
    proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
      fix i assume "i \<in> I"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
      have "locally_path_connected_space (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
        by (simp add: R \<open>i \<in> I\<close>)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
      moreover have "openin (X i) (W i) " "z i \<in> W i"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
        using \<open>z \<in> Pi\<^sub>E I W\<close> opeW \<open>i \<in> I\<close> by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
        using \<open>i \<in> I\<close> by (force simp: locally_path_connected_space_def neighbourhood_base_of)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
      show "?\<Phi> i"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
      proof (cases "W i = topspace (X i) \<and> path_connected_space(X i)")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
        case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
        then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
          using \<open>z i \<in> W i\<close> path_connectedin_topspace by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
      next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
        case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
        then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
          by (meson UC)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
    then obtain U C where
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   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>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
                        (W i = topspace (X i) \<and> path_connected_space (X i)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
                         \<longrightarrow> (U i) = topspace (X i) \<and> (C i) = topspace (X i))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
      by metis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
    let ?A = "{i \<in> I. \<not> path_connected_space (X i)} \<union> {i \<in> I. W i \<noteq> topspace (X i)}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
    have "{i \<in> I. U i \<noteq> topspace (X i)} \<subseteq> ?A"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
      by (clarsimp simp add: "*")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
    moreover have "finite ?A"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
      by (simp add: that finW)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
    ultimately have "finite {i \<in> I. U i \<noteq> topspace (X i)}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
      using finite_subset by auto
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
    then have "openin (product_topology X I) (Pi\<^sub>E I U)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
      using * by (simp add: openin_PiE_gen)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
    then show "\<exists>U. openin (product_topology X I) U \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
            (\<exists>V. path_connectedin (product_topology X I) V \<and> z \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> F)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
      apply (rule_tac x="PiE I U" in exI, simp)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
      apply (rule_tac x="PiE I C" in exI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
      using \<open>z \<in> Pi\<^sub>E I W\<close> \<open>Pi\<^sub>E I W \<subseteq> F\<close> *
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
      apply (simp add: path_connectedin_PiE subset_PiE PiE_iff PiE_mono dual_order.trans)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
      done
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
  ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
    using False by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   451
lemma locally_path_connected_is_realinterval:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   452
  assumes "is_interval S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   453
  shows "locally_path_connected_space(subtopology euclideanreal S)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   454
  unfolding locally_path_connected_space_def
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   455
proof (clarsimp simp add: neighbourhood_base_of openin_subtopology_alt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   456
  fix a U
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   457
  assume "a \<in> S" and "a \<in> U" and "open U"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   458
  then obtain r where "r > 0" and r: "ball a r \<subseteq> U"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   459
    by (metis open_contains_ball_eq)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   460
  show "\<exists>W. open W \<and> (\<exists>V. path_connectedin (top_of_set S) V \<and> a \<in> W \<and> S \<inter> W \<subseteq> V \<and> V \<subseteq> S \<and> V \<subseteq> U)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   461
  proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   462
    show "path_connectedin (top_of_set S) (S \<inter> ball a r)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   463
      by (simp add: assms is_interval_Int is_interval_ball_real is_interval_path_connected path_connectedin_subtopology)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   464
    show "a \<in> ball a r"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   465
      by (simp add: \<open>0 < r\<close>)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   466
  qed (use \<open>0 < r\<close> r in auto)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   467
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   468
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   469
lemma locally_path_connected_real_interval:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   470
 "locally_path_connected_space (subtopology euclideanreal{a..b})"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   471
  "locally_path_connected_space (subtopology euclideanreal{a<..<b})"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   472
  using locally_path_connected_is_realinterval
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   473
  by (auto simp add: is_interval_1)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   474
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   475
lemma locally_path_connected_space_prod_topology:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   476
   "locally_path_connected_space (prod_topology X Y) \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   477
      topspace (prod_topology X Y) = {} \<or>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   478
      locally_path_connected_space X \<and> locally_path_connected_space Y" (is "?lhs=?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   479
proof (cases "topspace(prod_topology X Y) = {}")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   480
  case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   481
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   482
    by (metis equals0D locally_path_connected_space_def neighbourhood_base_of_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   483
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   484
  case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   485
  then have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   486
    by simp_all
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   487
  show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   488
  proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   489
    assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   490
      by (metis ne locally_path_connected_space_retraction_map_image retraction_map_fst retraction_map_snd)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   491
  next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   492
    assume ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   493
    with False have X: "locally_path_connected_space X" and Y: "locally_path_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   494
      by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   495
    show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   496
      unfolding locally_path_connected_space_def neighbourhood_base_of
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   497
    proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   498
      fix UV x y
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   499
      assume UV: "openin (prod_topology X Y) UV" and "(x,y) \<in> UV"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   500
      obtain A B where W12: "openin X A \<and> openin Y B \<and> x \<in> A \<and> y \<in> B \<and> A \<times> B \<subseteq> UV"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   501
        using X Y by (metis UV \<open>(x,y) \<in> UV\<close> openin_prod_topology_alt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   502
      then obtain C D K L
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   503
        where "openin X C" "path_connectedin X K" "x \<in> C" "C \<subseteq> K" "K \<subseteq> A"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   504
          "openin Y D" "path_connectedin Y L" "y \<in> D" "D \<subseteq> L" "L \<subseteq> B"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   505
        by (metis X Y locally_path_connected_space)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   506
      with W12 \<open>openin X C\<close> \<open>openin Y D\<close>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   507
      show "\<exists>U V. openin (prod_topology X Y) U \<and> path_connectedin (prod_topology X Y) V \<and> (x, y) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> UV"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   508
        apply (rule_tac x="C \<times> D" in exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   509
        apply (rule_tac x="K \<times> L" in exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   510
        apply (auto simp: openin_prod_Times_iff path_connectedin_Times)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   511
        done
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   512
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   513
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   514
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   515
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   516
lemma locally_path_connected_space_sum_topology:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   517
   "locally_path_connected_space(sum_topology X I) \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   518
    (\<forall>i \<in> I. locally_path_connected_space (X i))" (is "?lhs=?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   519
proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   520
  assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   521
    by (smt (verit) homeomorphic_locally_path_connected_space locally_path_connected_space_open_subset topological_property_of_sum_component)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   522
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   523
  assume R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   524
  show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   525
  proof (clarsimp simp add: locally_path_connected_space_def neighbourhood_base_of forall_openin_sum_topology imp_conjL)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   526
    fix W i x
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   527
    assume ope: "\<forall>i\<in>I. openin (X i) (W i)" 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   528
      and "i \<in> I" and "x \<in> W i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   529
    then obtain U V where U: "openin (X i) U" and V: "path_connectedin (X i) V" 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   530
           and "x \<in> U" "U \<subseteq> V" "V \<subseteq> W i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   531
      by (metis R \<open>i \<in> I\<close> \<open>x \<in> W i\<close> locally_path_connected_space)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   532
    show "\<exists>U. openin (sum_topology X I) U \<and> (\<exists>V. path_connectedin (sum_topology X I) V \<and> (i, x) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> Sigma I W)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   533
    proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   534
      show "openin (sum_topology X I) (Pair i ` U)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   535
        by (meson U \<open>i \<in> I\<close> open_map_component_injection open_map_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   536
      show "path_connectedin (sum_topology X I) (Pair i ` V)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   537
        by (meson V \<open>i \<in> I\<close> continuous_map_component_injection path_connectedin_continuous_map_image)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   538
      show "Pair i ` V \<subseteq> Sigma I W"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   539
        using \<open>V \<subseteq> W i\<close> \<open>i \<in> I\<close> by force
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   540
    qed (use \<open>x \<in> U\<close> \<open>U \<subseteq> V\<close> in auto)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   541
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   542
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   543
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   544
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   545
subsection\<open>Locally connected spaces\<close>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   546
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   547
definition weakly_locally_connected_at 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   548
  where "weakly_locally_connected_at x X \<equiv> neighbourhood_base_at x (connectedin X) X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   549
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   550
definition locally_connected_at 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   551
  where "locally_connected_at x X \<equiv>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   552
           neighbourhood_base_at x (\<lambda>U. openin X U \<and> connectedin X U ) X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   553
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   554
definition locally_connected_space 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   555
  where "locally_connected_space X \<equiv> neighbourhood_base_of (connectedin X) X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   556
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   557
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   558
lemma locally_connected_A: "(\<forall>U x. openin X U \<and> x \<in> U
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   559
              \<longrightarrow> openin X (connected_component_of_set (subtopology X U) x))
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   560
       \<Longrightarrow> neighbourhood_base_of (\<lambda>U. openin X U \<and> connectedin X U) X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   561
  by (smt (verit, best) connected_component_of_refl connectedin_connected_component_of connectedin_subtopology mem_Collect_eq neighbourhood_base_of openin_subset topspace_subtopology_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   562
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   563
lemma locally_connected_B: "locally_connected_space X \<Longrightarrow> 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   564
          (\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (connected_component_of_set (subtopology X U) x))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   565
  unfolding locally_connected_space_def neighbourhood_base_of
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   566
  apply (erule all_forward)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   567
  apply clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   568
  apply (subst openin_subopen)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   569
  by (smt (verit, ccfv_threshold) Ball_Collect connected_component_of_def connected_component_of_equiv connectedin_subtopology in_mono mem_Collect_eq)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   570
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   571
lemma locally_connected_C: "neighbourhood_base_of (\<lambda>U. openin X U \<and> connectedin X U) X \<Longrightarrow> locally_connected_space X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   572
  using locally_connected_space_def neighbourhood_base_of_mono by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   573
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   574
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   575
lemma locally_connected_space_alt: 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   576
  "locally_connected_space X \<longleftrightarrow> neighbourhood_base_of (\<lambda>U. openin X U \<and> connectedin X U) X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   577
  using locally_connected_A locally_connected_B locally_connected_C by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   578
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   579
lemma locally_connected_space_eq_open_connected_component_of:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   580
  "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   581
        (\<forall>U x. openin X U \<and> x \<in> U
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   582
              \<longrightarrow> openin X (connected_component_of_set (subtopology X U) x))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   583
  by (meson locally_connected_A locally_connected_B locally_connected_C)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   584
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   585
lemma locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   586
   "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   587
     (\<forall>V x. openin X V \<and> x \<in> V \<longrightarrow> (\<exists>U. openin X U \<and> connectedin X U \<and> x \<in> U \<and> U \<subseteq> V))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   588
  by (simp add: locally_connected_space_alt open_neighbourhood_base_of)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   589
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   590
lemma locally_path_connected_imp_locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   591
   "locally_path_connected_space X \<Longrightarrow> locally_connected_space X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   592
  by (simp add: locally_connected_space_def locally_path_connected_space_def neighbourhood_base_of_mono path_connectedin_imp_connectedin)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   593
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   594
lemma locally_connected_space_open_connected_components:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   595
  "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   596
   (\<forall>U C. openin X U \<and> C \<in> connected_components_of(subtopology X U) \<longrightarrow> openin X C)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   597
  apply (simp add: locally_connected_space_eq_open_connected_component_of connected_components_of_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   598
  by (smt (verit) imageE image_eqI inf.orderE inf_commute openin_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   599
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   600
lemma openin_connected_component_of_locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   601
   "locally_connected_space X \<Longrightarrow> openin X (connected_component_of_set X x)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   602
  by (metis connected_component_of_eq_empty locally_connected_space_eq_open_connected_component_of openin_empty openin_topspace subtopology_topspace)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   603
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   604
lemma openin_connected_components_of_locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   605
   "\<lbrakk>locally_connected_space X; C \<in> connected_components_of X\<rbrakk> \<Longrightarrow> openin X C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   606
  by (metis locally_connected_space_open_connected_components openin_topspace subtopology_topspace)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   607
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   608
lemma weakly_locally_connected_at:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   609
   "weakly_locally_connected_at x X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   610
    (\<forall>V. openin X V \<and> x \<in> V
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   611
       \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> U \<subseteq> V \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   612
                (\<forall>y \<in> U. \<exists>C. connectedin X C \<and> C \<subseteq> V \<and> x \<in> C \<and> y \<in> C)))" (is "?lhs=?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   613
proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   614
  assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   615
    unfolding neighbourhood_base_at_def weakly_locally_connected_at_def
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   616
    by (meson subsetD subset_trans)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   617
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   618
  assume R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   619
  show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   620
    unfolding neighbourhood_base_at_def weakly_locally_connected_at_def
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   621
  proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   622
    fix V
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   623
    assume "openin X V" and "x \<in> V"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   624
    then obtain U where "openin X U" "x \<in> U" "U \<subseteq> V" 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   625
                  and U: "\<forall>y\<in>U. \<exists>C. connectedin X C \<and> C \<subseteq> V \<and> x \<in> C \<and> y \<in> C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   626
      using R by force
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   627
    show "\<exists>A B. openin X A \<and> connectedin X B \<and> x \<in> A \<and> A \<subseteq> B \<and> B \<subseteq> V"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   628
    proof (intro conjI exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   629
      show "connectedin X (connected_component_of_set (subtopology X V) x)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   630
        by (meson connectedin_connected_component_of connectedin_subtopology)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   631
      show "U \<subseteq> connected_component_of_set (subtopology X V) x"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   632
        using connected_component_of_maximal U
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   633
        by (simp add: connected_component_of_def connectedin_subtopology subsetI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   634
      show "connected_component_of_set (subtopology X V) x \<subseteq> V"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   635
        using connected_component_of_subset_topspace by fastforce
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   636
    qed (auto simp: \<open>x \<in> U\<close> \<open>openin X U\<close>)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   637
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   638
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   639
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   640
lemma locally_connected_space_iff_weak:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   641
  "locally_connected_space X \<longleftrightarrow> (\<forall>x \<in> topspace X. weakly_locally_connected_at x X)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   642
  by (simp add: locally_connected_space_def neighbourhood_base_of_def weakly_locally_connected_at_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   643
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   644
lemma locally_connected_space_im_kleinen:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   645
   "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   646
    (\<forall>V x. openin X V \<and> x \<in> V
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   647
          \<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> U \<subseteq> V \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   648
                    (\<forall>y \<in> U. \<exists>C. connectedin X C \<and> C \<subseteq> V \<and> x \<in> C \<and> y \<in> C)))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   649
  unfolding locally_connected_space_iff_weak weakly_locally_connected_at
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   650
  using openin_subset subsetD by fastforce
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   651
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   652
lemma locally_connected_space_open_subset:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   653
   "\<lbrakk>locally_connected_space X; openin X S\<rbrakk> \<Longrightarrow> locally_connected_space (subtopology X S)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   654
  apply (simp add: locally_connected_space_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   655
  by (smt (verit, ccfv_threshold) connectedin_subtopology neighbourhood_base_of openin_open_subtopology subset_trans)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   656
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   657
lemma locally_connected_space_quotient_map_image:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   658
  assumes X: "locally_connected_space X" and f: "quotient_map X Y f"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   659
  shows "locally_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   660
  unfolding locally_connected_space_open_connected_components
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   661
proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   662
  fix V C
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   663
  assume "openin Y V" and C: "C \<in> connected_components_of (subtopology Y V)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   664
  then have "C \<subseteq> topspace Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   665
    using connected_components_of_subset by force
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   666
  have ope1: "openin X {a \<in> topspace X. f a \<in> V}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   667
    using \<open>openin Y V\<close> f openin_continuous_map_preimage quotient_imp_continuous_map by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   668
  define Vf where "Vf \<equiv> {z \<in> topspace X. f z \<in> V}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   669
  have "openin X {x \<in> topspace X. f x \<in> C}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   670
  proof (clarsimp simp: openin_subopen [where S = "{x \<in> topspace X. f x \<in> C}"])
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   671
    fix x
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   672
    assume "x \<in> topspace X" and "f x \<in> C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   673
    show "\<exists>T. openin X T \<and> x \<in> T \<and> T \<subseteq> {x \<in> topspace X. f x \<in> C}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   674
    proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   675
      show "openin X (connected_component_of_set (subtopology X Vf) x)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   676
        by (metis Vf_def X connected_component_of_eq_empty locally_connected_B ope1 openin_empty
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   677
                  openin_subset topspace_subtopology_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   678
      show x_in_conn: "x \<in> connected_component_of_set (subtopology X Vf) x"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   679
        using C Vf_def \<open>f x \<in> C\<close> \<open>x \<in> topspace X\<close> connected_component_of_refl connected_components_of_subset by fastforce
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   680
      have "connected_component_of_set (subtopology X Vf) x \<subseteq> topspace X \<inter> Vf"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   681
        using connected_component_of_subset_topspace by fastforce
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   682
      moreover
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   683
      have "f ` connected_component_of_set (subtopology X Vf) x \<subseteq> C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   684
      proof (rule connected_components_of_maximal [where X = "subtopology Y V"])
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   685
        show "C \<in> connected_components_of (subtopology Y V)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   686
          by (simp add: C)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   687
        have \<section>: "quotient_map (subtopology X Vf) (subtopology Y V) f"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   688
          by (simp add: Vf_def \<open>openin Y V\<close> f quotient_map_restriction)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   689
        then show "connectedin (subtopology Y V) (f ` connected_component_of_set (subtopology X Vf) x)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   690
          by (metis connectedin_connected_component_of connectedin_continuous_map_image quotient_imp_continuous_map)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   691
        show "\<not> disjnt C (f ` connected_component_of_set (subtopology X Vf) x)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   692
          using \<open>f x \<in> C\<close> x_in_conn by (auto simp: disjnt_iff)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   693
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   694
      ultimately
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   695
      show "connected_component_of_set (subtopology X Vf) x \<subseteq> {x \<in> topspace X. f x \<in> C}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   696
        by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   697
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   698
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   699
  then show "openin Y C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   700
    using \<open>C \<subseteq> topspace Y\<close> f quotient_map_def by fastforce
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   701
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   702
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   703
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   704
lemma locally_connected_space_retraction_map_image:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   705
   "\<lbrakk>retraction_map X Y r; locally_connected_space X\<rbrakk>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   706
        \<Longrightarrow> locally_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   707
  using locally_connected_space_quotient_map_image retraction_imp_quotient_map by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   708
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   709
lemma homeomorphic_locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   710
   "X homeomorphic_space Y \<Longrightarrow> locally_connected_space X \<longleftrightarrow> locally_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   711
  by (meson homeomorphic_map_def homeomorphic_space homeomorphic_space_sym locally_connected_space_quotient_map_image)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   712
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   713
lemma locally_connected_space_euclideanreal: "locally_connected_space euclideanreal"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   714
  by (simp add: locally_path_connected_imp_locally_connected_space locally_path_connected_space_euclideanreal)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   715
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   716
lemma locally_connected_is_realinterval:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   717
   "is_interval S \<Longrightarrow> locally_connected_space(subtopology euclideanreal S)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   718
  by (simp add: locally_path_connected_imp_locally_connected_space locally_path_connected_is_realinterval)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   719
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   720
lemma locally_connected_real_interval:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   721
    "locally_connected_space (subtopology euclideanreal{a..b})"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   722
    "locally_connected_space (subtopology euclideanreal{a<..<b})"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   723
  using connected_Icc is_interval_connected_1 locally_connected_is_realinterval by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   724
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   725
lemma locally_connected_space_discrete_topology:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   726
   "locally_connected_space (discrete_topology U)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   727
  by (simp add: locally_path_connected_imp_locally_connected_space locally_path_connected_space_discrete_topology)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   728
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   729
lemma locally_path_connected_imp_locally_connected_at:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   730
   "locally_path_connected_at x X \<Longrightarrow> locally_connected_at x X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   731
  by (simp add: locally_connected_at_def locally_path_connected_at_def neighbourhood_base_at_mono path_connectedin_imp_connectedin)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   732
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   733
lemma weakly_locally_path_connected_imp_weakly_locally_connected_at:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   734
   "weakly_locally_path_connected_at x X
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   735
             \<Longrightarrow> weakly_locally_connected_at x X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   736
  by (simp add: neighbourhood_base_at_mono path_connectedin_imp_connectedin weakly_locally_connected_at_def weakly_locally_path_connected_at_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   737
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   738
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   739
lemma interior_of_locally_connected_subspace_component:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   740
  assumes X: "locally_connected_space X"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   741
    and C: "C \<in> connected_components_of (subtopology X S)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   742
  shows "X interior_of C = C \<inter> X interior_of S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   743
proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   744
  obtain Csub: "C \<subseteq> topspace X" "C \<subseteq> S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   745
    by (meson C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   746
  show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   747
  proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   748
    show "X interior_of C \<subseteq> C \<inter> X interior_of S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   749
      by (simp add: Csub interior_of_mono interior_of_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   750
    have eq: "X interior_of S = \<Union> (connected_components_of (subtopology X (X interior_of S)))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   751
      by (metis Union_connected_components_of interior_of_subset_topspace topspace_subtopology_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   752
    moreover have "C \<inter> D \<subseteq> X interior_of C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   753
      if "D \<in> connected_components_of (subtopology X (X interior_of S))" for D
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   754
    proof (cases "C \<inter> D = {}")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   755
      case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   756
      have "D \<subseteq> X interior_of C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   757
      proof (rule interior_of_maximal)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   758
        have "connectedin (subtopology X S) D"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   759
          by (meson connectedin_connected_components_of connectedin_subtopology interior_of_subset subset_trans that)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   760
        then show "D \<subseteq> C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   761
          by (meson C False connected_components_of_maximal disjnt_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   762
        show "openin X D"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   763
          using X locally_connected_space_open_connected_components openin_interior_of that by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   764
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   765
      then show ?thesis 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   766
        by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   767
    qed auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   768
    ultimately show "C \<inter> X interior_of S \<subseteq> X interior_of C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   769
      by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   770
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   771
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   772
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   773
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   774
lemma frontier_of_locally_connected_subspace_component:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   775
  assumes X: "locally_connected_space X" and "closedin X S" 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   776
    and C: "C \<in> connected_components_of (subtopology X S)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   777
  shows "X frontier_of C = C \<inter> X frontier_of S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   778
proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   779
  obtain Csub: "C \<subseteq> topspace X" "C \<subseteq> S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   780
    by (meson C connectedin_connected_components_of connectedin_subset_topspace connectedin_subtopology)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   781
  then have "X closure_of C - X interior_of C = C \<inter> X closure_of S - C \<inter> X interior_of S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   782
    using assms
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   783
    apply (simp add: closure_of_closedin flip: interior_of_locally_connected_subspace_component)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   784
    by (metis closedin_connected_components_of closedin_trans_full closure_of_eq inf.orderE)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   785
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   786
    by (simp add: Diff_Int_distrib frontier_of_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   787
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   788
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   789
(*Similar proof to locally_connected_space_prod_topology*)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   790
lemma locally_connected_space_prod_topology:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   791
   "locally_connected_space (prod_topology X Y) \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   792
      topspace (prod_topology X Y) = {} \<or>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   793
      locally_connected_space X \<and> locally_connected_space Y" (is "?lhs=?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   794
proof (cases "topspace(prod_topology X Y) = {}")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   795
  case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   796
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   797
    using locally_connected_space_iff_weak by force
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   798
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   799
  case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   800
  then have ne: "topspace X \<noteq> {}" "topspace Y \<noteq> {}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   801
    by simp_all
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   802
  show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   803
  proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   804
    assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   805
      by (metis locally_connected_space_quotient_map_image ne quotient_map_fst quotient_map_snd)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   806
  next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   807
    assume ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   808
    with False have X: "locally_connected_space X" and Y: "locally_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   809
      by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   810
    show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   811
      unfolding locally_connected_space_def neighbourhood_base_of
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   812
    proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   813
      fix UV x y
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   814
      assume UV: "openin (prod_topology X Y) UV" and "(x,y) \<in> UV"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   815
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   816
     obtain A B where W12: "openin X A \<and> openin Y B \<and> x \<in> A \<and> y \<in> B \<and> A \<times> B \<subseteq> UV"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   817
        using X Y by (metis UV \<open>(x,y) \<in> UV\<close> openin_prod_topology_alt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   818
      then obtain C D K L
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   819
        where "openin X C" "connectedin X K" "x \<in> C" "C \<subseteq> K" "K \<subseteq> A"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   820
          "openin Y D" "connectedin Y L" "y \<in> D" "D \<subseteq> L" "L \<subseteq> B"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   821
        by (metis X Y locally_connected_space)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   822
      with W12 \<open>openin X C\<close> \<open>openin Y D\<close>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   823
      show "\<exists>U V. openin (prod_topology X Y) U \<and> connectedin (prod_topology X Y) V \<and> (x, y) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> UV"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   824
        apply (rule_tac x="C \<times> D" in exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   825
        apply (rule_tac x="K \<times> L" in exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   826
        apply (auto simp: openin_prod_Times_iff connectedin_Times)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   827
        done
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   828
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   829
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   830
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   831
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   832
(*Same proof as locally_path_connected_space_product_topology*)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   833
lemma locally_connected_space_product_topology:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   834
   "locally_connected_space(product_topology X I) \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   835
        topspace(product_topology X I) = {} \<or>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   836
        finite {i. i \<in> I \<and> ~connected_space(X i)} \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   837
        (\<forall>i \<in> I. locally_connected_space(X i))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   838
    (is "?lhs \<longleftrightarrow> ?empty \<or> ?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   839
proof (cases ?empty)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   840
  case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   841
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   842
    by (simp add: locally_connected_space_def neighbourhood_base_of openin_closedin_eq)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   843
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   844
  case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   845
  then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   846
    by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   847
  have ?rhs if L: ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   848
  proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   849
    obtain U C where U: "openin (product_topology X I) U"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   850
      and V: "connectedin (product_topology X I) C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   851
      and "z \<in> U" "U \<subseteq> C" and Csub: "C \<subseteq> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   852
      using L apply (clarsimp simp add: locally_connected_space_def neighbourhood_base_of)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   853
      by (metis openin_topspace topspace_product_topology z)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   854
    then obtain V where finV: "finite {i \<in> I. V i \<noteq> topspace (X i)}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   855
      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"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   856
      by (force simp: openin_product_topology_alt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   857
    show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   858
    proof (intro conjI ballI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   859
      have "connected_space (X i)" if "i \<in> I" "V i = topspace (X i)" for i
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   860
      proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   861
        have pc: "connectedin (X i) ((\<lambda>x. x i) ` C)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   862
          apply (rule connectedin_continuous_map_image [OF _ V])
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   863
          by (simp add: continuous_map_product_projection \<open>i \<in> I\<close>)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   864
        moreover have "((\<lambda>x. x i) ` C) = topspace (X i)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   865
        proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   866
          show "(\<lambda>x. x i) ` C \<subseteq> topspace (X i)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   867
            by (simp add: pc connectedin_subset_topspace)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   868
          have "V i \<subseteq> (\<lambda>x. x i) ` (\<Pi>\<^sub>E i\<in>I. V i)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   869
            by (metis \<open>z \<in> Pi\<^sub>E I V\<close> empty_iff image_projection_PiE order_refl that(1))
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   870
          also have "\<dots> \<subseteq> (\<lambda>x. x i) ` U"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   871
            using subU by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   872
          finally show "topspace (X i) \<subseteq> (\<lambda>x. x i) ` C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   873
            using \<open>U \<subseteq> C\<close> that by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   874
        qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   875
        ultimately show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   876
          by (simp add: connectedin_topspace)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   877
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   878
      then have "{i \<in> I. \<not> connected_space (X i)} \<subseteq> {i \<in> I. V i \<noteq> topspace (X i)}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   879
        by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   880
      with finV show "finite {i \<in> I. \<not> connected_space (X i)}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   881
        using finite_subset by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   882
    next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   883
      show "locally_connected_space (X i)" if "i \<in> I" for i
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   884
        by (meson False L locally_connected_space_quotient_map_image quotient_map_product_projection that)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   885
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   886
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   887
  moreover have ?lhs if R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   888
  proof (clarsimp simp add: locally_connected_space_def neighbourhood_base_of)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   889
    fix F z
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   890
    assume "openin (product_topology X I) F" and "z \<in> F"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   891
    then obtain W where finW: "finite {i \<in> I. W i \<noteq> topspace (X i)}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   892
            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"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   893
      by (auto simp: openin_product_topology_alt)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   894
    have "\<forall>i \<in> I. \<exists>U C. openin (X i) U \<and> connectedin (X i) C \<and> z i \<in> U \<and> U \<subseteq> C \<and> C \<subseteq> W i \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   895
                        (W i = topspace (X i) \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   896
                         connected_space (X i) \<longrightarrow> U = topspace (X i) \<and> C = topspace (X i))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   897
          (is "\<forall>i \<in> I. ?\<Phi> i")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   898
    proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   899
      fix i assume "i \<in> I"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   900
      have "locally_connected_space (X i)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   901
        by (simp add: R \<open>i \<in> I\<close>)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   902
      moreover have "openin (X i) (W i) " "z i \<in> W i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   903
        using \<open>z \<in> Pi\<^sub>E I W\<close> opeW \<open>i \<in> I\<close> by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   904
      ultimately obtain U C where UC: "openin (X i) U" "connectedin (X i) C" "z i \<in> U" "U \<subseteq> C" "C \<subseteq> W i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   905
        using \<open>i \<in> I\<close> by (force simp: locally_connected_space_def neighbourhood_base_of)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   906
      show "?\<Phi> i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   907
      proof (cases "W i = topspace (X i) \<and> connected_space(X i)")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   908
        case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   909
        then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   910
          using \<open>z i \<in> W i\<close> connectedin_topspace by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   911
      next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   912
        case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   913
        then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   914
          by (meson UC)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   915
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   916
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   917
    then obtain U C where
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   918
      *: "\<And>i. i \<in> I \<Longrightarrow> openin (X i) (U i) \<and> connectedin (X i) (C i) \<and> z i \<in> (U i) \<and> (U i) \<subseteq> (C i) \<and> (C i) \<subseteq> W i \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   919
                        (W i = topspace (X i) \<and> connected_space (X i)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   920
                         \<longrightarrow> (U i) = topspace (X i) \<and> (C i) = topspace (X i))"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   921
      by metis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   922
    let ?A = "{i \<in> I. \<not> connected_space (X i)} \<union> {i \<in> I. W i \<noteq> topspace (X i)}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   923
    have "{i \<in> I. U i \<noteq> topspace (X i)} \<subseteq> ?A"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   924
      by (clarsimp simp add: "*")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   925
    moreover have "finite ?A"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   926
      by (simp add: that finW)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   927
    ultimately have "finite {i \<in> I. U i \<noteq> topspace (X i)}"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   928
      using finite_subset by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   929
    then have "openin (product_topology X I) (Pi\<^sub>E I U)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   930
      using * by (simp add: openin_PiE_gen)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   931
    then show "\<exists>U. openin (product_topology X I) U \<and>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   932
            (\<exists>V. connectedin (product_topology X I) V \<and> z \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> F)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   933
      apply (rule_tac x="PiE I U" in exI, simp)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   934
      apply (rule_tac x="PiE I C" in exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   935
      using \<open>z \<in> Pi\<^sub>E I W\<close> \<open>Pi\<^sub>E I W \<subseteq> F\<close> *
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   936
      apply (simp add: connectedin_PiE subset_PiE PiE_iff PiE_mono dual_order.trans)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   937
      done
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   938
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   939
  ultimately show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   940
    using False by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   941
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   942
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   943
lemma locally_connected_space_sum_topology:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   944
   "locally_connected_space(sum_topology X I) \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   945
    (\<forall>i \<in> I. locally_connected_space (X i))" (is "?lhs=?rhs")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   946
proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   947
  assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   948
    by (smt (verit) homeomorphic_locally_connected_space locally_connected_space_open_subset topological_property_of_sum_component)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   949
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   950
  assume R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   951
  show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   952
  proof (clarsimp simp add: locally_connected_space_def neighbourhood_base_of forall_openin_sum_topology imp_conjL)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   953
    fix W i x
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   954
    assume ope: "\<forall>i\<in>I. openin (X i) (W i)" 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   955
      and "i \<in> I" and "x \<in> W i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   956
    then obtain U V where U: "openin (X i) U" and V: "connectedin (X i) V" 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   957
           and "x \<in> U" "U \<subseteq> V" "V \<subseteq> W i"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   958
      by (metis R \<open>i \<in> I\<close> \<open>x \<in> W i\<close> locally_connected_space)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   959
    show "\<exists>U. openin (sum_topology X I) U \<and> (\<exists>V. connectedin (sum_topology X I) V \<and> (i,x) \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> Sigma I W)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   960
    proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   961
      show "openin (sum_topology X I) (Pair i ` U)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   962
        by (meson U \<open>i \<in> I\<close> open_map_component_injection open_map_def)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   963
      show "connectedin (sum_topology X I) (Pair i ` V)"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   964
        by (meson V \<open>i \<in> I\<close> continuous_map_component_injection connectedin_continuous_map_image)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   965
      show "Pair i ` V \<subseteq> Sigma I W"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   966
        using \<open>V \<subseteq> W i\<close> \<open>i \<in> I\<close> by force
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   967
    qed (use \<open>x \<in> U\<close> \<open>U \<subseteq> V\<close> in auto)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   968
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   969
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   970
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   971
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
end