| author | wenzelm | 
| Mon, 01 Mar 2021 14:58:00 +0100 | |
| changeset 73330 | 0fb889c361e6 | 
| parent 72228 | aa7cb84983e9 | 
| child 73932 | fd21b4a93043 | 
| permissions | -rw-r--r-- | 
| 66835 | 1  | 
(* Author: L C Paulson, University of Cambridge  | 
2  | 
Material split off from Topology_Euclidean_Space  | 
|
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  | 
|
| 69617 | 5  | 
section \<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 | 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 | 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> {})"
 | 
| 
 
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
 | 
23  | 
unfolding connected_def openin_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
 | 
24  | 
by safe 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
 | 
25  | 
|
| 
 
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  | 
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
 | 
27  | 
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
 | 
28  | 
shows "(\<exists>S. P (- S)) \<longleftrightarrow> (\<exists>S. P 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
 | 
29  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
| 
 
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  | 
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
 | 
31  | 
have ?rhs if ?lhs  | 
| 
 
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
 | 
32  | 
using that 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
 | 
33  | 
moreover have "P (- (- S))" if "P S" for 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
 | 
34  | 
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
 | 
35  | 
have "S = - (- S)" 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
 | 
36  | 
with that show ?thesis 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
 | 
37  | 
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
 | 
38  | 
ultimately show ?thesis 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
 | 
39  | 
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
 | 
40  | 
|
| 
 
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
 | 
41  | 
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
 | 
42  | 
(\<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
 | 
43  | 
     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
 | 
44  | 
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
 | 
45  | 
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
 | 
46  | 
    (\<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
 | 
47  | 
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
 | 
48  | 
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
 | 
49  | 
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
 | 
50  | 
    \<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
 | 
51  | 
(is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 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
 | 
52  | 
by (simp add: closed_def) 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
 | 
53  | 
  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
 | 
54  | 
(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
 | 
55  | 
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
 | 
56  | 
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
 | 
57  | 
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
 | 
58  | 
    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
 | 
59  | 
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
 | 
60  | 
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
 | 
61  | 
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
 | 
62  | 
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
 | 
63  | 
then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 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
 | 
64  | 
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
 | 
65  | 
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
 | 
66  | 
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
 | 
67  | 
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
 | 
68  | 
|
| 
 
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
 | 
69  | 
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
 | 
70  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
71  | 
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
 | 
72  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
73  | 
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
 | 
74  | 
|
| 
 
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
 | 
75  | 
lemma connected_componentI:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
76  | 
"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
 | 
77  | 
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
 | 
78  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
79  | 
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
 | 
80  | 
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
 | 
81  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
82  | 
lemma connected_component_refl: "x \<in> S \<Longrightarrow> connected_component S x 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
 | 
83  | 
by (auto simp: connected_component_def) (use connected_sing in 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
 | 
84  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
85  | 
lemma connected_component_refl_eq [simp]: "connected_component S x x \<longleftrightarrow> 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
 | 
86  | 
by (auto simp: connected_component_refl) (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
 | 
87  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
88  | 
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
 | 
89  | 
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
 | 
90  | 
|
| 
 
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  | 
lemma connected_component_trans:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
92  | 
"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
 | 
93  | 
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
 | 
94  | 
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
 | 
95  | 
|
| 
 
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
 | 
96  | 
lemma connected_component_of_subset:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
97  | 
"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
 | 
98  | 
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
 | 
99  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
100  | 
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
 | 
101  | 
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
 | 
102  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
103  | 
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
 | 
104  | 
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
 | 
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_eq_connected_component_set:  | 
| 
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. connected_component_set S x = S)"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
108  | 
proof (cases "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
 | 
109  | 
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
 | 
110  | 
then show ?thesis 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
 | 
111  | 
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
 | 
112  | 
case False  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
113  | 
then obtain x where "x \<in> S" by 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
 | 
114  | 
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
 | 
115  | 
proof  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
116  | 
assume "connected S"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
117  | 
then show "\<forall>x \<in> S. 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
 | 
118  | 
by (force 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
 | 
119  | 
next  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
120  | 
assume "\<forall>x \<in> S. connected_component_set S x = S"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
121  | 
then show "connected S"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
122  | 
by (metis \<open>x \<in> S\<close> connected_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
 | 
123  | 
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
 | 
124  | 
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
 | 
125  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
126  | 
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
 | 
127  | 
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
 | 
128  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
129  | 
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
 | 
130  | 
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
 | 
131  | 
|
| 
 
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
 | 
132  | 
lemma connected_iff_connected_component:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
133  | 
"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
 | 
134  | 
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
 | 
135  | 
|
| 
 
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  | 
lemma connected_component_maximal:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
137  | 
"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
 | 
138  | 
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
 | 
139  | 
|
| 
 
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  | 
lemma connected_component_mono:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
141  | 
"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
 | 
142  | 
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
 | 
143  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
144  | 
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
 | 
145  | 
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
 | 
146  | 
|
| 
 
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
 | 
147  | 
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
 | 
148  | 
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
 | 
149  | 
|
| 
 
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
 | 
150  | 
lemma connected_component_eq:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
151  | 
"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
 | 
152  | 
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
 | 
153  | 
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
 | 
154  | 
|
| 
 
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
 | 
155  | 
lemma closed_connected_component:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
156  | 
assumes S: "closed S"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
157  | 
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
 | 
158  | 
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
 | 
159  | 
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
 | 
160  | 
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
 | 
161  | 
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
 | 
162  | 
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
 | 
163  | 
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
 | 
164  | 
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
 | 
165  | 
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
 | 
166  | 
proof  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
167  | 
show "closure (connected_component_set S x) \<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
 | 
168  | 
apply (rule connected_component_maximal)  | 
| 
 
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  | 
apply (simp add: closure_def 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
 | 
170  | 
apply (simp add: connected_imp_connected_closure)  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
171  | 
apply (simp add: S closure_minimal 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
 | 
172  | 
done  | 
| 
 
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
 | 
173  | 
next  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
174  | 
show "connected_component_set S x \<subseteq> 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
 | 
175  | 
by (simp add: closure_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
 | 
176  | 
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
 | 
177  | 
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
 | 
178  | 
|
| 
 
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  | 
lemma connected_component_disjoint:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
180  | 
  "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
 | 
181  | 
a \<notin> 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
 | 
182  | 
apply (auto simp: 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
 | 
183  | 
using connected_component_eq connected_component_sym  | 
| 
 
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
 | 
184  | 
apply 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
 | 
185  | 
done  | 
| 
 
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
 | 
186  | 
|
| 
 
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
 | 
187  | 
lemma connected_component_nonoverlap:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
188  | 
  "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
 | 
189  | 
a \<notin> S \<or> b \<notin> S \<or> connected_component_set S a \<noteq> 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
 | 
190  | 
apply (auto 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
 | 
191  | 
using connected_component_refl_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
 | 
192  | 
apply 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
 | 
193  | 
apply (metis connected_component_eq 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
 | 
194  | 
apply (metis connected_component_eq 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
 | 
195  | 
done  | 
| 
 
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  | 
|
| 
 
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
 | 
197  | 
lemma connected_component_overlap:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
198  | 
  "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
 | 
199  | 
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
 | 
200  | 
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
 | 
201  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
202  | 
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
 | 
203  | 
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
 | 
204  | 
|
| 
 
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
 | 
205  | 
lemma connected_component_eq_eq:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
206  | 
"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
 | 
207  | 
x \<notin> S \<and> y \<notin> S \<or> x \<in> S \<and> y \<in> S \<and> connected_component S x y"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
208  | 
apply (cases "y \<in> S", simp)  | 
| 
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
 | 
209  | 
apply (metis connected_component_eq connected_component_eq_empty connected_component_refl_eq mem_Collect_eq)  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
210  | 
apply (cases "x \<in> S", simp)  | 
| 
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
 | 
211  | 
apply (metis 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
 | 
212  | 
using 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
 | 
213  | 
apply 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
 | 
214  | 
done  | 
| 
 
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
 | 
215  | 
|
| 
 
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
 | 
216  | 
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
 | 
217  | 
"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
 | 
218  | 
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
 | 
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_idemp:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
221  | 
"connected_component_set (connected_component_set S x) x = 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
 | 
222  | 
apply (rule subset_antisym)  | 
| 
 
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
 | 
223  | 
apply (simp add: connected_component_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
 | 
224  | 
apply (metis connected_component_eq_empty connected_component_maximal  | 
| 
 
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
 | 
225  | 
connected_component_refl_eq connected_connected_component mem_Collect_eq set_eq_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
 | 
226  | 
done  | 
| 
 
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
 | 
227  | 
|
| 
 
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
 | 
228  | 
lemma connected_component_unique:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
229  | 
"\<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
 | 
230  | 
\<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
 | 
231  | 
\<Longrightarrow> connected_component_set S x = c"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
232  | 
apply (rule subset_antisym)  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
233  | 
apply (meson connected_component_maximal connected_component_subset connected_connected_component contra_subsetD)  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
234  | 
by (simp add: connected_component_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
 | 
235  | 
|
| 
 
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
 | 
236  | 
lemma joinable_connected_component_eq:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
237  | 
"\<lbrakk>connected T; T \<subseteq> S;  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
238  | 
    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
 | 
239  | 
    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
 | 
240  | 
\<Longrightarrow> 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
 | 
241  | 
apply (simp add: ex_in_conv [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
 | 
242  | 
apply (rule 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
 | 
243  | 
by (metis (no_types, hide_lams) connected_component_eq_eq connected_component_in connected_component_maximal subsetD 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
 | 
244  | 
|
| 
 
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
 | 
245  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
246  | 
lemma Union_connected_component: "\<Union>(connected_component_set S ` 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
 | 
247  | 
apply (rule subset_antisym)  | 
| 
 
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
 | 
248  | 
apply (simp add: SUP_least connected_component_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
 | 
249  | 
using connected_component_refl_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
 | 
250  | 
by force  | 
| 
 
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
 | 
251  | 
|
| 
 
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
 | 
252  | 
|
| 
 
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  | 
lemma complement_connected_component_unions:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
254  | 
"S - connected_component_set S x =  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
255  | 
     \<Union>(connected_component_set S ` S - {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
 | 
256  | 
apply (subst Union_connected_component [symmetric], 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
 | 
257  | 
apply (metis connected_component_eq_eq 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
 | 
258  | 
by (metis connected_component_eq 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
 | 
259  | 
|
| 
 
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  | 
lemma connected_component_intermediate_subset:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
261  | 
"\<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
 | 
262  | 
\<Longrightarrow> connected_component_set T a = connected_component_set U a"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
263  | 
apply (case_tac "a \<in> 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
 | 
264  | 
apply (simp add: connected_component_maximal connected_component_mono subset_antisym)  | 
| 
 
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  | 
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
 | 
266  | 
|
| 
 
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
 | 
267  | 
|
| 
 
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  | 
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
 | 
269  | 
|
| 70136 | 270  | 
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
 | 
271  | 
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
 | 
272  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
273  | 
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
 | 
274  | 
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
 | 
275  | 
|
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
276  | 
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
 | 
277  | 
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
 | 
278  | 
|
| 
 
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  | 
lemma componentsE:  | 
| 
72228
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
280  | 
assumes "S \<in> components U"  | 
| 
 
aa7cb84983e9
minor tidying, also s->S and t->T
 
paulson <lp15@cam.ac.uk> 
parents: 
71172 
diff
changeset
 | 
281  | 
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
 | 
282  | 
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
 | 
283  | 
|
| 
 
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  | 
lemma Union_components [simp]: "\<Union>(components u) = u"  | 
| 
 
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  | 
apply (rule subset_antisym)  | 
| 
 
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  | 
using Union_connected_component components_def apply fastforce  | 
| 
 
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
 | 
287  | 
apply (metis Union_connected_component components_def set_eq_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
 | 
288  | 
done  | 
| 
 
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
 | 
289  | 
|
| 
 
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
 | 
290  | 
lemma pairwise_disjoint_components: "pairwise (\<lambda>X Y. X \<inter> Y = {}) (components u)"
 | 
| 
 
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
 | 
291  | 
apply (simp add: pairwise_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
 | 
292  | 
apply (auto simp: 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
 | 
293  | 
apply (metis connected_component_eq_eq 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
 | 
294  | 
done  | 
| 
 
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
 | 
295  | 
|
| 
 
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
 | 
296  | 
lemma in_components_nonempty: "c \<in> components s \<Longrightarrow> c \<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
 | 
297  | 
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
 | 
298  | 
|
| 
 
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
 | 
299  | 
lemma in_components_subset: "c \<in> components s \<Longrightarrow> c \<subseteq> 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
 | 
300  | 
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
 | 
301  | 
|
| 
 
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  | 
lemma in_components_connected: "c \<in> components s \<Longrightarrow> connected c"  | 
| 
 
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  | 
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
 | 
304  | 
|
| 
 
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  | 
lemma in_components_maximal:  | 
| 
 
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  | 
"c \<in> components 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
 | 
307  | 
    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)"
 | 
| 
 
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  | 
apply (rule iffI)  | 
| 
 
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  | 
apply (simp add: in_components_nonempty in_components_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
 | 
310  | 
apply (metis (full_types) components_iff connected_component_eq_self connected_component_intermediate_subset connected_component_refl in_components_subset mem_Collect_eq rev_subsetD)  | 
| 
 
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
 | 
311  | 
apply (metis bot.extremum_uniqueI components_iff connected_component_eq_empty connected_component_maximal connected_component_subset connected_connected_component subset_emptyI)  | 
| 
 
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
 | 
312  | 
done  | 
| 
 
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  | 
|
| 
 
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
 | 
314  | 
lemma joinable_components_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
 | 
315  | 
  "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"
 | 
| 
 
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  | 
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
 | 
317  | 
|
| 
 
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  | 
lemma closed_components: "\<lbrakk>closed s; c \<in> components s\<rbrakk> \<Longrightarrow> closed c"  | 
| 
 
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  | 
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
 | 
320  | 
|
| 
 
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  | 
lemma 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
 | 
322  | 
    "\<lbrakk>c \<in> components s; c' \<in> components s\<rbrakk> \<Longrightarrow> (c \<inter> c' = {}) \<longleftrightarrow> (c \<noteq> c')"
 | 
| 
 
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
 | 
323  | 
apply (auto simp: in_components_nonempty 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
 | 
324  | 
using connected_component_refl apply 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
 | 
325  | 
apply (metis connected_component_eq_eq 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
 | 
326  | 
by (metis connected_component_eq 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
 | 
327  | 
|
| 
 
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
 | 
328  | 
lemma components_eq: "\<lbrakk>c \<in> components s; c' \<in> components s\<rbrakk> \<Longrightarrow> (c = c' \<longleftrightarrow> c \<inter> c' \<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
 | 
329  | 
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
 | 
330  | 
|
| 
 
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  | 
lemma components_eq_empty [simp]: "components s = {} \<longleftrightarrow> 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
 | 
332  | 
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
 | 
333  | 
|
| 
 
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  | 
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
 | 
335  | 
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
 | 
336  | 
|
| 
 
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  | 
lemma connected_eq_connected_components_eq: "connected s \<longleftrightarrow> (\<forall>c \<in> components s. \<forall>c' \<in> components s. c = c')"  | 
| 
 
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
 | 
338  | 
by (metis (no_types, hide_lams) components_iff 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
 | 
339  | 
|
| 
 
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  | 
lemma components_eq_sing_iff: "components s = {s} \<longleftrightarrow> connected s \<and> 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
 | 
341  | 
apply (rule iffI)  | 
| 
 
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  | 
using in_components_connected apply fastforce  | 
| 
 
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  | 
apply safe  | 
| 
 
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
 | 
344  | 
using Union_components apply fastforce  | 
| 
 
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
 | 
345  | 
apply (metis components_iff connected_component_eq_self)  | 
| 
 
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
 | 
346  | 
using in_components_maximal  | 
| 
 
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  | 
apply 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
 | 
348  | 
done  | 
| 
 
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
 | 
349  | 
|
| 
 
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  | 
lemma components_eq_sing_exists: "(\<exists>a. components s = {a}) \<longleftrightarrow> connected s \<and> 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
 | 
351  | 
apply (rule iffI)  | 
| 
 
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
 | 
352  | 
using connected_eq_connected_components_eq apply fastforce  | 
| 
 
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  | 
apply (metis components_eq_sing_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
 | 
354  | 
done  | 
| 
 
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  | 
|
| 
 
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  | 
lemma connected_eq_components_subset_sing: "connected s \<longleftrightarrow> components s \<subseteq> {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
 | 
357  | 
by (metis Union_components components_empty components_eq_sing_iff connected_empty insert_subset order_refl subset_singletonD)  | 
| 
 
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
 | 
358  | 
|
| 
 
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
 | 
359  | 
lemma connected_eq_components_subset_sing_exists: "connected s \<longleftrightarrow> (\<exists>a. components s \<subseteq> {a})"
 | 
| 
 
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
 | 
360  | 
by (metis components_eq_sing_exists connected_eq_components_subset_sing empty_iff subset_iff subset_singletonD)  | 
| 
 
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
 | 
361  | 
|
| 
 
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  | 
lemma in_components_self: "s \<in> components s \<longleftrightarrow> connected s \<and> 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
 | 
363  | 
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
 | 
364  | 
|
| 
 
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  | 
lemma components_maximal: "\<lbrakk>c \<in> components s; connected t; t \<subseteq> s; c \<inter> t \<noteq> {}\<rbrakk> \<Longrightarrow> t \<subseteq> c"
 | 
| 
 
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  | 
apply (simp add: components_def ex_in_conv [symmetric], clarify)  | 
| 
 
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 (meson connected_component_def connected_component_trans)  | 
| 
 
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  | 
|
| 
 
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  | 
lemma exists_component_superset: "\<lbrakk>t \<subseteq> s; s \<noteq> {}; connected t\<rbrakk> \<Longrightarrow> \<exists>c. c \<in> components s \<and> t \<subseteq> c"
 | 
| 
 
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
 | 
370  | 
  apply (cases "t = {}", force)
 | 
| 
 
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
 | 
371  | 
apply (metis components_def ex_in_conv connected_component_maximal contra_subsetD image_eqI)  | 
| 
 
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
 | 
372  | 
done  | 
| 
 
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  | 
|
| 
 
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
 | 
374  | 
lemma components_intermediate_subset: "\<lbrakk>s \<in> components u; s \<subseteq> t; t \<subseteq> u\<rbrakk> \<Longrightarrow> s \<in> components 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
 | 
375  | 
apply (auto simp: 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
 | 
376  | 
apply (metis connected_component_eq_empty connected_component_intermediate_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
 | 
377  | 
done  | 
| 
 
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
 | 
378  | 
|
| 
 
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
 | 
379  | 
lemma in_components_unions_complement: "c \<in> components s \<Longrightarrow> s - c = \<Union>(components s - {c})"
 | 
| 
 
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
 | 
380  | 
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
 | 
381  | 
|
| 
 
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  | 
lemma connected_intermediate_closure:  | 
| 
 
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  | 
assumes cs: "connected s" and st: "s \<subseteq> t" and ts: "t \<subseteq> closure 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
 | 
384  | 
shows "connected 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
 | 
385  | 
proof (rule connectedI)  | 
| 
 
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
 | 
386  | 
fix A B  | 
| 
 
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
 | 
387  | 
  assume A: "open A" and B: "open B" and Alap: "A \<inter> t \<noteq> {}" and Blap: "B \<inter> t \<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
 | 
388  | 
    and disj: "A \<inter> B \<inter> t = {}" and cover: "t \<subseteq> A \<union> B"
 | 
| 
 
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  | 
  have disjs: "A \<inter> B \<inter> 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
 | 
390  | 
using disj st 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
 | 
391  | 
  have "A \<inter> closure 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
 | 
392  | 
using Alap Int_absorb1 ts 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
 | 
393  | 
  then have Alaps: "A \<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
 | 
394  | 
by (simp add: A open_Int_closure_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
 | 
395  | 
  have "B \<inter> closure 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
 | 
396  | 
using Blap Int_absorb1 ts 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
 | 
397  | 
  then have Blaps: "B \<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
 | 
398  | 
by (simp add: B open_Int_closure_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
 | 
399  | 
then show 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
 | 
400  | 
using cs [unfolded connected_def] A B disjs Alaps Blaps cover st  | 
| 
 
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  | 
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
 | 
402  | 
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
 | 
403  | 
|
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
404  | 
lemma closedin_connected_component: "closedin (top_of_set s) (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
 | 
405  | 
proof (cases "connected_component_set s 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
 | 
406  | 
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
 | 
407  | 
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
 | 
408  | 
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
 | 
409  | 
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
 | 
410  | 
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
 | 
411  | 
then obtain y where y: "connected_component s x y"  | 
| 
 
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
 | 
412  | 
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
 | 
413  | 
have *: "connected_component_set s x \<subseteq> s \<inter> closure (connected_component_set s 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
 | 
414  | 
by (auto simp: closure_def 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
 | 
415  | 
have "connected_component s x y \<Longrightarrow> s \<inter> closure (connected_component_set s x) \<subseteq> connected_component_set s 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
 | 
416  | 
apply (rule connected_component_maximal, 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
 | 
417  | 
using closure_subset connected_component_in apply fastforce  | 
| 
 
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
 | 
418  | 
using * connected_intermediate_closure apply 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
 | 
419  | 
done  | 
| 
 
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
 | 
420  | 
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
 | 
421  | 
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
 | 
422  | 
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
 | 
423  | 
|
| 
66939
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
424  | 
lemma closedin_component:  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
425  | 
"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
 | 
426  | 
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
 | 
427  | 
|
| 
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
 | 
428  | 
|
| 70136 | 429  | 
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
 | 
430  | 
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
 | 
431  | 
|
| 
 
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
 | 
432  | 
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
 | 
433  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 
 
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
 | 
434  | 
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
435  | 
        openin (top_of_set s) {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
 | 
436  | 
\<Longrightarrow> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"  | 
| 
 
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
 | 
437  | 
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
 | 
438  | 
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
 | 
439  | 
|
| 
 
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
 | 
440  | 
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
 | 
441  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 
 
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
 | 
442  | 
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
443  | 
        openin (top_of_set s) {x \<in> s. f x = a} \<Longrightarrow>
 | 
| 
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
 | 
444  | 
(\<exists>x \<in> s. f x = a) \<Longrightarrow> (\<forall>x \<in> s. f x = a)"  | 
| 
 
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
 | 
445  | 
using continuous_levelset_openin_cases[of s f ]  | 
| 
 
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
 | 
446  | 
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
 | 
447  | 
|
| 
 
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
 | 
448  | 
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
 | 
449  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 
 
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
 | 
450  | 
assumes "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
 | 
451  | 
and "continuous_on s f"  | 
| 
 
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
 | 
452  | 
    and "open {x \<in> s. f x = a}"
 | 
| 
 
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
 | 
453  | 
and "\<exists>x \<in> s. f x = a"  | 
| 
 
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
 | 
454  | 
shows "\<forall>x \<in> s. f x = a"  | 
| 
 
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
 | 
455  | 
using continuous_levelset_openin[OF assms(1,2), of a, unfolded openin_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
 | 
456  | 
using assms (3,4)  | 
| 
 
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
 | 
457  | 
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
 | 
458  | 
|
| 
 
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
 | 
459  | 
|
| 70136 | 460  | 
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
 | 
461  | 
|
| 
 
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
 | 
462  | 
lemma homeomorphic_connectedness:  | 
| 
 
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
 | 
463  | 
assumes "s homeomorphic t"  | 
| 
 
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
 | 
464  | 
shows "connected s \<longleftrightarrow> connected t"  | 
| 
 
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
 | 
465  | 
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
 | 
466  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
467  | 
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
 | 
468  | 
assumes "connected T"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
469  | 
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
 | 
470  | 
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
 | 
471  | 
\<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
 | 
472  | 
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
 | 
473  | 
      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
 | 
474  | 
shows "connected S"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
475  | 
proof (rule connectedI)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
476  | 
fix U V  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
477  | 
  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
 | 
478  | 
    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
 | 
479  | 
moreover  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
480  | 
  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
 | 
481  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
482  | 
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
 | 
483  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
484  | 
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
 | 
485  | 
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
 | 
486  | 
show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
487  | 
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
 | 
488  | 
              \<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
 | 
489  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
490  | 
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
 | 
491  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
492  | 
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
 | 
493  | 
by auto  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
494  | 
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
 | 
495  | 
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
 | 
496  | 
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
 | 
497  | 
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
 | 
498  | 
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
 | 
499  | 
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
 | 
500  | 
then show False  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
501  | 
    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
 | 
502  | 
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
 | 
503  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
504  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
505  | 
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
 | 
506  | 
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
 | 
507  | 
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
 | 
508  | 
    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
 | 
509  | 
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
 | 
510  | 
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
 | 
511  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
512  | 
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
 | 
513  | 
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
 | 
514  | 
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
 | 
515  | 
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
 | 
516  | 
show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
517  | 
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
 | 
518  | 
    show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
519  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
520  | 
      have "S \<inter> f -` C \<inter> f -` {y} = 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
 | 
521  | 
using that by blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
522  | 
      moreover have "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
 | 
523  | 
using \<open>C \<subseteq> T\<close> connT that by blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
524  | 
ultimately show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
525  | 
by metis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
526  | 
qed  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
527  | 
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
 | 
528  | 
\<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
 | 
529  | 
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
 | 
530  | 
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
 | 
531  | 
\<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
 | 
532  | 
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
 | 
533  | 
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
 | 
534  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
535  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
536  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
537  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
538  | 
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
 | 
539  | 
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
 | 
540  | 
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
 | 
541  | 
    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
 | 
542  | 
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
 | 
543  | 
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
 | 
544  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
545  | 
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
 | 
546  | 
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
 | 
547  | 
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
 | 
548  | 
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
 | 
549  | 
show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
550  | 
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
 | 
551  | 
    show "connected (S \<inter> f -` C \<inter> f -` {y})" if "y \<in> C" for y
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
552  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
553  | 
      have "S \<inter> f -` C \<inter> f -` {y} = 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
 | 
554  | 
using that by blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
555  | 
      moreover have "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
 | 
556  | 
using \<open>C \<subseteq> T\<close> connT that by blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
557  | 
ultimately show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
558  | 
by metis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
559  | 
qed  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
560  | 
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
 | 
561  | 
\<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
 | 
562  | 
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
 | 
563  | 
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
 | 
564  | 
\<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
 | 
565  | 
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
 | 
566  | 
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
 | 
567  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
568  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
569  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
570  | 
|
| 71137 | 571  | 
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
 | 
572  | 
|
| 71137 | 573  | 
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
 | 
574  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
575  | 
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
 | 
576  | 
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
 | 
577  | 
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
 | 
578  | 
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
 | 
579  | 
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
 | 
580  | 
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
 | 
581  | 
proof -  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
582  | 
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 | 583  | 
\<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
 | 
584  | 
by metis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
585  | 
show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
586  | 
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
 | 
587  | 
proof (rule *)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
588  | 
fix H1 H2  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
589  | 
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
 | 
590  | 
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
 | 
591  | 
               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
 | 
592  | 
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
 | 
593  | 
"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
 | 
594  | 
by (metis Un_upper1 closedin_closed_subset inf_commute)+  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
595  | 
have Seq: "S \<inter> (H1 \<union> H2) = S"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
596  | 
by (simp add: H)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
597  | 
have "S \<inter> ((S \<union> T) \<inter> H1) \<union> S \<inter> ((S \<union> T) \<inter> H2) = S"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
598  | 
using Seq by auto  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
599  | 
    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
 | 
600  | 
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
 | 
601  | 
    ultimately have "S \<inter> H1 = {} \<or> S \<inter> H2 = {}"
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
602  | 
by (metis (no_types) H Int_assoc \<open>S \<inter> (H1 \<union> H2) = S\<close> \<open>connected S\<close>  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
603  | 
clo Seq connected_closedin inf_bot_right inf_le1)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
604  | 
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
 | 
605  | 
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
 | 
606  | 
next  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
607  | 
fix H1 H2  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
608  | 
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
 | 
609  | 
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
 | 
610  | 
               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
 | 
611  | 
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
 | 
612  | 
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
 | 
613  | 
by auto  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
614  | 
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
 | 
615  | 
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
 | 
616  | 
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
 | 
617  | 
by auto  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
618  | 
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
 | 
619  | 
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
 | 
620  | 
show "openin (top_of_set (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
 | 
621  | 
using \<open>H2 \<subseteq> T\<close> apply (auto simp: openin_closedin_eq)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
622  | 
by (metis Diff_Diff_Int Diff_disjoint Diff_partition Diff_subset H Int_absorb1 Un_Diff)  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
623  | 
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
 | 
624  | 
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
 | 
625  | 
qed  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
626  | 
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
 | 
627  | 
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
 | 
628  | 
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
 | 
629  | 
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
 | 
630  | 
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
 | 
631  | 
qed (simp add: H)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
632  | 
ultimately  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
633  | 
    have H2: "H2 = {} \<or> H2 = U"
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
634  | 
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
 | 
635  | 
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
 | 
636  | 
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
 | 
637  | 
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
 | 
638  | 
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
 | 
639  | 
ultimately show False  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
640  | 
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
 | 
641  | 
qed blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
642  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
643  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
644  | 
|
| 
68607
 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 
immler 
parents: 
68527 
diff
changeset
 | 
645  | 
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
 | 
646  | 
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
 | 
647  | 
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
 | 
648  | 
shows "connected(U - C)"  | 
| 
68607
 
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
 
immler 
parents: 
68527 
diff
changeset
 | 
649  | 
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
 | 
650  | 
proof clarify  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
651  | 
fix H3 H4  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
652  | 
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
 | 
653  | 
and clo4: "closedin (top_of_set (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
 | 
654  | 
    and "H3 \<union> H4 = U - C" and "H3 \<inter> H4 = {}" and "H3 \<noteq> {}" and "H4 \<noteq> {}"
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
655  | 
and * [rule_format]:  | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
656  | 
"\<forall>H1 H2. \<not> closedin (top_of_set S) H1 \<or>  | 
| 
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
657  | 
\<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
 | 
658  | 
                      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
 | 
659  | 
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
 | 
660  | 
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
 | 
661  | 
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
 | 
662  | 
  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
 | 
663  | 
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
 | 
664  | 
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
 | 
665  | 
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
 | 
666  | 
show "openin (top_of_set (U - C)) H3"  | 
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
667  | 
apply (simp add: openin_closedin_eq \<open>H3 \<subseteq> U - C\<close>)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
668  | 
apply (simp add: closedin_subtopology)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
669  | 
      by (metis Diff_cancel Diff_triv Un_Diff clo4 \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> closedin_closed inf_commute sup_bot.left_neutral)
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
670  | 
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
 | 
671  | 
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
 | 
672  | 
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
 | 
673  | 
show "openin (top_of_set (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
 | 
674  | 
apply (simp add: openin_closedin_eq \<open>H4 \<subseteq> U - C\<close>)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
675  | 
apply (simp add: closedin_subtopology)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
676  | 
      by (metis Diff_cancel Int_commute Un_Diff Un_Diff_Int \<open>H3 \<inter> H4 = {}\<close> \<open>H3 \<union> H4 = U - C\<close> clo3 closedin_closed)
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
677  | 
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
 | 
678  | 
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
 | 
679  | 
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
 | 
680  | 
  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
 | 
681  | 
    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
 | 
682  | 
  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
 | 
683  | 
    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
 | 
684  | 
ultimately show False  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
685  | 
    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
 | 
686  | 
by auto  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66835 
diff
changeset
 | 
687  | 
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
 | 
688  | 
|
| 
 
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
 | 
689  | 
|
| 70136 | 690  | 
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
 | 
691  | 
|
| 
 
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
 | 
692  | 
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
 | 
693  | 
|
| 
 
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
 | 
694  | 
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
 | 
695  | 
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
 | 
696  | 
and conf: "continuous_on S f"  | 
| 
 
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
 | 
697  | 
and fim: "f ` S \<subseteq> 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
 | 
698  | 
      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
 | 
699  | 
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
 | 
700  | 
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
 | 
701  | 
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
 | 
702  | 
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
 | 
703  | 
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
 | 
704  | 
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
 | 
705  | 
  { fix x assume "x \<in> 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
 | 
706  | 
    then have "f ` S \<subseteq> {f 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
 | 
707  | 
by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI S cct)  | 
| 
 
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
 | 
708  | 
}  | 
| 
 
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
 | 
709  | 
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
 | 
710  | 
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
 | 
711  | 
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
 | 
712  | 
|
| 
 
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
 | 
713  | 
|
| 
 
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
 | 
714  | 
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
 | 
715  | 
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
 | 
716  | 
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
 | 
717  | 
\<lbrakk>continuous_on S f; finite(f ` S)\<rbrakk> \<Longrightarrow> 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
 | 
718  | 
shows "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
 | 
719  | 
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
 | 
720  | 
  { fix t u
 | 
| 
69922
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
721  | 
assume clt: "closedin (top_of_set S) t"  | 
| 
 
4a9167f377b0
new material about topology, etc.; also fixes for yesterday's
 
paulson <lp15@cam.ac.uk> 
parents: 
69617 
diff
changeset
 | 
722  | 
and clu: "closedin (top_of_set S) 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
 | 
723  | 
       and tue: "t \<inter> u = {}" and tus: "t \<union> u = 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
 | 
724  | 
have conif: "continuous_on S (\<lambda>x. if x \<in> t then 0 else 1)"  | 
| 
 
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
 | 
725  | 
apply (subst tus [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
 | 
726  | 
apply (rule continuous_on_cases_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
 | 
727  | 
using clt clu tue  | 
| 71172 | 728  | 
apply (auto simp: tus)  | 
| 
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
 | 
729  | 
done  | 
| 
 
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
 | 
730  | 
have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` 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
 | 
731  | 
      by (rule finite_subset [of _ "{0,1}"]) 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
 | 
732  | 
    have "t = {} \<or> u = {}"
 | 
| 
 
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
 | 
733  | 
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
 | 
734  | 
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
 | 
735  | 
}  | 
| 
 
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
 | 
736  | 
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
 | 
737  | 
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
 | 
738  | 
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
 | 
739  | 
|
| 69617 | 740  | 
end  |