src/HOL/Analysis/Locally.thy
author paulson <lp15@cam.ac.uk>
Thu, 17 Apr 2025 22:57:26 +0100
changeset 82522 62afd98e3f3e
parent 80914 d97fdabd9e2b
permissions -rw-r--r--
more tidying
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))"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
    47
  by (smt (verit) neighbourhood_base_of subsetD)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
lemma neighbourhood_base_of_open_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
   "\<lbrakk>neighbourhood_base_of P X; openin X S\<rbrakk>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
        \<Longrightarrow> neighbourhood_base_of P (subtopology X S)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
    52
  by (smt (verit, ccfv_SIG) neighbourhood_base_of openin_open_subtopology subset_trans)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
lemma neighbourhood_base_of_topology_base:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
   "openin X = arbitrary union_of (\<lambda>W. W \<in> \<B>)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
        \<Longrightarrow> neighbourhood_base_of P X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
             (\<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))"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
    58
  by (smt (verit, del_insts) neighbourhood_base_of openin_topology_base_unique subset_trans)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
lemma neighbourhood_base_at_unlocalized:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
  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
    62
  shows "neighbourhood_base_at x P X
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
     \<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
    64
         (is "?lhs = ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  assume R: ?rhs show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
    unfolding neighbourhood_base_at_def
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
    fix W
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
    assume "openin X W" "x \<in> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
    then have "x \<in> topspace X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
      using openin_subset by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
    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
    74
      by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
    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
    76
      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
    77
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
qed (simp add: neighbourhood_base_at_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
lemma neighbourhood_base_at_with_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
   "\<lbrakk>openin X U; x \<in> U\<rbrakk>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
        \<Longrightarrow> (neighbourhood_base_at x P X \<longleftrightarrow> neighbourhood_base_at x (\<lambda>T. T \<subseteq> U \<and> P T) X)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
    83
  unfolding neighbourhood_base_at_def by (metis IntI Int_subset_iff openin_Int)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
lemma neighbourhood_base_of_with_subset:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
   "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
    87
  using neighbourhood_base_at_with_subset
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
  by (fastforce simp add: neighbourhood_base_of_def)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
subsection\<open>Locally path-connected spaces\<close>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
definition weakly_locally_path_connected_at
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
  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
    94
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
definition locally_path_connected_at
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  where "locally_path_connected_at x X \<equiv>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
    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
    98
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
definition locally_path_connected_space
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  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
   101
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
lemma locally_path_connected_space_alt:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
  "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
   104
  (is "?P = ?Q")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  and locally_path_connected_space_eq_open_path_component_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
  "locally_path_connected_space X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
        (\<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
   108
  (is "?P = ?R")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  have ?P if ?Q
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
    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
   112
  moreover have ?R if P: ?P
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
  proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
    show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
    proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
      fix U y
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
      assume "openin X U" "y \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
      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
   119
        if "path_component_of (subtopology X U) y x" for x
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
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
        have "x \<in> U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
          using path_component_of_equiv that topspace_subtopology by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
        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
   125
      using P
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
      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
   127
        then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
          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
   129
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
      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
   131
        using openin_subopen by force
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  moreover have ?Q if ?R
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   135
    by (smt (verit) mem_Collect_eq open_neighbourhood_base_of openin_subset path_component_of_refl 
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   136
        path_connectedin_path_component_of path_connectedin_subtopology that topspace_subtopology_subset)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  ultimately show "?P = ?Q" "?P = ?R"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
    by blast+
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
lemma locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
   "locally_path_connected_space X
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
   \<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
   144
  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
   145
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
lemma locally_path_connected_space_open_path_components:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
   "locally_path_connected_space X \<longleftrightarrow>
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   148
        (\<forall>U C. openin X U \<and> C \<in> path_components_of(subtopology X U) \<longrightarrow> openin X C)"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   149
  apply (simp add: locally_path_connected_space_eq_open_path_component_of path_components_of_def)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   150
  by (smt (verit, ccfv_threshold) Int_iff image_iff openin_subset subset_iff)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
lemma openin_path_component_of_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
   "locally_path_connected_space X \<Longrightarrow> openin X (Collect (path_component_of X x))"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   154
  using locally_path_connected_space_eq_open_path_component_of openin_subopen path_component_of_eq_empty 
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   155
  by fastforce
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
lemma openin_path_components_of_locally_path_connected_space:
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   158
   "\<lbrakk>locally_path_connected_space X; C \<in> path_components_of X\<rbrakk> \<Longrightarrow> openin X C"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   159
  using locally_path_connected_space_open_path_components by force
69945
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 closedin_path_components_of_locally_path_connected_space:
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   162
   "\<lbrakk>locally_path_connected_space X; C \<in> path_components_of X\<rbrakk> \<Longrightarrow> closedin X C"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   163
  unfolding closedin_def
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   164
  by (metis Diff_iff complement_path_components_of_Union openin_clauses(3) openin_closedin_eq 
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   165
      openin_path_components_of_locally_path_connected_space)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
lemma closedin_path_component_of_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  assumes "locally_path_connected_space X"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  shows "closedin X (Collect (path_component_of X x))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
proof (cases "x \<in> topspace X")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  case True
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
    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
   174
next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
  case False
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
  then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
    by (metis closedin_empty path_component_of_eq_empty)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
lemma weakly_locally_path_connected_at:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
   "weakly_locally_path_connected_at x X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
    (\<forall>V. openin X V \<and> x \<in> V
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
          \<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
   184
                  (\<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
   185
         (is "?lhs = ?rhs")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
  assume ?lhs then show ?rhs
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   188
    by (smt (verit) neighbourhood_base_at_def subset_iff weakly_locally_path_connected_at_def)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
  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
   191
    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
   192
    for W U
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
  proof (intro exI conjI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
    let ?V = "(Collect (path_component_of (subtopology X W) x))"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
      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
   196
        by (meson path_connectedin_path_component_of path_connectedin_subtopology)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
      show "U \<subseteq> ?V"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
        by (auto simp: path_component_of path_connectedin_subtopology that)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
      show "?V \<subseteq> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
        by (meson path_connectedin_path_component_of path_connectedin_subtopology)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
  assume ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  then show ?lhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    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
   205
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
lemma locally_path_connected_space_im_kleinen:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
   "locally_path_connected_space X \<longleftrightarrow>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
      (\<forall>V x. openin X V \<and> x \<in> V
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
             \<longrightarrow> (\<exists>U. openin X U \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
                    x \<in> U \<and> U \<subseteq> V \<and>
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   212
                    (\<forall>y \<in> U. \<exists>C. path_connectedin X C \<and>
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   213
                                C \<subseteq> V \<and> x \<in> C \<and> y \<in> C)))"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   214
         (is "?lhs = ?rhs")
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   215
proof
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   216
  show "?lhs \<Longrightarrow> ?rhs"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   217
    by (metis locally_path_connected_space)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   218
  assume ?rhs
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   219
  then show ?lhs
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   220
    unfolding locally_path_connected_space_def neighbourhood_base_of
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   221
    by (metis neighbourhood_base_at_def weakly_locally_path_connected_at weakly_locally_path_connected_at_def)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   222
qed
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
lemma locally_path_connected_space_open_subset:
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   225
   "\<lbrakk>locally_path_connected_space X; openin X S\<rbrakk>
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   226
        \<Longrightarrow> locally_path_connected_space (subtopology X S)"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   227
  by (smt (verit, best) locally_path_connected_space openin_open_subtopology path_connectedin_subtopology subset_trans)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
lemma locally_path_connected_space_quotient_map_image:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  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
   231
  shows "locally_path_connected_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
  unfolding locally_path_connected_space_open_path_components
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
proof clarify
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
  fix V C
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
  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
   236
  then have sub: "C \<subseteq> topspace Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
    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
   238
  have "openin X {x \<in> topspace X. f x \<in> C}"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
  proof (subst openin_subopen, clarify)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
    fix x
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
    assume x: "x \<in> topspace X" and "f x \<in> C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
    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
   243
    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
   244
    proof (intro exI conjI)
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   245
      have *: "\<exists>U. openin X U \<and> ?T \<in> path_components_of (subtopology X U)"
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
      proof (intro exI conjI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
        show "openin X ({z \<in> topspace X. f z \<in> V})"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
          using V f openin_subset quotient_map_def by fastforce
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   249
        have "x \<in> topspace (subtopology X {z \<in> topspace X. f z \<in> V})"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   250
          using \<open>f x \<in> C\<close> c path_components_of_subset x by force
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   251
        then show "Collect (path_component_of (subtopology X {z \<in> topspace X. f z \<in> V}) x)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   252
            \<in> path_components_of (subtopology X {z \<in> topspace X. f z \<in> V})"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   253
          by (meson path_component_in_path_components_of)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
      with X show "openin X ?T"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
        using locally_path_connected_space_open_path_components by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
      show x: "x \<in> ?T"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   258
        using * nonempty_path_components_of path_component_of_eq path_component_of_eq_empty by fastforce
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
      have "f ` ?T \<subseteq> C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
      proof (rule path_components_of_maximal)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
        show "C \<in> path_components_of (subtopology Y V)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
          by (simp add: c)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
        show "path_connectedin (subtopology Y V) (f ` ?T)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
        proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
          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
   266
            by (simp add: V f quotient_map_restriction)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
          then show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
            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
   269
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
        show "\<not> disjnt C (f ` ?T)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
          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
   272
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
      then show "?T \<subseteq> {x \<in> topspace X. f x \<in> C}"
71172
nipkow
parents: 71137
diff changeset
   274
        by (force simp: path_component_of_equiv)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
    qed
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 "openin Y C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
    using f sub by (simp add: quotient_map_def)
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
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
lemma homeomorphic_locally_path_connected_space:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
  assumes "X homeomorphic_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
  shows "locally_path_connected_space X \<longleftrightarrow> locally_path_connected_space Y"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   284
  using assms
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   285
    unfolding homeomorphic_space_def homeomorphic_map_def homeomorphic_maps_map
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   286
    by (metis locally_path_connected_space_quotient_map_image)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
lemma locally_path_connected_space_retraction_map_image:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
   "\<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
   290
        \<Longrightarrow> locally_path_connected_space Y"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
  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
   292
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
lemma locally_path_connected_space_euclideanreal: "locally_path_connected_space euclideanreal"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
  unfolding locally_path_connected_space_def neighbourhood_base_of
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
  proof clarsimp
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
  fix W and x :: "real"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
  assume "open W" "x \<in> W"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
  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
   299
    by (auto simp: open_real)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
  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
   301
    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
   302
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
lemma locally_path_connected_space_discrete_topology:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
   "locally_path_connected_space (discrete_topology U)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
  using locally_path_connected_space_open_path_components by fastforce
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
lemma path_component_eq_connected_component_of:
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
  assumes "locally_path_connected_space X"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   310
  shows "path_component_of_set X x = connected_component_of_set X x"
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
proof (cases "x \<in> topspace X")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
  case True
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   313
  have "path_component_of_set X x \<subseteq> connected_component_of_set X x"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   314
    by (simp add: path_component_subset_connected_component_of)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   315
  moreover have "closedin X (path_component_of_set X x)"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   316
    by (simp add: assms closedin_path_component_of_locally_path_connected_space)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   317
  moreover have "openin X (path_component_of_set X x)"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   318
    by (simp add: assms openin_path_component_of_locally_path_connected_space)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   319
  moreover have "path_component_of_set X x \<noteq> {}"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   320
    by (meson True path_component_of_eq_empty)
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   321
  ultimately show ?thesis
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   322
    using connectedin_connected_component_of [of X x] unfolding connectedin_def
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   323
    by (metis closedin_subset_topspace connected_space_clopen_in  
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   324
        subset_openin_subtopology topspace_subtopology_subset)
69945
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>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   342
        (product_topology X I) = trivial_topology \<or>
69945
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))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   353
    using discrete_topology_unique_derived_set by (fastforce iff: null_topspace_iff_trivial)
69945
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))"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   359
      by (metis L locally_path_connected_space openin_topspace topspace_product_topology z)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
    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
   361
      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
   362
      by (force simp: openin_product_topology_alt)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
    show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
    proof (intro conjI ballI)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
      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
   366
      proof -
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
        have pc: "path_connectedin (X i) ((\<lambda>x. x i) ` C)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   368
          by (metis V continuous_map_product_projection path_connectedin_continuous_map_image that(1))
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
        moreover have "((\<lambda>x. x i) ` C) = topspace (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
        proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
          show "(\<lambda>x. x i) ` C \<subseteq> topspace (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
            by (simp add: pc path_connectedin_subset_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
          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
   374
            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
   375
          also have "\<dots> \<subseteq> (\<lambda>x. x i) ` U"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
            using subU by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
          finally show "topspace (X i) \<subseteq> (\<lambda>x. x i) ` C"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
            using \<open>U \<subseteq> C\<close> that by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
        qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
        ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
          by (simp add: path_connectedin_topspace)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
      qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
      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
   384
        by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
      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
   386
        using finite_subset by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
    next
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
      show "locally_path_connected_space (X i)" if "i \<in> I" for i
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   389
        by (meson False L locally_path_connected_space_quotient_map_image quotient_map_product_projection that)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
  moreover have ?lhs if R: ?rhs
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
  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
   394
    fix F z
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
    assume "openin (product_topology X I) F" and "z \<in> F"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
    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
   397
            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
   398
      by (auto simp: openin_product_topology_alt)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
    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
   400
                        (W i = topspace (X i) \<and>
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
                         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
   402
          (is "\<forall>i \<in> I. ?\<Phi> i")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
    proof
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
      fix i assume "i \<in> I"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
      have "locally_path_connected_space (X i)"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
        by (simp add: R \<open>i \<in> I\<close>)
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   407
      moreover have *:"openin (X i) (W i) " "z i \<in> W i"
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
        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
   409
      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
   410
        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
   411
      show "?\<Phi> i"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   412
        by (metis UC * openin_subset path_connectedin_topspace)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
    qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
    then obtain U C where
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
      *: "\<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
   416
                        (W i = topspace (X i) \<and> path_connected_space (X i)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
                         \<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
   418
      by metis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
    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
   420
    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
   421
      by (clarsimp simp add: "*")
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
    moreover have "finite ?A"
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
      by (simp add: that finW)
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
    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
   425
      using finite_subset by auto
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   426
    with * have "openin (product_topology X I) (Pi\<^sub>E I U)"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   427
      by (simp add: openin_PiE_gen)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
    then show "\<exists>U. openin (product_topology X I) U \<and>
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   429
              (\<exists>V. path_connectedin (product_topology X I) V \<and> z \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> F)"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   430
      by (metis (no_types, opaque_lifting) subsetI \<open>z \<in> Pi\<^sub>E I W\<close> \<open>Pi\<^sub>E I W \<subseteq> F\<close> * path_connectedin_PiE 
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   431
          PiE_iff PiE_mono order.trans)
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
  qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
  ultimately show ?thesis
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
    using False by blast
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
qed
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   437
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
   438
  assumes "is_interval S"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   439
  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
   440
  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
   441
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
   442
  fix a U
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   443
  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
   444
  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
   445
    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
   446
  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
   447
  proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   448
    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
   449
      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
   450
    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
   451
      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
   452
  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
   453
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   454
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   455
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
   456
 "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
   457
  "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
   458
  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
   459
  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
   460
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   461
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
   462
   "locally_path_connected_space (prod_topology X Y) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   463
      (prod_topology X Y) = trivial_topology \<or>
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   464
      locally_path_connected_space X \<and> locally_path_connected_space Y" (is "?lhs=?rhs")
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   465
proof (cases "(prod_topology X Y) = trivial_topology")
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   466
  case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   467
  then show ?thesis
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   468
    using locally_path_connected_space_discrete_topology by force
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   469
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   470
  case False
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   471
  then have ne: "X \<noteq> trivial_topology" "Y \<noteq> trivial_topology"
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   472
    by simp_all
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   473
  show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   474
  proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   475
    assume ?lhs then show ?rhs
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   476
      by (meson locally_path_connected_space_quotient_map_image ne(1) ne(2) quotient_map_fst quotient_map_snd)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   477
  next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   478
    assume ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   479
    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
   480
      by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   481
    show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   482
      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
   483
    proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   484
      fix UV x y
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   485
      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
   486
      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
   487
        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
   488
      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
   489
        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
   490
          "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
   491
        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
   492
      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
   493
      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
   494
        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
   495
        apply (rule_tac x="K \<times> L" in exI)
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   496
        apply (fastforce simp: openin_prod_Times_iff path_connectedin_Times)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   497
        done
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   498
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   499
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   500
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   501
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   502
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
   503
   "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
   504
    (\<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
   505
proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   506
  assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   507
    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
   508
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   509
  assume R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   510
  show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   511
  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
   512
    fix W i x
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   513
    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
   514
      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
   515
    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
   516
           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
   517
      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
   518
    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
   519
    proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   520
      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
   521
        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
   522
      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
   523
        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
   524
      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
   525
        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
   526
    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
   527
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   528
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   529
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   530
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   531
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
   532
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   533
definition weakly_locally_connected_at 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   534
  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
   535
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   536
definition locally_connected_at 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   537
  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
   538
           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
   539
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   540
definition locally_connected_space 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   541
  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
   542
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   543
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   544
lemma locally_connected_A: "(\<forall>U x. openin X U \<and> x \<in> U \<longrightarrow> openin X (connected_component_of_set (subtopology X U) x))
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   545
       \<Longrightarrow> neighbourhood_base_of (\<lambda>U. openin X U \<and> connectedin X U) X"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   546
  unfolding neighbourhood_base_of
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   547
  by (metis (full_types) connected_component_of_refl connectedin_connected_component_of connectedin_subtopology mem_Collect_eq openin_subset topspace_subtopology_subset)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   548
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   549
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
   550
          (\<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
   551
  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
   552
  apply (erule all_forward)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   553
  apply clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   554
  apply (subst openin_subopen)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   555
  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
   556
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   557
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
   558
  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
   559
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   560
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   561
lemma locally_connected_space_alt: 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   562
  "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
   563
  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
   564
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   565
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
   566
  "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   567
        (\<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
   568
              \<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
   569
  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
   570
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   571
lemma locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   572
   "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   573
     (\<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
   574
  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
   575
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   576
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
   577
   "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
   578
  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
   579
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   580
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
   581
  "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   582
   (\<forall>U C. openin X U \<and> C \<in> connected_components_of(subtopology X U) \<longrightarrow> openin X C)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   583
  unfolding connected_components_of_def locally_connected_space_eq_open_connected_component_of
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   584
  by (smt (verit, best) image_iff openin_subset topspace_subtopology_subset)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   585
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   586
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
   587
   "locally_connected_space X \<Longrightarrow> openin X (connected_component_of_set X x)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   588
  by (metis connected_component_of_eq_empty locally_connected_B openin_empty openin_topspace subtopology_topspace)
78037
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 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
   591
   "\<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
   592
  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
   593
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   594
lemma weakly_locally_connected_at:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   595
   "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
   596
    (\<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
   597
       \<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
   598
                (\<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
   599
proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   600
  assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   601
    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
   602
    by (meson subsetD subset_trans)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   603
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   604
  assume R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   605
  show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   606
    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
   607
  proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   608
    fix V
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   609
    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
   610
    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
   611
                  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
   612
      using R by force
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   613
    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
   614
    proof (intro conjI exI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   615
      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
   616
        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
   617
      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
   618
        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
   619
        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
   620
      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
   621
        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
   622
    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
   623
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   624
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   625
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   626
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
   627
  "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
   628
  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
   629
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   630
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
   631
   "locally_connected_space X \<longleftrightarrow>
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   632
    (\<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
   633
          \<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
   634
                    (\<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
   635
  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
   636
  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
   637
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   638
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
   639
   "\<lbrakk>locally_connected_space X; openin X S\<rbrakk> \<Longrightarrow> locally_connected_space (subtopology X S)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   640
  unfolding locally_connected_space_def neighbourhood_base_of
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   641
  by (smt (verit) connectedin_subtopology openin_open_subtopology subset_trans)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   642
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   643
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
   644
  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
   645
  shows "locally_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   646
  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
   647
proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   648
  fix V C
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   649
  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
   650
  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
   651
    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
   652
  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
   653
    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
   654
  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
   655
  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
   656
  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
   657
    fix x
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   658
    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
   659
    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
   660
    proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   661
      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
   662
        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
   663
                  openin_subset topspace_subtopology_subset)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   664
      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
   665
        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
   666
      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
   667
        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
   668
      moreover
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   669
      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
   670
      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
   671
        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
   672
          by (simp add: C)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   673
        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
   674
          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
   675
        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
   676
          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
   677
        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
   678
          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
   679
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   680
      ultimately
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   681
      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
   682
        by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   683
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   684
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   685
  then show "openin Y C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   686
    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
   687
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   688
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   689
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   690
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
   691
   "\<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
   692
        \<Longrightarrow> locally_connected_space Y"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   693
  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
   694
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   695
lemma homeomorphic_locally_connected_space:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   696
   "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
   697
  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
   698
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   699
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
   700
  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
   701
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   702
lemma locally_connected_is_realinterval:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   703
   "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
   704
  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
   705
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   706
lemma locally_connected_real_interval:
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   707
    "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
   708
    "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
   709
  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
   710
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   711
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
   712
   "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
   713
  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
   714
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   715
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
   716
   "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
   717
  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
   718
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   719
lemma weakly_locally_path_connected_imp_weakly_locally_connected_at:
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   720
   "weakly_locally_path_connected_at x X \<Longrightarrow> weakly_locally_connected_at x X"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   721
  by (metis path_connectedin_imp_connectedin weakly_locally_connected_at weakly_locally_path_connected_at)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   722
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   723
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   724
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
   725
  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
   726
    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
   727
  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
   728
proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   729
  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
   730
    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
   731
  show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   732
  proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   733
    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
   734
      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
   735
    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
   736
      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
   737
    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
   738
      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
   739
    proof (cases "C \<inter> D = {}")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   740
      case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   741
      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
   742
      proof (rule interior_of_maximal)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   743
        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
   744
          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
   745
        then show "D \<subseteq> C"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   746
          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
   747
        show "openin X D"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   748
          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
   749
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   750
      then show ?thesis 
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   751
        by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   752
    qed auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   753
    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
   754
      by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   755
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   756
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   757
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   758
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   759
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
   760
  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
   761
    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
   762
  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
   763
proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   764
  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
   765
    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
   766
  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
   767
    using assms
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   768
    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
   769
    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
   770
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   771
    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
   772
qed
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
(*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
   775
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
   776
   "locally_connected_space (prod_topology X Y) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   777
      (prod_topology X Y) = trivial_topology \<or>
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   778
      locally_connected_space X \<and> locally_connected_space Y" (is "?lhs=?rhs")
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   779
proof (cases "(prod_topology X Y) = trivial_topology")
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   780
  case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   781
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   782
    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
   783
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   784
  case False
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   785
  then have ne: "X \<noteq> trivial_topology" "Y \<noteq> trivial_topology"
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   786
    by simp_all
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   787
  show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   788
  proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   789
    assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   790
      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
   791
  next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   792
    assume ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   793
    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
   794
      by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   795
    show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   796
      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
   797
    proof clarify
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   798
      fix UV x y
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   799
      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
   800
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   801
     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
   802
        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
   803
      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
   804
        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
   805
          "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
   806
        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
   807
      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
   808
      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
   809
        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
   810
        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
   811
        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
   812
        done
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   813
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   814
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   815
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   816
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   817
(*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
   818
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
   819
   "locally_connected_space(product_topology X I) \<longleftrightarrow>
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   820
        (product_topology X I) = trivial_topology \<or>
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   821
        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
   822
        (\<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
   823
    (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
   824
proof (cases ?empty)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   825
  case True
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   826
  then show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   827
    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
   828
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   829
  case False
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   830
  then obtain z where z: "z \<in> (\<Pi>\<^sub>E i\<in>I. topspace (X i))"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   831
    using discrete_topology_unique_derived_set by (fastforce iff: null_topspace_iff_trivial)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   832
  have ?rhs if L: ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   833
  proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   834
    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
   835
      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
   836
      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
   837
      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
   838
      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
   839
    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
   840
      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
   841
      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
   842
    show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   843
    proof (intro conjI ballI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   844
      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
   845
      proof -
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   846
        have pc: "connectedin (X i) ((\<lambda>x. x i) ` C)"
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   847
          by (metis V connectedin_continuous_map_image continuous_map_product_projection that(1))
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   848
        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
   849
        proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   850
          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
   851
            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
   852
          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
   853
            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
   854
          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
   855
            using subU by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   856
          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
   857
            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
   858
        qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   859
        ultimately show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   860
          by (simp add: connectedin_topspace)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   861
      qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   862
      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
   863
        by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   864
      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
   865
        using finite_subset by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   866
    next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   867
      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
   868
        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
   869
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   870
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   871
  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
   872
  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
   873
    fix F z
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   874
    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
   875
    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
   876
            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
   877
      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
   878
    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
   879
                        (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
   880
                         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
   881
          (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
   882
    proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   883
      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
   884
      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
   885
        by (simp add: R \<open>i \<in> I\<close>)
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   886
      moreover have *: "openin (X i) (W i)" "z i \<in> W i"
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   887
        using \<open>z \<in> Pi\<^sub>E I W\<close> opeW \<open>i \<in> I\<close> by auto
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   888
      ultimately obtain U C where "openin (X i) U" "connectedin (X i) C" "z i \<in> U" "U \<subseteq> C" "C \<subseteq> W i"
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   889
        using \<open>i \<in> I\<close> by (force simp: locally_connected_space_def neighbourhood_base_of)
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   890
      then show "?\<Phi> i"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   891
        by (metis * connectedin_topspace openin_subset)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   892
    qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   893
    then obtain U C where
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   894
      *: "\<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
   895
                        (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
   896
                         \<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
   897
      by metis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   898
    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
   899
    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
   900
      by (clarsimp simp add: "*")
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   901
    moreover have "finite ?A"
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   902
      by (simp add: that finW)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   903
    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
   904
      using finite_subset by auto
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   905
    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
   906
      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
   907
    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
   908
            (\<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
   909
      using \<open>z \<in> Pi\<^sub>E I W\<close> \<open>Pi\<^sub>E I W \<subseteq> F\<close> *
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   910
      by (metis (no_types, opaque_lifting) PiE_iff PiE_mono connectedin_PiE subset_iff)
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   911
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   912
  ultimately show ?thesis
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   913
    using False by blast
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   914
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   915
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   916
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
   917
   "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
   918
    (\<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
   919
proof
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   920
  assume ?lhs then show ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   921
    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
   922
next
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   923
  assume R: ?rhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   924
  show ?lhs
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   925
  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
   926
    fix W i x
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   927
    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
   928
      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
   929
    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
   930
           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
   931
      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
   932
    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
   933
    proof (intro exI conjI)
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   934
      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
   935
        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
   936
      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
   937
        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
   938
      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
   939
        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
   940
    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
   941
  qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   942
qed
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   943
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   944
subsection \<open>Dimension of a topological space\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   945
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   946
text\<open>Basic definition of the small inductive dimension relation. Works in any topological space.\<close>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   947
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 78480
diff changeset
   948
inductive dimension_le :: "['a topology, int] \<Rightarrow> bool" (infix \<open>dim'_le\<close> 50) 
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   949
  where "\<lbrakk>-1 \<le> n;
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   950
        \<And>V a. \<lbrakk>openin X V; a \<in> V\<rbrakk> \<Longrightarrow> \<exists>U. a \<in> U \<and> U \<subseteq> V \<and> openin X U \<and> (subtopology X (X frontier_of U)) dim_le (n-1)\<rbrakk>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   951
              \<Longrightarrow> X dim_le (n::int)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   952
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   953
lemma dimension_le_neighbourhood_base:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   954
   "X dim_le n \<longleftrightarrow>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   955
   -1 \<le> n \<and> neighbourhood_base_of (\<lambda>U. openin X U \<and> (subtopology X (X frontier_of U)) dim_le (n-1)) X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   956
  by (smt (verit, best) dimension_le.simps open_neighbourhood_base_of)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   957
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   958
lemma dimension_le_bound: "X dim_le n \<Longrightarrow>-1 \<le> n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   959
  using dimension_le.simps by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   960
  
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   961
lemma dimension_le_mono [rule_format]:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   962
  assumes "X dim_le m"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   963
  shows "m \<le> n \<longrightarrow> X dim_le n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   964
  using assms
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   965
proof (induction arbitrary: n rule: dimension_le.induct)
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   966
qed (smt (verit) dimension_le.simps)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   967
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   968
inductive_simps dim_le_minus2 [simp]: "X dim_le -2"
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   969
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   970
lemma dimension_le_eq_empty [simp]:
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   971
   "X dim_le -1 \<longleftrightarrow> X = trivial_topology"
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   972
proof
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   973
  show "X dim_le (-1) \<Longrightarrow> X = trivial_topology"
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   974
    by (force intro: dimension_le.cases)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   975
next
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
   976
  show "X = trivial_topology \<Longrightarrow> X dim_le (-1)"
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   977
    using dimension_le.simps openin_subset by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   978
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   979
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   980
lemma dimension_le_0_neighbourhood_base_of_clopen:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   981
  "X dim_le 0 \<longleftrightarrow> neighbourhood_base_of (\<lambda>U. closedin X U \<and> openin X U) X"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   982
proof -
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   983
  have "(subtopology X (X frontier_of U) dim_le -1) = closedin X U" 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   984
      if "openin X U" for U
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   985
    using that clopenin_eq_frontier_of openin_subset 
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78250
diff changeset
   986
    by (fastforce simp add: subtopology_trivial_iff frontier_of_subset_topspace Int_absorb1)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   987
  then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   988
    by (smt (verit, del_insts) dimension_le.simps open_neighbourhood_base_of)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   989
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   990
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   991
lemma dimension_le_subtopology:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   992
  "X dim_le n \<Longrightarrow> subtopology X S dim_le n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   993
proof (induction arbitrary: S rule: dimension_le.induct)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   994
  case (1 n X)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   995
  show ?case 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   996
  proof (intro dimension_le.intros)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   997
    show "- 1 \<le> n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   998
      by (simp add: "1.hyps")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
   999
    fix U' a
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1000
    assume U': "openin (subtopology X S) U'" and "a \<in> U'"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1001
    then obtain U where U: "openin X U" "U' = U \<inter> S"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1002
      by (meson openin_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1003
    then obtain V where "a \<in> V" "V \<subseteq> U" "openin X V" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1004
      and subV: "subtopology X (X frontier_of V) dim_le n-1" 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1005
      and dimV: "\<And>T. subtopology X (X frontier_of V \<inter> T) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1006
      by (metis "1.IH" Int_iff \<open>a \<in> U'\<close> subtopology_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1007
    show "\<exists>W. a \<in> W \<and> W \<subseteq> U' \<and> openin (subtopology X S) W \<and> subtopology (subtopology X S) (subtopology X S frontier_of W) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1008
    proof (intro exI conjI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1009
      show "a \<in> S \<inter> V" "S \<inter> V \<subseteq> U'"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1010
        using \<open>U' = U \<inter> S\<close> \<open>a \<in> U'\<close> \<open>a \<in> V\<close> \<open>V \<subseteq> U\<close> by blast+
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1011
      show "openin (subtopology X S) (S \<inter> V)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1012
        by (simp add: \<open>openin X V\<close> openin_subtopology_Int2)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1013
      have "S \<inter> subtopology X S frontier_of V \<subseteq> X frontier_of V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1014
        by (simp add: frontier_of_subtopology_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1015
      then show "subtopology (subtopology X S) (subtopology X S frontier_of (S \<inter> V)) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1016
        by (metis dimV frontier_of_restrict inf.absorb_iff2 inf_left_idem subtopology_subtopology topspace_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1017
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1018
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1019
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1020
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1021
lemma dimension_le_subtopologies:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1022
   "\<lbrakk>subtopology X T dim_le n; S \<subseteq> T\<rbrakk> \<Longrightarrow> (subtopology X S) dim_le n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1023
  by (metis dimension_le_subtopology inf.absorb_iff2 subtopology_subtopology)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1024
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1025
lemma dimension_le_eq_subtopology:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1026
   "(subtopology X S) dim_le n \<longleftrightarrow>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1027
    -1 \<le> n \<and>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1028
    (\<forall>V a. openin X V \<and> a \<in> V \<and> a \<in> S
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1029
           \<longrightarrow> (\<exists>U. a \<in> U \<and> U \<subseteq> V \<and> openin X U \<and>
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1030
                    subtopology X (subtopology X S frontier_of (S \<inter> U)) dim_le (n-1)))"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1031
proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1032
  have *: "(\<exists>T. a \<in> T \<and> T \<inter> S \<subseteq> V \<inter> S \<and> openin X T \<and> subtopology X (S \<inter> (subtopology X S frontier_of (T \<inter> S))) dim_le n-1)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1033
       \<longleftrightarrow> (\<exists>U. a \<in> U \<and> U \<subseteq> V \<and> openin X U \<and> subtopology X (subtopology X S frontier_of (S \<inter> U)) dim_le n-1)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1034
    if "a \<in> V" "a \<in> S" "openin X V" for a V
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1035
  proof -
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1036
    have "\<exists>U. a \<in> U \<and> U \<subseteq> V \<and> openin X U \<and> subtopology X (subtopology X S frontier_of (S \<inter> U)) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1037
      if "a \<in> T" and sub: "T \<inter> S \<subseteq> V \<inter> S" and "openin X T"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1038
        and dim: "subtopology X (S \<inter> subtopology X S frontier_of (T \<inter> S)) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1039
      for T 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1040
    proof (intro exI conjI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1041
      show "openin X (T \<inter> V)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1042
        using \<open>openin X V\<close> \<open>openin X T\<close> by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1043
      show "subtopology X (subtopology X S frontier_of (S \<inter> (T \<inter> V))) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1044
        by (metis dim frontier_of_subset_subtopology inf.boundedE inf_absorb2 inf_assoc inf_commute sub)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1045
    qed (use \<open>a \<in> V\<close> \<open>a \<in> T\<close> in auto)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1046
    moreover have "\<exists>T. a \<in> T \<and> T \<inter> S \<subseteq> V \<inter> S \<and> openin X T \<and> subtopology X (S \<inter> subtopology X S frontier_of (T \<inter> S)) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1047
      if "a \<in> U" and "U \<subseteq> V" and "openin X U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1048
        and dim: "subtopology X (subtopology X S frontier_of (S \<inter> U)) dim_le n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1049
      for U
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1050
      by (metis that frontier_of_subset_subtopology inf_absorb2 inf_commute inf_le1 le_inf_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1051
    ultimately show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1052
      by safe
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1053
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1054
  show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1055
    apply (simp add: dimension_le.simps [of _ n] subtopology_subtopology openin_subtopology flip: *)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1056
    by (safe; metis Int_iff inf_le2 le_inf_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1057
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1058
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1059
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1060
lemma homeomorphic_space_dimension_le_aux:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1061
  assumes "X homeomorphic_space Y" "X dim_le of_nat n - 1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1062
  shows "Y dim_le of_nat n - 1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1063
  using assms
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1064
proof (induction n arbitrary: X Y)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1065
  case 0
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1066
  then show ?case
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1067
    by (simp add: dimension_le_eq_empty homeomorphic_empty_space)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1068
next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1069
  case (Suc n)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1070
  then have X_dim_n: "X dim_le n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1071
    by simp
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1072
  show ?case 
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1073
  proof (clarsimp simp add: dimension_le.simps [of Y n])
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1074
    fix V b
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1075
    assume "openin Y V" and "b \<in> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1076
    obtain f g where fg: "homeomorphic_maps X Y f g"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1077
      using \<open>X homeomorphic_space Y\<close> homeomorphic_space_def by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1078
    then have "openin X (g ` V)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1079
      using \<open>openin Y V\<close> homeomorphic_map_openness_eq homeomorphic_maps_map by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1080
    then obtain U where "g b \<in> U" "openin X U" and gim: "U \<subseteq> g ` V" and sub: "subtopology X (X frontier_of U) dim_le int n - int 1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1081
      using X_dim_n unfolding dimension_le.simps [of X n] by (metis \<open>b \<in> V\<close> imageI of_nat_eq_1_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1082
    show "\<exists>U. b \<in> U \<and> U \<subseteq> V \<and> openin Y U \<and> subtopology Y (Y frontier_of U) dim_le int n - 1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1083
    proof (intro conjI exI)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1084
      show "b \<in> f ` U"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1085
        by (metis (no_types, lifting) \<open>b \<in> V\<close> \<open>g b \<in> U\<close> \<open>openin Y V\<close> fg homeomorphic_maps_map image_iff openin_subset subsetD)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1086
      show "f ` U \<subseteq> V"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1087
        by (smt (verit, ccfv_threshold) \<open>openin Y V\<close> fg gim homeomorphic_maps_map image_iff openin_subset subset_iff)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1088
      show "openin Y (f ` U)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1089
        using \<open>openin X U\<close> fg homeomorphic_map_openness_eq homeomorphic_maps_map by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1090
      show "subtopology Y (Y frontier_of f ` U) dim_le int n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1091
      proof (rule Suc.IH)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1092
        have "homeomorphic_maps (subtopology X (X frontier_of U)) (subtopology Y (Y frontier_of f ` U)) f g"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1093
          using \<open>openin X U\<close> fg
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1094
          by (metis frontier_of_subset_topspace homeomorphic_map_frontier_of homeomorphic_maps_map homeomorphic_maps_subtopologies openin_subset topspace_subtopology topspace_subtopology_subset)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1095
        then show "subtopology X (X frontier_of U) homeomorphic_space subtopology Y (Y frontier_of f ` U)"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1096
          using homeomorphic_space_def by blast
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1097
        show "subtopology X (X frontier_of U) dim_le int n-1"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1098
          using sub by fastforce
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1099
      qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1100
    qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1101
  qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1102
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1103
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1104
lemma homeomorphic_space_dimension_le:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1105
  assumes "X homeomorphic_space Y"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1106
  shows "X dim_le n \<longleftrightarrow> Y dim_le n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1107
proof (cases "n \<ge> -1")
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1108
  case True
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1109
  then show ?thesis
78480
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1110
    using homeomorphic_space_dimension_le_aux [of _ _ "nat(n+1)"] 
b22f39c54e8c Tidied up more messy proofs
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1111
    by (smt (verit) assms homeomorphic_space_sym nat_eq_iff)
78250
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1112
next
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1113
  case False
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1114
  then show ?thesis
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1115
    by (metis dimension_le_bound)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1116
qed
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1117
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1118
lemma dimension_le_retraction_map_image:
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1119
   "\<lbrakk>retraction_map X Y r; X dim_le n\<rbrakk> \<Longrightarrow> Y dim_le n"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1120
  by (meson dimension_le_subtopology homeomorphic_space_dimension_le retraction_map_def retraction_maps_section_image2)
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1121
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1122
lemma dimension_le_discrete_topology [simp]: "(discrete_topology U) dim_le 0"
400aecdfd71f Another tranche of HOL Light material on metric and topological spaces
paulson <lp15@cam.ac.uk>
parents: 78037
diff changeset
  1123
  using dimension_le.simps dimension_le_eq_empty by fastforce
78037
37894dff0111 More material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
  1124
69945
35ba13ac6e5c New abstract topological material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
end