src/HOL/Analysis/Connected.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 82538 4b132ea7d575
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
66835
ecc99a5a1ab8 fixed markup
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
     1
(*  Author:     L C Paulson, University of Cambridge
ecc99a5a1ab8 fixed markup
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
     2
    Material split off from Topology_Euclidean_Space
ecc99a5a1ab8 fixed markup
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
     3
*)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
     5
chapter \<open>Connected Components\<close>
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Connected
69544
5aa5a8d6e5b5 split off theorems involving classes below metric_space and real_normed_vector
immler
parents: 69529
diff changeset
     8
  imports
69617
63ee37c519a3 reduced dependencies of Connected.thy
immler
parents: 69616
diff changeset
     9
    Abstract_Topology_2
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
    12
subsection\<^marker>\<open>tag unimportant\<close> \<open>Connectedness\<close>
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
lemma connected_local:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
 "connected S \<longleftrightarrow>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
  \<not> (\<exists>e1 e2.
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
    17
      openin (top_of_set S) e1 \<and>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
    18
      openin (top_of_set S) e2 \<and>
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
      S \<subseteq> e1 \<union> e2 \<and>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
      e1 \<inter> e2 = {} \<and>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
      e1 \<noteq> {} \<and>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
      e2 \<noteq> {})"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
    23
  using connected_openin by blast
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
lemma exists_diff:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  fixes P :: "'a set \<Rightarrow> bool"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
  shows "(\<exists>S. P (- S)) \<longleftrightarrow> (\<exists>S. P S)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
    28
  by (metis boolean_algebra_class.boolean_algebra.double_compl)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
lemma connected_clopen: "connected S \<longleftrightarrow>
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
    31
  (\<forall>T. openin (top_of_set S) T \<and>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
    32
     closedin (top_of_set S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
proof -
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  have "\<not> connected S \<longleftrightarrow>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
    (\<exists>e1 e2. open e1 \<and> open (- e2) \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
    unfolding connected_def openin_open closedin_closed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
    by (metis double_complement)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  then have th0: "connected S \<longleftrightarrow>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
    \<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
    (is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)")
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
    41
    unfolding closed_def by metis
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    (is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
    unfolding connected_def openin_open closedin_closed by auto
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
  have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" for e2
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
  proof -
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
    have "?P e2 e1 \<longleftrightarrow> (\<exists>t. closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t \<noteq> S)" for e1
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
      by auto
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
    then show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
      by metis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
  qed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  then show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
    by (simp add: th0 th1)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
qed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
subsection \<open>Connected components, considered as a connectedness relation or a set\<close>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    58
definition\<^marker>\<open>tag important\<close> "connected_component S x y \<equiv> \<exists>T. connected T \<and> T \<subseteq> S \<and> x \<in> T \<and> y \<in> T"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    60
abbreviation "connected_component_set S x \<equiv> Collect (connected_component S x)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
lemma connected_componentI:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    63
  "connected T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> x \<in> T \<Longrightarrow> y \<in> T \<Longrightarrow> connected_component S x y"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  by (auto simp: connected_component_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    66
lemma connected_component_in: "connected_component S x y \<Longrightarrow> x \<in> S \<and> y \<in> S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  by (auto simp: connected_component_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    69
lemma connected_component_refl: "x \<in> S \<Longrightarrow> connected_component S x x"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
    70
  using connected_component_def connected_sing by blast
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
    71
 
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    72
lemma connected_component_refl_eq [simp]: "connected_component S x x \<longleftrightarrow> x \<in> S"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
    73
  using connected_component_in connected_component_refl by blast
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    75
lemma connected_component_sym: "connected_component S x y \<Longrightarrow> connected_component S y x"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
  by (auto simp: connected_component_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
lemma connected_component_trans:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    79
  "connected_component S x y \<Longrightarrow> connected_component S y z \<Longrightarrow> connected_component S x z"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  unfolding connected_component_def
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  by (metis Int_iff Un_iff Un_subset_iff equals0D connected_Un)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
lemma connected_component_of_subset:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    84
  "connected_component S x y \<Longrightarrow> S \<subseteq> T \<Longrightarrow> connected_component T x y"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  by (auto simp: connected_component_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    87
lemma connected_component_Union: "connected_component_set S x = \<Union>{T. connected T \<and> x \<in> T \<and> T \<subseteq> S}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
  by (auto simp: connected_component_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    90
lemma connected_connected_component [iff]: "connected (connected_component_set S x)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  by (auto simp: connected_component_Union intro: connected_Union)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
lemma connected_iff_eq_connected_component_set:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
    94
  "connected S \<longleftrightarrow> (\<forall>x \<in> S. connected_component_set S x = S)"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
    95
proof 
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
    96
  show "\<forall>x\<in>S. connected_component_set S x = S \<Longrightarrow> connected S"
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
    97
    by (metis connectedI_const connected_connected_component)
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
    98
qed (auto simp: connected_component_def)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   100
lemma connected_component_subset: "connected_component_set S x \<subseteq> S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  using connected_component_in by blast
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   103
lemma connected_component_eq_self: "connected S \<Longrightarrow> x \<in> S \<Longrightarrow> connected_component_set S x = S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  by (simp add: connected_iff_eq_connected_component_set)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
lemma connected_iff_connected_component:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   107
  "connected S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. connected_component S x y)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  using connected_component_in by (auto simp: connected_iff_eq_connected_component_set)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
lemma connected_component_maximal:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   111
  "x \<in> T \<Longrightarrow> connected T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> T \<subseteq> (connected_component_set S x)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  using connected_component_eq_self connected_component_of_subset by blast
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
lemma connected_component_mono:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   115
  "S \<subseteq> T \<Longrightarrow> connected_component_set S x \<subseteq> connected_component_set T x"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  by (simp add: Collect_mono connected_component_of_subset)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   118
lemma connected_component_eq_empty [simp]: "connected_component_set S x = {} \<longleftrightarrow> x \<notin> S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
  using connected_component_refl by (fastforce simp: connected_component_in)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
lemma connected_component_set_empty [simp]: "connected_component_set {} x = {}"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  using connected_component_eq_empty by blast
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
lemma connected_component_eq:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   125
  "y \<in> connected_component_set S x \<Longrightarrow> (connected_component_set S y = connected_component_set S x)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  by (metis (no_types, lifting)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
      Collect_cong connected_component_sym connected_component_trans mem_Collect_eq)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
lemma closed_connected_component:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   130
  assumes S: "closed S"
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   131
  shows "closed (connected_component_set S x)"
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   132
proof (cases "x \<in> S")
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  case False
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  then show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
    by (metis connected_component_eq_empty closed_empty)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
next
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  case True
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
  show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
    unfolding closure_eq [symmetric]
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  proof
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   141
    show "closure (connected_component_set S x) \<subseteq> connected_component_set S x"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   142
    proof (rule connected_component_maximal)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   143
      show "x \<in> closure (connected_component_set S x)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   144
        by (simp add: closure_def True)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   145
      show "connected (closure (connected_component_set S x))"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   146
        by (simp add: connected_imp_connected_closure)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   147
      show "closure (connected_component_set S x) \<subseteq> S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   148
        by (simp add: S closure_minimal connected_component_subset)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   149
    qed  
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   150
  qed (simp add: closure_subset)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
qed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
lemma connected_component_disjoint:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   154
  "connected_component_set S a \<inter> connected_component_set S b = {} \<longleftrightarrow>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   155
    a \<notin> connected_component_set S b"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   156
  using connected_component_eq connected_component_sym by fastforce
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
lemma connected_component_nonoverlap:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   159
  "connected_component_set S a \<inter> connected_component_set S b = {} \<longleftrightarrow>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   160
    a \<notin> S \<or> b \<notin> S \<or> connected_component_set S a \<noteq> connected_component_set S b"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   161
  by (metis connected_component_disjoint connected_component_eq connected_component_eq_empty inf.idem)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
lemma connected_component_overlap:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   164
  "connected_component_set S a \<inter> connected_component_set S b \<noteq> {} \<longleftrightarrow>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   165
    a \<in> S \<and> b \<in> S \<and> connected_component_set S a = connected_component_set S b"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
  by (auto simp: connected_component_nonoverlap)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   168
lemma connected_component_sym_eq: "connected_component S x y \<longleftrightarrow> connected_component S y x"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  using connected_component_sym by blast
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
lemma connected_component_eq_eq:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   172
  "connected_component_set S x = connected_component_set S y \<longleftrightarrow>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   173
    x \<notin> S \<and> y \<notin> S \<or> x \<in> S \<and> y \<in> S \<and> connected_component S x y"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   174
  by (metis connected_component_eq connected_component_eq_empty connected_component_refl mem_Collect_eq)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   175
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
lemma connected_iff_connected_component_eq:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   178
  "connected S \<longleftrightarrow> (\<forall>x \<in> S. \<forall>y \<in> S. connected_component_set S x = connected_component_set S y)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  by (simp add: connected_component_eq_eq connected_iff_connected_component)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
lemma connected_component_idemp:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   182
  "connected_component_set (connected_component_set S x) x = connected_component_set S x"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   183
proof
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   184
  show "connected_component_set S x \<subseteq> connected_component_set (connected_component_set S x) x"
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   185
    by (metis connected_component_eq_empty connected_component_maximal order.refl
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   186
        connected_component_refl connected_connected_component mem_Collect_eq)
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   187
qed (simp add: connected_component_subset)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
lemma connected_component_unique:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   190
  "\<lbrakk>x \<in> c; c \<subseteq> S; connected c;
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   191
    \<And>c'. \<lbrakk>x \<in> c'; c' \<subseteq> S; connected c'\<rbrakk> \<Longrightarrow> c' \<subseteq> c\<rbrakk>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   192
        \<Longrightarrow> connected_component_set S x = c"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   193
  by (simp add: connected_component_maximal connected_component_subset subsetD
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   194
      subset_antisym)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
lemma joinable_connected_component_eq:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   197
  "\<lbrakk>connected T; T \<subseteq> S;
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   198
    connected_component_set S x \<inter> T \<noteq> {};
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   199
    connected_component_set S y \<inter> T \<noteq> {}\<rbrakk>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   200
    \<Longrightarrow> connected_component_set S x = connected_component_set S y"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   201
  by (metis (full_types) subsetD connected_component_eq connected_component_maximal disjoint_iff_not_equal)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   203
lemma Union_connected_component: "\<Union>(connected_component_set S ` S) = S"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   204
proof
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   205
  show "\<Union>(connected_component_set S ` S) \<subseteq> S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   206
    by (simp add: SUP_least connected_component_subset)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   207
qed (use connected_component_refl_eq in force)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
lemma complement_connected_component_unions:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   210
    "S - connected_component_set S x =
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   211
     \<Union>(connected_component_set S ` S - {connected_component_set S x})"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   212
    (is "?lhs = ?rhs")
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   213
proof
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   214
  show "?lhs \<subseteq> ?rhs"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   215
    by (metis Diff_subset Diff_subset_conv Sup_insert Union_connected_component insert_Diff_single)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   216
  show "?rhs \<subseteq> ?lhs"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   217
    by clarsimp (metis connected_component_eq_eq connected_component_in)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   218
qed
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
lemma connected_component_intermediate_subset:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   221
        "\<lbrakk>connected_component_set U a \<subseteq> T; T \<subseteq> U\<rbrakk>
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   222
        \<Longrightarrow> connected_component_set T a = connected_component_set U a"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   223
  by (metis connected_component_idemp connected_component_mono subset_antisym)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   224
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   225
lemma connected_component_homeomorphismI:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   226
  assumes "homeomorphism A B f g" "connected_component A x y"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   227
  shows   "connected_component B (f x) (f y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   228
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   229
  from assms obtain T where T: "connected T" "T \<subseteq> A" "x \<in> T" "y \<in> T"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   230
    unfolding connected_component_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   231
  have "connected (f ` T)" "f ` T \<subseteq> B" "f x \<in> f ` T" "f y \<in> f ` T"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   232
    using assms T continuous_on_subset[of A f T]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   233
    by (auto intro!: connected_continuous_image simp: homeomorphism_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   234
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   235
    unfolding connected_component_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   236
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   237
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   238
lemma connected_component_homeomorphism_iff:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   239
  assumes "homeomorphism A B f g"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   240
  shows   "connected_component A x y \<longleftrightarrow> x \<in> A \<and> y \<in> A \<and> connected_component B (f x) (f y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   241
  by (metis assms connected_component_homeomorphismI connected_component_in homeomorphism_apply1 homeomorphism_sym)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   242
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   243
lemma connected_component_set_homeomorphism:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   244
  assumes "homeomorphism A B f g" "x \<in> A"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   245
  shows   "connected_component_set B (f x) = f ` connected_component_set A x"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   246
proof -
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   247
  have "\<And>y. connected_component B (f x) y
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   248
          \<Longrightarrow> \<exists>u. u \<in> A \<and> connected_component B (f x) (f u) \<and> y = f u"
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   249
    using assms by (metis connected_component_in homeomorphism_def image_iff)
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   250
  with assms show ?thesis
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   251
    by (auto simp: image_iff connected_component_homeomorphism_iff)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   252
qed
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
subsection \<open>The set of connected components of a set\<close>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   256
definition\<^marker>\<open>tag important\<close> components:: "'a::topological_space set \<Rightarrow> 'a set set"
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   257
  where "components S \<equiv> connected_component_set S ` S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   259
lemma components_iff: "S \<in> components U \<longleftrightarrow> (\<exists>x. x \<in> U \<and> S = connected_component_set U x)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
  by (auto simp: components_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   262
lemma componentsI: "x \<in> U \<Longrightarrow> connected_component_set U x \<in> components U"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
  by (auto simp: components_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
lemma componentsE:
72228
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   266
  assumes "S \<in> components U"
aa7cb84983e9 minor tidying, also s->S and t->T
paulson <lp15@cam.ac.uk>
parents: 71172
diff changeset
   267
  obtains x where "x \<in> U" "S = connected_component_set U x"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  using assms by (auto simp: components_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   270
lemma Union_components [simp]: "\<Union>(components U) = U"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   271
  by (simp add: Union_connected_component components_def)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   273
lemma pairwise_disjoint_components: "pairwise (\<lambda>X Y. X \<inter> Y = {}) (components U)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   274
  unfolding pairwise_def
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   275
  by (metis (full_types) components_iff connected_component_nonoverlap)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   277
lemma in_components_nonempty: "C \<in> components S \<Longrightarrow> C \<noteq> {}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
    by (metis components_iff connected_component_eq_empty)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   280
lemma in_components_subset: "C \<in> components S \<Longrightarrow> C \<subseteq> S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
  using Union_components by blast
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   283
lemma in_components_connected: "C \<in> components S \<Longrightarrow> connected C"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
  by (metis components_iff connected_connected_component)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
lemma in_components_maximal:
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   287
  "C \<in> components S \<longleftrightarrow>
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   288
    C \<noteq> {} \<and> C \<subseteq> S \<and> connected C \<and> (\<forall>D. D \<noteq> {} \<and> C \<subseteq> D \<and> D \<subseteq> S \<and> connected D \<longrightarrow> D = C)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   289
(is "?lhs \<longleftrightarrow> ?rhs")
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   290
proof
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   291
  assume L: ?lhs
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   292
  then have "C \<subseteq> S" "connected C"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   293
    by (simp_all add: in_components_subset in_components_connected)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   294
  then show ?rhs
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   295
    by (metis (full_types) L components_iff connected_component_maximal connected_component_refl empty_iff mem_Collect_eq subsetD subset_antisym)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   296
next
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   297
  show "?rhs \<Longrightarrow> ?lhs"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   298
    by (metis bot.extremum componentsI connected_component_maximal connected_component_subset
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   299
        connected_connected_component order_antisym_conv subset_iff)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   300
qed
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   301
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
lemma joinable_components_eq:
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   304
  "connected T \<and> T \<subseteq> S \<and> c1 \<in> components S \<and> c2 \<in> components S \<and> c1 \<inter> T \<noteq> {} \<and> c2 \<inter> T \<noteq> {} \<Longrightarrow> c1 = c2"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
  by (metis (full_types) components_iff joinable_connected_component_eq)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   307
lemma closed_components: "\<lbrakk>closed S; C \<in> components S\<rbrakk> \<Longrightarrow> closed C"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
  by (metis closed_connected_component components_iff)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
lemma components_nonoverlap:
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   311
    "\<lbrakk>C \<in> components S; C' \<in> components S\<rbrakk> \<Longrightarrow> (C \<inter> C' = {}) \<longleftrightarrow> (C \<noteq> C')"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   312
  by (metis (full_types) components_iff connected_component_nonoverlap)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   314
lemma components_eq: "\<lbrakk>C \<in> components S; C' \<in> components S\<rbrakk> \<Longrightarrow> (C = C' \<longleftrightarrow> C \<inter> C' \<noteq> {})"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
  by (metis components_nonoverlap)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   317
lemma components_eq_empty [simp]: "components S = {} \<longleftrightarrow> S = {}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
  by (simp add: components_def)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
lemma components_empty [simp]: "components {} = {}"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
  by simp
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   323
lemma connected_eq_connected_components_eq: "connected S \<longleftrightarrow> (\<forall>C \<in> components S. \<forall>C' \<in> components S. C = C')"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 72228
diff changeset
   324
  by (metis (no_types, opaque_lifting) components_iff connected_component_eq_eq connected_iff_connected_component)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   326
lemma components_eq_sing_iff: "components S = {S} \<longleftrightarrow> connected S \<and> S \<noteq> {}" (is "?lhs \<longleftrightarrow> ?rhs")
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   327
proof
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   328
  show "?rhs \<Longrightarrow> ?lhs"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   329
    by (metis components_iff connected_component_eq_self equals0I insert_iff mk_disjoint_insert)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   330
qed (use in_components_connected in fastforce)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   332
lemma components_eq_sing_exists: "(\<exists>a. components S = {a}) \<longleftrightarrow> connected S \<and> S \<noteq> {}"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   333
  by (metis Union_components ccpo_Sup_singleton components_eq_sing_iff)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   335
lemma connected_eq_components_subset_sing: "connected S \<longleftrightarrow> components S \<subseteq> {S}"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   336
  by (metis components_eq_empty components_eq_sing_iff connected_empty subset_singleton_iff)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   338
lemma connected_eq_components_subset_sing_exists: "connected S \<longleftrightarrow> (\<exists>a. components S \<subseteq> {a})"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   339
  by (metis components_eq_sing_exists connected_eq_components_subset_sing subset_singleton_iff)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   341
lemma in_components_self: "S \<in> components S \<longleftrightarrow> connected S \<and> S \<noteq> {}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
  by (metis components_empty components_eq_sing_iff empty_iff in_components_connected insertI1)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   344
lemma components_maximal: "\<lbrakk>C \<in> components S; connected T; T \<subseteq> S; C \<inter> T \<noteq> {}\<rbrakk> \<Longrightarrow> T \<subseteq> C"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   345
  by (metis (lifting) ext Int_Un_eq(4) Int_absorb Un_upper1 bot_eq_sup_iff connected_Un
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   346
      in_components_maximal sup.mono sup.orderI)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   348
lemma exists_component_superset: "\<lbrakk>T \<subseteq> S; S \<noteq> {}; connected T\<rbrakk> \<Longrightarrow> \<exists>C. C \<in> components S \<and> T \<subseteq> C"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   349
  by (meson componentsI connected_component_maximal equals0I subset_eq)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   351
lemma components_intermediate_subset: "\<lbrakk>S \<in> components U; S \<subseteq> T; T \<subseteq> U\<rbrakk> \<Longrightarrow> S \<in> components T"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   352
  by (smt (verit, best) dual_order.trans in_components_maximal)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   354
lemma in_components_unions_complement: "C \<in> components S \<Longrightarrow> S - C = \<Union>(components S - {C})"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  by (metis complement_connected_component_unions components_def components_iff)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
lemma connected_intermediate_closure:
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   358
  assumes cs: "connected S" and st: "S \<subseteq> T" and ts: "T \<subseteq> closure S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   359
  shows "connected T"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   360
  using assms unfolding connected_def
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   361
  by (smt (verit) Int_assoc inf.absorb_iff2 inf_bot_left open_Int_closure_eq_empty)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   363
lemma closedin_connected_component: "closedin (top_of_set S) (connected_component_set S x)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   364
proof (cases "connected_component_set S x = {}")
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
  case True
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
  then show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
    by (metis closedin_empty)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
next
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
  case False
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   370
  then obtain y where y: "connected_component S x y" and "x \<in> S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   371
    using connected_component_eq_empty by blast
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   372
  have *: "connected_component_set S x \<subseteq> S \<inter> closure (connected_component_set S x)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
    by (auto simp: closure_def connected_component_in)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   374
  have **: "x \<in> closure (connected_component_set S x)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   375
    by (simp add: \<open>x \<in> S\<close> closure_def)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   376
  have "S \<inter> closure (connected_component_set S x) \<subseteq> connected_component_set S x" if "connected_component S x y"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   377
  proof (rule connected_component_maximal)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   378
    show "connected (S \<inter> closure (connected_component_set S x))"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   379
      using "*" connected_intermediate_closure by blast
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   380
  qed (use \<open>x \<in> S\<close> ** in auto)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
  with y * show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
    by (auto simp: closedin_closed)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
qed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
66939
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   385
lemma closedin_component:
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   386
   "C \<in> components S \<Longrightarrow> closedin (top_of_set S) C"
66939
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   387
  using closedin_connected_component componentsE by blast
04678058308f New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents: 66884
diff changeset
   388
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   390
subsection\<^marker>\<open>tag unimportant\<close> \<open>Proving a function is constant on a connected set
69615
e502cd4d7062 moved material from Connected.thy to more appropriate places
immler
parents: 69614
diff changeset
   391
  by proving that a level set is open\<close>
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
lemma continuous_levelset_openin_cases:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   395
  shows "connected S \<Longrightarrow> continuous_on S f \<Longrightarrow>
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   396
        openin (top_of_set S) {x \<in> S. f x = a}
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   397
        \<Longrightarrow> (\<forall>x \<in> S. f x \<noteq> a) \<or> (\<forall>x \<in> S. f x = a)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
  unfolding connected_clopen
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
  using continuous_closedin_preimage_constant by auto
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
lemma continuous_levelset_openin:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   403
  shows "connected S \<Longrightarrow> continuous_on S f \<Longrightarrow>
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   404
        openin (top_of_set S) {x \<in> S. f x = a} \<Longrightarrow>
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   405
        (\<exists>x \<in> S. f x = a)  \<Longrightarrow> (\<forall>x \<in> S. f x = a)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   406
  using continuous_levelset_openin_cases[of S f ]
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
  by meson
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
lemma continuous_levelset_open:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
  fixes f :: "_ \<Rightarrow> 'b::t1_space"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   411
  assumes S: "connected S" "continuous_on S f"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   412
    and a: "open {x \<in> S. f x = a}" "\<exists>x \<in> S.  f x = a"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   413
  shows "\<forall>x \<in> S. f x = a"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   414
  using a continuous_levelset_openin[OF S, of a, unfolded openin_open]
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
  by fast
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   418
subsection\<^marker>\<open>tag unimportant\<close> \<open>Preservation of Connectedness\<close>
67727
ce3e87a51488 moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents: 67707
diff changeset
   419
ce3e87a51488 moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents: 67707
diff changeset
   420
lemma homeomorphic_connectedness:
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   421
  assumes "S homeomorphic T"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   422
  shows "connected S \<longleftrightarrow> connected T"
67727
ce3e87a51488 moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents: 67707
diff changeset
   423
using assms unfolding homeomorphic_def homeomorphism_def by (metis connected_continuous_image)
ce3e87a51488 moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents: 67707
diff changeset
   424
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   425
lemma connected_monotone_quotient_preimage:
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   426
  assumes "connected T"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   427
      and contf: "continuous_on S f" and fim: "f ` S = T"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   428
      and opT: "\<And>U. U \<subseteq> T
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   429
                 \<Longrightarrow> openin (top_of_set S) (S \<inter> f -` U) \<longleftrightarrow>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   430
                     openin (top_of_set T) U"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   431
      and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   432
    shows "connected S"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   433
proof (rule connectedI)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   434
  fix U V
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   435
  assume "open U" and "open V" and "U \<inter> S \<noteq> {}" and "V \<inter> S \<noteq> {}"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   436
    and "U \<inter> V \<inter> S = {}" and "S \<subseteq> U \<union> V"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   437
  moreover
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   438
  have disjoint: "f ` (S \<inter> U) \<inter> f ` (S \<inter> V) = {}"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   439
  proof -
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   440
    have False if "y \<in> f ` (S \<inter> U) \<inter> f ` (S \<inter> V)" for y
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   441
    proof -
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   442
      have "y \<in> T"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   443
        using fim that by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   444
      show ?thesis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   445
        using connectedD [OF connT [OF \<open>y \<in> T\<close>] \<open>open U\<close> \<open>open V\<close>]
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   446
              \<open>S \<subseteq> U \<union> V\<close> \<open>U \<inter> V \<inter> S = {}\<close> that by fastforce
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   447
    qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   448
    then show ?thesis by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   449
  qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   450
  ultimately have UU: "(S \<inter> f -` f ` (S \<inter> U)) = S \<inter> U" and VV: "(S \<inter> f -` f ` (S \<inter> V)) = S \<inter> V"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   451
    by auto
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   452
  have opeU: "openin (top_of_set T) (f ` (S \<inter> U))"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   453
    by (metis UU \<open>open U\<close> fim image_Int_subset le_inf_iff opT openin_open_Int)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   454
  have opeV: "openin (top_of_set T) (f ` (S \<inter> V))"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   455
    by (metis opT fim VV \<open>open V\<close> openin_open_Int image_Int_subset inf.bounded_iff)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   456
  have "T \<subseteq> f ` (S \<inter> U) \<union> f ` (S \<inter> V)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   457
    using \<open>S \<subseteq> U \<union> V\<close> fim by auto
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   458
  then show False
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   459
    using \<open>connected T\<close> disjoint opeU opeV \<open>U \<inter> S \<noteq> {}\<close> \<open>V \<inter> S \<noteq> {}\<close>
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   460
    by (auto simp: connected_openin)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   461
qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   462
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   463
lemma connected_open_monotone_preimage:
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   464
  assumes contf: "continuous_on S f" and fim: "f ` S = T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   465
    and ST: "\<And>C. openin (top_of_set S) C \<Longrightarrow> openin (top_of_set T) (f ` C)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   466
    and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   467
    and "connected C" "C \<subseteq> T"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   468
  shows "connected (S \<inter> f -` C)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   469
proof -
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   470
  have contf': "continuous_on (S \<inter> f -` C) f"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   471
    by (meson contf continuous_on_subset inf_le1)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   472
  have eqC: "f ` (S \<inter> f -` C) = C"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   473
    using \<open>C \<subseteq> T\<close> fim by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   474
  show ?thesis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   475
  proof (rule connected_monotone_quotient_preimage [OF \<open>connected C\<close> contf' eqC])
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   476
    show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   477
      by (metis Int_assoc Int_empty_right Int_insert_right_if1 assms(6) connT in_mono that vimage_Int)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   478
    have "\<And>U. openin (top_of_set (S \<inter> f -` C)) U
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   479
               \<Longrightarrow> openin (top_of_set C) (f ` U)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   480
      using open_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   481
    then show "\<And>D. D \<subseteq> C
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   482
          \<Longrightarrow> openin (top_of_set (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) =
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   483
              openin (top_of_set C) D"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   484
      using open_map_imp_quotient_map [of "(S \<inter> f -` C)" f] contf' by (simp add: eqC)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   485
  qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   486
qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   487
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   488
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   489
lemma connected_closed_monotone_preimage:
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   490
  assumes contf: "continuous_on S f" and fim: "f ` S = T"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   491
    and ST: "\<And>C. closedin (top_of_set S) C \<Longrightarrow> closedin (top_of_set T) (f ` C)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   492
    and connT: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   493
    and "connected C" "C \<subseteq> T"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   494
  shows "connected (S \<inter> f -` C)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   495
proof -
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   496
  have contf': "continuous_on (S \<inter> f -` C) f"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   497
    by (meson contf continuous_on_subset inf_le1)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   498
  have eqC: "f ` (S \<inter> f -` C) = C"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   499
    using \<open>C \<subseteq> T\<close> fim by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   500
  show ?thesis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   501
  proof (rule connected_monotone_quotient_preimage [OF \<open>connected C\<close> contf' eqC])
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   502
    show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   503
      by (metis Int_assoc Int_empty_right Int_insert_right_if1 \<open>C \<subseteq> T\<close> connT subsetD that vimage_Int)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   504
    have "\<And>U. closedin (top_of_set (S \<inter> f -` C)) U
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   505
               \<Longrightarrow> closedin (top_of_set C) (f ` U)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   506
      using closed_map_restrict [OF _ ST \<open>C \<subseteq> T\<close>] by metis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   507
    then show "\<And>D. D \<subseteq> C
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   508
          \<Longrightarrow> openin (top_of_set (S \<inter> f -` C)) (S \<inter> f -` C \<inter> f -` D) =
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   509
              openin (top_of_set C) D"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   510
      using closed_map_imp_quotient_map [of "(S \<inter> f -` C)" f] contf' by (simp add: eqC)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   511
  qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   512
qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   513
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   514
71137
nipkow
parents: 70136
diff changeset
   515
subsection\<open>Lemmas about components\<close>
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   516
71137
nipkow
parents: 70136
diff changeset
   517
text  \<open>See Newman IV, 3.3 and 3.4.\<close>
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   518
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   519
lemma connected_Un_clopen_in_complement:
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   520
  fixes S U :: "'a::metric_space set"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   521
  assumes "connected S" "connected U" "S \<subseteq> U" 
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   522
      and opeT: "openin (top_of_set (U - S)) T" 
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   523
      and cloT: "closedin (top_of_set (U - S)) T"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   524
    shows "connected (S \<union> T)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   525
proof -
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   526
  have *: "\<lbrakk>\<And>x y. P x y \<longleftrightarrow> P y x; \<And>x y. P x y \<Longrightarrow> S \<subseteq> x \<or> S \<subseteq> y;
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69325
diff changeset
   527
            \<And>x y. \<lbrakk>P x y; S \<subseteq> x\<rbrakk> \<Longrightarrow> False\<rbrakk> \<Longrightarrow> \<not>(\<exists>x y. (P x y))" for P
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   528
    by metis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   529
  show ?thesis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   530
    unfolding connected_closedin_eq
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   531
  proof (rule *)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   532
    fix H1 H2
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   533
    assume H: "closedin (top_of_set (S \<union> T)) H1 \<and> 
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   534
               closedin (top_of_set (S \<union> T)) H2 \<and>
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   535
               H1 \<union> H2 = S \<union> T \<and> H1 \<inter> H2 = {} \<and> H1 \<noteq> {} \<and> H2 \<noteq> {}"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   536
    then have clo: "closedin (top_of_set S) (S \<inter> H1)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   537
                   "closedin (top_of_set S) (S \<inter> H2)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   538
      by (metis Un_upper1 closedin_closed_subset inf_commute)+
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   539
    moreover have "S \<inter> ((S \<union> T) \<inter> H1) \<union> S \<inter> ((S \<union> T) \<inter> H2) = S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   540
      using H by blast
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   541
    moreover have "H1 \<inter> (S \<inter> ((S \<union> T) \<inter> H2)) = {}"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   542
      using H by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   543
    ultimately have "S \<inter> H1 = {} \<or> S \<inter> H2 = {}"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   544
      by (smt (verit) Int_assoc \<open>connected S\<close> connected_closedin_eq inf_commute inf_sup_absorb)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   545
    then show "S \<subseteq> H1 \<or> S \<subseteq> H2"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   546
      using H \<open>connected S\<close> unfolding connected_closedin by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   547
  next
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   548
    fix H1 H2
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   549
    assume H: "closedin (top_of_set (S \<union> T)) H1 \<and>
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   550
               closedin (top_of_set (S \<union> T)) H2 \<and>
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   551
               H1 \<union> H2 = S \<union> T \<and> H1 \<inter> H2 = {} \<and> H1 \<noteq> {} \<and> H2 \<noteq> {}" 
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   552
       and "S \<subseteq> H1"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   553
    then have H2T: "H2 \<subseteq> T"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   554
      by auto
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   555
    have "T \<subseteq> U"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   556
      using Diff_iff opeT openin_imp_subset by auto
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   557
    with \<open>S \<subseteq> U\<close> have Ueq: "U = (U - S) \<union> (S \<union> T)" 
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   558
      by auto
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   559
    have "openin (top_of_set ((U - S) \<union> (S \<union> T))) H2"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   560
    proof (rule openin_subtopology_Un)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   561
      show "openin (top_of_set (S \<union> T)) H2"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   562
        by (metis Diff_cancel H Un_Diff Un_Diff_Int closedin_subset openin_closedin_eq topspace_euclidean_subtopology)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   563
      then show "openin (top_of_set (U - S)) H2"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   564
        by (meson H2T Un_upper2 opeT openin_subset_trans openin_trans)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   565
    qed
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   566
    moreover have "closedin (top_of_set ((U - S) \<union> (S \<union> T))) H2"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   567
    proof (rule closedin_subtopology_Un)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   568
      show "closedin (top_of_set (U - S)) H2"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   569
        using H H2T cloT closedin_subset_trans 
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   570
        by (blast intro: closedin_subtopology_Un closedin_trans)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   571
    qed (simp add: H)
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   572
    ultimately have H2: "H2 = {} \<or> H2 = U"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   573
      using Ueq \<open>connected U\<close> unfolding connected_clopen by metis   
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   574
    then have "H2 \<subseteq> S"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   575
      by (metis Diff_partition H Un_Diff_cancel Un_subset_iff \<open>H2 \<subseteq> T\<close> assms(3) inf.orderE opeT openin_imp_subset)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   576
    moreover have "T \<subseteq> H2 - S"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   577
      by (metis (no_types) H2 H opeT openin_closedin_eq topspace_euclidean_subtopology)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   578
    ultimately show False
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   579
      using H \<open>S \<subseteq> H1\<close> by blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   580
  qed blast
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   581
qed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   582
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   583
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68527
diff changeset
   584
proposition component_diff_connected:
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   585
  fixes S :: "'a::metric_space set"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   586
  assumes "connected S" "connected U" "S \<subseteq> U" and C: "C \<in> components (U - S)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   587
  shows "connected(U - C)"
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68527
diff changeset
   588
  using \<open>connected S\<close> unfolding connected_closedin_eq not_ex de_Morgan_conj
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   589
proof clarify
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   590
  fix H3 H4 
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   591
  assume clo3: "closedin (top_of_set (U - C)) H3" 
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   592
    and clo4: "closedin (top_of_set (U - C)) H4" 
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   593
    and H34: "H3 \<union> H4 = U - C" "H3 \<inter> H4 = {}" and "H3 \<noteq> {}" and "H4 \<noteq> {}"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   594
    and * [rule_format]: "\<forall>H1 H2. \<not> closedin (top_of_set S) H1 \<or>
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   595
                      \<not> closedin (top_of_set S) H2 \<or>
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   596
                      H1 \<union> H2 \<noteq> S \<or> H1 \<inter> H2 \<noteq> {} \<or> \<not> H1 \<noteq> {} \<or> \<not> H2 \<noteq> {}"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   597
  then have "H3 \<subseteq> U-C" and ope3: "openin (top_of_set (U - C)) (U - C - H3)"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   598
    and "H4 \<subseteq> U-C" and ope4: "openin (top_of_set (U - C)) (U - C - H4)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   599
    by (auto simp: closedin_def)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   600
  have "C \<noteq> {}" "C \<subseteq> U-S" "connected C"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   601
    using C in_components_nonempty in_components_subset in_components_maximal by blast+
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   602
  have cCH3: "connected (C \<union> H3)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   603
  proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo3])
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   604
    show "openin (top_of_set (U - C)) H3"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   605
      by (metis Diff_cancel Un_Diff Un_Diff_Int \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> ope4)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   606
  qed (use clo3 \<open>C \<subseteq> U - S\<close> in auto)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   607
  have cCH4: "connected (C \<union> H4)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   608
  proof (rule connected_Un_clopen_in_complement [OF \<open>connected C\<close> \<open>connected U\<close> _ _ clo4])
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   609
    show "openin (top_of_set (U - C)) H4"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   610
      by (metis Diff_cancel Diff_triv Int_Un_eq(2) Un_Diff H34 inf_commute ope3)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   611
  qed (use clo4 \<open>C \<subseteq> U - S\<close> in auto)
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69617
diff changeset
   612
  have "closedin (top_of_set S) (S \<inter> H3)" "closedin (top_of_set S) (S \<inter> H4)"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   613
    using clo3 clo4 \<open>S \<subseteq> U\<close> \<open>C \<subseteq> U - S\<close> by (auto simp: closedin_closed)
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   614
  moreover have "S \<inter> H3 \<noteq> {}"      
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   615
    using components_maximal [OF C cCH3] \<open>C \<noteq> {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H3 \<noteq> {}\<close> \<open>H3 \<subseteq> U - C\<close> by auto
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   616
  moreover have "S \<inter> H4 \<noteq> {}"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   617
    using components_maximal [OF C cCH4] \<open>C \<noteq> {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H4 \<noteq> {}\<close> \<open>H4 \<subseteq> U - C\<close> by auto
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   618
  ultimately show False
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   619
    using * [of "S \<inter> H3" "S \<inter> H4"] \<open>H3 \<inter> H4 = {}\<close> \<open>C \<subseteq> U - S\<close> \<open>H3 \<union> H4 = U - C\<close> \<open>S \<subseteq> U\<close> 
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   620
    by auto
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   621
qed
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   624
subsection\<^marker>\<open>tag unimportant\<close>\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
lemma continuous_disconnected_range_constant:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
  assumes S: "connected S"
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
      and conf: "continuous_on S f"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   631
      and fim: "f \<in> S \<rightarrow> T"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   632
      and cct: "\<And>y. y \<in> T \<Longrightarrow> connected_component_set T y = {y}"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   633
    shows "f constant_on S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
proof (cases "S = {}")
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   635
  case True then show ?thesis
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   636
    by (simp add: constant_on_def)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
next
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
  case False
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   639
  then have "f ` S \<subseteq> {f x}" if "x \<in> S" for x
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   640
    by (metis PiE S cct connected_component_maximal connected_continuous_image [OF conf] fim image_eqI 
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   641
        image_subset_iff that)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
  with False show ?thesis
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   643
    unfolding constant_on_def by blast
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
qed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
text\<open>This proof requires the existence of two separate values of the range type.\<close>
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
lemma finite_range_constant_imp_connected:
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
  assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   650
              \<lbrakk>continuous_on S f; finite(f ` S)\<rbrakk> \<Longrightarrow> f constant_on S"
82538
4b132ea7d575 More tidying and some variable renaming
paulson <lp15@cam.ac.uk>
parents: 78475
diff changeset
   651
  shows "connected S"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
proof -
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   653
  { fix T U
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   654
    assume clt: "closedin (top_of_set S) T"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   655
       and clu: "closedin (top_of_set S) U"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   656
       and tue: "T \<inter> U = {}" and tus: "T \<union> U = S"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   657
    have "continuous_on (T \<union> U) (\<lambda>x. if x \<in> T then 0 else 1)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   658
      using clt clu tue by (intro continuous_on_cases_local) (auto simp: tus)
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   659
    then have conif: "continuous_on S (\<lambda>x. if x \<in> T then 0 else 1)"
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   660
      using tus by blast
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   661
    have fi: "finite ((\<lambda>x. if x \<in> T then 0 else 1) ` S)"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
      by (rule finite_subset [of _ "{0,1}"]) auto
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   663
    have "T = {} \<or> U = {}"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
      using assms [OF conif fi] tus [symmetric]
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66835
diff changeset
   665
      by (auto simp: Ball_def constant_on_def) (metis IntI empty_iff one_neq_zero tue)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
  }
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
  then show ?thesis
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
    by (simp add: connected_closedin_eq)
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
qed
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
69617
63ee37c519a3 reduced dependencies of Connected.thy
immler
parents: 69616
diff changeset
   671
end