| author | nipkow | 
| Wed, 10 Jan 2018 15:25:09 +0100 | |
| changeset 67399 | eab6ce8368fa | 
| parent 66955 | 289f390c4e57 | 
| child 67968 | a5ad4c015d1c | 
| permissions | -rw-r--r-- | 
| 64122 | 1  | 
section \<open>Extending Continous Maps, Invariance of Domain, etc..\<close>  | 
| 
64006
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
3  | 
text\<open>Ported from HOL Light (moretop.ml) by L C Paulson\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
4  | 
|
| 
64289
 
42f28160bad9
HOL-Analysis: move Function Topology from AFP/Ergodict_Theory; HOL-Probability: move Essential Supremum from AFP/Lp
 
hoelzl 
parents: 
64287 
diff
changeset
 | 
5  | 
theory Further_Topology  | 
| 64291 | 6  | 
imports Equivalence_Lebesgue_Henstock_Integration Weierstrass_Theorems Polytope Complex_Transcendental  | 
| 
64006
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
7  | 
begin  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
9  | 
subsection\<open>A map from a sphere to a higher dimensional sphere is nullhomotopic\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
10  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
11  | 
lemma spheremap_lemma1:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
12  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
13  | 
assumes "subspace S" "subspace T" and dimST: "dim S < dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
14  | 
and "S \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
15  | 
and diff_f: "f differentiable_on sphere 0 1 \<inter> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
16  | 
shows "f ` (sphere 0 1 \<inter> S) \<noteq> sphere 0 1 \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
17  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
18  | 
assume fim: "f ` (sphere 0 1 \<inter> S) = sphere 0 1 \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
19  | 
have inS: "\<And>x. \<lbrakk>x \<in> S; x \<noteq> 0\<rbrakk> \<Longrightarrow> (x /\<^sub>R norm x) \<in> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
20  | 
using subspace_mul \<open>subspace S\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
21  | 
  have subS01: "(\<lambda>x. x /\<^sub>R norm x) ` (S - {0}) \<subseteq> sphere 0 1 \<inter> S"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
22  | 
using \<open>subspace S\<close> subspace_mul by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
23  | 
  then have diff_f': "f differentiable_on (\<lambda>x. x /\<^sub>R norm x) ` (S - {0})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
24  | 
by (rule differentiable_on_subset [OF diff_f])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
25  | 
define g where "g \<equiv> \<lambda>x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
26  | 
  have gdiff: "g differentiable_on S - {0}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
27  | 
unfolding g_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
28  | 
by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
29  | 
  have geq: "g ` (S - {0}) = T - {0}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
30  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
31  | 
    have "g ` (S - {0}) \<subseteq> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
32  | 
apply (auto simp: g_def subspace_mul [OF \<open>subspace T\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
33  | 
apply (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \<open>subspace T\<close>] fim image_subset_iff inf_le2 singletonD)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
34  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
35  | 
    moreover have "g ` (S - {0}) \<subseteq> UNIV - {0}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
36  | 
proof (clarsimp simp: g_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
37  | 
fix y  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
38  | 
assume "y \<in> S" and f0: "f (y /\<^sub>R norm y) = 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
39  | 
then have "y \<noteq> 0 \<Longrightarrow> y /\<^sub>R norm y \<in> sphere 0 1 \<inter> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
40  | 
by (auto simp: subspace_mul [OF \<open>subspace S\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
41  | 
then show "y = 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
42  | 
by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
43  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
44  | 
    ultimately show "g ` (S - {0}) \<subseteq> T - {0}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
45  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
46  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
47  | 
have *: "sphere 0 1 \<inter> T \<subseteq> f ` (sphere 0 1 \<inter> S)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
48  | 
using fim by (simp add: image_subset_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
49  | 
    have "x \<in> (\<lambda>x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
50  | 
if "x \<in> T" "x \<noteq> 0" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
51  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
52  | 
have "x /\<^sub>R norm x \<in> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
53  | 
using \<open>subspace T\<close> subspace_mul that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
54  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
55  | 
using * [THEN subsetD, of "x /\<^sub>R norm x"] that apply clarsimp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
56  | 
apply (rule_tac x="norm x *\<^sub>R xa" in image_eqI, simp)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
57  | 
apply (metis norm_eq_zero right_inverse scaleR_one scaleR_scaleR)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
58  | 
using \<open>subspace S\<close> subspace_mul apply force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
59  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
60  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
61  | 
    then have "T - {0} \<subseteq> (\<lambda>x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
62  | 
by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
63  | 
    then show "T - {0} \<subseteq> g ` (S - {0})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
64  | 
by (simp add: g_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
65  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
66  | 
  define T' where "T' \<equiv> {y. \<forall>x \<in> T. orthogonal x y}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
67  | 
have "subspace T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
68  | 
by (simp add: subspace_orthogonal_to_vectors T'_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
69  | 
  have dim_eq: "dim T' + dim T = DIM('a)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
70  | 
using dim_subspace_orthogonal_to_vectors [of T UNIV] \<open>subspace T\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
71  | 
by (simp add: dim_UNIV T'_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
72  | 
have "\<exists>v1 v2. v1 \<in> span T \<and> (\<forall>w \<in> span T. orthogonal v2 w) \<and> x = v1 + v2" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
73  | 
by (force intro: orthogonal_subspace_decomp_exists [of T x])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
74  | 
then obtain p1 p2 where p1span: "p1 x \<in> span T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
75  | 
and "\<And>w. w \<in> span T \<Longrightarrow> orthogonal (p2 x) w"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
76  | 
and eq: "p1 x + p2 x = x" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
77  | 
by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
78  | 
then have p1: "\<And>z. p1 z \<in> T" and ortho: "\<And>w. w \<in> T \<Longrightarrow> orthogonal (p2 x) w" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
79  | 
using span_eq \<open>subspace T\<close> by blast+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
80  | 
then have p2: "\<And>z. p2 z \<in> T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
81  | 
by (simp add: T'_def orthogonal_commute)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
82  | 
have p12_eq: "\<And>x y. \<lbrakk>x \<in> T; y \<in> T'\<rbrakk> \<Longrightarrow> p1(x + y) = x \<and> p2(x + y) = y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
83  | 
proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T'])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
84  | 
show "\<And>x y. \<lbrakk>x \<in> T; y \<in> T'\<rbrakk> \<Longrightarrow> p2 (x + y) \<in> span T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
85  | 
using span_eq p2 \<open>subspace T'\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
86  | 
show "\<And>a b. \<lbrakk>a \<in> T; b \<in> T'\<rbrakk> \<Longrightarrow> orthogonal a b"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
87  | 
using T'_def by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
88  | 
qed (auto simp: span_superset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
89  | 
then have "\<And>c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \<and> p2 (c *\<^sub>R x) = c *\<^sub>R p2 x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
90  | 
by (metis eq \<open>subspace T\<close> \<open>subspace T'\<close> p1 p2 scaleR_add_right subspace_mul)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
91  | 
moreover have lin_add: "\<And>x y. p1 (x + y) = p1 x + p1 y \<and> p2 (x + y) = p2 x + p2 y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
92  | 
proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T'])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
93  | 
show "\<And>x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
94  | 
by (simp add: add.assoc add.left_commute eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
95  | 
show "\<And>a b. \<lbrakk>a \<in> T; b \<in> T'\<rbrakk> \<Longrightarrow> orthogonal a b"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
96  | 
using T'_def by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
97  | 
qed (auto simp: p1span p2 span_superset subspace_add)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
98  | 
ultimately have "linear p1" "linear p2"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
99  | 
by unfold_locales auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
100  | 
  have "(\<lambda>z. g (p1 z)) differentiable_on {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
101  | 
apply (rule differentiable_on_compose [where f=g])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
102  | 
apply (rule linear_imp_differentiable_on [OF \<open>linear p1\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
103  | 
apply (rule differentiable_on_subset [OF gdiff])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
104  | 
using p12_eq \<open>S \<subseteq> T\<close> apply auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
105  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
106  | 
  then have diff: "(\<lambda>x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
107  | 
by (intro derivative_intros linear_imp_differentiable_on [OF \<open>linear p2\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
108  | 
  have "dim {x + y |x y. x \<in> S - {0} \<and> y \<in> T'} \<le> dim {x + y |x y. x \<in> S  \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
109  | 
by (blast intro: dim_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
110  | 
also have "... = dim S + dim T' - dim (S \<inter> T')"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
111  | 
using dim_sums_Int [OF \<open>subspace S\<close> \<open>subspace T'\<close>]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
112  | 
by (simp add: algebra_simps)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
113  | 
  also have "... < DIM('a)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
114  | 
using dimST dim_eq by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
115  | 
  finally have neg: "negligible {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
116  | 
by (rule negligible_lowdim)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
117  | 
  have "negligible ((\<lambda>x. g (p1 x) + p2 x) ` {x + y |x y. x \<in> S - {0} \<and> y \<in> T'})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
118  | 
by (rule negligible_differentiable_image_negligible [OF order_refl neg diff])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
119  | 
  then have "negligible {x + y |x y. x \<in> g ` (S - {0}) \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
120  | 
proof (rule negligible_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
121  | 
have "\<lbrakk>t' \<in> T'; s \<in> S; s \<noteq> 0\<rbrakk>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
122  | 
\<Longrightarrow> g s + t' \<in> (\<lambda>x. g (p1 x) + p2 x) `  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
123  | 
                         {x + t' |x t'. x \<in> S \<and> x \<noteq> 0 \<and> t' \<in> T'}" for t' s
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
124  | 
apply (rule_tac x="s + t'" in image_eqI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
125  | 
using \<open>S \<subseteq> T\<close> p12_eq by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
126  | 
    then show "{x + y |x y. x \<in> g ` (S - {0}) \<and> y \<in> T'}
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
127  | 
          \<subseteq> (\<lambda>x. g (p1 x) + p2 x) ` {x + y |x y. x \<in> S - {0} \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
128  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
129  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
130  | 
  moreover have "- T' \<subseteq> {x + y |x y. x \<in> g ` (S - {0}) \<and> y \<in> T'}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
131  | 
proof clarsimp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
132  | 
fix z assume "z \<notin> T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
133  | 
    show "\<exists>x y. z = x + y \<and> x \<in> g ` (S - {0}) \<and> y \<in> T'"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
134  | 
apply (rule_tac x="p1 z" in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
135  | 
apply (rule_tac x="p2 z" in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
136  | 
apply (simp add: p1 eq p2 geq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
137  | 
by (metis \<open>z \<notin> T'\<close> add.left_neutral eq p2)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
138  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
139  | 
ultimately have "negligible (-T')"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
140  | 
using negligible_subset by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
141  | 
moreover have "negligible T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
142  | 
using negligible_lowdim  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
143  | 
by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
144  | 
ultimately have "negligible (-T' \<union> T')"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
145  | 
by (metis negligible_Un_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
146  | 
then show False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
147  | 
using negligible_Un_eq non_negligible_UNIV by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
148  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
151  | 
lemma spheremap_lemma2:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
152  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
153  | 
assumes ST: "subspace S" "subspace T" "dim S < dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
154  | 
and "S \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
155  | 
and contf: "continuous_on (sphere 0 1 \<inter> S) f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
156  | 
and fim: "f ` (sphere 0 1 \<inter> S) \<subseteq> sphere 0 1 \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
157  | 
shows "\<exists>c. homotopic_with (\<lambda>x. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) f (\<lambda>x. c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
158  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
159  | 
have [simp]: "\<And>x. \<lbrakk>norm x = 1; x \<in> S\<rbrakk> \<Longrightarrow> norm (f x) = 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
160  | 
using fim by (simp add: image_subset_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
161  | 
have "compact (sphere 0 1 \<inter> S)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
162  | 
by (simp add: \<open>subspace S\<close> closed_subspace compact_Int_closed)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
163  | 
then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \<inter> S) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
164  | 
and g12: "\<And>x. x \<in> sphere 0 1 \<inter> S \<Longrightarrow> norm(f x - g x) < 1/2"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
165  | 
apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \<open>subspace T\<close>, of "1/2"])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
166  | 
using fim apply auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
167  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
168  | 
have gnz: "g x \<noteq> 0" if "x \<in> sphere 0 1 \<inter> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
169  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
170  | 
have "norm (f x) = 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
171  | 
using fim that by (simp add: image_subset_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
172  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
173  | 
using g12 [OF that] by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
174  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
175  | 
have diffg: "g differentiable_on sphere 0 1 \<inter> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
176  | 
by (metis pfg differentiable_on_polynomial_function)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
177  | 
define h where "h \<equiv> \<lambda>x. inverse(norm(g x)) *\<^sub>R g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
178  | 
have h: "x \<in> sphere 0 1 \<inter> S \<Longrightarrow> h x \<in> sphere 0 1 \<inter> T" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
179  | 
unfolding h_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
180  | 
using gnz [of x]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
181  | 
by (auto simp: subspace_mul [OF \<open>subspace T\<close>] subsetD [OF gim])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
182  | 
have diffh: "h differentiable_on sphere 0 1 \<inter> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
183  | 
unfolding h_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
184  | 
apply (intro derivative_intros diffg differentiable_on_compose [OF diffg])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
185  | 
using gnz apply auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
186  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
187  | 
  have homfg: "homotopic_with (\<lambda>z. True) (sphere 0 1 \<inter> S) (T - {0}) f g"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
188  | 
proof (rule homotopic_with_linear [OF contf])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
189  | 
show "continuous_on (sphere 0 1 \<inter> S) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
190  | 
using pfg by (simp add: differentiable_imp_continuous_on diffg)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
191  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
192  | 
have non0fg: "0 \<notin> closed_segment (f x) (g x)" if "norm x = 1" "x \<in> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
193  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
194  | 
have "f x \<in> sphere 0 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
195  | 
using fim that by (simp add: image_subset_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
196  | 
moreover have "norm(f x - g x) < 1/2"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
197  | 
apply (rule g12)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
198  | 
using that by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
199  | 
ultimately show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
200  | 
by (auto simp: norm_minus_commute dest: segment_bound)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
201  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
202  | 
    show "\<And>x. x \<in> sphere 0 1 \<inter> S \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> T - {0}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
203  | 
apply (simp add: subset_Diff_insert non0fg)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
204  | 
apply (simp add: segment_convex_hull)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
205  | 
apply (rule hull_minimal)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
206  | 
using fim image_eqI gim apply force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
207  | 
apply (rule subspace_imp_convex [OF \<open>subspace T\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
208  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
209  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
210  | 
obtain d where d: "d \<in> (sphere 0 1 \<inter> T) - h ` (sphere 0 1 \<inter> S)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
211  | 
using h spheremap_lemma1 [OF ST \<open>S \<subseteq> T\<close> diffh] by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
212  | 
then have non0hd: "0 \<notin> closed_segment (h x) (- d)" if "norm x = 1" "x \<in> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
213  | 
using midpoint_between [of 0 "h x" "-d"] that h [of x]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
214  | 
by (auto simp: between_mem_segment midpoint_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
215  | 
have conth: "continuous_on (sphere 0 1 \<inter> S) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
216  | 
using differentiable_imp_continuous_on diffh by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
217  | 
  have hom_hd: "homotopic_with (\<lambda>z. True) (sphere 0 1 \<inter> S) (T - {0}) h (\<lambda>x. -d)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
218  | 
apply (rule homotopic_with_linear [OF conth continuous_on_const])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
219  | 
apply (simp add: subset_Diff_insert non0hd)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
220  | 
apply (simp add: segment_convex_hull)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
221  | 
apply (rule hull_minimal)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
222  | 
using h d apply (force simp: subspace_neg [OF \<open>subspace T\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
223  | 
apply (rule subspace_imp_convex [OF \<open>subspace T\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
224  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
225  | 
  have conT0: "continuous_on (T - {0}) (\<lambda>y. inverse(norm y) *\<^sub>R y)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
226  | 
by (intro continuous_intros) auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
227  | 
  have sub0T: "(\<lambda>y. y /\<^sub>R norm y) ` (T - {0}) \<subseteq> sphere 0 1 \<inter> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
228  | 
by (fastforce simp: assms(2) subspace_mul)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
229  | 
obtain c where homhc: "homotopic_with (\<lambda>z. True) (sphere 0 1 \<inter> S) (sphere 0 1 \<inter> T) h (\<lambda>x. c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
230  | 
apply (rule_tac c="-d" in that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
231  | 
apply (rule homotopic_with_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
232  | 
apply (rule homotopic_compose_continuous_left [OF hom_hd conT0 sub0T])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
233  | 
using d apply (auto simp: h_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
234  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
235  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
236  | 
apply (rule_tac x=c in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
237  | 
apply (rule homotopic_with_trans [OF _ homhc])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
238  | 
apply (rule homotopic_with_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
239  | 
apply (rule homotopic_compose_continuous_left [OF homfg conT0 sub0T])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
240  | 
apply (auto simp: h_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
241  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
242  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
243  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
244  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
245  | 
lemma spheremap_lemma3:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
246  | 
assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \<le> dim U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
247  | 
  obtains T where "subspace T" "T \<subseteq> U" "S \<noteq> {} \<Longrightarrow> aff_dim T = aff_dim S"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
248  | 
"(rel_frontier S) homeomorphic (sphere 0 1 \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
249  | 
proof (cases "S = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
250  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
251  | 
with \<open>subspace U\<close> subspace_0 show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
252  | 
    by (rule_tac T = "{0}" in that) auto
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
253  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
254  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
255  | 
then obtain a where "a \<in> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
256  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
257  | 
then have affS: "aff_dim S = int (dim ((\<lambda>x. -a+x) ` S))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
258  | 
by (metis hull_inc aff_dim_eq_dim)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
259  | 
with affSU have "dim ((\<lambda>x. -a+x) ` S) \<le> dim U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
260  | 
by linarith  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
261  | 
with choose_subspace_of_subspace  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
262  | 
obtain T where "subspace T" "T \<subseteq> span U" and dimT: "dim T = dim ((\<lambda>x. -a+x) ` S)" .  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
263  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
264  | 
proof (rule that [OF \<open>subspace T\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
265  | 
show "T \<subseteq> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
266  | 
using span_eq \<open>subspace U\<close> \<open>T \<subseteq> span U\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
267  | 
show "aff_dim T = aff_dim S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
268  | 
using dimT \<open>subspace T\<close> affS aff_dim_subspace by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
269  | 
show "rel_frontier S homeomorphic sphere 0 1 \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
270  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
271  | 
have "aff_dim (ball 0 1 \<inter> T) = aff_dim (T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
272  | 
by (metis IntI interior_ball \<open>subspace T\<close> aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
273  | 
then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
274  | 
using \<open>aff_dim T = aff_dim S\<close> by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
275  | 
have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
276  | 
apply (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \<open>convex S\<close> \<open>bounded S\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
277  | 
apply (simp add: \<open>subspace T\<close> convex_Int subspace_imp_convex)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
278  | 
apply (simp add: bounded_Int)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
279  | 
apply (rule affS_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
280  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
281  | 
also have "... = frontier (ball 0 1) \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
282  | 
apply (rule convex_affine_rel_frontier_Int [OF convex_ball])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
283  | 
apply (simp add: \<open>subspace T\<close> subspace_imp_affine)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
284  | 
using \<open>subspace T\<close> subspace_0 by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
285  | 
also have "... = sphere 0 1 \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
286  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
287  | 
finally show ?thesis .  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
288  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
289  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
290  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
291  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
292  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
293  | 
proposition inessential_spheremap_lowdim_gen:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
294  | 
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
295  | 
assumes "convex S" "bounded S" "convex T" "bounded T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
296  | 
and affST: "aff_dim S < aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
297  | 
and contf: "continuous_on (rel_frontier S) f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
298  | 
and fim: "f ` (rel_frontier S) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
299  | 
obtains c where "homotopic_with (\<lambda>z. True) (rel_frontier S) (rel_frontier T) f (\<lambda>x. c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
300  | 
proof (cases "S = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
301  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
302  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
303  | 
by (simp add: that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
304  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
305  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
306  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
307  | 
  proof (cases "T = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
308  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
309  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
310  | 
using fim that by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
311  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
312  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
313  | 
obtain T':: "'a set"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
314  | 
where "subspace T'" and affT': "aff_dim T' = aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
315  | 
and homT: "rel_frontier T homeomorphic sphere 0 1 \<inter> T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
316  | 
apply (rule spheremap_lemma3 [OF \<open>bounded T\<close> \<open>convex T\<close> subspace_UNIV, where 'b='a])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
317  | 
apply (simp add: dim_UNIV aff_dim_le_DIM)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
318  | 
      using \<open>T \<noteq> {}\<close> by blast
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
319  | 
with homeomorphic_imp_homotopy_eqv  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
320  | 
have relT: "sphere 0 1 \<inter> T' homotopy_eqv rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
321  | 
using homotopy_eqv_sym by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
322  | 
have "aff_dim S \<le> int (dim T')"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
323  | 
using affT' \<open>subspace T'\<close> affST aff_dim_subspace by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
324  | 
    with spheremap_lemma3 [OF \<open>bounded S\<close> \<open>convex S\<close> \<open>subspace T'\<close>] \<open>S \<noteq> {}\<close>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
325  | 
obtain S':: "'a set" where "subspace S'" "S' \<subseteq> T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
326  | 
and affS': "aff_dim S' = aff_dim S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
327  | 
and homT: "rel_frontier S homeomorphic sphere 0 1 \<inter> S'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
328  | 
by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
329  | 
with homeomorphic_imp_homotopy_eqv  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
330  | 
have relS: "sphere 0 1 \<inter> S' homotopy_eqv rel_frontier S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
331  | 
using homotopy_eqv_sym by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
332  | 
have dimST': "dim S' < dim T'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
333  | 
by (metis \<open>S' \<subseteq> T'\<close> \<open>subspace S'\<close> \<open>subspace T'\<close> affS' affST affT' less_irrefl not_le subspace_dim_equal)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
334  | 
have "\<exists>c. homotopic_with (\<lambda>z. True) (rel_frontier S) (rel_frontier T) f (\<lambda>x. c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
335  | 
apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
336  | 
apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
337  | 
apply (metis dimST' \<open>subspace S'\<close> \<open>subspace T'\<close> \<open>S' \<subseteq> T'\<close> spheremap_lemma2, blast)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
338  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
339  | 
with that show ?thesis by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
340  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
341  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
342  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
343  | 
lemma inessential_spheremap_lowdim:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
344  | 
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
345  | 
assumes  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
346  | 
   "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \<subseteq> (sphere b s)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
347  | 
obtains c where "homotopic_with (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
348  | 
proof (cases "s \<le> 0")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
349  | 
case True then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
350  | 
by (meson nullhomotopic_into_contractible f contractible_sphere that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
351  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
352  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
353  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
354  | 
proof (cases "r \<le> 0")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
355  | 
case True then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
356  | 
by (meson f nullhomotopic_from_contractible contractible_sphere that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
357  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
358  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
359  | 
with \<open>~ s \<le> 0\<close> have "r > 0" "s > 0" by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
360  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
361  | 
apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
362  | 
using \<open>0 < r\<close> \<open>0 < s\<close> assms(1)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
363  | 
apply (simp_all add: f aff_dim_cball)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
364  | 
using that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
365  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
366  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
367  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
368  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
369  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
370  | 
subsection\<open> Some technical lemmas about extending maps from cell complexes.\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
371  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
372  | 
lemma extending_maps_Union_aux:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
373  | 
assumes fin: "finite \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
374  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
375  | 
and "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>; S \<noteq> T\<rbrakk> \<Longrightarrow> S \<inter> T \<subseteq> K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
376  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
377  | 
shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
378  | 
using assms  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
379  | 
proof (induction \<F>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
380  | 
case empty show ?case by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
381  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
382  | 
case (insert S \<F>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
383  | 
then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \<subseteq> T" and feq: "\<forall>x \<in> S \<inter> K. f x = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
384  | 
by (meson insertI1)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
385  | 
obtain g where contg: "continuous_on (\<Union>\<F>) g" and gim: "g ` \<Union>\<F> \<subseteq> T" and geq: "\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
386  | 
using insert by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
387  | 
have fg: "f x = g x" if "x \<in> T" "T \<in> \<F>" "x \<in> S" for x T  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
388  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
389  | 
have "T \<inter> S \<subseteq> K \<or> S = T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
390  | 
using that by (metis (no_types) insert.prems(2) insertCI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
391  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
392  | 
using UnionI feq geq \<open>S \<notin> \<F>\<close> subsetD that by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
393  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
394  | 
show ?case  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
395  | 
apply (rule_tac x="\<lambda>x. if x \<in> S then f x else g x" in exI, simp)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
396  | 
apply (intro conjI continuous_on_cases)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
397  | 
apply (simp_all add: insert closed_Union contf contg)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
398  | 
using fim gim feq geq  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
399  | 
apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
400  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
401  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
402  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
403  | 
lemma extending_maps_Union:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
404  | 
assumes fin: "finite \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
405  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>g. continuous_on S g \<and> g ` S \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
406  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
407  | 
and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; ~ X \<subseteq> Y; ~ Y \<subseteq> X\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
408  | 
shows "\<exists>g. continuous_on (\<Union>\<F>) g \<and> g ` (\<Union>\<F>) \<subseteq> T \<and> (\<forall>x \<in> \<Union>\<F> \<inter> K. g x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
409  | 
apply (simp add: Union_maximal_sets [OF fin, symmetric])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
410  | 
apply (rule extending_maps_Union_aux)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
411  | 
apply (simp_all add: Union_maximal_sets [OF fin] assms)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
412  | 
by (metis K psubsetI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
413  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
414  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
415  | 
lemma extend_map_lemma:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
416  | 
assumes "finite \<F>" "\<G> \<subseteq> \<F>" "convex T" "bounded T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
417  | 
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
418  | 
and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X < aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
419  | 
and face: "\<And>S T. \<lbrakk>S \<in> \<F>; T \<in> \<F>\<rbrakk> \<Longrightarrow> (S \<inter> T) face_of S \<and> (S \<inter> T) face_of T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
420  | 
and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
421  | 
obtains g where "continuous_on (\<Union>\<F>) g" "g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
422  | 
proof (cases "\<F> - \<G> = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
423  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
424  | 
then have "\<Union>\<F> \<subseteq> \<Union>\<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
425  | 
by (simp add: Union_mono)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
426  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
427  | 
apply (rule_tac g=f in that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
428  | 
using contf continuous_on_subset apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
429  | 
using fim apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
430  | 
by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
431  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
432  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
433  | 
then have "0 \<le> aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
434  | 
by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
435  | 
then obtain i::nat where i: "int i = aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
436  | 
by (metis nonneg_eq_int)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
437  | 
  have Union_empty_eq: "\<Union>{D. D = {} \<and> P D} = {}" for P :: "'a set \<Rightarrow> bool"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
438  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
439  | 
  have extendf: "\<exists>g. continuous_on (\<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < i})) g \<and>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
440  | 
                     g ` (\<Union> (\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < i})) \<subseteq> rel_frontier T \<and>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
441  | 
(\<forall>x \<in> \<Union>\<G>. g x = f x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
442  | 
if "i \<le> aff_dim T" for i::nat  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
443  | 
using that  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
444  | 
proof (induction i)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
445  | 
case 0 then show ?case  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
446  | 
apply (simp add: Union_empty_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
447  | 
apply (rule_tac x=f in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
448  | 
apply (intro conjI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
449  | 
using contf continuous_on_subset apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
450  | 
using fim apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
451  | 
by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
452  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
453  | 
case (Suc p)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
454  | 
    with \<open>bounded T\<close> have "rel_frontier T \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
455  | 
by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
456  | 
then obtain t where t: "t \<in> rel_frontier T" by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
457  | 
have ple: "int p \<le> aff_dim T" using Suc.prems by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
458  | 
    obtain h where conth: "continuous_on (\<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p})) h"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
459  | 
               and him: "h ` (\<Union> (\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}))
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
460  | 
\<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
461  | 
and heq: "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> h x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
462  | 
using Suc.IH [OF ple] by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
463  | 
    let ?Faces = "{D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D \<le> p}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
464  | 
have extendh: "\<exists>g. continuous_on D g \<and>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
465  | 
g ` D \<subseteq> rel_frontier T \<and>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
466  | 
                       (\<forall>x \<in> D \<inter> \<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}). g x = h x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
467  | 
if D: "D \<in> \<G> \<union> ?Faces" for D  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
468  | 
    proof (cases "D \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p})")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
469  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
470  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
471  | 
apply (rule_tac x=h in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
472  | 
apply (intro conjI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
473  | 
apply (blast intro: continuous_on_subset [OF conth])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
474  | 
using him apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
475  | 
by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
476  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
477  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
478  | 
note notDsub = False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
479  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
480  | 
      proof (cases "\<exists>a. D = {a}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
481  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
482  | 
        then obtain a where "D = {a}" by auto
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
483  | 
with notDsub t show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
484  | 
by (rule_tac x="\<lambda>x. t" in exI) simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
485  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
486  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
487  | 
        have "D \<noteq> {}" using notDsub by auto
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
488  | 
        have Dnotin: "D \<notin> \<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
489  | 
using notDsub by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
490  | 
then have "D \<notin> \<G>" by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
491  | 
        have "D \<in> ?Faces - {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < p}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
492  | 
using Dnotin that by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
493  | 
then obtain C where "C \<in> \<F>" "D face_of C" and affD: "aff_dim D = int p"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
494  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
495  | 
then have "bounded D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
496  | 
using face_of_polytope_polytope poly polytope_imp_bounded by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
497  | 
then have [simp]: "\<not> affine D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
498  | 
          using affine_bounded_eq_trivial False \<open>D \<noteq> {}\<close> \<open>bounded D\<close> by blast
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
499  | 
        have "{F. F facet_of D} \<subseteq> {E. E face_of C \<and> aff_dim E < int p}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
500  | 
apply clarify  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
501  | 
apply (metis \<open>D face_of C\<close> affD eq_iff face_of_trans facet_of_def zle_diff1_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
502  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
503  | 
moreover have "polyhedron D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
504  | 
using \<open>C \<in> \<F>\<close> \<open>D face_of C\<close> face_of_polytope_polytope poly polytope_imp_polyhedron by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
505  | 
        ultimately have relf_sub: "rel_frontier D \<subseteq> \<Union> {E. E face_of C \<and> aff_dim E < p}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
506  | 
by (simp add: rel_frontier_of_polyhedron Union_mono)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
507  | 
then have him_relf: "h ` rel_frontier D \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
508  | 
using \<open>C \<in> \<F>\<close> him by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
509  | 
have "convex D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
510  | 
by (simp add: \<open>polyhedron D\<close> polyhedron_imp_convex)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
511  | 
have affD_lessT: "aff_dim D < aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
512  | 
using Suc.prems affD by linarith  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
513  | 
have contDh: "continuous_on (rel_frontier D) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
514  | 
using \<open>C \<in> \<F>\<close> relf_sub by (blast intro: continuous_on_subset [OF conth])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
515  | 
then have *: "(\<exists>c. homotopic_with (\<lambda>x. True) (rel_frontier D) (rel_frontier T) h (\<lambda>x. c)) =  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
516  | 
(\<exists>g. continuous_on UNIV g \<and> range g \<subseteq> rel_frontier T \<and>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
517  | 
(\<forall>x\<in>rel_frontier D. g x = h x))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
518  | 
apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
519  | 
apply (simp_all add: assms rel_frontier_eq_empty him_relf)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
520  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
521  | 
have "(\<exists>c. homotopic_with (\<lambda>x. True) (rel_frontier D)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
522  | 
(rel_frontier T) h (\<lambda>x. c))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
523  | 
by (metis inessential_spheremap_lowdim_gen  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
524  | 
[OF \<open>convex D\<close> \<open>bounded D\<close> \<open>convex T\<close> \<open>bounded T\<close> affD_lessT contDh him_relf])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
525  | 
then obtain g where contg: "continuous_on UNIV g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
526  | 
and gim: "range g \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
527  | 
and gh: "\<And>x. x \<in> rel_frontier D \<Longrightarrow> g x = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
528  | 
by (metis *)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
529  | 
have "D \<inter> E \<subseteq> rel_frontier D"  | 
| 67399 | 530  | 
             if "E \<in> \<G> \<union> {D. Bex \<F> ((face_of) D) \<and> aff_dim D < int p}" for E
 | 
| 
64006
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
531  | 
proof (rule face_of_subset_rel_frontier)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
532  | 
show "D \<inter> E face_of D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
533  | 
using that \<open>C \<in> \<F>\<close> \<open>D face_of C\<close> face  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
534  | 
apply auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
535  | 
apply (meson face_of_Int_subface \<open>\<G> \<subseteq> \<F>\<close> face_of_refl_eq poly polytope_imp_convex subsetD)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
536  | 
using face_of_Int_subface apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
537  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
538  | 
show "D \<inter> E \<noteq> D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
539  | 
using that notDsub by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
540  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
541  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
542  | 
apply (rule_tac x=g in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
543  | 
apply (intro conjI ballI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
544  | 
using continuous_on_subset contg apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
545  | 
using gim apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
546  | 
using gh by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
547  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
548  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
549  | 
have intle: "i < 1 + int j \<longleftrightarrow> i \<le> int j" for i j  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
550  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
551  | 
have "finite \<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
552  | 
using \<open>finite \<F>\<close> \<open>\<G> \<subseteq> \<F>\<close> rev_finite_subset by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
553  | 
then have fin: "finite (\<G> \<union> ?Faces)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
554  | 
apply simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
555  | 
      apply (rule_tac B = "\<Union>{{D. D face_of C}| C. C \<in> \<F>}" in finite_subset)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
556  | 
by (auto simp: \<open>finite \<F>\<close> finite_polytope_faces poly)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
557  | 
have clo: "closed S" if "S \<in> \<G> \<union> ?Faces" for S  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
558  | 
using that \<open>\<G> \<subseteq> \<F>\<close> face_of_polytope_polytope poly polytope_imp_closed by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
559  | 
    have K: "X \<inter> Y \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int p})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
560  | 
if "X \<in> \<G> \<union> ?Faces" "Y \<in> \<G> \<union> ?Faces" "~ Y \<subseteq> X" for X Y  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
561  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
562  | 
have ff: "X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
563  | 
if XY: "X face_of D" "Y face_of E" and DE: "D \<in> \<F>" "E \<in> \<F>" for D E  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
564  | 
apply (rule face_of_Int_subface [OF _ _ XY])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
565  | 
apply (auto simp: face DE)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
566  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
567  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
568  | 
using that  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
569  | 
apply auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
570  | 
apply (drule_tac x="X \<inter> Y" in spec, safe)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
571  | 
using ff face_of_imp_convex [of X] face_of_imp_convex [of Y]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
572  | 
apply (fastforce dest: face_of_aff_dim_lt)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
573  | 
by (meson face_of_trans ff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
574  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
575  | 
obtain g where "continuous_on (\<Union>(\<G> \<union> ?Faces)) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
576  | 
"g ` \<Union>(\<G> \<union> ?Faces) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
577  | 
"(\<forall>x \<in> \<Union>(\<G> \<union> ?Faces) \<inter>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
578  | 
                          \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < p}). g x = h x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
579  | 
apply (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
580  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
581  | 
then show ?case  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
582  | 
apply (simp add: intle local.heq [symmetric], blast)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
583  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
584  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
585  | 
  have eq: "\<Union>(\<G> \<union> {D. \<exists>C \<in> \<F>. D face_of C \<and> aff_dim D < i}) = \<Union>\<F>"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
586  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
587  | 
    show "\<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < int i}) \<subseteq> \<Union>\<F>"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
588  | 
apply (rule Union_subsetI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
589  | 
using \<open>\<G> \<subseteq> \<F>\<close> face_of_imp_subset apply force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
590  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
591  | 
    show "\<Union>\<F> \<subseteq> \<Union>(\<G> \<union> {D. \<exists>C\<in>\<F>. D face_of C \<and> aff_dim D < i})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
592  | 
apply (rule Union_mono)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
593  | 
using face apply (fastforce simp: aff i)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
594  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
595  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
596  | 
have "int i \<le> aff_dim T" by (simp add: i)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
597  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
598  | 
using extendf [of i] unfolding eq by (metis that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
599  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
600  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
601  | 
lemma extend_map_lemma_cofinite0:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
602  | 
assumes "finite \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
603  | 
and "pairwise (\<lambda>S T. S \<inter> T \<subseteq> K) \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
604  | 
      and "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
605  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
606  | 
shows "\<exists>C g. finite C \<and> disjnt C U \<and> card C \<le> card \<F> \<and>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
607  | 
continuous_on (\<Union>\<F> - C) g \<and> g ` (\<Union>\<F> - C) \<subseteq> T  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
608  | 
\<and> (\<forall>x \<in> (\<Union>\<F> - C) \<inter> K. g x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
609  | 
using assms  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
610  | 
proof induction  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
611  | 
case empty then show ?case  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
612  | 
by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
613  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
614  | 
case (insert X \<F>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
615  | 
then have "closed X" and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
616  | 
        and \<F>: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (S - {a}) g \<and> g ` (S - {a}) \<subseteq> T \<and> (\<forall>x \<in> S \<inter> K. g x = h x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
617  | 
and pwX: "\<And>Y. Y \<in> \<F> \<and> Y \<noteq> X \<longrightarrow> X \<inter> Y \<subseteq> K \<and> Y \<inter> X \<subseteq> K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
618  | 
and pwF: "pairwise (\<lambda> S T. S \<inter> T \<subseteq> K) \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
619  | 
by (simp_all add: pairwise_insert)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
620  | 
obtain C g where C: "finite C" "disjnt C U" "card C \<le> card \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
621  | 
and contg: "continuous_on (\<Union>\<F> - C) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
622  | 
and gim: "g ` (\<Union>\<F> - C) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
623  | 
and gh: "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
624  | 
using insert.IH [OF pwF \<F> clo] by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
625  | 
obtain a f where "a \<notin> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
626  | 
               and contf: "continuous_on (X - {a}) f"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
627  | 
               and fim: "f ` (X - {a}) \<subseteq> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
628  | 
and fh: "(\<forall>x \<in> X \<inter> K. f x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
629  | 
using insert.prems by (meson insertI1)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
630  | 
show ?case  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
631  | 
proof (intro exI conjI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
632  | 
show "finite (insert a C)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
633  | 
by (simp add: C)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
634  | 
show "disjnt (insert a C) U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
635  | 
using C \<open>a \<notin> U\<close> by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
636  | 
show "card (insert a C) \<le> card (insert X \<F>)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
637  | 
by (simp add: C card_insert_if insert.hyps le_SucI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
638  | 
have "closed (\<Union>\<F>)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
639  | 
using clo insert.hyps by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
640  | 
have "continuous_on (X - insert a C \<union> (\<Union>\<F> - insert a C)) (\<lambda>x. if x \<in> X then f x else g x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
641  | 
apply (rule continuous_on_cases_local)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
642  | 
apply (simp_all add: closedin_closed)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
643  | 
using \<open>closed X\<close> apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
644  | 
using \<open>closed (\<Union>\<F>)\<close> apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
645  | 
using contf apply (force simp: elim: continuous_on_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
646  | 
using contg apply (force simp: elim: continuous_on_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
647  | 
using fh gh insert.hyps pwX by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
648  | 
then show "continuous_on (\<Union>insert X \<F> - insert a C) (\<lambda>a. if a \<in> X then f a else g a)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
649  | 
by (blast intro: continuous_on_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
650  | 
show "\<forall>x\<in>(\<Union>insert X \<F> - insert a C) \<inter> K. (if x \<in> X then f x else g x) = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
651  | 
using gh by (auto simp: fh)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
652  | 
show "(\<lambda>a. if a \<in> X then f a else g a) ` (\<Union>insert X \<F> - insert a C) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
653  | 
using fim gim by auto force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
654  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
655  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
656  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
657  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
658  | 
lemma extend_map_lemma_cofinite1:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
659  | 
assumes "finite \<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
660  | 
    and \<F>: "\<And>X. X \<in> \<F> \<Longrightarrow> \<exists>a g. a \<notin> U \<and> continuous_on (X - {a}) g \<and> g ` (X - {a}) \<subseteq> T \<and> (\<forall>x \<in> X \<inter> K. g x = h x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
661  | 
and clo: "\<And>X. X \<in> \<F> \<Longrightarrow> closed X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
662  | 
and K: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>; ~(X \<subseteq> Y); ~(Y \<subseteq> X)\<rbrakk> \<Longrightarrow> X \<inter> Y \<subseteq> K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
663  | 
obtains C g where "finite C" "disjnt C U" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
664  | 
"g ` (\<Union>\<F> - C) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
665  | 
"\<And>x. x \<in> (\<Union>\<F> - C) \<inter> K \<Longrightarrow> g x = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
666  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
667  | 
  let ?\<F> = "{X \<in> \<F>. \<forall>Y\<in>\<F>. \<not> X \<subset> Y}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
668  | 
have [simp]: "\<Union>?\<F> = \<Union>\<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
669  | 
by (simp add: Union_maximal_sets assms)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
670  | 
have fin: "finite ?\<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
671  | 
by (force intro: finite_subset [OF _ \<open>finite \<F>\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
672  | 
have pw: "pairwise (\<lambda> S T. S \<inter> T \<subseteq> K) ?\<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
673  | 
by (simp add: pairwise_def) (metis K psubsetI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
674  | 
  have "card {X \<in> \<F>. \<forall>Y\<in>\<F>. \<not> X \<subset> Y} \<le> card \<F>"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
675  | 
by (simp add: \<open>finite \<F>\<close> card_mono)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
676  | 
moreover  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
677  | 
obtain C g where "finite C \<and> disjnt C U \<and> card C \<le> card ?\<F> \<and>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
678  | 
continuous_on (\<Union>?\<F> - C) g \<and> g ` (\<Union>?\<F> - C) \<subseteq> T  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
679  | 
\<and> (\<forall>x \<in> (\<Union>?\<F> - C) \<inter> K. g x = h x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
680  | 
apply (rule exE [OF extend_map_lemma_cofinite0 [OF fin pw, of U T h]])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
681  | 
apply (fastforce intro!: clo \<F>)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
682  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
683  | 
ultimately show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
684  | 
by (rule_tac C=C and g=g in that) auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
685  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
686  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
687  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
688  | 
lemma extend_map_lemma_cofinite:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
689  | 
assumes "finite \<F>" "\<G> \<subseteq> \<F>" and T: "convex T" "bounded T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
690  | 
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
691  | 
and contf: "continuous_on (\<Union>\<G>) f" and fim: "f ` (\<Union>\<G>) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
692  | 
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
693  | 
and aff: "\<And>X. X \<in> \<F> - \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
694  | 
obtains C g where  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
695  | 
"finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
696  | 
"g ` (\<Union> \<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
697  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
698  | 
  define \<H> where "\<H> \<equiv> \<G> \<union> {D. \<exists>C \<in> \<F> - \<G>. D face_of C \<and> aff_dim D < aff_dim T}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
699  | 
have "finite \<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
700  | 
using assms finite_subset by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
701  | 
  moreover have "finite (\<Union>{{D. D face_of C} |C. C \<in> \<F>})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
702  | 
apply (rule finite_Union)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
703  | 
apply (simp add: \<open>finite \<F>\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
704  | 
using finite_polytope_faces poly by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
705  | 
ultimately have "finite \<H>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
706  | 
apply (simp add: \<H>_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
707  | 
    apply (rule finite_subset [of _ "\<Union> {{D. D face_of C} | C. C \<in> \<F>}"], auto)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
708  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
709  | 
have *: "\<And>X Y. \<lbrakk>X \<in> \<H>; Y \<in> \<H>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
710  | 
unfolding \<H>_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
711  | 
apply (elim UnE bexE CollectE DiffE)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
712  | 
using subsetD [OF \<open>\<G> \<subseteq> \<F>\<close>] apply (simp_all add: face)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
713  | 
apply (meson subsetD [OF \<open>\<G> \<subseteq> \<F>\<close>] face face_of_Int_subface face_of_imp_subset face_of_refl poly polytope_imp_convex)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
714  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
715  | 
obtain h where conth: "continuous_on (\<Union>\<H>) h" and him: "h ` (\<Union>\<H>) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
716  | 
and hf: "\<And>x. x \<in> \<Union>\<G> \<Longrightarrow> h x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
717  | 
using \<open>finite \<H>\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
718  | 
unfolding \<H>_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
719  | 
apply (rule extend_map_lemma [OF _ Un_upper1 T _ _ _ contf fim])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
720  | 
using \<open>\<G> \<subseteq> \<F>\<close> face_of_polytope_polytope poly apply fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
721  | 
using * apply (auto simp: \<H>_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
722  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
723  | 
have "bounded (\<Union>\<G>)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
724  | 
using \<open>finite \<G>\<close> \<open>\<G> \<subseteq> \<F>\<close> poly polytope_imp_bounded by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
725  | 
then have "\<Union>\<G> \<noteq> UNIV"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
726  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
727  | 
then obtain a where a: "a \<notin> \<Union>\<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
728  | 
by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
729  | 
  have \<F>: "\<exists>a g. a \<notin> \<Union>\<G> \<and> continuous_on (D - {a}) g \<and>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
730  | 
                  g ` (D - {a}) \<subseteq> rel_frontier T \<and> (\<forall>x \<in> D \<inter> \<Union>\<H>. g x = h x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
731  | 
if "D \<in> \<F>" for D  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
732  | 
proof (cases "D \<subseteq> \<Union>\<H>")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
733  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
734  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
735  | 
apply (rule_tac x=a in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
736  | 
apply (rule_tac x=h in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
737  | 
using him apply (blast intro!: \<open>a \<notin> \<Union>\<G>\<close> continuous_on_subset [OF conth]) +  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
738  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
739  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
740  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
741  | 
note D_not_subset = False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
742  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
743  | 
proof (cases "D \<in> \<G>")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
744  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
745  | 
with D_not_subset show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
746  | 
by (auto simp: \<H>_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
747  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
748  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
749  | 
then have affD: "aff_dim D \<le> aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
750  | 
by (simp add: \<open>D \<in> \<F>\<close> aff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
751  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
752  | 
      proof (cases "rel_interior D = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
753  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
754  | 
with \<open>D \<in> \<F>\<close> poly a show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
755  | 
by (force simp: rel_interior_eq_empty polytope_imp_convex)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
756  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
757  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
758  | 
then obtain b where brelD: "b \<in> rel_interior D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
759  | 
by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
760  | 
have "polyhedron D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
761  | 
by (simp add: poly polytope_imp_polyhedron that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
762  | 
        have "rel_frontier D retract_of affine hull D - {b}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
763  | 
by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
764  | 
        then obtain r where relfD: "rel_frontier D \<subseteq> affine hull D - {b}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
765  | 
                        and contr: "continuous_on (affine hull D - {b}) r"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
766  | 
                        and rim: "r ` (affine hull D - {b}) \<subseteq> rel_frontier D"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
767  | 
and rid: "\<And>x. x \<in> rel_frontier D \<Longrightarrow> r x = x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
768  | 
by (auto simp: retract_of_def retraction_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
769  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
770  | 
proof (intro exI conjI ballI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
771  | 
show "b \<notin> \<Union>\<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
772  | 
proof clarify  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
773  | 
fix E  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
774  | 
assume "b \<in> E" "E \<in> \<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
775  | 
then have "E \<inter> D face_of E \<and> E \<inter> D face_of D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
776  | 
using \<open>\<G> \<subseteq> \<F>\<close> face that by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
777  | 
with face_of_subset_rel_frontier \<open>E \<in> \<G>\<close> \<open>b \<in> E\<close> brelD rel_interior_subset [of D]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
778  | 
D_not_subset rel_frontier_def \<H>_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
779  | 
show False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
780  | 
by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
781  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
782  | 
          have "r ` (D - {b}) \<subseteq> r ` (affine hull D - {b})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
783  | 
by (simp add: Diff_mono hull_subset image_mono)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
784  | 
also have "... \<subseteq> rel_frontier D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
785  | 
by (rule rim)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
786  | 
          also have "... \<subseteq> \<Union>{E. E face_of D \<and> aff_dim E < aff_dim T}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
787  | 
using affD  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
788  | 
by (force simp: rel_frontier_of_polyhedron [OF \<open>polyhedron D\<close>] facet_of_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
789  | 
also have "... \<subseteq> \<Union>(\<H>)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
790  | 
using D_not_subset \<H>_def that by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
791  | 
          finally have rsub: "r ` (D - {b}) \<subseteq> \<Union>(\<H>)" .
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
792  | 
          show "continuous_on (D - {b}) (h \<circ> r)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
793  | 
apply (intro conjI \<open>b \<notin> \<Union>\<G>\<close> continuous_on_compose)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
794  | 
apply (rule continuous_on_subset [OF contr])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
795  | 
apply (simp add: Diff_mono hull_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
796  | 
apply (rule continuous_on_subset [OF conth rsub])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
797  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
798  | 
          show "(h \<circ> r) ` (D - {b}) \<subseteq> rel_frontier T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
799  | 
using brelD him rsub by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
800  | 
show "(h \<circ> r) x = h x" if x: "x \<in> D \<inter> \<Union>\<H>" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
801  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
802  | 
consider A where "x \<in> D" "A \<in> \<G>" "x \<in> A"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
803  | 
| A B where "x \<in> D" "A face_of B" "B \<in> \<F>" "B \<notin> \<G>" "aff_dim A < aff_dim T" "x \<in> A"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
804  | 
using x by (auto simp: \<H>_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
805  | 
then have xrel: "x \<in> rel_frontier D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
806  | 
proof cases  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
807  | 
case 1 show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
808  | 
proof (rule face_of_subset_rel_frontier [THEN subsetD])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
809  | 
show "D \<inter> A face_of D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
810  | 
using \<open>A \<in> \<G>\<close> \<open>\<G> \<subseteq> \<F>\<close> face \<open>D \<in> \<F>\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
811  | 
show "D \<inter> A \<noteq> D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
812  | 
using \<open>A \<in> \<G>\<close> D_not_subset \<H>_def by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
813  | 
qed (auto simp: 1)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
814  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
815  | 
case 2 show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
816  | 
proof (rule face_of_subset_rel_frontier [THEN subsetD])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
817  | 
show "D \<inter> A face_of D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
818  | 
apply (rule face_of_Int_subface [of D B _ A, THEN conjunct1])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
819  | 
apply (simp_all add: 2 \<open>D \<in> \<F>\<close> face)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
820  | 
apply (simp add: \<open>polyhedron D\<close> polyhedron_imp_convex face_of_refl)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
821  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
822  | 
show "D \<inter> A \<noteq> D"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
823  | 
using "2" D_not_subset \<H>_def by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
824  | 
qed (auto simp: 2)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
825  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
826  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
827  | 
by (simp add: rid xrel)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
828  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
829  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
830  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
831  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
832  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
833  | 
have clo: "\<And>S. S \<in> \<F> \<Longrightarrow> closed S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
834  | 
by (simp add: poly polytope_imp_closed)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
835  | 
obtain C g where "finite C" "disjnt C (\<Union>\<G>)" "card C \<le> card \<F>" "continuous_on (\<Union>\<F> - C) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
836  | 
"g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
837  | 
and gh: "\<And>x. x \<in> (\<Union>\<F> - C) \<inter> \<Union>\<H> \<Longrightarrow> g x = h x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
838  | 
proof (rule extend_map_lemma_cofinite1 [OF \<open>finite \<F>\<close> \<F> clo])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
839  | 
show "X \<inter> Y \<subseteq> \<Union>\<H>" if XY: "X \<in> \<F>" "Y \<in> \<F>" and "\<not> X \<subseteq> Y" "\<not> Y \<subseteq> X" for X Y  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
840  | 
proof (cases "X \<in> \<G>")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
841  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
842  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
843  | 
by (auto simp: \<H>_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
844  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
845  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
846  | 
have "X \<inter> Y \<noteq> X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
847  | 
using \<open>\<not> X \<subseteq> Y\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
848  | 
with XY  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
849  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
850  | 
by (clarsimp simp: \<H>_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
851  | 
(metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
852  | 
not_le poly polytope_imp_convex)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
853  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
854  | 
qed (blast)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
855  | 
with \<open>\<G> \<subseteq> \<F>\<close> show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
856  | 
apply (rule_tac C=C and g=g in that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
857  | 
apply (auto simp: disjnt_def hf [symmetric] \<H>_def intro!: gh)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
858  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
859  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
860  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
861  | 
text\<open>The next two proofs are similar\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
862  | 
theorem extend_map_cell_complex_to_sphere:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
863  | 
assumes "finite \<F>" and S: "S \<subseteq> \<Union>\<F>" "closed S" and T: "convex T" "bounded T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
864  | 
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
865  | 
and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X < aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
866  | 
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
867  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
868  | 
obtains g where "continuous_on (\<Union>\<F>) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
869  | 
"g ` (\<Union>\<F>) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
870  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
871  | 
obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g ` V \<subseteq> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
872  | 
using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
873  | 
have "compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
874  | 
by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
875  | 
then obtain d where "d > 0" and d: "\<And>x y. \<lbrakk>x \<in> S; y \<in> - V\<rbrakk> \<Longrightarrow> d \<le> dist x y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
876  | 
using separate_compact_closed [of S "-V"] \<open>open V\<close> \<open>S \<subseteq> V\<close> by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
877  | 
obtain \<G> where "finite \<G>" "\<Union>\<G> = \<Union>\<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
878  | 
and diaG: "\<And>X. X \<in> \<G> \<Longrightarrow> diameter X < d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
879  | 
and polyG: "\<And>X. X \<in> \<G> \<Longrightarrow> polytope X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
880  | 
and affG: "\<And>X. X \<in> \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T - 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
881  | 
and faceG: "\<And>X Y. \<lbrakk>X \<in> \<G>; Y \<in> \<G>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
882  | 
proof (rule cell_complex_subdivision_exists [OF \<open>d>0\<close> \<open>finite \<F>\<close> poly _ face])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
883  | 
show "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> aff_dim T - 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
884  | 
by (simp add: aff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
885  | 
qed auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
886  | 
obtain h where conth: "continuous_on (\<Union>\<G>) h" and him: "h ` \<Union>\<G> \<subseteq> rel_frontier T" and hg: "\<And>x. x \<in> \<Union>(\<G> \<inter> Pow V) \<Longrightarrow> h x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
887  | 
proof (rule extend_map_lemma [of \<G> "\<G> \<inter> Pow V" T g])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
888  | 
show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
889  | 
by (metis Union_Int_subset Union_Pow_eq \<open>continuous_on V g\<close> continuous_on_subset le_inf_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
890  | 
qed (use \<open>finite \<G>\<close> T polyG affG faceG gim in fastforce)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
891  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
892  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
893  | 
show "continuous_on (\<Union>\<F>) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
894  | 
using \<open>\<Union>\<G> = \<Union>\<F>\<close> conth by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
895  | 
show "h ` \<Union>\<F> \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
896  | 
using \<open>\<Union>\<G> = \<Union>\<F>\<close> him by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
897  | 
show "h x = f x" if "x \<in> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
898  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
899  | 
have "x \<in> \<Union>\<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
900  | 
using \<open>\<Union>\<G> = \<Union>\<F>\<close> \<open>S \<subseteq> \<Union>\<F>\<close> that by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
901  | 
then obtain X where "x \<in> X" "X \<in> \<G>" by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
902  | 
then have "diameter X < d" "bounded X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
903  | 
by (auto simp: diaG \<open>X \<in> \<G>\<close> polyG polytope_imp_bounded)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
904  | 
then have "X \<subseteq> V" using d [OF \<open>x \<in> S\<close>] diameter_bounded_bound [OF \<open>bounded X\<close> \<open>x \<in> X\<close>]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
905  | 
by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
906  | 
have "h x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
907  | 
apply (rule hg)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
908  | 
using \<open>X \<in> \<G>\<close> \<open>X \<subseteq> V\<close> \<open>x \<in> X\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
909  | 
also have "... = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
910  | 
by (simp add: gf that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
911  | 
finally show "h x = f x" .  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
912  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
913  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
914  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
915  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
916  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
917  | 
theorem extend_map_cell_complex_to_sphere_cofinite:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
918  | 
assumes "finite \<F>" and S: "S \<subseteq> \<Union>\<F>" "closed S" and T: "convex T" "bounded T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
919  | 
and poly: "\<And>X. X \<in> \<F> \<Longrightarrow> polytope X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
920  | 
and aff: "\<And>X. X \<in> \<F> \<Longrightarrow> aff_dim X \<le> aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
921  | 
and face: "\<And>X Y. \<lbrakk>X \<in> \<F>; Y \<in> \<F>\<rbrakk> \<Longrightarrow> (X \<inter> Y) face_of X \<and> (X \<inter> Y) face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
922  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
923  | 
obtains C g where "finite C" "disjnt C S" "continuous_on (\<Union>\<F> - C) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
924  | 
"g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
925  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
926  | 
obtain V g where "S \<subseteq> V" "open V" "continuous_on V g" and gim: "g ` V \<subseteq> rel_frontier T" and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
927  | 
using neighbourhood_extension_into_ANR [OF contf fim _ \<open>closed S\<close>] ANR_rel_frontier_convex T by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
928  | 
have "compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
929  | 
by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
930  | 
then obtain d where "d > 0" and d: "\<And>x y. \<lbrakk>x \<in> S; y \<in> - V\<rbrakk> \<Longrightarrow> d \<le> dist x y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
931  | 
using separate_compact_closed [of S "-V"] \<open>open V\<close> \<open>S \<subseteq> V\<close> by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
932  | 
obtain \<G> where "finite \<G>" "\<Union>\<G> = \<Union>\<F>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
933  | 
and diaG: "\<And>X. X \<in> \<G> \<Longrightarrow> diameter X < d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
934  | 
and polyG: "\<And>X. X \<in> \<G> \<Longrightarrow> polytope X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
935  | 
and affG: "\<And>X. X \<in> \<G> \<Longrightarrow> aff_dim X \<le> aff_dim T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
936  | 
and faceG: "\<And>X Y. \<lbrakk>X \<in> \<G>; Y \<in> \<G>\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
937  | 
by (rule cell_complex_subdivision_exists [OF \<open>d>0\<close> \<open>finite \<F>\<close> poly aff face]) auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
938  | 
obtain C h where "finite C" and dis: "disjnt C (\<Union>(\<G> \<inter> Pow V))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
939  | 
and card: "card C \<le> card \<G>" and conth: "continuous_on (\<Union>\<G> - C) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
940  | 
and him: "h ` (\<Union>\<G> - C) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
941  | 
and hg: "\<And>x. x \<in> \<Union>(\<G> \<inter> Pow V) \<Longrightarrow> h x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
942  | 
proof (rule extend_map_lemma_cofinite [of \<G> "\<G> \<inter> Pow V" T g])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
943  | 
show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
944  | 
by (metis Union_Int_subset Union_Pow_eq \<open>continuous_on V g\<close> continuous_on_subset le_inf_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
945  | 
show "g ` \<Union>(\<G> \<inter> Pow V) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
946  | 
using gim by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
947  | 
qed (auto intro: \<open>finite \<G>\<close> T polyG affG dest: faceG)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
948  | 
have Ssub: "S \<subseteq> \<Union>(\<G> \<inter> Pow V)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
949  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
950  | 
fix x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
951  | 
assume "x \<in> S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
952  | 
then have "x \<in> \<Union>\<G>"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
953  | 
using \<open>\<Union>\<G> = \<Union>\<F>\<close> \<open>S \<subseteq> \<Union>\<F>\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
954  | 
then obtain X where "x \<in> X" "X \<in> \<G>" by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
955  | 
then have "diameter X < d" "bounded X"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
956  | 
by (auto simp: diaG \<open>X \<in> \<G>\<close> polyG polytope_imp_bounded)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
957  | 
then have "X \<subseteq> V" using d [OF \<open>x \<in> S\<close>] diameter_bounded_bound [OF \<open>bounded X\<close> \<open>x \<in> X\<close>]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
958  | 
by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
959  | 
then show "x \<in> \<Union>(\<G> \<inter> Pow V)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
960  | 
using \<open>X \<in> \<G>\<close> \<open>x \<in> X\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
961  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
962  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
963  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
964  | 
show "continuous_on (\<Union>\<F>-C) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
965  | 
using \<open>\<Union>\<G> = \<Union>\<F>\<close> conth by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
966  | 
show "h ` (\<Union>\<F> - C) \<subseteq> rel_frontier T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
967  | 
using \<open>\<Union>\<G> = \<Union>\<F>\<close> him by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
968  | 
show "h x = f x" if "x \<in> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
969  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
970  | 
have "h x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
971  | 
apply (rule hg)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
972  | 
using Ssub that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
973  | 
also have "... = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
974  | 
by (simp add: gf that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
975  | 
finally show "h x = f x" .  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
976  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
977  | 
show "disjnt C S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
978  | 
using dis Ssub by (meson disjnt_iff subset_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
979  | 
qed (intro \<open>finite C\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
980  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
981  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
982  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
983  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
984  | 
subsection\<open> Special cases and corollaries involving spheres.\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
985  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
986  | 
lemma disjnt_Diff1: "X \<subseteq> Y' \<Longrightarrow> disjnt (X - Y) (X' - Y')"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
987  | 
by (auto simp: disjnt_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
988  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
989  | 
proposition extend_map_affine_to_sphere_cofinite_simple:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
990  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
991  | 
assumes "compact S" "convex U" "bounded U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
992  | 
and aff: "aff_dim T \<le> aff_dim U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
993  | 
and "S \<subseteq> T" and contf: "continuous_on S f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
994  | 
and fim: "f ` S \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
995  | 
obtains K g where "finite K" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
996  | 
"g ` (T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
997  | 
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
998  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
999  | 
have "\<exists>K g. finite K \<and> disjnt K S \<and> continuous_on (T - K) g \<and>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1000  | 
g ` (T - K) \<subseteq> rel_frontier U \<and> (\<forall>x \<in> S. g x = f x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1001  | 
if "affine T" "S \<subseteq> T" and aff: "aff_dim T \<le> aff_dim U" for T  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1002  | 
  proof (cases "S = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1003  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1004  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1005  | 
    proof (cases "rel_frontier U = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1006  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1007  | 
with \<open>bounded U\<close> have "aff_dim U \<le> 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1008  | 
using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1009  | 
with aff have "aff_dim T \<le> 0" by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1010  | 
      then obtain a where "T \<subseteq> {a}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1011  | 
using \<open>affine T\<close> affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1012  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1013  | 
        using \<open>S = {}\<close> fim
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1014  | 
by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1015  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1016  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1017  | 
then obtain a where "a \<in> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1018  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1019  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1020  | 
        using continuous_on_const [of _ a] \<open>S = {}\<close> by force
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1021  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1022  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1023  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1024  | 
have "bounded S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1025  | 
by (simp add: \<open>compact S\<close> compact_imp_bounded)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1026  | 
then obtain b where b: "S \<subseteq> cbox (-b) b"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1027  | 
using bounded_subset_cbox_symmetric by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1028  | 
define bbox where "bbox \<equiv> cbox (-(b+One)) (b+One)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1029  | 
have "cbox (-b) b \<subseteq> bbox"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1030  | 
by (auto simp: bbox_def algebra_simps intro!: subset_box_imp)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1031  | 
with b \<open>S \<subseteq> T\<close> have "S \<subseteq> bbox \<inter> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1032  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1033  | 
    then have Ssub: "S \<subseteq> \<Union>{bbox \<inter> T}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1034  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1035  | 
then have "aff_dim (bbox \<inter> T) \<le> aff_dim U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1036  | 
by (metis aff aff_dim_subset inf_commute inf_le1 order_trans)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1037  | 
obtain K g where K: "finite K" "disjnt K S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1038  | 
                 and contg: "continuous_on (\<Union>{bbox \<inter> T} - K) g"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1039  | 
                 and gim: "g ` (\<Union>{bbox \<inter> T} - K) \<subseteq> rel_frontier U"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1040  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1041  | 
proof (rule extend_map_cell_complex_to_sphere_cofinite  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1042  | 
[OF _ Ssub _ \<open>convex U\<close> \<open>bounded U\<close> _ _ _ contf fim])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1043  | 
show "closed S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1044  | 
using \<open>compact S\<close> compact_eq_bounded_closed by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1045  | 
      show poly: "\<And>X. X \<in> {bbox \<inter> T} \<Longrightarrow> polytope X"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1046  | 
by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \<open>affine T\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1047  | 
      show "\<And>X Y. \<lbrakk>X \<in> {bbox \<inter> T}; Y \<in> {bbox \<inter> T}\<rbrakk> \<Longrightarrow> X \<inter> Y face_of X \<and> X \<inter> Y face_of Y"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1048  | 
by (simp add:poly face_of_refl polytope_imp_convex)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1049  | 
      show "\<And>X. X \<in> {bbox \<inter> T} \<Longrightarrow> aff_dim X \<le> aff_dim U"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1050  | 
by (simp add: \<open>aff_dim (bbox \<inter> T) \<le> aff_dim U\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1051  | 
qed auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1052  | 
define fro where "fro \<equiv> \<lambda>d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1053  | 
obtain d where d12: "1/2 \<le> d" "d \<le> 1" and dd: "disjnt K (fro d)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1054  | 
proof (rule disjoint_family_elem_disjnt [OF _ \<open>finite K\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1055  | 
      show "infinite {1/2..1::real}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1056  | 
by (simp add: infinite_Icc)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1057  | 
have dis1: "disjnt (fro x) (fro y)" if "x<y" for x y  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1058  | 
by (auto simp: algebra_simps that subset_box_imp disjnt_Diff1 frontier_def fro_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1059  | 
      then show "disjoint_family_on fro {1/2..1}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1060  | 
by (auto simp: disjoint_family_on_def disjnt_def neq_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1061  | 
qed auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1062  | 
define c where "c \<equiv> b + d *\<^sub>R One"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1063  | 
have cbsub: "cbox (-b) b \<subseteq> box (-c) c" "cbox (-b) b \<subseteq> cbox (-c) c" "cbox (-c) c \<subseteq> bbox"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1064  | 
using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1065  | 
have clo_cbT: "closed (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1066  | 
by (simp add: affine_closed closed_Int closed_cbox \<open>affine T\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1067  | 
    have cpT_ne: "cbox (- c) c \<inter> T \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1068  | 
      using \<open>S \<noteq> {}\<close> b cbsub(2) \<open>S \<subseteq> T\<close> by fastforce
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1069  | 
have "closest_point (cbox (- c) c \<inter> T) x \<notin> K" if "x \<in> T" "x \<notin> K" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1070  | 
proof (cases "x \<in> cbox (-c) c")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1071  | 
case True with that show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1072  | 
by (simp add: closest_point_self)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1073  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1074  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1075  | 
      have int_ne: "interior (cbox (-c) c) \<inter> T \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1076  | 
        using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b \<open>cbox (- b) b \<subseteq> box (- c) c\<close> by force
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1077  | 
have "convex T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1078  | 
by (meson \<open>affine T\<close> affine_imp_convex)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1079  | 
then have "x \<in> affine hull (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1080  | 
          by (metis Int_commute Int_iff \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> cbsub(1) \<open>x \<in> T\<close> affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1081  | 
then have "x \<in> affine hull (cbox (- c) c \<inter> T) - rel_interior (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1082  | 
by (meson DiffI False Int_iff rel_interior_subset subsetCE)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1083  | 
then have "closest_point (cbox (- c) c \<inter> T) x \<in> rel_frontier (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1084  | 
by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1085  | 
moreover have "(rel_frontier (cbox (- c) c \<inter> T)) \<subseteq> fro d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1086  | 
apply (subst convex_affine_rel_frontier_Int [OF _ \<open>affine T\<close> int_ne])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1087  | 
apply (auto simp: fro_def c_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1088  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1089  | 
ultimately show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1090  | 
using dd by (force simp: disjnt_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1091  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1092  | 
    then have cpt_subset: "closest_point (cbox (- c) c \<inter> T) ` (T - K) \<subseteq> \<Union>{bbox \<inter> T} - K"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1093  | 
using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1094  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1095  | 
proof (intro conjI ballI exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1096  | 
have "continuous_on (T - K) (closest_point (cbox (- c) c \<inter> T))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1097  | 
apply (rule continuous_on_closest_point)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1098  | 
        using \<open>S \<noteq> {}\<close> cbsub(2) b that
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1099  | 
by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \<open>affine T\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1100  | 
then show "continuous_on (T - K) (g \<circ> closest_point (cbox (- c) c \<inter> T))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1101  | 
by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1102  | 
      have "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K) \<subseteq> g ` (\<Union>{bbox \<inter> T} - K)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1103  | 
by (metis image_comp image_mono cpt_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1104  | 
also have "... \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1105  | 
by (rule gim)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1106  | 
finally show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K) \<subseteq> rel_frontier U" .  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1107  | 
show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) x = f x" if "x \<in> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1108  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1109  | 
have "(g \<circ> closest_point (cbox (- c) c \<inter> T)) x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1110  | 
unfolding o_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1111  | 
by (metis IntI \<open>S \<subseteq> T\<close> b cbsub(2) closest_point_self subset_eq that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1112  | 
also have "... = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1113  | 
by (simp add: that gf)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1114  | 
finally show ?thesis .  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1115  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1116  | 
qed (auto simp: K)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1117  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1118  | 
then obtain K g where "finite K" "disjnt K S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1119  | 
and contg: "continuous_on (affine hull T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1120  | 
and gim: "g ` (affine hull T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1121  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1122  | 
by (metis aff affine_affine_hull aff_dim_affine_hull  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1123  | 
order_trans [OF \<open>S \<subseteq> T\<close> hull_subset [of T affine]])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1124  | 
then obtain K g where "finite K" "disjnt K S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1125  | 
and contg: "continuous_on (T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1126  | 
and gim: "g ` (T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1127  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1128  | 
by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1129  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1130  | 
by (rule_tac K="K \<inter> T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1131  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1132  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1133  | 
subsection\<open>Extending maps to spheres\<close>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1134  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1135  | 
(*Up to extend_map_affine_to_sphere_cofinite_gen*)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1136  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1137  | 
lemma extend_map_affine_to_sphere1:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1138  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::topological_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1139  | 
assumes "finite K" "affine U" and contf: "continuous_on (U - K) f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1140  | 
and fim: "f ` (U - K) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1141  | 
      and comps: "\<And>C. \<lbrakk>C \<in> components(U - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1142  | 
and clo: "closedin (subtopology euclidean U) S" and K: "disjnt K S" "K \<subseteq> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1143  | 
obtains g where "continuous_on (U - L) g" "g ` (U - L) \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1144  | 
proof (cases "K = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1145  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1146  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1147  | 
by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1148  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1149  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1150  | 
have "S \<subseteq> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1151  | 
using clo closedin_limpt by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1152  | 
  then have "(U - S) \<inter> K \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1153  | 
by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1154  | 
  then have "\<Union>(components (U - S)) \<inter> K \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1155  | 
using Union_components by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1156  | 
  then obtain C0 where C0: "C0 \<in> components (U - S)" "C0 \<inter> K \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1157  | 
by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1158  | 
have "convex U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1159  | 
by (simp add: affine_imp_convex \<open>affine U\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1160  | 
then have "locally connected U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1161  | 
by (rule convex_imp_locally_connected)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1162  | 
  have "\<exists>a g. a \<in> C \<and> a \<in> L \<and> continuous_on (S \<union> (C - {a})) g \<and>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1163  | 
              g ` (S \<union> (C - {a})) \<subseteq> T \<and> (\<forall>x \<in> S. g x = f x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1164  | 
       if C: "C \<in> components (U - S)" and CK: "C \<inter> K \<noteq> {}" for C
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1165  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1166  | 
    have "C \<subseteq> U-S" "C \<inter> L \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1167  | 
by (simp_all add: in_components_subset comps that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1168  | 
then obtain a where a: "a \<in> C" "a \<in> L" by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1169  | 
have opeUC: "openin (subtopology euclidean U) C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1170  | 
proof (rule openin_trans)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1171  | 
show "openin (subtopology euclidean (U-S)) C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1172  | 
by (simp add: \<open>locally connected U\<close> clo locally_diff_closed openin_components_locally_connected [OF _ C])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1173  | 
show "openin (subtopology euclidean U) (U - S)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1174  | 
by (simp add: clo openin_diff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1175  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1176  | 
then obtain d where "C \<subseteq> U" "0 < d" and d: "cball a d \<inter> U \<subseteq> C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1177  | 
using openin_contains_cball by (metis \<open>a \<in> C\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1178  | 
then have "ball a d \<inter> U \<subseteq> C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1179  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1180  | 
obtain h k where homhk: "homeomorphism (S \<union> C) (S \<union> C) h k"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1181  | 
                 and subC: "{x. (~ (h x = x \<and> k x = x))} \<subseteq> C"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1182  | 
                 and bou: "bounded {x. (~ (h x = x \<and> k x = x))}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1183  | 
and hin: "\<And>x. x \<in> C \<inter> K \<Longrightarrow> h x \<in> ball a d \<inter> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1184  | 
proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \<inter> U" "C \<inter> K" "S \<union> C"])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1185  | 
show "openin (subtopology euclidean C) (ball a d \<inter> 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: 
65064 
diff
changeset
 | 
1186  | 
by (metis open_ball \<open>C \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology)  | 
| 
64006
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1187  | 
show "openin (subtopology euclidean (affine hull C)) C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1188  | 
by (metis \<open>a \<in> C\<close> \<open>openin (subtopology euclidean U) C\<close> affine_hull_eq affine_hull_openin all_not_in_conv \<open>affine U\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1189  | 
      show "ball a d \<inter> U \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1190  | 
using \<open>0 < d\<close> \<open>C \<subseteq> U\<close> \<open>a \<in> C\<close> by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1191  | 
show "finite (C \<inter> K)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1192  | 
by (simp add: \<open>finite K\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1193  | 
show "S \<union> C \<subseteq> affine hull C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1194  | 
by (metis \<open>C \<subseteq> U\<close> \<open>S \<subseteq> U\<close> \<open>a \<in> C\<close> opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1195  | 
show "connected C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1196  | 
by (metis C in_components_connected)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1197  | 
qed auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1198  | 
have a_BU: "a \<in> ball a d \<inter> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1199  | 
using \<open>0 < d\<close> \<open>C \<subseteq> U\<close> \<open>a \<in> C\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1200  | 
    have "rel_frontier (cball a d \<inter> U) retract_of (affine hull (cball a d \<inter> U) - {a})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1201  | 
apply (rule rel_frontier_retract_of_punctured_affine_hull)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1202  | 
apply (auto simp: \<open>convex U\<close> convex_Int)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1203  | 
by (metis \<open>affine U\<close> convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1204  | 
moreover have "rel_frontier (cball a d \<inter> U) = frontier (cball a d) \<inter> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1205  | 
apply (rule convex_affine_rel_frontier_Int)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1206  | 
using a_BU by (force simp: \<open>affine U\<close>)+  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1207  | 
moreover have "affine hull (cball a d \<inter> U) = U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1208  | 
by (metis \<open>convex U\<close> a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \<open>affine U\<close> equals0D inf.commute interior_cball)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1209  | 
    ultimately have "frontier (cball a d) \<inter> U retract_of (U - {a})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1210  | 
by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1211  | 
    then obtain r where contr: "continuous_on (U - {a}) r"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1212  | 
                    and rim: "r ` (U - {a}) \<subseteq> sphere a d"  "r ` (U - {a}) \<subseteq> U"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1213  | 
and req: "\<And>x. x \<in> sphere a d \<inter> U \<Longrightarrow> r x = x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1214  | 
using \<open>affine U\<close> by (auto simp: retract_of_def retraction_def hull_same)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1215  | 
define j where "j \<equiv> \<lambda>x. if x \<in> ball a d then r x else x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1216  | 
have kj: "\<And>x. x \<in> S \<Longrightarrow> k (j x) = x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1217  | 
using \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> j_def subC by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1218  | 
    have Uaeq: "U - {a} = (cball a d - {a}) \<inter> U \<union> (U - ball a d)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1219  | 
using \<open>0 < d\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1220  | 
    have jim: "j ` (S \<union> (C - {a})) \<subseteq> (S \<union> C) - ball a d"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1221  | 
proof clarify  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1222  | 
      fix y  assume "y \<in> S \<union> (C - {a})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1223  | 
      then have "y \<in> U - {a}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1224  | 
using \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>a \<in> C\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1225  | 
then have "r y \<in> sphere a d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1226  | 
using rim by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1227  | 
then show "j y \<in> S \<union> C - ball a d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1228  | 
apply (simp add: j_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1229  | 
        using \<open>r y \<in> sphere a d\<close> \<open>y \<in> U - {a}\<close> \<open>y \<in> S \<union> (C - {a})\<close> d rim by fastforce
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1230  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1231  | 
    have contj: "continuous_on (U - {a}) j"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1232  | 
unfolding j_def Uaeq  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1233  | 
proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1234  | 
      show "\<exists>T. closed T \<and> (cball a d - {a}) \<inter> U = (U - {a}) \<inter> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1235  | 
apply (rule_tac x="(cball a d) \<inter> U" in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1236  | 
using affine_closed \<open>affine U\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1237  | 
      show "\<exists>T. closed T \<and> U - ball a d = (U - {a}) \<inter> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1238  | 
apply (rule_tac x="U - ball a d" in exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1239  | 
using \<open>0 < d\<close> by (force simp: affine_closed \<open>affine U\<close> closed_Diff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1240  | 
      show "continuous_on ((cball a d - {a}) \<inter> U) r"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1241  | 
by (force intro: continuous_on_subset [OF contr])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1242  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1243  | 
have fT: "x \<in> U - K \<Longrightarrow> f x \<in> T" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1244  | 
using fim by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1245  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1246  | 
proof (intro conjI exI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1247  | 
      show "continuous_on (S \<union> (C - {a})) (f \<circ> k \<circ> j)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1248  | 
proof (intro continuous_on_compose)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1249  | 
        show "continuous_on (S \<union> (C - {a})) j"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1250  | 
apply (rule continuous_on_subset [OF contj])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1251  | 
using \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>a \<in> C\<close> by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1252  | 
        show "continuous_on (j ` (S \<union> (C - {a}))) k"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1253  | 
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1254  | 
using jim \<open>C \<subseteq> U - S\<close> \<open>S \<subseteq> U\<close> \<open>ball a d \<inter> U \<subseteq> C\<close> j_def by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1255  | 
        show "continuous_on (k ` j ` (S \<union> (C - {a}))) f"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1256  | 
proof (clarify intro!: continuous_on_subset [OF contf])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1257  | 
          fix y  assume "y \<in> S \<union> (C - {a})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1258  | 
have ky: "k y \<in> S \<union> C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1259  | 
            using homeomorphism_image2 [OF homhk] \<open>y \<in> S \<union> (C - {a})\<close> by blast
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1260  | 
have jy: "j y \<in> S \<union> C - ball a d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1261  | 
            using Un_iff \<open>y \<in> S \<union> (C - {a})\<close> jim by auto
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1262  | 
show "k (j y) \<in> U - K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1263  | 
apply safe  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1264  | 
using \<open>C \<subseteq> U\<close> \<open>S \<subseteq> U\<close> homeomorphism_image2 [OF homhk] jy apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1265  | 
by (metis DiffD1 DiffD2 Int_iff Un_iff \<open>disjnt K S\<close> disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1266  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1267  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1268  | 
have ST: "\<And>x. x \<in> S \<Longrightarrow> (f \<circ> k \<circ> j) x \<in> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1269  | 
apply (simp add: kj)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1270  | 
apply (metis DiffI \<open>S \<subseteq> U\<close> \<open>disjnt K S\<close> subsetD disjnt_iff fim image_subset_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1271  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1272  | 
moreover have "(f \<circ> k \<circ> j) x \<in> T" if "x \<in> C" "x \<noteq> a" "x \<notin> S" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1273  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1274  | 
have rx: "r x \<in> sphere a d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1275  | 
using \<open>C \<subseteq> U\<close> rim that by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1276  | 
have jj: "j x \<in> S \<union> C - ball a d"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1277  | 
using jim that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1278  | 
have "k (j x) = j x \<longrightarrow> k (j x) \<in> C \<or> j x \<in> C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1279  | 
by (metis Diff_iff Int_iff Un_iff \<open>S \<subseteq> U\<close> subsetD d j_def jj rx sphere_cball that(1))  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1280  | 
then have "k (j x) \<in> C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1281  | 
using homeomorphism_apply2 [OF homhk, of "j x"] \<open>C \<subseteq> U\<close> \<open>S \<subseteq> U\<close> a rx  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1282  | 
by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1283  | 
with jj \<open>C \<subseteq> U\<close> show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1284  | 
apply safe  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1285  | 
using ST j_def apply fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1286  | 
apply (auto simp: not_less intro!: fT)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1287  | 
by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1288  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1289  | 
      ultimately show "(f \<circ> k \<circ> j) ` (S \<union> (C - {a})) \<subseteq> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1290  | 
by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1291  | 
show "\<forall>x\<in>S. (f \<circ> k \<circ> j) x = f x" using kj by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1292  | 
qed (auto simp: a)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1293  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1294  | 
then obtain a h where  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1295  | 
    ah: "\<And>C. \<lbrakk>C \<in> components (U - S); C \<inter> K \<noteq> {}\<rbrakk>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1296  | 
           \<Longrightarrow> a C \<in> C \<and> a C \<in> L \<and> continuous_on (S \<union> (C - {a C})) (h C) \<and>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1297  | 
               h C ` (S \<union> (C - {a C})) \<subseteq> T \<and> (\<forall>x \<in> S. h C x = f x)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1298  | 
using that by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1299  | 
  define F where "F \<equiv> {C \<in> components (U - S). C \<inter> K \<noteq> {}}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1300  | 
  define G where "G \<equiv> {C \<in> components (U - S). C \<inter> K = {}}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1301  | 
  define UF where "UF \<equiv> (\<Union>C\<in>F. C - {a C})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1302  | 
have "C0 \<in> F"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1303  | 
by (auto simp: F_def C0)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1304  | 
have "finite F"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1305  | 
proof (subst finite_image_iff [of "\<lambda>C. C \<inter> K" F, symmetric])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1306  | 
show "inj_on (\<lambda>C. C \<inter> K) F"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1307  | 
unfolding F_def inj_on_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1308  | 
using components_nonoverlap by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1309  | 
show "finite ((\<lambda>C. C \<inter> K) ` F)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1310  | 
unfolding F_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1311  | 
by (rule finite_subset [of _ "Pow K"]) (auto simp: \<open>finite K\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1312  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1313  | 
obtain g where contg: "continuous_on (S \<union> UF) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1314  | 
             and gh: "\<And>x i. \<lbrakk>i \<in> F; x \<in> (S \<union> UF) \<inter> (S \<union> (i - {a i}))\<rbrakk>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1315  | 
\<Longrightarrow> g x = h i x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1316  | 
  proof (rule pasting_lemma_exists_closed [OF \<open>finite F\<close>, of "S \<union> UF" "\<lambda>C. S \<union> (C - {a C})" h])
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1317  | 
    show "S \<union> UF \<subseteq> (\<Union>C\<in>F. S \<union> (C - {a C}))"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1318  | 
using \<open>C0 \<in> F\<close> by (force simp: UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1319  | 
    show "closedin (subtopology euclidean (S \<union> UF)) (S \<union> (C - {a C}))"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1320  | 
if "C \<in> F" for C  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1321  | 
proof (rule closedin_closed_subset [of U "S \<union> C"])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1322  | 
show "closedin (subtopology euclidean U) (S \<union> C)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1323  | 
apply (rule closedin_Un_complement_component [OF \<open>locally connected U\<close> clo])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1324  | 
using F_def that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1325  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1326  | 
have "x = a C'" if "C' \<in> F" "x \<in> C'" "x \<notin> U" for x C'  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1327  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1328  | 
have "\<forall>A. x \<in> \<Union>A \<or> C' \<notin> A"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1329  | 
using \<open>x \<in> C'\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1330  | 
with that show "x = a C'"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1331  | 
by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1332  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1333  | 
then show "S \<union> UF \<subseteq> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1334  | 
using \<open>S \<subseteq> U\<close> by (force simp: UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1335  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1336  | 
      show "S \<union> (C - {a C}) = (S \<union> C) \<inter> (S \<union> UF)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1337  | 
using F_def UF_def components_nonoverlap that by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1338  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1339  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1340  | 
    show "continuous_on (S \<union> (C' - {a C'})) (h C')" if "C' \<in> F" for C'
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1341  | 
using ah F_def that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1342  | 
show "\<And>i j x. \<lbrakk>i \<in> F; j \<in> F;  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1343  | 
                   x \<in> (S \<union> UF) \<inter> (S \<union> (i - {a i})) \<inter> (S \<union> (j - {a j}))\<rbrakk>
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1344  | 
\<Longrightarrow> h i x = h j x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1345  | 
using components_eq by (fastforce simp: components_eq F_def ah)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1346  | 
qed blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1347  | 
have SU': "S \<union> \<Union>G \<union> (S \<union> UF) \<subseteq> U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1348  | 
using \<open>S \<subseteq> U\<close> in_components_subset by (auto simp: F_def G_def UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1349  | 
have clo1: "closedin (subtopology euclidean (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> \<Union>G)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1350  | 
proof (rule closedin_closed_subset [OF _ SU'])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1351  | 
have *: "\<And>C. C \<in> F \<Longrightarrow> openin (subtopology euclidean U) C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1352  | 
unfolding F_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1353  | 
by clarify (metis (no_types, lifting) \<open>locally connected U\<close> clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1354  | 
show "closedin (subtopology euclidean U) (U - UF)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1355  | 
unfolding UF_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1356  | 
by (force intro: openin_delete *)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1357  | 
show "S \<union> \<Union>G = (U - UF) \<inter> (S \<union> \<Union>G \<union> (S \<union> UF))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1358  | 
using \<open>S \<subseteq> U\<close> apply (auto simp: F_def G_def UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1359  | 
apply (metis Diff_iff UnionI Union_components)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1360  | 
apply (metis DiffD1 UnionI Union_components)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1361  | 
by (metis (no_types, lifting) IntI components_nonoverlap empty_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1362  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1363  | 
have clo2: "closedin (subtopology euclidean (S \<union> \<Union>G \<union> (S \<union> UF))) (S \<union> UF)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1364  | 
proof (rule closedin_closed_subset [OF _ SU'])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1365  | 
show "closedin (subtopology euclidean U) (\<Union>C\<in>F. S \<union> C)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1366  | 
apply (rule closedin_Union)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1367  | 
apply (simp add: \<open>finite F\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1368  | 
using F_def \<open>locally connected U\<close> clo closedin_Un_complement_component by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1369  | 
show "S \<union> UF = (\<Union>C\<in>F. S \<union> C) \<inter> (S \<union> \<Union>G \<union> (S \<union> UF))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1370  | 
using \<open>S \<subseteq> U\<close> apply (auto simp: F_def G_def UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1371  | 
using C0 apply blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1372  | 
by (metis components_nonoverlap disjnt_def disjnt_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1373  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1374  | 
have SUG: "S \<union> \<Union>G \<subseteq> U - K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1375  | 
using \<open>S \<subseteq> U\<close> K apply (auto simp: G_def disjnt_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1376  | 
by (meson Diff_iff subsetD in_components_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1377  | 
then have contf': "continuous_on (S \<union> \<Union>G) f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1378  | 
by (rule continuous_on_subset [OF contf])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1379  | 
have contg': "continuous_on (S \<union> UF) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1380  | 
apply (rule continuous_on_subset [OF contg])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1381  | 
using \<open>S \<subseteq> U\<close> by (auto simp: F_def G_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1382  | 
have "\<And>x. \<lbrakk>S \<subseteq> U; x \<in> S\<rbrakk> \<Longrightarrow> f x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1383  | 
by (subst gh) (auto simp: ah C0 intro: \<open>C0 \<in> F\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1384  | 
then have f_eq_g: "\<And>x. x \<in> S \<union> UF \<and> x \<in> S \<union> \<Union>G \<Longrightarrow> f x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1385  | 
using \<open>S \<subseteq> U\<close> apply (auto simp: F_def G_def UF_def dest: in_components_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1386  | 
using components_eq by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1387  | 
have cont: "continuous_on (S \<union> \<Union>G \<union> (S \<union> UF)) (\<lambda>x. if x \<in> S \<union> \<Union>G then f x else g x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1388  | 
by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\<lambda>x. x \<in> S \<union> \<Union>G"])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1389  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1390  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1391  | 
have UF: "\<Union>F - L \<subseteq> UF"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1392  | 
unfolding F_def UF_def using ah by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1393  | 
have "U - S - L = \<Union>(components (U - S)) - L"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1394  | 
by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1395  | 
also have "... = \<Union>F \<union> \<Union>G - L"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1396  | 
unfolding F_def G_def by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1397  | 
also have "... \<subseteq> UF \<union> \<Union>G"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1398  | 
using UF by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1399  | 
finally have "U - L \<subseteq> S \<union> \<Union>G \<union> (S \<union> UF)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1400  | 
by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1401  | 
then show "continuous_on (U - L) (\<lambda>x. if x \<in> S \<union> \<Union>G then f x else g x)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1402  | 
by (rule continuous_on_subset [OF cont])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1403  | 
    have "((U - L) \<inter> {x. x \<notin> S \<and> (\<forall>xa\<in>G. x \<notin> xa)}) \<subseteq>  ((U - L) \<inter> (-S \<inter> UF))"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1404  | 
using \<open>U - L \<subseteq> S \<union> \<Union>G \<union> (S \<union> UF)\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1405  | 
moreover have "g ` ((U - L) \<inter> (-S \<inter> UF)) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1406  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1407  | 
have "g x \<in> T" if "x \<in> U" "x \<notin> L" "x \<notin> S" "C \<in> F" "x \<in> C" "x \<noteq> a C" for x C  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1408  | 
proof (subst gh)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1409  | 
        show "x \<in> (S \<union> UF) \<inter> (S \<union> (C - {a C}))"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1410  | 
using that by (auto simp: UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1411  | 
show "h C x \<in> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1412  | 
using ah that by (fastforce simp add: F_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1413  | 
qed (rule that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1414  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1415  | 
by (force simp: UF_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1416  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1417  | 
    ultimately have "g ` ((U - L) \<inter> {x. x \<notin> S \<and> (\<forall>xa\<in>G. x \<notin> xa)}) \<subseteq> T"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1418  | 
using image_mono order_trans by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1419  | 
moreover have "f ` ((U - L) \<inter> (S \<union> \<Union>G)) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1420  | 
using fim SUG by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1421  | 
ultimately show "(\<lambda>x. if x \<in> S \<union> \<Union>G then f x else g x) ` (U - L) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1422  | 
by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1423  | 
show "\<And>x. x \<in> S \<Longrightarrow> (if x \<in> S \<union> \<Union>G then f x else g x) = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1424  | 
by (simp add: F_def G_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1425  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1426  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1427  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1428  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1429  | 
lemma extend_map_affine_to_sphere2:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1430  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1431  | 
assumes "compact S" "convex U" "bounded U" "affine T" "S \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1432  | 
and affTU: "aff_dim T \<le> aff_dim U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1433  | 
and contf: "continuous_on S f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1434  | 
and fim: "f ` S \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1435  | 
      and ovlap: "\<And>C. C \<in> components(T - S) \<Longrightarrow> C \<inter> L \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1436  | 
obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1437  | 
"continuous_on (T - K) g" "g ` (T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1438  | 
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1439  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1440  | 
obtain K g where K: "finite K" "K \<subseteq> T" "disjnt K S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1441  | 
and contg: "continuous_on (T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1442  | 
and gim: "g ` (T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1443  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1444  | 
using assms extend_map_affine_to_sphere_cofinite_simple by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1445  | 
have "(\<exists>y C. C \<in> components (T - S) \<and> x \<in> C \<and> y \<in> C \<and> y \<in> L)" if "x \<in> K" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1446  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1447  | 
have "x \<in> T-S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1448  | 
using \<open>K \<subseteq> T\<close> \<open>disjnt K S\<close> disjnt_def that by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1449  | 
then obtain C where "C \<in> components(T - S)" "x \<in> C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1450  | 
by (metis UnionE Union_components)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1451  | 
with ovlap [of C] show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1452  | 
by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1453  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1454  | 
then obtain \<xi> where \<xi>: "\<And>x. x \<in> K \<Longrightarrow> \<exists>C. C \<in> components (T - S) \<and> x \<in> C \<and> \<xi> x \<in> C \<and> \<xi> x \<in> L"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1455  | 
by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1456  | 
obtain h where conth: "continuous_on (T - \<xi> ` K) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1457  | 
and him: "h ` (T - \<xi> ` K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1458  | 
and hg: "\<And>x. x \<in> S \<Longrightarrow> h x = g x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1459  | 
proof (rule extend_map_affine_to_sphere1 [OF \<open>finite K\<close> \<open>affine T\<close> contg gim, of S "\<xi> ` K"])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1460  | 
show cloTS: "closedin (subtopology euclidean T) S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1461  | 
by (simp add: \<open>compact S\<close> \<open>S \<subseteq> T\<close> closed_subset compact_imp_closed)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1462  | 
    show "\<And>C. \<lbrakk>C \<in> components (T - S); C \<inter> K \<noteq> {}\<rbrakk> \<Longrightarrow> C \<inter> \<xi> ` K \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1463  | 
using \<xi> components_eq by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1464  | 
qed (use K in auto)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1465  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1466  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1467  | 
show *: "\<xi> ` K \<subseteq> L"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1468  | 
using \<xi> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1469  | 
show "finite (\<xi> ` K)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1470  | 
by (simp add: K)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1471  | 
show "\<xi> ` K \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1472  | 
by clarify (meson \<xi> Diff_iff contra_subsetD in_components_subset)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1473  | 
show "continuous_on (T - \<xi> ` K) h"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1474  | 
by (rule conth)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1475  | 
show "disjnt (\<xi> ` K) S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1476  | 
using K  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1477  | 
apply (auto simp: disjnt_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1478  | 
by (metis \<xi> DiffD2 UnionI Union_components)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1479  | 
qed (simp_all add: him hg gf)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1480  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1481  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1482  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1483  | 
proposition extend_map_affine_to_sphere_cofinite_gen:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1484  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1485  | 
assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1486  | 
and aff: "aff_dim T \<le> aff_dim U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1487  | 
and contf: "continuous_on S f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1488  | 
and fim: "f ` S \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1489  | 
      and dis: "\<And>C. \<lbrakk>C \<in> components(T - S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1490  | 
obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1491  | 
"g ` (T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1492  | 
"\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1493  | 
proof (cases "S = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1494  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1495  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1496  | 
  proof (cases "rel_frontier U = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1497  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1498  | 
with aff have "aff_dim T \<le> 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1499  | 
apply (simp add: rel_frontier_eq_empty)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1500  | 
using affine_bounded_eq_lowdim \<open>bounded U\<close> order_trans by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1501  | 
with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1502  | 
by linarith  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1503  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1504  | 
proof cases  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1505  | 
assume "aff_dim T = -1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1506  | 
      then have "T = {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1507  | 
by (simp add: aff_dim_empty)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1508  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1509  | 
        by (rule_tac K="{}" in that) auto
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1510  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1511  | 
assume "aff_dim T = 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1512  | 
      then obtain a where "T = {a}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1513  | 
using aff_dim_eq_0 by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1514  | 
then have "a \<in> L"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1515  | 
        using dis [of "{a}"] \<open>S = {}\<close> by (auto simp: in_components_self)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1516  | 
      with \<open>S = {}\<close> \<open>T = {a}\<close> show ?thesis
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1517  | 
        by (rule_tac K="{a}" and g=f in that) auto
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1518  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1519  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1520  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1521  | 
then obtain y where "y \<in> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1522  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1523  | 
    with \<open>S = {}\<close> show ?thesis
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1524  | 
      by (rule_tac K="{}" and g="\<lambda>x. y" in that)  (auto simp: continuous_on_const)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1525  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1526  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1527  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1528  | 
have "bounded S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1529  | 
by (simp add: assms compact_imp_bounded)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1530  | 
then obtain b where b: "S \<subseteq> cbox (-b) b"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1531  | 
using bounded_subset_cbox_symmetric by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1532  | 
  define LU where "LU \<equiv> L \<union> (\<Union> {C \<in> components (T - S). ~bounded C} - cbox (-(b+One)) (b+One))"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1533  | 
obtain K g where "finite K" "K \<subseteq> LU" "K \<subseteq> T" "disjnt K S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1534  | 
and contg: "continuous_on (T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1535  | 
and gim: "g ` (T - K) \<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1536  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1537  | 
proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1538  | 
    show "C \<inter> LU \<noteq> {}" if "C \<in> components (T - S)" for C
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1539  | 
proof (cases "bounded C")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1540  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1541  | 
with dis that show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1542  | 
unfolding LU_def by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1543  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1544  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1545  | 
      then have "\<not> bounded (\<Union>{C \<in> components (T - S). \<not> bounded C})"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1546  | 
by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1547  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1548  | 
apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1549  | 
by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1550  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1551  | 
qed blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1552  | 
have *: False if "x \<in> cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1553  | 
"x \<notin> box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1554  | 
"0 \<le> m" "m < n" "n \<le> 1" for m n x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1555  | 
using that by (auto simp: mem_box algebra_simps)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1556  | 
  have "disjoint_family_on (\<lambda>d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1557  | 
by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1558  | 
then obtain d where d12: "1/2 \<le> d" "d \<le> 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1559  | 
and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1560  | 
    using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\<lambda>d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"]
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1561  | 
by (auto simp: \<open>finite K\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1562  | 
define c where "c \<equiv> b + d *\<^sub>R One"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1563  | 
have cbsub: "cbox (-b) b \<subseteq> box (-c) c"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1564  | 
"cbox (-b) b \<subseteq> cbox (-c) c"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1565  | 
"cbox (-c) c \<subseteq> cbox (-(b+One)) (b+One)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1566  | 
using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1567  | 
have clo_cT: "closed (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1568  | 
using affine_closed \<open>affine T\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1569  | 
  have cT_ne: "cbox (- c) c \<inter> T \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1570  | 
    using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b cbsub by fastforce
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1571  | 
have S_sub_cc: "S \<subseteq> cbox (- c) c"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1572  | 
using \<open>cbox (- b) b \<subseteq> cbox (- c) c\<close> b by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1573  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1574  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1575  | 
show "finite (K \<inter> cbox (-(b+One)) (b+One))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1576  | 
using \<open>finite K\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1577  | 
show "K \<inter> cbox (- (b + One)) (b + One) \<subseteq> L"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1578  | 
using \<open>K \<subseteq> LU\<close> by (auto simp: LU_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1579  | 
show "K \<inter> cbox (- (b + One)) (b + One) \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1580  | 
using \<open>K \<subseteq> T\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1581  | 
show "disjnt (K \<inter> cbox (- (b + One)) (b + One)) S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1582  | 
using \<open>disjnt K S\<close> by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1583  | 
have cloTK: "closest_point (cbox (- c) c \<inter> T) x \<in> T - K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1584  | 
if "x \<in> T" and Knot: "x \<in> K \<longrightarrow> x \<notin> cbox (- b - One) (b + One)" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1585  | 
proof (cases "x \<in> cbox (- c) c")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1586  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1587  | 
with \<open>x \<in> T\<close> show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1588  | 
using cbsub(3) Knot by (force simp: closest_point_self)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1589  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1590  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1591  | 
have clo_in_rf: "closest_point (cbox (- c) c \<inter> T) x \<in> rel_frontier (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1592  | 
proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1593  | 
        have "T \<inter> interior (cbox (- c) c) \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1594  | 
          using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b cbsub(1) by fastforce
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1595  | 
then show "x \<in> affine hull (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1596  | 
by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \<open>affine T\<close> hull_inc that(1))  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1597  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1598  | 
show "False" if "x \<in> rel_interior (cbox (- c) c \<inter> T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1599  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1600  | 
          have "interior (cbox (- c) c) \<inter> T \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1601  | 
            using \<open>S \<noteq> {}\<close> \<open>S \<subseteq> T\<close> b cbsub(1) by fastforce
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1602  | 
then have "affine hull (T \<inter> cbox (- c) c) = T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1603  | 
using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1604  | 
by (simp add: affine_imp_convex \<open>affine T\<close> inf_commute)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1605  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1606  | 
by (meson subsetD le_inf_iff rel_interior_subset that False)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1607  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1608  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1609  | 
have "closest_point (cbox (- c) c \<inter> T) x \<notin> K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1610  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1611  | 
assume inK: "closest_point (cbox (- c) c \<inter> T) x \<in> K"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1612  | 
have "\<And>x. x \<in> K \<Longrightarrow> x \<notin> frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1613  | 
by (metis ddis disjnt_iff)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1614  | 
then show False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1615  | 
by (metis DiffI Int_iff \<open>affine T\<close> cT_ne c_def clo_cT clo_in_rf closest_point_in_set  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1616  | 
convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1617  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1618  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1619  | 
using cT_ne clo_cT closest_point_in_set by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1620  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1621  | 
show "continuous_on (T - K \<inter> cbox (- (b + One)) (b + One)) (g \<circ> closest_point (cbox (-c) c \<inter> T))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1622  | 
apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1623  | 
apply (simp_all add: clo_cT affine_imp_convex \<open>affine T\<close> convex_Int cT_ne)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1624  | 
using cloTK by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1625  | 
have "g (closest_point (cbox (- c) c \<inter> T) x) \<in> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1626  | 
if "x \<in> T" "x \<in> K \<longrightarrow> x \<notin> cbox (- b - One) (b + One)" for x  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1627  | 
apply (rule gim [THEN subsetD])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1628  | 
using that cloTK by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1629  | 
then show "(g \<circ> closest_point (cbox (- c) c \<inter> T)) ` (T - K \<inter> cbox (- (b + One)) (b + One))  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1630  | 
\<subseteq> rel_frontier U"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1631  | 
by force  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1632  | 
show "\<And>x. x \<in> S \<Longrightarrow> (g \<circ> closest_point (cbox (- c) c \<inter> T)) x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1633  | 
by simp (metis (mono_tags, lifting) IntI \<open>S \<subseteq> T\<close> cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1634  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1635  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1636  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1637  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1638  | 
corollary extend_map_affine_to_sphere_cofinite:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1639  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1640  | 
assumes SUT: "compact S" "affine T" "S \<subseteq> T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1641  | 
      and aff: "aff_dim T \<le> DIM('b)" and "0 \<le> r"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1642  | 
and contf: "continuous_on S f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1643  | 
and fim: "f ` S \<subseteq> sphere a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1644  | 
      and dis: "\<And>C. \<lbrakk>C \<in> components(T - S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1645  | 
obtains K g where "finite K" "K \<subseteq> L" "K \<subseteq> T" "disjnt K S" "continuous_on (T - K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1646  | 
"g ` (T - K) \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1647  | 
proof (cases "r = 0")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1648  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1649  | 
with fim show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1650  | 
    by (rule_tac K="{}" and g = "\<lambda>x. a" in that) (auto simp: continuous_on_const)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1651  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1652  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1653  | 
with assms have "0 < r" by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1654  | 
then have "aff_dim T \<le> aff_dim (cball a r)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1655  | 
by (simp add: aff aff_dim_cball)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1656  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1657  | 
apply (rule extend_map_affine_to_sphere_cofinite_gen  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1658  | 
[OF \<open>compact S\<close> convex_cball bounded_cball \<open>affine T\<close> \<open>S \<subseteq> T\<close> _ contf])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1659  | 
using fim apply (auto simp: assms False that dest: dis)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1660  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1661  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1662  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1663  | 
corollary extend_map_UNIV_to_sphere_cofinite:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1664  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1665  | 
  assumes aff: "DIM('a) \<le> DIM('b)" and "0 \<le> r"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1666  | 
and SUT: "compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1667  | 
and contf: "continuous_on S f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1668  | 
and fim: "f ` S \<subseteq> sphere a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1669  | 
      and dis: "\<And>C. \<lbrakk>C \<in> components(- S); bounded C\<rbrakk> \<Longrightarrow> C \<inter> L \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1670  | 
obtains K g where "finite K" "K \<subseteq> L" "disjnt K S" "continuous_on (- K) g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1671  | 
"g ` (- K) \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1672  | 
apply (rule extend_map_affine_to_sphere_cofinite  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1673  | 
[OF \<open>compact S\<close> affine_UNIV subset_UNIV _ \<open>0 \<le> r\<close> contf fim dis])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1674  | 
apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1675  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1676  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1677  | 
corollary extend_map_UNIV_to_sphere_no_bounded_component:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1678  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1679  | 
  assumes aff: "DIM('a) \<le> DIM('b)" and "0 \<le> r"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1680  | 
and SUT: "compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1681  | 
and contf: "continuous_on S f"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1682  | 
and fim: "f ` S \<subseteq> sphere a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1683  | 
and dis: "\<And>C. C \<in> components(- S) \<Longrightarrow> \<not> bounded C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1684  | 
obtains g where "continuous_on UNIV g" "g ` UNIV \<subseteq> sphere a r" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1685  | 
apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \<open>0 \<le> r\<close> \<open>compact S\<close> contf fim, of "{}"])
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1686  | 
apply (auto simp: that dest: dis)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1687  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1688  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1689  | 
theorem Borsuk_separation_theorem_gen:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1690  | 
fixes S :: "'a::euclidean_space set"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1691  | 
assumes "compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1692  | 
shows "(\<forall>c \<in> components(- S). ~bounded c) \<longleftrightarrow>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1693  | 
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1694  | 
\<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1695  | 
(is "?lhs = ?rhs")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1696  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1697  | 
assume L [rule_format]: ?lhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1698  | 
show ?rhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1699  | 
proof clarify  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1700  | 
fix f :: "'a \<Rightarrow> 'a"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1701  | 
assume contf: "continuous_on S f" and fim: "f ` S \<subseteq> sphere 0 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1702  | 
obtain g where contg: "continuous_on UNIV g" and gim: "range g \<subseteq> sphere 0 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1703  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1704  | 
by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \<open>compact S\<close> contf fim L]) auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1705  | 
then show "\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1706  | 
using nullhomotopic_from_contractible [OF contg gim]  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1707  | 
by (metis assms compact_imp_closed contf empty_iff fim homotopic_with_equal nullhomotopic_into_sphere_extension)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1708  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1709  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1710  | 
assume R [rule_format]: ?rhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1711  | 
show ?lhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1712  | 
unfolding components_def  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1713  | 
proof clarify  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1714  | 
fix a  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1715  | 
assume "a \<notin> S" and a: "bounded (connected_component_set (- S) a)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1716  | 
have cont: "continuous_on S (\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1717  | 
apply (intro continuous_intros)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1718  | 
using \<open>a \<notin> S\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1719  | 
have im: "(\<lambda>x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \<subseteq> sphere 0 1"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1720  | 
by clarsimp (metis \<open>a \<notin> S\<close> eq_iff_diff_eq_0 left_inverse norm_eq_zero)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1721  | 
show False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1722  | 
using R cont im Borsuk_map_essential_bounded_component [OF \<open>compact S\<close> \<open>a \<notin> S\<close>] a by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1723  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1724  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1725  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1726  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1727  | 
corollary Borsuk_separation_theorem:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1728  | 
fixes S :: "'a::euclidean_space set"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1729  | 
  assumes "compact S" and 2: "2 \<le> DIM('a)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1730  | 
shows "connected(- S) \<longleftrightarrow>  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1731  | 
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::'a) 1  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1732  | 
\<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) f (\<lambda>x. c)))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1733  | 
(is "?lhs = ?rhs")  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1734  | 
proof  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1735  | 
assume L: ?lhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1736  | 
show ?rhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1737  | 
  proof (cases "S = {}")
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1738  | 
case True  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1739  | 
then show ?thesis by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1740  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1741  | 
case False  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1742  | 
then have "(\<forall>c\<in>components (- S). \<not> bounded c)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1743  | 
by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1744  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1745  | 
by (simp add: Borsuk_separation_theorem_gen [OF \<open>compact S\<close>])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1746  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1747  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1748  | 
assume R: ?rhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1749  | 
then show ?lhs  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1750  | 
apply (simp add: Borsuk_separation_theorem_gen [OF \<open>compact S\<close>, symmetric])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1751  | 
apply (auto simp: components_def connected_iff_eq_connected_component_set)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1752  | 
using connected_component_in apply fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1753  | 
using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \<open>compact S\<close> compact_eq_bounded_closed by fastforce  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1754  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1755  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1756  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1757  | 
lemma homotopy_eqv_separation:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1758  | 
fixes S :: "'a::euclidean_space set" and T :: "'a set"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1759  | 
assumes "S homotopy_eqv T" and "compact S" and "compact T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1760  | 
shows "connected(- S) \<longleftrightarrow> connected(- T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1761  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1762  | 
  consider "DIM('a) = 1" | "2 \<le> DIM('a)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1763  | 
by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1764  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1765  | 
proof cases  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1766  | 
case 1  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1767  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1768  | 
using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1769  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1770  | 
case 2  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1771  | 
with assms show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1772  | 
by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1773  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1774  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1775  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1776  | 
lemma Jordan_Brouwer_separation:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1777  | 
fixes S :: "'a::euclidean_space set" and a::'a  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1778  | 
assumes hom: "S homeomorphic sphere a r" and "0 < r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1779  | 
shows "\<not> connected(- S)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1780  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1781  | 
  have "- sphere a r \<inter> ball a r \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1782  | 
using \<open>0 < r\<close> by (simp add: Int_absorb1 subset_eq)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1783  | 
moreover  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1784  | 
have eq: "- sphere a r - ball a r = - cball a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1785  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1786  | 
  have "- cball a r \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1787  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1788  | 
    have "frontier (cball a r) \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1789  | 
using \<open>0 < r\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1790  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1791  | 
by (metis frontier_complement frontier_empty)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1792  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1793  | 
  with eq have "- sphere a r - ball a r \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1794  | 
by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1795  | 
moreover  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1796  | 
have "connected (- S) = connected (- sphere a r)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1797  | 
proof (rule homotopy_eqv_separation)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1798  | 
show "S homotopy_eqv sphere a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1799  | 
using hom homeomorphic_imp_homotopy_eqv by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1800  | 
show "compact (sphere a r)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1801  | 
by simp  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1802  | 
then show " compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1803  | 
using hom homeomorphic_compactness by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1804  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1805  | 
ultimately show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1806  | 
using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \<open>0 < r\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1807  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1808  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1809  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1810  | 
lemma Jordan_Brouwer_frontier:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1811  | 
fixes S :: "'a::euclidean_space set" and a::'a  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1812  | 
  assumes S: "S homeomorphic sphere a r" and T: "T \<in> components(- S)" and 2: "2 \<le> DIM('a)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1813  | 
shows "frontier T = S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1814  | 
proof (cases r rule: linorder_cases)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1815  | 
assume "r < 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1816  | 
with S T show ?thesis by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1817  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1818  | 
assume "r = 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1819  | 
  with S T card_eq_SucD obtain b where "S = {b}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1820  | 
    by (auto simp: homeomorphic_finite [of "{a}" S])
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1821  | 
  have "components (- {b}) = { -{b}}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1822  | 
    using T \<open>S = {b}\<close> by (auto simp: components_eq_sing_iff connected_punctured_universe 2)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1823  | 
with T show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1824  | 
    by (metis \<open>S = {b}\<close> cball_trivial frontier_cball frontier_complement singletonD sphere_trivial)
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1825  | 
next  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1826  | 
assume "r > 0"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1827  | 
have "compact S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1828  | 
using homeomorphic_compactness compact_sphere S by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1829  | 
show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1830  | 
proof (rule frontier_minimal_separating_closed)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1831  | 
show "closed S"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1832  | 
using \<open>compact S\<close> compact_eq_bounded_closed by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1833  | 
show "\<not> connected (- S)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1834  | 
using Jordan_Brouwer_separation S \<open>0 < r\<close> by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1835  | 
obtain f g where hom: "homeomorphism S (sphere a r) f g"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1836  | 
using S by (auto simp: homeomorphic_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1837  | 
show "connected (- T)" if "closed T" "T \<subset> S" for T  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1838  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1839  | 
have "f ` T \<subseteq> sphere a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1840  | 
using \<open>T \<subset> S\<close> hom homeomorphism_image1 by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1841  | 
moreover have "f ` T \<noteq> sphere a r"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1842  | 
using \<open>T \<subset> S\<close> hom  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1843  | 
by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1844  | 
ultimately have "f ` T \<subset> sphere a r" by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1845  | 
then have "connected (- f ` T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1846  | 
by (rule psubset_sphere_Compl_connected [OF _ \<open>0 < r\<close> 2])  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1847  | 
moreover have "compact T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1848  | 
using \<open>compact S\<close> bounded_subset compact_eq_bounded_closed that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1849  | 
moreover then have "compact (f ` T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1850  | 
by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \<open>T \<subset> S\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1851  | 
moreover have "T homotopy_eqv f ` T"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1852  | 
by (meson \<open>f ` T \<subseteq> sphere a r\<close> dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \<open>T \<subset> S\<close>)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1853  | 
ultimately show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1854  | 
using homotopy_eqv_separation [of T "f`T"] by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1855  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1856  | 
qed (rule T)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1857  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1858  | 
|
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1859  | 
lemma Jordan_Brouwer_nonseparation:  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1860  | 
fixes S :: "'a::euclidean_space set" and a::'a  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1861  | 
  assumes S: "S homeomorphic sphere a r" and "T \<subset> S" and 2: "2 \<le> DIM('a)"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1862  | 
shows "connected(- T)"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1863  | 
proof -  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1864  | 
have *: "connected(C \<union> (S - T))" if "C \<in> components(- S)" for C  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1865  | 
proof (rule connected_intermediate_closure)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1866  | 
show "connected C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1867  | 
using in_components_connected that by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1868  | 
have "S = frontier C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1869  | 
using "2" Jordan_Brouwer_frontier S that by blast  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1870  | 
with closure_subset show "C \<union> (S - T) \<subseteq> closure C"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1871  | 
by (auto simp: frontier_def)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1872  | 
qed auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1873  | 
  have "components(- S) \<noteq> {}"
 | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1874  | 
by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1875  | 
components_eq_empty homeomorphic_compactness)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1876  | 
then have "- T = (\<Union>C \<in> components(- S). C \<union> (S - T))"  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1877  | 
using Union_components [of "-S"] \<open>T \<subset> S\<close> by auto  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1878  | 
then show ?thesis  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1879  | 
apply (rule ssubst)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1880  | 
apply (rule connected_Union)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1881  | 
using \<open>T \<subset> S\<close> apply (auto simp: *)  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1882  | 
done  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1883  | 
qed  | 
| 
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1884  | 
|
| 64122 | 1885  | 
subsection\<open> Invariance of domain and corollaries\<close>  | 
1886  | 
||
1887  | 
lemma invariance_of_domain_ball:  | 
|
1888  | 
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"  | 
|
1889  | 
assumes contf: "continuous_on (cball a r) f" and "0 < r"  | 
|
1890  | 
and inj: "inj_on f (cball a r)"  | 
|
1891  | 
shows "open(f ` ball a r)"  | 
|
1892  | 
proof (cases "DIM('a) = 1")
 | 
|
1893  | 
case True  | 
|
1894  | 
obtain h::"'a\<Rightarrow>real" and k  | 
|
1895  | 
where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV"  | 
|
1896  | 
"\<And>x. norm(h x) = norm x" "\<And>x. norm(k x) = norm x"  | 
|
1897  | 
"\<And>x. k(h x) = x" "\<And>x. h(k x) = x"  | 
|
1898  | 
apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real])  | 
|
1899  | 
using True  | 
|
1900  | 
apply force  | 
|
1901  | 
by (metis UNIV_I UNIV_eq_I imageI)  | 
|
1902  | 
have cont: "continuous_on S h" "continuous_on T k" for S T  | 
|
1903  | 
by (simp_all add: \<open>linear h\<close> \<open>linear k\<close> linear_continuous_on linear_linear)  | 
|
1904  | 
have "continuous_on (h ` cball a r) (h \<circ> f \<circ> k)"  | 
|
1905  | 
apply (intro continuous_on_compose cont continuous_on_subset [OF contf])  | 
|
1906  | 
apply (auto simp: \<open>\<And>x. k (h x) = x\<close>)  | 
|
1907  | 
done  | 
|
1908  | 
moreover have "is_interval (h ` cball a r)"  | 
|
1909  | 
by (simp add: is_interval_connected_1 \<open>linear h\<close> linear_continuous_on linear_linear connected_continuous_image)  | 
|
1910  | 
moreover have "inj_on (h \<circ> f \<circ> k) (h ` cball a r)"  | 
|
1911  | 
using inj by (simp add: inj_on_def) (metis \<open>\<And>x. k (h x) = x\<close>)  | 
|
1912  | 
ultimately have *: "\<And>T. \<lbrakk>open T; T \<subseteq> h ` cball a r\<rbrakk> \<Longrightarrow> open ((h \<circ> f \<circ> k) ` T)"  | 
|
1913  | 
using injective_eq_1d_open_map_UNIV by blast  | 
|
1914  | 
have "open ((h \<circ> f \<circ> k) ` (h ` ball a r))"  | 
|
1915  | 
by (rule *) (auto simp: \<open>linear h\<close> \<open>range h = UNIV\<close> open_surjective_linear_image)  | 
|
1916  | 
then have "open ((h \<circ> f) ` ball a r)"  | 
|
1917  | 
by (simp add: image_comp \<open>\<And>x. k (h x) = x\<close> cong: image_cong)  | 
|
1918  | 
then show ?thesis  | 
|
1919  | 
apply (simp add: image_comp [symmetric])  | 
|
1920  | 
apply (metis open_bijective_linear_image_eq \<open>linear h\<close> \<open>\<And>x. k (h x) = x\<close> \<open>range h = UNIV\<close> bijI inj_on_def)  | 
|
1921  | 
done  | 
|
1922  | 
next  | 
|
1923  | 
case False  | 
|
1924  | 
  then have 2: "DIM('a) \<ge> 2"
 | 
|
1925  | 
by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq)  | 
|
1926  | 
have fimsub: "f ` ball a r \<subseteq> - f ` sphere a r"  | 
|
1927  | 
using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl)  | 
|
1928  | 
have hom: "f ` sphere a r homeomorphic sphere a r"  | 
|
1929  | 
by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball)  | 
|
1930  | 
then have nconn: "\<not> connected (- f ` sphere a r)"  | 
|
1931  | 
by (rule Jordan_Brouwer_separation) (auto simp: \<open>0 < r\<close>)  | 
|
1932  | 
obtain C where C: "C \<in> components (- f ` sphere a r)" and "bounded C"  | 
|
1933  | 
apply (rule cobounded_has_bounded_component [OF _ nconn])  | 
|
1934  | 
apply (simp_all add: 2)  | 
|
1935  | 
by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball)  | 
|
1936  | 
moreover have "f ` (ball a r) = C"  | 
|
1937  | 
proof  | 
|
1938  | 
    have "C \<noteq> {}"
 | 
|
1939  | 
by (rule in_components_nonempty [OF C])  | 
|
1940  | 
show "C \<subseteq> f ` ball a r"  | 
|
1941  | 
proof (rule ccontr)  | 
|
1942  | 
assume nonsub: "\<not> C \<subseteq> f ` ball a r"  | 
|
1943  | 
have "- f ` cball a r \<subseteq> C"  | 
|
1944  | 
proof (rule components_maximal [OF C])  | 
|
1945  | 
have "f ` cball a r homeomorphic cball a r"  | 
|
1946  | 
using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast  | 
|
1947  | 
then show "connected (- f ` cball a r)"  | 
|
1948  | 
by (auto intro: connected_complement_homeomorphic_convex_compact 2)  | 
|
1949  | 
show "- f ` cball a r \<subseteq> - f ` sphere a r"  | 
|
1950  | 
by auto  | 
|
1951  | 
        then show "C \<inter> - f ` cball a r \<noteq> {}"
 | 
|
1952  | 
          using \<open>C \<noteq> {}\<close> in_components_subset [OF C] nonsub
 | 
|
1953  | 
using image_iff by fastforce  | 
|
1954  | 
qed  | 
|
1955  | 
then have "bounded (- f ` cball a r)"  | 
|
1956  | 
using bounded_subset \<open>bounded C\<close> by auto  | 
|
1957  | 
then have "\<not> bounded (f ` cball a r)"  | 
|
1958  | 
using cobounded_imp_unbounded by blast  | 
|
1959  | 
then show "False"  | 
|
1960  | 
using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast  | 
|
1961  | 
qed  | 
|
1962  | 
    with \<open>C \<noteq> {}\<close> have "C \<inter> f ` ball a r \<noteq> {}"
 | 
|
1963  | 
by (simp add: inf.absorb_iff1)  | 
|
1964  | 
then show "f ` ball a r \<subseteq> C"  | 
|
1965  | 
by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset)  | 
|
1966  | 
qed  | 
|
1967  | 
moreover have "open (- f ` sphere a r)"  | 
|
1968  | 
using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast  | 
|
1969  | 
ultimately show ?thesis  | 
|
1970  | 
using open_components by blast  | 
|
1971  | 
qed  | 
|
1972  | 
||
1973  | 
||
1974  | 
text\<open>Proved by L. E. J. Brouwer (1912)\<close>  | 
|
1975  | 
theorem invariance_of_domain:  | 
|
1976  | 
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"  | 
|
1977  | 
assumes "continuous_on S f" "open S" "inj_on f S"  | 
|
1978  | 
shows "open(f ` S)"  | 
|
1979  | 
unfolding open_subopen [of "f`S"]  | 
|
1980  | 
proof clarify  | 
|
1981  | 
fix a  | 
|
1982  | 
assume "a \<in> S"  | 
|
1983  | 
obtain \<delta> where "\<delta> > 0" and \<delta>: "cball a \<delta> \<subseteq> S"  | 
|
1984  | 
using \<open>open S\<close> \<open>a \<in> S\<close> open_contains_cball_eq by blast  | 
|
1985  | 
show "\<exists>T. open T \<and> f a \<in> T \<and> T \<subseteq> f ` S"  | 
|
1986  | 
proof (intro exI conjI)  | 
|
1987  | 
show "open (f ` (ball a \<delta>))"  | 
|
1988  | 
by (meson \<delta> \<open>0 < \<delta>\<close> assms continuous_on_subset inj_on_subset invariance_of_domain_ball)  | 
|
1989  | 
show "f a \<in> f ` ball a \<delta>"  | 
|
1990  | 
by (simp add: \<open>0 < \<delta>\<close>)  | 
|
1991  | 
show "f ` ball a \<delta> \<subseteq> f ` S"  | 
|
1992  | 
using \<delta> ball_subset_cball by blast  | 
|
1993  | 
qed  | 
|
1994  | 
qed  | 
|
1995  | 
||
1996  | 
lemma inv_of_domain_ss0:  | 
|
1997  | 
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"  | 
|
1998  | 
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"  | 
|
1999  | 
      and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)"
 | 
|
2000  | 
and ope: "openin (subtopology euclidean S) U"  | 
|
2001  | 
shows "openin (subtopology euclidean S) (f ` U)"  | 
|
2002  | 
proof -  | 
|
2003  | 
have "U \<subseteq> S"  | 
|
2004  | 
using ope openin_imp_subset by blast  | 
|
2005  | 
have "(UNIV::'b set) homeomorphic S"  | 
|
2006  | 
by (simp add: \<open>subspace S\<close> dimS dim_UNIV homeomorphic_subspaces)  | 
|
2007  | 
then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k"  | 
|
2008  | 
using homeomorphic_def by blast  | 
|
2009  | 
have homkh: "homeomorphism S (k ` S) k h"  | 
|
2010  | 
using homhk homeomorphism_image2 homeomorphism_sym by fastforce  | 
|
2011  | 
have "open ((k \<circ> f \<circ> h) ` k ` U)"  | 
|
2012  | 
proof (rule invariance_of_domain)  | 
|
2013  | 
show "continuous_on (k ` U) (k \<circ> f \<circ> h)"  | 
|
2014  | 
proof (intro continuous_intros)  | 
|
2015  | 
show "continuous_on (k ` U) h"  | 
|
2016  | 
by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest)  | 
|
2017  | 
show "continuous_on (h ` k ` U) f"  | 
|
2018  | 
apply (rule continuous_on_subset [OF contf], clarify)  | 
|
2019  | 
apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD)  | 
|
2020  | 
done  | 
|
2021  | 
show "continuous_on (f ` h ` k ` U) k"  | 
|
2022  | 
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]])  | 
|
2023  | 
using fim homhk homeomorphism_apply2 ope openin_subset by fastforce  | 
|
2024  | 
qed  | 
|
2025  | 
have ope_iff: "\<And>T. open T \<longleftrightarrow> openin (subtopology euclidean (k ` S)) T"  | 
|
2026  | 
using homhk homeomorphism_image2 open_openin by fastforce  | 
|
2027  | 
show "open (k ` U)"  | 
|
2028  | 
by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope])  | 
|
2029  | 
show "inj_on (k \<circ> f \<circ> h) (k ` U)"  | 
|
2030  | 
apply (clarsimp simp: inj_on_def)  | 
|
2031  | 
by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \<open>U \<subseteq> S\<close>)  | 
|
2032  | 
qed  | 
|
2033  | 
moreover  | 
|
2034  | 
have eq: "f ` U = h ` (k \<circ> f \<circ> h \<circ> k) ` U"  | 
|
2035  | 
apply (auto simp: image_comp [symmetric])  | 
|
2036  | 
apply (metis homkh \<open>U \<subseteq> S\<close> fim homeomorphism_image2 homeomorphism_of_subsets homhk imageI subset_UNIV)  | 
|
2037  | 
by (metis \<open>U \<subseteq> S\<close> subsetD fim homeomorphism_def homhk image_eqI)  | 
|
2038  | 
ultimately show ?thesis  | 
|
2039  | 
by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV)  | 
|
2040  | 
qed  | 
|
2041  | 
||
2042  | 
lemma inv_of_domain_ss1:  | 
|
2043  | 
fixes f :: "'a \<Rightarrow> 'a::euclidean_space"  | 
|
2044  | 
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S"  | 
|
2045  | 
and "subspace S"  | 
|
2046  | 
and ope: "openin (subtopology euclidean S) U"  | 
|
2047  | 
shows "openin (subtopology euclidean S) (f ` U)"  | 
|
2048  | 
proof -  | 
|
2049  | 
  define S' where "S' \<equiv> {y. \<forall>x \<in> S. orthogonal x y}"
 | 
|
2050  | 
have "subspace S'"  | 
|
2051  | 
by (simp add: S'_def subspace_orthogonal_to_vectors)  | 
|
2052  | 
define g where "g \<equiv> \<lambda>z::'a*'a. ((f \<circ> fst)z, snd z)"  | 
|
2053  | 
have "openin (subtopology euclidean (S \<times> S')) (g ` (U \<times> S'))"  | 
|
2054  | 
proof (rule inv_of_domain_ss0)  | 
|
2055  | 
show "continuous_on (U \<times> S') g"  | 
|
2056  | 
apply (simp add: g_def)  | 
|
2057  | 
apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto)  | 
|
2058  | 
done  | 
|
2059  | 
show "g ` (U \<times> S') \<subseteq> S \<times> S'"  | 
|
2060  | 
using fim by (auto simp: g_def)  | 
|
2061  | 
show "inj_on g (U \<times> S')"  | 
|
2062  | 
using injf by (auto simp: g_def inj_on_def)  | 
|
2063  | 
show "subspace (S \<times> S')"  | 
|
2064  | 
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> subspace_Times)  | 
|
2065  | 
show "openin (subtopology euclidean (S \<times> S')) (U \<times> S')"  | 
|
2066  | 
by (simp add: openin_Times [OF ope])  | 
|
2067  | 
have "dim (S \<times> S') = dim S + dim S'"  | 
|
2068  | 
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> dim_Times)  | 
|
2069  | 
    also have "... = DIM('a)"
 | 
|
2070  | 
using dim_subspace_orthogonal_to_vectors [OF \<open>subspace S\<close> subspace_UNIV]  | 
|
2071  | 
by (simp add: add.commute S'_def)  | 
|
2072  | 
    finally show "dim (S \<times> S') = DIM('a)" .
 | 
|
2073  | 
qed  | 
|
2074  | 
moreover have "g ` (U \<times> S') = f ` U \<times> S'"  | 
|
2075  | 
by (auto simp: g_def image_iff)  | 
|
2076  | 
moreover have "0 \<in> S'"  | 
|
2077  | 
using \<open>subspace S'\<close> subspace_affine by blast  | 
|
2078  | 
ultimately show ?thesis  | 
|
2079  | 
by (auto simp: openin_Times_eq)  | 
|
2080  | 
qed  | 
|
2081  | 
||
2082  | 
||
2083  | 
corollary invariance_of_domain_subspaces:  | 
|
2084  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2085  | 
assumes ope: "openin (subtopology euclidean U) S"  | 
|
2086  | 
and "subspace U" "subspace V" and VU: "dim V \<le> dim U"  | 
|
2087  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"  | 
|
2088  | 
and injf: "inj_on f S"  | 
|
2089  | 
shows "openin (subtopology euclidean V) (f ` S)"  | 
|
2090  | 
proof -  | 
|
2091  | 
obtain V' where "subspace V'" "V' \<subseteq> U" "dim V' = dim V"  | 
|
2092  | 
using choose_subspace_of_subspace [OF VU]  | 
|
2093  | 
by (metis span_eq \<open>subspace U\<close>)  | 
|
2094  | 
then have "V homeomorphic V'"  | 
|
2095  | 
by (simp add: \<open>subspace V\<close> homeomorphic_subspaces)  | 
|
2096  | 
then obtain h k where homhk: "homeomorphism V V' h k"  | 
|
2097  | 
using homeomorphic_def by blast  | 
|
2098  | 
have eq: "f ` S = k ` (h \<circ> f) ` S"  | 
|
2099  | 
proof -  | 
|
2100  | 
have "k ` h ` f ` S = f ` S"  | 
|
2101  | 
by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl)  | 
|
2102  | 
then show ?thesis  | 
|
2103  | 
by (simp add: image_comp)  | 
|
2104  | 
qed  | 
|
2105  | 
show ?thesis  | 
|
2106  | 
unfolding eq  | 
|
2107  | 
proof (rule homeomorphism_imp_open_map)  | 
|
2108  | 
show homkh: "homeomorphism V' V k h"  | 
|
2109  | 
by (simp add: homeomorphism_symD homhk)  | 
|
2110  | 
have hfV': "(h \<circ> f) ` S \<subseteq> V'"  | 
|
2111  | 
using fim homeomorphism_image1 homhk by fastforce  | 
|
2112  | 
moreover have "openin (subtopology euclidean U) ((h \<circ> f) ` S)"  | 
|
2113  | 
proof (rule inv_of_domain_ss1)  | 
|
2114  | 
show "continuous_on S (h \<circ> f)"  | 
|
2115  | 
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)  | 
|
2116  | 
show "inj_on (h \<circ> f) S"  | 
|
2117  | 
apply (clarsimp simp: inj_on_def)  | 
|
2118  | 
by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf)  | 
|
2119  | 
show "(h \<circ> f) ` S \<subseteq> U"  | 
|
2120  | 
using \<open>V' \<subseteq> U\<close> hfV' by auto  | 
|
2121  | 
qed (auto simp: assms)  | 
|
2122  | 
ultimately show "openin (subtopology euclidean V') ((h \<circ> f) ` S)"  | 
|
2123  | 
using openin_subset_trans \<open>V' \<subseteq> U\<close> by force  | 
|
2124  | 
qed  | 
|
2125  | 
qed  | 
|
2126  | 
||
2127  | 
corollary invariance_of_dimension_subspaces:  | 
|
2128  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2129  | 
assumes ope: "openin (subtopology euclidean U) S"  | 
|
2130  | 
and "subspace U" "subspace V"  | 
|
2131  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"  | 
|
2132  | 
      and injf: "inj_on f S" and "S \<noteq> {}"
 | 
|
2133  | 
shows "dim U \<le> dim V"  | 
|
2134  | 
proof -  | 
|
2135  | 
have "False" if "dim V < dim U"  | 
|
2136  | 
proof -  | 
|
2137  | 
obtain T where "subspace T" "T \<subseteq> U" "dim T = dim V"  | 
|
2138  | 
using choose_subspace_of_subspace [of "dim V" U]  | 
|
2139  | 
by (metis span_eq \<open>subspace U\<close> \<open>dim V < dim U\<close> linear not_le)  | 
|
2140  | 
then have "V homeomorphic T"  | 
|
2141  | 
by (simp add: \<open>subspace V\<close> homeomorphic_subspaces)  | 
|
2142  | 
then obtain h k where homhk: "homeomorphism V T h k"  | 
|
2143  | 
using homeomorphic_def by blast  | 
|
2144  | 
have "continuous_on S (h \<circ> f)"  | 
|
2145  | 
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)  | 
|
2146  | 
moreover have "(h \<circ> f) ` S \<subseteq> U"  | 
|
2147  | 
using \<open>T \<subseteq> U\<close> fim homeomorphism_image1 homhk by fastforce  | 
|
2148  | 
moreover have "inj_on (h \<circ> f) S"  | 
|
2149  | 
apply (clarsimp simp: inj_on_def)  | 
|
2150  | 
by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf)  | 
|
2151  | 
ultimately have ope_hf: "openin (subtopology euclidean U) ((h \<circ> f) ` S)"  | 
|
2152  | 
using invariance_of_domain_subspaces [OF ope \<open>subspace U\<close> \<open>subspace U\<close>] by auto  | 
|
2153  | 
have "(h \<circ> f) ` S \<subseteq> T"  | 
|
2154  | 
using fim homeomorphism_image1 homhk by fastforce  | 
|
2155  | 
then show ?thesis  | 
|
2156  | 
      by (metis dim_openin \<open>dim T = dim V\<close> ope_hf \<open>subspace U\<close> \<open>S \<noteq> {}\<close> dim_subset image_is_empty not_le that)
 | 
|
2157  | 
qed  | 
|
2158  | 
then show ?thesis  | 
|
2159  | 
using not_less by blast  | 
|
2160  | 
qed  | 
|
2161  | 
||
2162  | 
corollary invariance_of_domain_affine_sets:  | 
|
2163  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2164  | 
assumes ope: "openin (subtopology euclidean U) S"  | 
|
2165  | 
and aff: "affine U" "affine V" "aff_dim V \<le> aff_dim U"  | 
|
2166  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"  | 
|
2167  | 
and injf: "inj_on f S"  | 
|
2168  | 
shows "openin (subtopology euclidean V) (f ` S)"  | 
|
2169  | 
proof (cases "S = {}")
 | 
|
2170  | 
case True  | 
|
2171  | 
then show ?thesis by auto  | 
|
2172  | 
next  | 
|
2173  | 
case False  | 
|
2174  | 
obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"  | 
|
2175  | 
using False fim ope openin_contains_cball by fastforce  | 
|
| 67399 | 2176  | 
have "openin (subtopology euclidean ((+) (- b) ` V)) (((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S)"  | 
| 64122 | 2177  | 
proof (rule invariance_of_domain_subspaces)  | 
| 67399 | 2178  | 
show "openin (subtopology euclidean ((+) (- a) ` U)) ((+) (- a) ` S)"  | 
| 64122 | 2179  | 
by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois)  | 
| 67399 | 2180  | 
show "subspace ((+) (- a) ` U)"  | 
| 64122 | 2181  | 
by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace \<open>affine U\<close>)  | 
| 67399 | 2182  | 
show "subspace ((+) (- b) ` V)"  | 
| 64122 | 2183  | 
by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace \<open>affine V\<close>)  | 
| 67399 | 2184  | 
show "dim ((+) (- b) ` V) \<le> dim ((+) (- a) ` U)"  | 
| 64122 | 2185  | 
by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)  | 
| 67399 | 2186  | 
show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"  | 
| 64122 | 2187  | 
by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)  | 
| 67399 | 2188  | 
show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"  | 
| 64122 | 2189  | 
using fim by auto  | 
| 67399 | 2190  | 
show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"  | 
| 64122 | 2191  | 
by (auto simp: inj_on_def) (meson inj_onD injf)  | 
2192  | 
qed  | 
|
2193  | 
then show ?thesis  | 
|
2194  | 
by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois)  | 
|
2195  | 
qed  | 
|
2196  | 
||
2197  | 
corollary invariance_of_dimension_affine_sets:  | 
|
2198  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2199  | 
assumes ope: "openin (subtopology euclidean U) S"  | 
|
2200  | 
and aff: "affine U" "affine V"  | 
|
2201  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"  | 
|
2202  | 
      and injf: "inj_on f S" and "S \<noteq> {}"
 | 
|
2203  | 
shows "aff_dim U \<le> aff_dim V"  | 
|
2204  | 
proof -  | 
|
2205  | 
obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"  | 
|
2206  | 
    using \<open>S \<noteq> {}\<close> fim ope openin_contains_cball by fastforce
 | 
|
| 67399 | 2207  | 
have "dim ((+) (- a) ` U) \<le> dim ((+) (- b) ` V)"  | 
| 64122 | 2208  | 
proof (rule invariance_of_dimension_subspaces)  | 
| 67399 | 2209  | 
show "openin (subtopology euclidean ((+) (- a) ` U)) ((+) (- a) ` S)"  | 
| 64122 | 2210  | 
by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois)  | 
| 67399 | 2211  | 
show "subspace ((+) (- a) ` U)"  | 
| 64122 | 2212  | 
by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace \<open>affine U\<close>)  | 
| 67399 | 2213  | 
show "subspace ((+) (- b) ` V)"  | 
| 64122 | 2214  | 
by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace \<open>affine V\<close>)  | 
| 67399 | 2215  | 
show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"  | 
| 64122 | 2216  | 
by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)  | 
| 67399 | 2217  | 
show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"  | 
| 64122 | 2218  | 
using fim by auto  | 
| 67399 | 2219  | 
show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"  | 
| 64122 | 2220  | 
by (auto simp: inj_on_def) (meson inj_onD injf)  | 
2221  | 
  qed (use \<open>S \<noteq> {}\<close> in auto)
 | 
|
2222  | 
then show ?thesis  | 
|
2223  | 
by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)  | 
|
2224  | 
qed  | 
|
2225  | 
||
2226  | 
corollary invariance_of_dimension:  | 
|
2227  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2228  | 
assumes contf: "continuous_on S f" and "open S"  | 
|
2229  | 
      and injf: "inj_on f S" and "S \<noteq> {}"
 | 
|
2230  | 
    shows "DIM('a) \<le> DIM('b)"
 | 
|
2231  | 
using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms  | 
|
2232  | 
by auto  | 
|
2233  | 
||
2234  | 
||
2235  | 
corollary continuous_injective_image_subspace_dim_le:  | 
|
2236  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2237  | 
assumes "subspace S" "subspace T"  | 
|
2238  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"  | 
|
2239  | 
and injf: "inj_on f S"  | 
|
2240  | 
shows "dim S \<le> dim T"  | 
|
2241  | 
apply (rule invariance_of_dimension_subspaces [of S S _ f])  | 
|
2242  | 
using assms by (auto simp: subspace_affine)  | 
|
2243  | 
||
2244  | 
lemma invariance_of_dimension_convex_domain:  | 
|
2245  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2246  | 
assumes "convex S"  | 
|
2247  | 
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> affine hull T"  | 
|
2248  | 
and injf: "inj_on f S"  | 
|
2249  | 
shows "aff_dim S \<le> aff_dim T"  | 
|
2250  | 
proof (cases "S = {}")
 | 
|
2251  | 
case True  | 
|
2252  | 
then show ?thesis by (simp add: aff_dim_geq)  | 
|
2253  | 
next  | 
|
2254  | 
case False  | 
|
2255  | 
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"  | 
|
2256  | 
proof (rule invariance_of_dimension_affine_sets)  | 
|
2257  | 
show "openin (subtopology euclidean (affine hull S)) (rel_interior S)"  | 
|
2258  | 
by (simp add: openin_rel_interior)  | 
|
2259  | 
show "continuous_on (rel_interior S) f"  | 
|
2260  | 
using contf continuous_on_subset rel_interior_subset by blast  | 
|
2261  | 
show "f ` rel_interior S \<subseteq> affine hull T"  | 
|
2262  | 
using fim rel_interior_subset by blast  | 
|
2263  | 
show "inj_on f (rel_interior S)"  | 
|
2264  | 
using inj_on_subset injf rel_interior_subset by blast  | 
|
2265  | 
    show "rel_interior S \<noteq> {}"
 | 
|
2266  | 
by (simp add: False \<open>convex S\<close> rel_interior_eq_empty)  | 
|
2267  | 
qed auto  | 
|
2268  | 
then show ?thesis  | 
|
2269  | 
by simp  | 
|
2270  | 
qed  | 
|
2271  | 
||
2272  | 
||
2273  | 
lemma homeomorphic_convex_sets_le:  | 
|
2274  | 
assumes "convex S" "S homeomorphic T"  | 
|
2275  | 
shows "aff_dim S \<le> aff_dim T"  | 
|
2276  | 
proof -  | 
|
2277  | 
obtain h k where homhk: "homeomorphism S T h k"  | 
|
2278  | 
using homeomorphic_def assms by blast  | 
|
2279  | 
show ?thesis  | 
|
2280  | 
proof (rule invariance_of_dimension_convex_domain [OF \<open>convex S\<close>])  | 
|
2281  | 
show "continuous_on S h"  | 
|
2282  | 
using homeomorphism_def homhk by blast  | 
|
2283  | 
show "h ` S \<subseteq> affine hull T"  | 
|
2284  | 
by (metis homeomorphism_def homhk hull_subset)  | 
|
2285  | 
show "inj_on h S"  | 
|
2286  | 
by (meson homeomorphism_apply1 homhk inj_on_inverseI)  | 
|
2287  | 
qed  | 
|
2288  | 
qed  | 
|
2289  | 
||
2290  | 
lemma homeomorphic_convex_sets:  | 
|
2291  | 
assumes "convex S" "convex T" "S homeomorphic T"  | 
|
2292  | 
shows "aff_dim S = aff_dim T"  | 
|
2293  | 
by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym)  | 
|
2294  | 
||
2295  | 
lemma homeomorphic_convex_compact_sets_eq:  | 
|
2296  | 
assumes "convex S" "compact S" "convex T" "compact T"  | 
|
2297  | 
shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"  | 
|
2298  | 
by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets)  | 
|
2299  | 
||
2300  | 
lemma invariance_of_domain_gen:  | 
|
2301  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2302  | 
  assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
 | 
|
2303  | 
shows "open(f ` S)"  | 
|
2304  | 
using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto  | 
|
2305  | 
||
2306  | 
lemma injective_into_1d_imp_open_map_UNIV:  | 
|
2307  | 
fixes f :: "'a::euclidean_space \<Rightarrow> real"  | 
|
2308  | 
assumes "open T" "continuous_on S f" "inj_on f S" "T \<subseteq> S"  | 
|
2309  | 
shows "open (f ` T)"  | 
|
2310  | 
apply (rule invariance_of_domain_gen [OF \<open>open T\<close>])  | 
|
2311  | 
using assms apply (auto simp: elim: continuous_on_subset subset_inj_on)  | 
|
2312  | 
done  | 
|
2313  | 
||
2314  | 
lemma continuous_on_inverse_open:  | 
|
2315  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2316  | 
  assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
 | 
|
2317  | 
shows "continuous_on (f ` S) g"  | 
|
2318  | 
proof (clarsimp simp add: continuous_openin_preimage_eq)  | 
|
2319  | 
fix T :: "'a set"  | 
|
2320  | 
assume "open T"  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
2321  | 
have eq: "f ` S \<inter> g -` T = f ` (S \<inter> T)"  | 
| 64122 | 2322  | 
by (auto simp: gf)  | 
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
2323  | 
show "openin (subtopology euclidean (f ` S)) (f ` S \<inter> g -` T)"  | 
| 64122 | 2324  | 
apply (subst eq)  | 
2325  | 
apply (rule open_openin_trans)  | 
|
2326  | 
apply (rule invariance_of_domain_gen)  | 
|
2327  | 
using assms  | 
|
2328  | 
apply auto  | 
|
2329  | 
using inj_on_inverseI apply auto[1]  | 
|
2330  | 
by (metis \<open>open T\<close> continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq)  | 
|
2331  | 
qed  | 
|
2332  | 
||
2333  | 
lemma invariance_of_domain_homeomorphism:  | 
|
2334  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2335  | 
  assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
 | 
|
2336  | 
obtains g where "homeomorphism S (f ` S) f g"  | 
|
2337  | 
proof  | 
|
2338  | 
show "homeomorphism S (f ` S) f (inv_into S f)"  | 
|
2339  | 
by (simp add: assms continuous_on_inverse_open homeomorphism_def)  | 
|
2340  | 
qed  | 
|
2341  | 
||
2342  | 
corollary invariance_of_domain_homeomorphic:  | 
|
2343  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2344  | 
  assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
 | 
|
2345  | 
shows "S homeomorphic (f ` S)"  | 
|
2346  | 
using invariance_of_domain_homeomorphism [OF assms]  | 
|
2347  | 
by (meson homeomorphic_def)  | 
|
2348  | 
||
| 64287 | 2349  | 
lemma continuous_image_subset_interior:  | 
2350  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2351  | 
  assumes "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
 | 
|
2352  | 
shows "f ` (interior S) \<subseteq> interior(f ` S)"  | 
|
2353  | 
apply (rule interior_maximal)  | 
|
2354  | 
apply (simp add: image_mono interior_subset)  | 
|
2355  | 
apply (rule invariance_of_domain_gen)  | 
|
2356  | 
using assms  | 
|
2357  | 
apply (auto simp: subset_inj_on interior_subset continuous_on_subset)  | 
|
2358  | 
done  | 
|
2359  | 
||
2360  | 
lemma homeomorphic_interiors_same_dimension:  | 
|
2361  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2362  | 
  assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)"
 | 
|
2363  | 
shows "(interior S) homeomorphic (interior T)"  | 
|
2364  | 
using assms [unfolded homeomorphic_minimal]  | 
|
2365  | 
unfolding homeomorphic_def  | 
|
2366  | 
proof (clarify elim!: ex_forward)  | 
|
2367  | 
fix f g  | 
|
2368  | 
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"  | 
|
2369  | 
and contf: "continuous_on S f" and contg: "continuous_on T g"  | 
|
2370  | 
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"  | 
|
2371  | 
by (auto simp: inj_on_def intro: rev_image_eqI) metis+  | 
|
2372  | 
have fim: "f ` interior S \<subseteq> interior T"  | 
|
2373  | 
using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp  | 
|
2374  | 
have gim: "g ` interior T \<subseteq> interior S"  | 
|
2375  | 
using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp  | 
|
2376  | 
show "homeomorphism (interior S) (interior T) f g"  | 
|
2377  | 
unfolding homeomorphism_def  | 
|
2378  | 
proof (intro conjI ballI)  | 
|
2379  | 
show "\<And>x. x \<in> interior S \<Longrightarrow> g (f x) = x"  | 
|
2380  | 
by (meson \<open>\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x\<close> subsetD interior_subset)  | 
|
2381  | 
have "interior T \<subseteq> f ` interior S"  | 
|
2382  | 
proof  | 
|
2383  | 
fix x assume "x \<in> interior T"  | 
|
2384  | 
then have "g x \<in> interior S"  | 
|
2385  | 
using gim by blast  | 
|
2386  | 
then show "x \<in> f ` interior S"  | 
|
2387  | 
by (metis T \<open>x \<in> interior T\<close> image_iff interior_subset subsetCE)  | 
|
2388  | 
qed  | 
|
2389  | 
then show "f ` interior S = interior T"  | 
|
2390  | 
using fim by blast  | 
|
2391  | 
show "continuous_on (interior S) f"  | 
|
2392  | 
by (metis interior_subset continuous_on_subset contf)  | 
|
2393  | 
show "\<And>y. y \<in> interior T \<Longrightarrow> f (g y) = y"  | 
|
2394  | 
by (meson T subsetD interior_subset)  | 
|
2395  | 
have "interior S \<subseteq> g ` interior T"  | 
|
2396  | 
proof  | 
|
2397  | 
fix x assume "x \<in> interior S"  | 
|
2398  | 
then have "f x \<in> interior T"  | 
|
2399  | 
using fim by blast  | 
|
2400  | 
then show "x \<in> g ` interior T"  | 
|
2401  | 
by (metis S \<open>x \<in> interior S\<close> image_iff interior_subset subsetCE)  | 
|
2402  | 
qed  | 
|
2403  | 
then show "g ` interior T = interior S"  | 
|
2404  | 
using gim by blast  | 
|
2405  | 
show "continuous_on (interior T) g"  | 
|
2406  | 
by (metis interior_subset continuous_on_subset contg)  | 
|
2407  | 
qed  | 
|
2408  | 
qed  | 
|
2409  | 
||
2410  | 
lemma homeomorphic_open_imp_same_dimension:  | 
|
2411  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2412  | 
  assumes "S homeomorphic T" "open S" "S \<noteq> {}" "open T" "T \<noteq> {}"
 | 
|
2413  | 
  shows "DIM('a) = DIM('b)"
 | 
|
2414  | 
using assms  | 
|
2415  | 
apply (simp add: homeomorphic_minimal)  | 
|
2416  | 
apply (rule order_antisym; metis inj_onI invariance_of_dimension)  | 
|
2417  | 
done  | 
|
2418  | 
||
2419  | 
lemma homeomorphic_interiors:  | 
|
2420  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2421  | 
  assumes "S homeomorphic T" "interior S = {} \<longleftrightarrow> interior T = {}"
 | 
|
2422  | 
shows "(interior S) homeomorphic (interior T)"  | 
|
2423  | 
proof (cases "interior T = {}")
 | 
|
2424  | 
case True  | 
|
2425  | 
with assms show ?thesis by auto  | 
|
2426  | 
next  | 
|
2427  | 
case False  | 
|
2428  | 
  then have "DIM('a) = DIM('b)"
 | 
|
2429  | 
using assms  | 
|
2430  | 
apply (simp add: homeomorphic_minimal)  | 
|
2431  | 
apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior)  | 
|
2432  | 
done  | 
|
2433  | 
then show ?thesis  | 
|
2434  | 
by (rule homeomorphic_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])  | 
|
2435  | 
qed  | 
|
2436  | 
||
2437  | 
lemma homeomorphic_frontiers_same_dimension:  | 
|
2438  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2439  | 
  assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)"
 | 
|
2440  | 
shows "(frontier S) homeomorphic (frontier T)"  | 
|
2441  | 
using assms [unfolded homeomorphic_minimal]  | 
|
2442  | 
unfolding homeomorphic_def  | 
|
2443  | 
proof (clarify elim!: ex_forward)  | 
|
2444  | 
fix f g  | 
|
2445  | 
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"  | 
|
2446  | 
and contf: "continuous_on S f" and contg: "continuous_on T g"  | 
|
2447  | 
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"  | 
|
2448  | 
by (auto simp: inj_on_def intro: rev_image_eqI) metis+  | 
|
2449  | 
have "g ` interior T \<subseteq> interior S"  | 
|
2450  | 
using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp  | 
|
2451  | 
then have fim: "f ` frontier S \<subseteq> frontier T"  | 
|
2452  | 
apply (simp add: frontier_def)  | 
|
2453  | 
using continuous_image_subset_interior assms(2) assms(3) S by auto  | 
|
2454  | 
have "f ` interior S \<subseteq> interior T"  | 
|
2455  | 
using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp  | 
|
2456  | 
then have gim: "g ` frontier T \<subseteq> frontier S"  | 
|
2457  | 
apply (simp add: frontier_def)  | 
|
2458  | 
using continuous_image_subset_interior T assms(2) assms(3) by auto  | 
|
2459  | 
show "homeomorphism (frontier S) (frontier T) f g"  | 
|
2460  | 
unfolding homeomorphism_def  | 
|
2461  | 
proof (intro conjI ballI)  | 
|
2462  | 
show gf: "\<And>x. x \<in> frontier S \<Longrightarrow> g (f x) = x"  | 
|
2463  | 
by (simp add: S assms(2) frontier_def)  | 
|
2464  | 
show fg: "\<And>y. y \<in> frontier T \<Longrightarrow> f (g y) = y"  | 
|
2465  | 
by (simp add: T assms(3) frontier_def)  | 
|
2466  | 
have "frontier T \<subseteq> f ` frontier S"  | 
|
2467  | 
proof  | 
|
2468  | 
fix x assume "x \<in> frontier T"  | 
|
2469  | 
then have "g x \<in> frontier S"  | 
|
2470  | 
using gim by blast  | 
|
2471  | 
then show "x \<in> f ` frontier S"  | 
|
2472  | 
by (metis fg \<open>x \<in> frontier T\<close> imageI)  | 
|
2473  | 
qed  | 
|
2474  | 
then show "f ` frontier S = frontier T"  | 
|
2475  | 
using fim by blast  | 
|
2476  | 
show "continuous_on (frontier S) f"  | 
|
2477  | 
by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def)  | 
|
2478  | 
have "frontier S \<subseteq> g ` frontier T"  | 
|
2479  | 
proof  | 
|
2480  | 
fix x assume "x \<in> frontier S"  | 
|
2481  | 
then have "f x \<in> frontier T"  | 
|
2482  | 
using fim by blast  | 
|
2483  | 
then show "x \<in> g ` frontier T"  | 
|
2484  | 
by (metis gf \<open>x \<in> frontier S\<close> imageI)  | 
|
2485  | 
qed  | 
|
2486  | 
then show "g ` frontier T = frontier S"  | 
|
2487  | 
using gim by blast  | 
|
2488  | 
show "continuous_on (frontier T) g"  | 
|
2489  | 
by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def)  | 
|
2490  | 
qed  | 
|
2491  | 
qed  | 
|
2492  | 
||
2493  | 
lemma homeomorphic_frontiers:  | 
|
2494  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2495  | 
assumes "S homeomorphic T" "closed S" "closed T"  | 
|
2496  | 
          "interior S = {} \<longleftrightarrow> interior T = {}"
 | 
|
2497  | 
shows "(frontier S) homeomorphic (frontier T)"  | 
|
2498  | 
proof (cases "interior T = {}")
 | 
|
2499  | 
case True  | 
|
2500  | 
then show ?thesis  | 
|
2501  | 
by (metis Diff_empty assms closure_eq frontier_def)  | 
|
2502  | 
next  | 
|
2503  | 
case False  | 
|
2504  | 
show ?thesis  | 
|
2505  | 
apply (rule homeomorphic_frontiers_same_dimension)  | 
|
2506  | 
apply (simp_all add: assms)  | 
|
2507  | 
using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast  | 
|
2508  | 
qed  | 
|
2509  | 
||
2510  | 
lemma continuous_image_subset_rel_interior:  | 
|
2511  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2512  | 
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"  | 
|
2513  | 
and TS: "aff_dim T \<le> aff_dim S"  | 
|
2514  | 
shows "f ` (rel_interior S) \<subseteq> rel_interior(f ` S)"  | 
|
2515  | 
proof (rule rel_interior_maximal)  | 
|
2516  | 
show "f ` rel_interior S \<subseteq> f ` S"  | 
|
2517  | 
by(simp add: image_mono rel_interior_subset)  | 
|
2518  | 
show "openin (subtopology euclidean (affine hull f ` S)) (f ` rel_interior S)"  | 
|
2519  | 
proof (rule invariance_of_domain_affine_sets)  | 
|
2520  | 
show "openin (subtopology euclidean (affine hull S)) (rel_interior S)"  | 
|
2521  | 
by (simp add: openin_rel_interior)  | 
|
2522  | 
show "aff_dim (affine hull f ` S) \<le> aff_dim (affine hull S)"  | 
|
2523  | 
by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans)  | 
|
2524  | 
show "f ` rel_interior S \<subseteq> affine hull f ` S"  | 
|
2525  | 
by (meson \<open>f ` rel_interior S \<subseteq> f ` S\<close> hull_subset order_trans)  | 
|
2526  | 
show "continuous_on (rel_interior S) f"  | 
|
2527  | 
using contf continuous_on_subset rel_interior_subset by blast  | 
|
2528  | 
show "inj_on f (rel_interior S)"  | 
|
2529  | 
using inj_on_subset injf rel_interior_subset by blast  | 
|
2530  | 
qed auto  | 
|
2531  | 
qed  | 
|
2532  | 
||
2533  | 
lemma homeomorphic_rel_interiors_same_dimension:  | 
|
2534  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2535  | 
assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"  | 
|
2536  | 
shows "(rel_interior S) homeomorphic (rel_interior T)"  | 
|
2537  | 
using assms [unfolded homeomorphic_minimal]  | 
|
2538  | 
unfolding homeomorphic_def  | 
|
2539  | 
proof (clarify elim!: ex_forward)  | 
|
2540  | 
fix f g  | 
|
2541  | 
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"  | 
|
2542  | 
and contf: "continuous_on S f" and contg: "continuous_on T g"  | 
|
2543  | 
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"  | 
|
2544  | 
by (auto simp: inj_on_def intro: rev_image_eqI) metis+  | 
|
2545  | 
have fim: "f ` rel_interior S \<subseteq> rel_interior T"  | 
|
2546  | 
by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)  | 
|
2547  | 
have gim: "g ` rel_interior T \<subseteq> rel_interior S"  | 
|
2548  | 
by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)  | 
|
2549  | 
show "homeomorphism (rel_interior S) (rel_interior T) f g"  | 
|
2550  | 
unfolding homeomorphism_def  | 
|
2551  | 
proof (intro conjI ballI)  | 
|
2552  | 
show gf: "\<And>x. x \<in> rel_interior S \<Longrightarrow> g (f x) = x"  | 
|
2553  | 
using S rel_interior_subset by blast  | 
|
2554  | 
show fg: "\<And>y. y \<in> rel_interior T \<Longrightarrow> f (g y) = y"  | 
|
2555  | 
using T mem_rel_interior_ball by blast  | 
|
2556  | 
have "rel_interior T \<subseteq> f ` rel_interior S"  | 
|
2557  | 
proof  | 
|
2558  | 
fix x assume "x \<in> rel_interior T"  | 
|
2559  | 
then have "g x \<in> rel_interior S"  | 
|
2560  | 
using gim by blast  | 
|
2561  | 
then show "x \<in> f ` rel_interior S"  | 
|
2562  | 
by (metis fg \<open>x \<in> rel_interior T\<close> imageI)  | 
|
2563  | 
qed  | 
|
2564  | 
moreover have "f ` rel_interior S \<subseteq> rel_interior T"  | 
|
2565  | 
by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)  | 
|
2566  | 
ultimately show "f ` rel_interior S = rel_interior T"  | 
|
2567  | 
by blast  | 
|
2568  | 
show "continuous_on (rel_interior S) f"  | 
|
2569  | 
using contf continuous_on_subset rel_interior_subset by blast  | 
|
2570  | 
have "rel_interior S \<subseteq> g ` rel_interior T"  | 
|
2571  | 
proof  | 
|
2572  | 
fix x assume "x \<in> rel_interior S"  | 
|
2573  | 
then have "f x \<in> rel_interior T"  | 
|
2574  | 
using fim by blast  | 
|
2575  | 
then show "x \<in> g ` rel_interior T"  | 
|
2576  | 
by (metis gf \<open>x \<in> rel_interior S\<close> imageI)  | 
|
2577  | 
qed  | 
|
2578  | 
then show "g ` rel_interior T = rel_interior S"  | 
|
2579  | 
using gim by blast  | 
|
2580  | 
show "continuous_on (rel_interior T) g"  | 
|
2581  | 
using contg continuous_on_subset rel_interior_subset by blast  | 
|
2582  | 
qed  | 
|
2583  | 
qed  | 
|
2584  | 
||
2585  | 
lemma homeomorphic_rel_interiors:  | 
|
2586  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2587  | 
  assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
 | 
|
2588  | 
shows "(rel_interior S) homeomorphic (rel_interior T)"  | 
|
2589  | 
proof (cases "rel_interior T = {}")
 | 
|
2590  | 
case True  | 
|
2591  | 
with assms show ?thesis by auto  | 
|
2592  | 
next  | 
|
2593  | 
case False  | 
|
2594  | 
obtain f g  | 
|
2595  | 
where S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"  | 
|
2596  | 
and contf: "continuous_on S f" and contg: "continuous_on T g"  | 
|
2597  | 
using assms [unfolded homeomorphic_minimal] by auto  | 
|
2598  | 
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"  | 
|
2599  | 
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f])  | 
|
2600  | 
apply (simp_all add: openin_rel_interior False assms)  | 
|
2601  | 
using contf continuous_on_subset rel_interior_subset apply blast  | 
|
2602  | 
apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)  | 
|
2603  | 
apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset)  | 
|
2604  | 
done  | 
|
2605  | 
moreover have "aff_dim (affine hull T) \<le> aff_dim (affine hull S)"  | 
|
2606  | 
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g])  | 
|
2607  | 
apply (simp_all add: openin_rel_interior False assms)  | 
|
2608  | 
using contg continuous_on_subset rel_interior_subset apply blast  | 
|
2609  | 
apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD)  | 
|
2610  | 
apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset)  | 
|
2611  | 
done  | 
|
2612  | 
ultimately have "aff_dim S = aff_dim T" by force  | 
|
2613  | 
then show ?thesis  | 
|
2614  | 
by (rule homeomorphic_rel_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])  | 
|
2615  | 
qed  | 
|
2616  | 
||
2617  | 
||
2618  | 
lemma homeomorphic_rel_boundaries_same_dimension:  | 
|
2619  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2620  | 
assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"  | 
|
2621  | 
shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"  | 
|
2622  | 
using assms [unfolded homeomorphic_minimal]  | 
|
2623  | 
unfolding homeomorphic_def  | 
|
2624  | 
proof (clarify elim!: ex_forward)  | 
|
2625  | 
fix f g  | 
|
2626  | 
assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"  | 
|
2627  | 
and contf: "continuous_on S f" and contg: "continuous_on T g"  | 
|
2628  | 
then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"  | 
|
2629  | 
by (auto simp: inj_on_def intro: rev_image_eqI) metis+  | 
|
2630  | 
have fim: "f ` rel_interior S \<subseteq> rel_interior T"  | 
|
2631  | 
by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)  | 
|
2632  | 
have gim: "g ` rel_interior T \<subseteq> rel_interior S"  | 
|
2633  | 
by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)  | 
|
2634  | 
show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g"  | 
|
2635  | 
unfolding homeomorphism_def  | 
|
2636  | 
proof (intro conjI ballI)  | 
|
2637  | 
show gf: "\<And>x. x \<in> S - rel_interior S \<Longrightarrow> g (f x) = x"  | 
|
2638  | 
using S rel_interior_subset by blast  | 
|
2639  | 
show fg: "\<And>y. y \<in> T - rel_interior T \<Longrightarrow> f (g y) = y"  | 
|
2640  | 
using T mem_rel_interior_ball by blast  | 
|
2641  | 
show "f ` (S - rel_interior S) = T - rel_interior T"  | 
|
2642  | 
using S fST fim gim by auto  | 
|
2643  | 
show "continuous_on (S - rel_interior S) f"  | 
|
2644  | 
using contf continuous_on_subset rel_interior_subset by blast  | 
|
2645  | 
show "g ` (T - rel_interior T) = S - rel_interior S"  | 
|
2646  | 
using T gTS gim fim by auto  | 
|
2647  | 
show "continuous_on (T - rel_interior T) g"  | 
|
2648  | 
using contg continuous_on_subset rel_interior_subset by blast  | 
|
2649  | 
qed  | 
|
2650  | 
qed  | 
|
2651  | 
||
2652  | 
lemma homeomorphic_rel_boundaries:  | 
|
2653  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2654  | 
  assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
 | 
|
2655  | 
shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"  | 
|
2656  | 
proof (cases "rel_interior T = {}")
 | 
|
2657  | 
case True  | 
|
2658  | 
with assms show ?thesis by auto  | 
|
2659  | 
next  | 
|
2660  | 
case False  | 
|
2661  | 
obtain f g  | 
|
2662  | 
where S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"  | 
|
2663  | 
and contf: "continuous_on S f" and contg: "continuous_on T g"  | 
|
2664  | 
using assms [unfolded homeomorphic_minimal] by auto  | 
|
2665  | 
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"  | 
|
2666  | 
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f])  | 
|
2667  | 
apply (simp_all add: openin_rel_interior False assms)  | 
|
2668  | 
using contf continuous_on_subset rel_interior_subset apply blast  | 
|
2669  | 
apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)  | 
|
2670  | 
apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset)  | 
|
2671  | 
done  | 
|
2672  | 
moreover have "aff_dim (affine hull T) \<le> aff_dim (affine hull S)"  | 
|
2673  | 
apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g])  | 
|
2674  | 
apply (simp_all add: openin_rel_interior False assms)  | 
|
2675  | 
using contg continuous_on_subset rel_interior_subset apply blast  | 
|
2676  | 
apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD)  | 
|
2677  | 
apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset)  | 
|
2678  | 
done  | 
|
2679  | 
ultimately have "aff_dim S = aff_dim T" by force  | 
|
2680  | 
then show ?thesis  | 
|
2681  | 
by (rule homeomorphic_rel_boundaries_same_dimension [OF \<open>S homeomorphic T\<close>])  | 
|
2682  | 
qed  | 
|
2683  | 
||
2684  | 
proposition uniformly_continuous_homeomorphism_UNIV_trivial:  | 
|
2685  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'a"  | 
|
2686  | 
assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g"  | 
|
2687  | 
shows "S = UNIV"  | 
|
2688  | 
proof (cases "S = {}")
 | 
|
2689  | 
case True  | 
|
2690  | 
then show ?thesis  | 
|
2691  | 
by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI)  | 
|
2692  | 
next  | 
|
2693  | 
case False  | 
|
2694  | 
have "inj g"  | 
|
2695  | 
by (metis UNIV_I hom homeomorphism_apply2 injI)  | 
|
2696  | 
then have "open (g ` UNIV)"  | 
|
2697  | 
by (blast intro: invariance_of_domain hom homeomorphism_cont2)  | 
|
2698  | 
then have "open S"  | 
|
2699  | 
using hom homeomorphism_image2 by blast  | 
|
2700  | 
moreover have "complete S"  | 
|
2701  | 
unfolding complete_def  | 
|
2702  | 
proof clarify  | 
|
2703  | 
fix \<sigma>  | 
|
2704  | 
assume \<sigma>: "\<forall>n. \<sigma> n \<in> S" and "Cauchy \<sigma>"  | 
|
2705  | 
have "Cauchy (f o \<sigma>)"  | 
|
2706  | 
using uniformly_continuous_imp_Cauchy_continuous \<open>Cauchy \<sigma>\<close> \<sigma> contf by blast  | 
|
2707  | 
then obtain l where "(f \<circ> \<sigma>) \<longlonglongrightarrow> l"  | 
|
2708  | 
by (auto simp: convergent_eq_Cauchy [symmetric])  | 
|
2709  | 
show "\<exists>l\<in>S. \<sigma> \<longlonglongrightarrow> l"  | 
|
2710  | 
proof  | 
|
2711  | 
show "g l \<in> S"  | 
|
2712  | 
using hom homeomorphism_image2 by blast  | 
|
2713  | 
have "(g \<circ> (f \<circ> \<sigma>)) \<longlonglongrightarrow> g l"  | 
|
2714  | 
by (meson UNIV_I \<open>(f \<circ> \<sigma>) \<longlonglongrightarrow> l\<close> continuous_on_sequentially hom homeomorphism_cont2)  | 
|
2715  | 
then show "\<sigma> \<longlonglongrightarrow> g l"  | 
|
2716  | 
proof -  | 
|
2717  | 
have "\<forall>n. \<sigma> n = (g \<circ> (f \<circ> \<sigma>)) n"  | 
|
2718  | 
by (metis (no_types) \<sigma> comp_eq_dest_lhs hom homeomorphism_apply1)  | 
|
2719  | 
then show ?thesis  | 
|
2720  | 
by (metis (no_types) LIMSEQ_iff \<open>(g \<circ> (f \<circ> \<sigma>)) \<longlonglongrightarrow> g l\<close>)  | 
|
2721  | 
qed  | 
|
2722  | 
qed  | 
|
2723  | 
qed  | 
|
2724  | 
then have "closed S"  | 
|
2725  | 
by (simp add: complete_eq_closed)  | 
|
2726  | 
ultimately show ?thesis  | 
|
2727  | 
using clopen [of S] False by simp  | 
|
2728  | 
qed  | 
|
2729  | 
||
| 64396 | 2730  | 
subsection\<open> Dimension-based conditions for various homeomorphisms.\<close>  | 
2731  | 
||
2732  | 
lemma homeomorphic_subspaces_eq:  | 
|
2733  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2734  | 
assumes "subspace S" "subspace T"  | 
|
2735  | 
shows "S homeomorphic T \<longleftrightarrow> dim S = dim T"  | 
|
2736  | 
proof  | 
|
2737  | 
assume "S homeomorphic T"  | 
|
2738  | 
then obtain f g where hom: "homeomorphism S T f g"  | 
|
2739  | 
using homeomorphic_def by blast  | 
|
2740  | 
show "dim S = dim T"  | 
|
2741  | 
proof (rule order_antisym)  | 
|
2742  | 
show "dim S \<le> dim T"  | 
|
2743  | 
by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le)  | 
|
2744  | 
show "dim T \<le> dim S"  | 
|
2745  | 
by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le)  | 
|
2746  | 
qed  | 
|
2747  | 
next  | 
|
2748  | 
assume "dim S = dim T"  | 
|
2749  | 
then show "S homeomorphic T"  | 
|
2750  | 
by (simp add: assms homeomorphic_subspaces)  | 
|
2751  | 
qed  | 
|
2752  | 
||
2753  | 
lemma homeomorphic_affine_sets_eq:  | 
|
2754  | 
fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"  | 
|
2755  | 
assumes "affine S" "affine T"  | 
|
2756  | 
shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"  | 
|
2757  | 
proof (cases "S = {} \<or> T = {}")
 | 
|
2758  | 
case True  | 
|
2759  | 
then show ?thesis  | 
|
2760  | 
using assms homeomorphic_affine_sets by force  | 
|
2761  | 
next  | 
|
2762  | 
case False  | 
|
2763  | 
then obtain a b where "a \<in> S" "b \<in> T"  | 
|
2764  | 
by blast  | 
|
| 67399 | 2765  | 
then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)"  | 
| 64396 | 2766  | 
using affine_diffs_subspace assms by blast+  | 
2767  | 
then show ?thesis  | 
|
2768  | 
by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets)  | 
|
2769  | 
qed  | 
|
2770  | 
||
2771  | 
lemma homeomorphic_hyperplanes_eq:  | 
|
2772  | 
fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space"  | 
|
2773  | 
assumes "a \<noteq> 0" "c \<noteq> 0"  | 
|
2774  | 
  shows "({x. a \<bullet> x = b} homeomorphic {x. c \<bullet> x = d} \<longleftrightarrow> DIM('a) = DIM('b))"
 | 
|
2775  | 
apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms)  | 
|
2776  | 
by (metis DIM_positive Suc_pred)  | 
|
2777  | 
||
2778  | 
lemma homeomorphic_UNIV_UNIV:  | 
|
2779  | 
shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \<longleftrightarrow>  | 
|
2780  | 
    DIM('a::euclidean_space) = DIM('b::euclidean_space)"
 | 
|
2781  | 
by (simp add: homeomorphic_subspaces_eq)  | 
|
2782  | 
||
2783  | 
lemma simply_connected_sphere_gen:  | 
|
2784  | 
assumes "convex S" "bounded S" and 3: "3 \<le> aff_dim S"  | 
|
2785  | 
shows "simply_connected(rel_frontier S)"  | 
|
2786  | 
proof -  | 
|
2787  | 
have pa: "path_connected (rel_frontier S)"  | 
|
2788  | 
using assms by (simp add: path_connected_sphere_gen)  | 
|
2789  | 
show ?thesis  | 
|
2790  | 
proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa)  | 
|
2791  | 
fix f  | 
|
2792  | 
assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \<subseteq> rel_frontier S"  | 
|
2793  | 
have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)"  | 
|
2794  | 
by simp  | 
|
2795  | 
have "convex (cball (0::complex) 1)"  | 
|
2796  | 
by (rule convex_cball)  | 
|
2797  | 
then obtain c where "homotopic_with (\<lambda>z. True) (sphere (0::complex) 1) (rel_frontier S) f (\<lambda>x. c)"  | 
|
2798  | 
apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \<open>convex S\<close> \<open>bounded S\<close>, where f=f])  | 
|
2799  | 
using f 3  | 
|
2800  | 
apply (auto simp: aff_dim_cball)  | 
|
2801  | 
done  | 
|
2802  | 
then show "\<exists>a. homotopic_with (\<lambda>h. True) (sphere 0 1) (rel_frontier S) f (\<lambda>x. a)"  | 
|
2803  | 
by blast  | 
|
2804  | 
qed  | 
|
2805  | 
qed  | 
|
2806  | 
||
2807  | 
subsection\<open>more invariance of domain\<close>  | 
|
2808  | 
||
2809  | 
proposition invariance_of_domain_sphere_affine_set_gen:  | 
|
2810  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2811  | 
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"  | 
|
2812  | 
and U: "bounded U" "convex U"  | 
|
2813  | 
and "affine T" and affTU: "aff_dim T < aff_dim U"  | 
|
2814  | 
and ope: "openin (subtopology euclidean (rel_frontier U)) S"  | 
|
2815  | 
shows "openin (subtopology euclidean T) (f ` S)"  | 
|
2816  | 
proof (cases "rel_frontier U = {}")
 | 
|
2817  | 
case True  | 
|
2818  | 
then show ?thesis  | 
|
2819  | 
using ope openin_subset by force  | 
|
2820  | 
next  | 
|
2821  | 
case False  | 
|
2822  | 
obtain b c where b: "b \<in> rel_frontier U" and c: "c \<in> rel_frontier U" and "b \<noteq> c"  | 
|
2823  | 
using \<open>bounded U\<close> rel_frontier_not_sing [of U] subset_singletonD False by fastforce  | 
|
2824  | 
obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1"  | 
|
2825  | 
proof (rule choose_affine_subset [OF affine_UNIV])  | 
|
2826  | 
show "- 1 \<le> aff_dim U - 1"  | 
|
2827  | 
by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le)  | 
|
2828  | 
show "aff_dim U - 1 \<le> aff_dim (UNIV::'a set)"  | 
|
2829  | 
by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq)  | 
|
2830  | 
qed auto  | 
|
2831  | 
have SU: "S \<subseteq> rel_frontier U"  | 
|
2832  | 
using ope openin_imp_subset by auto  | 
|
2833  | 
  have homb: "rel_frontier U - {b} homeomorphic V"
 | 
|
2834  | 
   and homc: "rel_frontier U - {c} homeomorphic V"
 | 
|
2835  | 
using homeomorphic_punctured_sphere_affine_gen [of U _ V]  | 
|
2836  | 
by (simp_all add: \<open>affine V\<close> affV U b c)  | 
|
2837  | 
then obtain g h j k  | 
|
2838  | 
           where gh: "homeomorphism (rel_frontier U - {b}) V g h"
 | 
|
2839  | 
             and jk: "homeomorphism (rel_frontier U - {c}) V j k"
 | 
|
2840  | 
by (auto simp: homeomorphic_def)  | 
|
2841  | 
  with SU have hgsub: "(h ` g ` (S - {b})) \<subseteq> S" and kjsub: "(k ` j ` (S - {c})) \<subseteq> S"
 | 
|
2842  | 
by (simp_all add: homeomorphism_def subset_eq)  | 
|
2843  | 
have [simp]: "aff_dim T \<le> aff_dim V"  | 
|
2844  | 
by (simp add: affTU affV)  | 
|
2845  | 
  have "openin (subtopology euclidean T) ((f \<circ> h) ` g ` (S - {b}))"
 | 
|
2846  | 
proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])  | 
|
2847  | 
    show "openin (subtopology euclidean V) (g ` (S - {b}))"
 | 
|
2848  | 
apply (rule homeomorphism_imp_open_map [OF gh])  | 
|
2849  | 
by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl)  | 
|
2850  | 
    show "continuous_on (g ` (S - {b})) (f \<circ> h)"
 | 
|
2851  | 
apply (rule continuous_on_compose)  | 
|
2852  | 
apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset)  | 
|
2853  | 
using contf continuous_on_subset hgsub by blast  | 
|
2854  | 
    show "inj_on (f \<circ> h) (g ` (S - {b}))"
 | 
|
2855  | 
using kjsub  | 
|
2856  | 
apply (clarsimp simp add: inj_on_def)  | 
|
2857  | 
by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD)  | 
|
2858  | 
    show "(f \<circ> h) ` g ` (S - {b}) \<subseteq> T"
 | 
|
2859  | 
by (metis fim image_comp image_mono hgsub subset_trans)  | 
|
2860  | 
qed (auto simp: assms)  | 
|
2861  | 
moreover  | 
|
2862  | 
  have "openin (subtopology euclidean T) ((f \<circ> k) ` j ` (S - {c}))"
 | 
|
2863  | 
proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])  | 
|
2864  | 
    show "openin (subtopology euclidean V) (j ` (S - {c}))"
 | 
|
2865  | 
apply (rule homeomorphism_imp_open_map [OF jk])  | 
|
2866  | 
by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl)  | 
|
2867  | 
    show "continuous_on (j ` (S - {c})) (f \<circ> k)"
 | 
|
2868  | 
apply (rule continuous_on_compose)  | 
|
2869  | 
apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset)  | 
|
2870  | 
using contf continuous_on_subset kjsub by blast  | 
|
2871  | 
    show "inj_on (f \<circ> k) (j ` (S - {c}))"
 | 
|
2872  | 
using kjsub  | 
|
2873  | 
apply (clarsimp simp add: inj_on_def)  | 
|
2874  | 
by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD)  | 
|
2875  | 
    show "(f \<circ> k) ` j ` (S - {c}) \<subseteq> T"
 | 
|
2876  | 
by (metis fim image_comp image_mono kjsub subset_trans)  | 
|
2877  | 
qed (auto simp: assms)  | 
|
2878  | 
  ultimately have "openin (subtopology euclidean T) ((f \<circ> h) ` g ` (S - {b}) \<union> ((f \<circ> k) ` j ` (S - {c})))"
 | 
|
2879  | 
by (rule openin_Un)  | 
|
2880  | 
  moreover have "(f \<circ> h) ` g ` (S - {b}) = f ` (S - {b})"
 | 
|
2881  | 
proof -  | 
|
2882  | 
    have "h ` g ` (S - {b}) = (S - {b})"
 | 
|
2883  | 
proof  | 
|
2884  | 
      show "h ` g ` (S - {b}) \<subseteq> S - {b}"
 | 
|
2885  | 
using homeomorphism_apply1 [OF gh] SU  | 
|
2886  | 
by (fastforce simp add: image_iff image_subset_iff)  | 
|
2887  | 
      show "S - {b} \<subseteq> h ` g ` (S - {b})"
 | 
|
2888  | 
apply clarify  | 
|
2889  | 
by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def)  | 
|
2890  | 
qed  | 
|
2891  | 
then show ?thesis  | 
|
2892  | 
by (metis image_comp)  | 
|
2893  | 
qed  | 
|
2894  | 
  moreover have "(f \<circ> k) ` j ` (S - {c}) = f ` (S - {c})"
 | 
|
2895  | 
proof -  | 
|
2896  | 
    have "k ` j ` (S - {c}) = (S - {c})"
 | 
|
2897  | 
proof  | 
|
2898  | 
      show "k ` j ` (S - {c}) \<subseteq> S - {c}"
 | 
|
2899  | 
using homeomorphism_apply1 [OF jk] SU  | 
|
2900  | 
by (fastforce simp add: image_iff image_subset_iff)  | 
|
2901  | 
      show "S - {c} \<subseteq> k ` j ` (S - {c})"
 | 
|
2902  | 
apply clarify  | 
|
2903  | 
by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def)  | 
|
2904  | 
qed  | 
|
2905  | 
then show ?thesis  | 
|
2906  | 
by (metis image_comp)  | 
|
2907  | 
qed  | 
|
2908  | 
  moreover have "f ` (S - {b}) \<union> f ` (S - {c}) = f ` (S)"
 | 
|
2909  | 
using \<open>b \<noteq> c\<close> by blast  | 
|
2910  | 
ultimately show ?thesis  | 
|
2911  | 
by simp  | 
|
2912  | 
qed  | 
|
2913  | 
||
2914  | 
||
2915  | 
lemma invariance_of_domain_sphere_affine_set:  | 
|
2916  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2917  | 
assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"  | 
|
2918  | 
      and "r \<noteq> 0" "affine T" and affTU: "aff_dim T < DIM('a)"
 | 
|
2919  | 
and ope: "openin (subtopology euclidean (sphere a r)) S"  | 
|
2920  | 
shows "openin (subtopology euclidean T) (f ` S)"  | 
|
2921  | 
proof (cases "sphere a r = {}")
 | 
|
2922  | 
case True  | 
|
2923  | 
then show ?thesis  | 
|
2924  | 
using ope openin_subset by force  | 
|
2925  | 
next  | 
|
2926  | 
case False  | 
|
2927  | 
show ?thesis  | 
|
2928  | 
proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \<open>affine T\<close>])  | 
|
2929  | 
show "aff_dim T < aff_dim (cball a r)"  | 
|
2930  | 
by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty)  | 
|
2931  | 
show "openin (subtopology euclidean (rel_frontier (cball a r))) S"  | 
|
2932  | 
by (simp add: \<open>r \<noteq> 0\<close> ope)  | 
|
2933  | 
qed  | 
|
2934  | 
qed  | 
|
2935  | 
||
2936  | 
lemma no_embedding_sphere_lowdim:  | 
|
2937  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
2938  | 
assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0"  | 
|
2939  | 
   shows "DIM('a) \<le> DIM('b)"
 | 
|
2940  | 
proof -  | 
|
2941  | 
  have "False" if "DIM('a) > DIM('b)"
 | 
|
2942  | 
proof -  | 
|
2943  | 
have "compact (f ` sphere a r)"  | 
|
2944  | 
using compact_continuous_image  | 
|
2945  | 
by (simp add: compact_continuous_image contf)  | 
|
2946  | 
then have "\<not> open (f ` sphere a r)"  | 
|
2947  | 
using compact_open  | 
|
2948  | 
by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty)  | 
|
2949  | 
then show False  | 
|
2950  | 
using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \<open>r > 0\<close>  | 
|
2951  | 
by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that)  | 
|
2952  | 
qed  | 
|
2953  | 
then show ?thesis  | 
|
2954  | 
using not_less by blast  | 
|
2955  | 
qed  | 
|
2956  | 
||
2957  | 
lemma simply_connected_sphere:  | 
|
2958  | 
fixes a :: "'a::euclidean_space"  | 
|
2959  | 
  assumes "3 \<le> DIM('a)"
 | 
|
2960  | 
shows "simply_connected(sphere a r)"  | 
|
2961  | 
proof (cases rule: linorder_cases [of r 0])  | 
|
2962  | 
case less  | 
|
2963  | 
then show ?thesis by simp  | 
|
2964  | 
next  | 
|
2965  | 
case equal  | 
|
2966  | 
then show ?thesis by (auto simp: convex_imp_simply_connected)  | 
|
2967  | 
next  | 
|
2968  | 
case greater  | 
|
2969  | 
then show ?thesis  | 
|
2970  | 
using simply_connected_sphere_gen [of "cball a r"] assms  | 
|
2971  | 
by (simp add: aff_dim_cball)  | 
|
2972  | 
qed  | 
|
2973  | 
||
| 
64789
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2974  | 
lemma simply_connected_sphere_eq:  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2975  | 
fixes a :: "'a::euclidean_space"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2976  | 
  shows "simply_connected(sphere a r) \<longleftrightarrow> 3 \<le> DIM('a) \<or> r \<le> 0"  (is "?lhs = ?rhs")
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2977  | 
proof (cases "r \<le> 0")  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2978  | 
case True  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2979  | 
have "simply_connected (sphere a r)"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2980  | 
apply (rule convex_imp_simply_connected)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2981  | 
using True less_eq_real_def by auto  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2982  | 
with True show ?thesis by auto  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2983  | 
next  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2984  | 
case False  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2985  | 
show ?thesis  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2986  | 
proof  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2987  | 
assume L: ?lhs  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2988  | 
    have "False" if "DIM('a) = 1 \<or> DIM('a) = 2"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2989  | 
using that  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2990  | 
proof  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2991  | 
      assume "DIM('a) = 1"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2992  | 
with L show False  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2993  | 
using connected_sphere_eq simply_connected_imp_connected  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2994  | 
by (metis False Suc_1 not_less_eq_eq order_refl)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2995  | 
next  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2996  | 
      assume "DIM('a) = 2"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2997  | 
then have "sphere a r homeomorphic sphere (0::complex) 1"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2998  | 
by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
2999  | 
then have "simply_connected(sphere (0::complex) 1)"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3000  | 
using L homeomorphic_simply_connected_eq by blast  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3001  | 
then obtain a::complex where "homotopic_with (\<lambda>h. True) (sphere 0 1) (sphere 0 1) id (\<lambda>x. a)"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3002  | 
apply (simp add: simply_connected_eq_contractible_circlemap)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3003  | 
by (metis continuous_on_id' id_apply image_id subset_refl)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3004  | 
then show False  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3005  | 
using contractible_sphere contractible_def not_one_le_zero by blast  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3006  | 
qed  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3007  | 
with False show ?rhs  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3008  | 
apply simp  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3009  | 
by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3010  | 
next  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3011  | 
assume ?rhs  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3012  | 
with False show ?lhs by (simp add: simply_connected_sphere)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3013  | 
qed  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3014  | 
qed  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3015  | 
|
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3016  | 
|
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3017  | 
lemma simply_connected_punctured_universe_eq:  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3018  | 
fixes a :: "'a::euclidean_space"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3019  | 
  shows "simply_connected(- {a}) \<longleftrightarrow> 3 \<le> DIM('a)"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3020  | 
proof -  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3021  | 
have [simp]: "a \<in> rel_interior (cball a 1)"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3022  | 
by (simp add: rel_interior_nonempty_interior)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3023  | 
  have [simp]: "affine hull cball a 1 - {a} = -{a}"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3024  | 
by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3025  | 
  have "simply_connected(- {a}) \<longleftrightarrow> simply_connected(sphere a 1)"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3026  | 
apply (rule sym)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3027  | 
apply (rule homotopy_eqv_simple_connectedness)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3028  | 
using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3029  | 
done  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3030  | 
  also have "...  \<longleftrightarrow> 3 \<le> DIM('a)"
 | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3031  | 
by (simp add: simply_connected_sphere_eq)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3032  | 
finally show ?thesis .  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3033  | 
qed  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3034  | 
|
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3035  | 
lemma not_simply_connected_circle:  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3036  | 
fixes a :: complex  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3037  | 
shows "0 < r \<Longrightarrow> \<not> simply_connected(sphere a r)"  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3038  | 
by (simp add: simply_connected_sphere_eq)  | 
| 
 
6440577e34ee
connectedness, circles not simply connected , punctured universe
 
paulson <lp15@cam.ac.uk> 
parents: 
64508 
diff
changeset
 | 
3039  | 
|
| 64847 | 3040  | 
|
| 
64790
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3041  | 
proposition simply_connected_punctured_convex:  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3042  | 
fixes a :: "'a::euclidean_space"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3043  | 
assumes "convex S" and 3: "3 \<le> aff_dim S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3044  | 
    shows "simply_connected(S - {a})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3045  | 
proof (cases "a \<in> rel_interior S")  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3046  | 
case True  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3047  | 
then obtain e where "a \<in> S" "0 < e" and e: "cball a e \<inter> affine hull S \<subseteq> S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3048  | 
by (auto simp: rel_interior_cball)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3049  | 
have con: "convex (cball a e \<inter> affine hull S)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3050  | 
by (simp add: convex_Int)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3051  | 
have bo: "bounded (cball a e \<inter> affine hull S)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3052  | 
by (simp add: bounded_Int)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3053  | 
  have "affine hull S \<inter> interior (cball a e) \<noteq> {}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3054  | 
using \<open>0 < e\<close> \<open>a \<in> S\<close> hull_subset by fastforce  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3055  | 
then have "3 \<le> aff_dim (affine hull S \<inter> cball a e)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3056  | 
by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3057  | 
also have "... = aff_dim (cball a e \<inter> affine hull S)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3058  | 
by (simp add: Int_commute)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3059  | 
finally have "3 \<le> aff_dim (cball a e \<inter> affine hull S)" .  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3060  | 
  moreover have "rel_frontier (cball a e \<inter> affine hull S) homotopy_eqv S - {a}"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3061  | 
proof (rule homotopy_eqv_rel_frontier_punctured_convex)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3062  | 
show "a \<in> rel_interior (cball a e \<inter> affine hull S)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3063  | 
by (meson IntI Int_mono \<open>a \<in> S\<close> \<open>0 < e\<close> e \<open>cball a e \<inter> affine hull S \<subseteq> S\<close> ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3064  | 
have "closed (cball a e \<inter> affine hull S)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3065  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3066  | 
then show "rel_frontier (cball a e \<inter> affine hull S) \<subseteq> S"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3067  | 
apply (simp add: rel_frontier_def)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3068  | 
using e by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3069  | 
show "S \<subseteq> affine hull (cball a e \<inter> affine hull S)"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3070  | 
by (metis (no_types, lifting) IntI \<open>a \<in> S\<close> \<open>0 < e\<close> affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3071  | 
qed (auto simp: assms con bo)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3072  | 
ultimately show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3073  | 
using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo]  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3074  | 
by blast  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3075  | 
next  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3076  | 
case False  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3077  | 
show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3078  | 
apply (rule contractible_imp_simply_connected)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3079  | 
apply (rule contractible_convex_tweak_boundary_points [OF \<open>convex S\<close>])  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3080  | 
apply (simp add: False rel_interior_subset subset_Diff_insert)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3081  | 
by (meson Diff_subset closure_subset subset_trans)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3082  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3083  | 
|
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3084  | 
corollary simply_connected_punctured_universe:  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3085  | 
fixes a :: "'a::euclidean_space"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3086  | 
  assumes "3 \<le> DIM('a)"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3087  | 
  shows "simply_connected(- {a})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3088  | 
proof -  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3089  | 
have [simp]: "affine hull cball a 1 = UNIV"  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3090  | 
apply auto  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3091  | 
by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3092  | 
  have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})"
 | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3093  | 
apply (rule homotopy_eqv_simple_connectedness)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3094  | 
apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3095  | 
apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3096  | 
done  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3097  | 
then show ?thesis  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3098  | 
using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV)  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3099  | 
qed  | 
| 
 
ed38f9a834d8
New theory of arcwise connected sets and other new material
 
paulson <lp15@cam.ac.uk> 
parents: 
64789 
diff
changeset
 | 
3100  | 
|
| 64396 | 3101  | 
|
| 64287 | 3102  | 
subsection\<open>The power, squaring and exponential functions as covering maps\<close>  | 
3103  | 
||
3104  | 
proposition covering_space_power_punctured_plane:  | 
|
3105  | 
assumes "0 < n"  | 
|
3106  | 
    shows "covering_space (- {0}) (\<lambda>z::complex. z^n) (- {0})"
 | 
|
3107  | 
proof -  | 
|
3108  | 
consider "n = 1" | "2 \<le> n" using assms by linarith  | 
|
3109  | 
then obtain e where "0 < e"  | 
|
3110  | 
and e: "\<And>w z. cmod(w - z) < e * cmod z \<Longrightarrow> (w^n = z^n \<longleftrightarrow> w = z)"  | 
|
3111  | 
proof cases  | 
|
3112  | 
assume "n = 1" then show ?thesis  | 
|
3113  | 
by (rule_tac e=1 in that) auto  | 
|
3114  | 
next  | 
|
3115  | 
assume "2 \<le> n"  | 
|
3116  | 
have eq_if_pow_eq:  | 
|
3117  | 
"w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z"  | 
|
3118  | 
and eq: "w^n = z^n" for w z  | 
|
3119  | 
proof (cases "z = 0")  | 
|
3120  | 
case True with eq assms show ?thesis by (auto simp: power_0_left)  | 
|
3121  | 
next  | 
|
3122  | 
case False  | 
|
3123  | 
then have "z \<noteq> 0" by auto  | 
|
3124  | 
have "(w/z)^n = 1"  | 
|
3125  | 
by (metis False divide_self_if eq power_divide power_one)  | 
|
3126  | 
then obtain j where j: "w / z = exp (2 * of_real pi * \<i> * j / n)" and "j < n"  | 
|
3127  | 
using Suc_leI assms \<open>2 \<le> n\<close> complex_roots_unity [THEN eqset_imp_iff, of n "w/z"]  | 
|
3128  | 
by force  | 
|
3129  | 
have "cmod (w/z - 1) < 2 * sin (pi / real n)"  | 
|
3130  | 
using lt assms \<open>z \<noteq> 0\<close> by (simp add: divide_simps norm_divide)  | 
|
3131  | 
then have "cmod (exp (\<i> * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)"  | 
|
3132  | 
by (simp add: j field_simps)  | 
|
3133  | 
then have "2 * \<bar>sin((2 * pi * j / n) / 2)\<bar> < 2 * sin (pi / real n)"  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
3134  | 
by (simp only: dist_exp_i_1)  | 
| 64287 | 3135  | 
then have sin_less: "sin((pi * j / n)) < sin (pi / real n)"  | 
3136  | 
by (simp add: field_simps)  | 
|
3137  | 
then have "w / z = 1"  | 
|
3138  | 
proof (cases "j = 0")  | 
|
3139  | 
case True then show ?thesis by (auto simp: j)  | 
|
3140  | 
next  | 
|
3141  | 
case False  | 
|
3142  | 
then have "sin (pi / real n) \<le> sin((pi * j / n))"  | 
|
3143  | 
proof (cases "j / n \<le> 1/2")  | 
|
3144  | 
case True  | 
|
3145  | 
show ?thesis  | 
|
3146  | 
apply (rule sin_monotone_2pi_le)  | 
|
3147  | 
using \<open>j \<noteq> 0 \<close> \<open>j < n\<close> True  | 
|
3148  | 
apply (auto simp: field_simps intro: order_trans [of _ 0])  | 
|
3149  | 
done  | 
|
3150  | 
next  | 
|
3151  | 
case False  | 
|
3152  | 
then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)"  | 
|
3153  | 
using \<open>j < n\<close> by (simp add: algebra_simps diff_divide_distrib of_nat_diff)  | 
|
3154  | 
show ?thesis  | 
|
3155  | 
apply (simp only: seq)  | 
|
3156  | 
apply (rule sin_monotone_2pi_le)  | 
|
3157  | 
using \<open>j < n\<close> False  | 
|
3158  | 
apply (auto simp: field_simps intro: order_trans [of _ 0])  | 
|
3159  | 
done  | 
|
3160  | 
qed  | 
|
3161  | 
with sin_less show ?thesis by force  | 
|
3162  | 
qed  | 
|
3163  | 
then show ?thesis by simp  | 
|
3164  | 
qed  | 
|
3165  | 
show ?thesis  | 
|
3166  | 
apply (rule_tac e = "2 * sin(pi / n)" in that)  | 
|
3167  | 
apply (force simp: \<open>2 \<le> n\<close> sin_pi_divide_n_gt_0)  | 
|
3168  | 
apply (meson eq_if_pow_eq)  | 
|
3169  | 
done  | 
|
3170  | 
qed  | 
|
3171  | 
  have zn1: "continuous_on (- {0}) (\<lambda>z::complex. z^n)"
 | 
|
3172  | 
by (rule continuous_intros)+  | 
|
3173  | 
  have zn2: "(\<lambda>z::complex. z^n) ` (- {0}) = - {0}"
 | 
|
3174  | 
using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n])  | 
|
3175  | 
have zn3: "\<exists>T. z^n \<in> T \<and> open T \<and> 0 \<notin> T \<and>  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3176  | 
               (\<exists>v. \<Union>v = -{0} \<inter> (\<lambda>z. z ^ n) -` T \<and>
 | 
| 64287 | 3177  | 
(\<forall>u\<in>v. open u \<and> 0 \<notin> u) \<and>  | 
3178  | 
pairwise disjnt v \<and>  | 
|
3179  | 
(\<forall>u\<in>v. Ex (homeomorphism u T (\<lambda>z. z^n))))"  | 
|
3180  | 
if "z \<noteq> 0" for z::complex  | 
|
3181  | 
proof -  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
3182  | 
define d where "d \<equiv> min (1/2) (e/4) * norm z"  | 
| 64287 | 3183  | 
have "0 < d"  | 
3184  | 
by (simp add: d_def \<open>0 < e\<close> \<open>z \<noteq> 0\<close>)  | 
|
3185  | 
have iff_x_eq_y: "x^n = y^n \<longleftrightarrow> x = y"  | 
|
3186  | 
if eq: "w^n = z^n" and x: "x \<in> ball w d" and y: "y \<in> ball w d" for w x y  | 
|
3187  | 
proof -  | 
|
3188  | 
have [simp]: "norm z = norm w" using that  | 
|
3189  | 
by (simp add: assms power_eq_imp_eq_norm)  | 
|
3190  | 
show ?thesis  | 
|
3191  | 
proof (cases "w = 0")  | 
|
3192  | 
case True with \<open>z \<noteq> 0\<close> assms eq  | 
|
3193  | 
show ?thesis by (auto simp: power_0_left)  | 
|
3194  | 
next  | 
|
3195  | 
case False  | 
|
3196  | 
have "cmod (x - y) < 2*d"  | 
|
3197  | 
using x y  | 
|
3198  | 
by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add)  | 
|
3199  | 
also have "... \<le> 2 * e / 4 * norm w"  | 
|
3200  | 
using \<open>e > 0\<close> by (simp add: d_def min_mult_distrib_right)  | 
|
3201  | 
also have "... = e * (cmod w / 2)"  | 
|
3202  | 
by simp  | 
|
3203  | 
also have "... \<le> e * cmod y"  | 
|
3204  | 
apply (rule mult_left_mono)  | 
|
3205  | 
using \<open>e > 0\<close> y  | 
|
3206  | 
apply (simp_all add: dist_norm d_def min_mult_distrib_right del: divide_const_simps)  | 
|
3207  | 
apply (metis dist_0_norm dist_complex_def dist_triangle_half_l linorder_not_less order_less_irrefl)  | 
|
3208  | 
done  | 
|
3209  | 
finally have "cmod (x - y) < e * cmod y" .  | 
|
3210  | 
then show ?thesis by (rule e)  | 
|
3211  | 
qed  | 
|
3212  | 
qed  | 
|
3213  | 
then have inj: "inj_on (\<lambda>w. w^n) (ball z d)"  | 
|
3214  | 
by (simp add: inj_on_def)  | 
|
3215  | 
have cont: "continuous_on (ball z d) (\<lambda>w. w ^ n)"  | 
|
3216  | 
by (intro continuous_intros)  | 
|
3217  | 
have noncon: "\<not> (\<lambda>w::complex. w^n) constant_on UNIV"  | 
|
3218  | 
by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power)  | 
|
3219  | 
have im_eq: "(\<lambda>w. w^n) ` ball z' d = (\<lambda>w. w^n) ` ball z d"  | 
|
3220  | 
if z': "z'^n = z^n" for z'  | 
|
3221  | 
proof -  | 
|
3222  | 
have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast  | 
|
3223  | 
have "(w \<in> (\<lambda>w. w^n) ` ball z' d) = (w \<in> (\<lambda>w. w^n) ` ball z d)" for w  | 
|
3224  | 
proof (cases "w=0")  | 
|
3225  | 
case True with assms show ?thesis  | 
|
3226  | 
by (simp add: image_def ball_def nz')  | 
|
3227  | 
next  | 
|
3228  | 
case False  | 
|
3229  | 
have "z' \<noteq> 0" using \<open>z \<noteq> 0\<close> nz' by force  | 
|
3230  | 
have [simp]: "(z*x / z')^n = x^n" if "x \<noteq> 0" for x  | 
|
3231  | 
using z' that by (simp add: field_simps \<open>z \<noteq> 0\<close>)  | 
|
3232  | 
have [simp]: "cmod (z - z * x / z') = cmod (z' - x)" if "x \<noteq> 0" for x  | 
|
3233  | 
proof -  | 
|
3234  | 
have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')"  | 
|
3235  | 
by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) complex_divide_def mult.right_neutral norm_mult right_diff_distrib')  | 
|
3236  | 
also have "... = cmod z' * cmod (1 - x / z')"  | 
|
3237  | 
by (simp add: nz')  | 
|
3238  | 
also have "... = cmod (z' - x)"  | 
|
3239  | 
by (simp add: \<open>z' \<noteq> 0\<close> diff_divide_eq_iff norm_divide)  | 
|
3240  | 
finally show ?thesis .  | 
|
3241  | 
qed  | 
|
3242  | 
have [simp]: "(z'*x / z)^n = x^n" if "x \<noteq> 0" for x  | 
|
3243  | 
using z' that by (simp add: field_simps \<open>z \<noteq> 0\<close>)  | 
|
3244  | 
have [simp]: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \<noteq> 0" for x  | 
|
3245  | 
proof -  | 
|
3246  | 
have "cmod (z * (1 - x * inverse z)) = cmod (z - x)"  | 
|
3247  | 
by (metis \<open>z \<noteq> 0\<close> diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7))  | 
|
3248  | 
then show ?thesis  | 
|
3249  | 
by (metis (no_types) mult.assoc complex_divide_def mult.right_neutral norm_mult nz' right_diff_distrib')  | 
|
3250  | 
qed  | 
|
3251  | 
show ?thesis  | 
|
3252  | 
unfolding image_def ball_def  | 
|
3253  | 
apply safe  | 
|
3254  | 
apply simp_all  | 
|
3255  | 
apply (rule_tac x="z/z' * x" in exI)  | 
|
3256  | 
using assms False apply (simp add: dist_norm)  | 
|
3257  | 
apply (rule_tac x="z'/z * x" in exI)  | 
|
3258  | 
using assms False apply (simp add: dist_norm)  | 
|
3259  | 
done  | 
|
3260  | 
qed  | 
|
3261  | 
then show ?thesis by blast  | 
|
3262  | 
qed  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3263  | 
|
| 64287 | 3264  | 
have ex_ball: "\<exists>B. (\<exists>z'. B = ball z' d \<and> z'^n = z^n) \<and> x \<in> B"  | 
3265  | 
if "x \<noteq> 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w  | 
|
3266  | 
proof -  | 
|
3267  | 
have "w \<noteq> 0" by (metis assms power_eq_0_iff that(1) that(2))  | 
|
3268  | 
have [simp]: "cmod x = cmod w"  | 
|
3269  | 
using assms power_eq_imp_eq_norm eq by blast  | 
|
3270  | 
have [simp]: "cmod (x * z / w - x) = cmod (z - w)"  | 
|
3271  | 
proof -  | 
|
3272  | 
have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)"  | 
|
3273  | 
by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right)  | 
|
3274  | 
also have "... = cmod w * cmod (z / w - 1)"  | 
|
3275  | 
by simp  | 
|
3276  | 
also have "... = cmod (z - w)"  | 
|
3277  | 
by (simp add: \<open>w \<noteq> 0\<close> divide_diff_eq_iff nonzero_norm_divide)  | 
|
3278  | 
finally show ?thesis .  | 
|
3279  | 
qed  | 
|
3280  | 
show ?thesis  | 
|
3281  | 
apply (rule_tac x="ball (z / w * x) d" in exI)  | 
|
3282  | 
using \<open>d > 0\<close> that  | 
|
3283  | 
apply (simp add: ball_eq_ball_iff)  | 
|
3284  | 
apply (simp add: \<open>z \<noteq> 0\<close> \<open>w \<noteq> 0\<close> field_simps)  | 
|
3285  | 
apply (simp add: dist_norm)  | 
|
3286  | 
done  | 
|
3287  | 
qed  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3288  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3289  | 
show ?thesis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3290  | 
proof (rule exI, intro conjI)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3291  | 
show "z ^ n \<in> (\<lambda>w. w ^ n) ` ball z d"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3292  | 
using \<open>d > 0\<close> by simp  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3293  | 
show "open ((\<lambda>w. w ^ n) ` ball z d)"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3294  | 
by (rule invariance_of_domain [OF cont open_ball inj])  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3295  | 
show "0 \<notin> (\<lambda>w. w ^ n) ` ball z d"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3296  | 
using \<open>z \<noteq> 0\<close> assms by (force simp: d_def)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3297  | 
      show "\<exists>v. \<Union>v = - {0} \<inter> (\<lambda>z. z ^ n) -` (\<lambda>w. w ^ n) ` ball z d \<and>
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3298  | 
(\<forall>u\<in>v. open u \<and> 0 \<notin> u) \<and>  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3299  | 
disjoint v \<and>  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3300  | 
(\<forall>u\<in>v. Ex (homeomorphism u ((\<lambda>w. w ^ n) ` ball z d) (\<lambda>z. z ^ n)))"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3301  | 
proof (rule exI, intro ballI conjI)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3302  | 
        show "\<Union>{ball z' d |z'. z'^n = z^n} = - {0} \<inter> (\<lambda>z. z ^ n) -` (\<lambda>w. w ^ n) ` ball z d" (is "?l = ?r")
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3303  | 
proof  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3304  | 
show "?l \<subseteq> ?r"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3305  | 
apply auto  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3306  | 
apply (simp add: assms d_def power_eq_imp_eq_norm that)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3307  | 
by (metis im_eq image_eqI mem_ball)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3308  | 
show "?r \<subseteq> ?l"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3309  | 
by auto (meson ex_ball)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3310  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3311  | 
        show "\<And>u. u \<in> {ball z' d |z'. z' ^ n = z ^ n} \<Longrightarrow> 0 \<notin> u"
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3312  | 
by (force simp add: assms d_def power_eq_imp_eq_norm that)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3313  | 
|
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3314  | 
        show "disjoint {ball z' d |z'. z' ^ n = z ^ n}"
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3315  | 
proof (clarsimp simp add: pairwise_def disjnt_iff)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3316  | 
fix \<xi> \<zeta> x  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3317  | 
assume "\<xi>^n = z^n" "\<zeta>^n = z^n" "ball \<xi> d \<noteq> ball \<zeta> d"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3318  | 
and "dist \<xi> x < d" "dist \<zeta> x < d"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3319  | 
then have "dist \<xi> \<zeta> < d+d"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3320  | 
using dist_triangle_less_add by blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3321  | 
then have "cmod (\<xi> - \<zeta>) < 2*d"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3322  | 
by (simp add: dist_norm)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3323  | 
also have "... \<le> e * cmod z"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3324  | 
using mult_right_mono \<open>0 < e\<close> that by (auto simp: d_def)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3325  | 
finally have "cmod (\<xi> - \<zeta>) < e * cmod z" .  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3326  | 
with e have "\<xi> = \<zeta>"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3327  | 
by (metis \<open>\<xi>^n = z^n\<close> \<open>\<zeta>^n = z^n\<close> assms power_eq_imp_eq_norm)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3328  | 
then show "False"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3329  | 
using \<open>ball \<xi> d \<noteq> ball \<zeta> d\<close> by blast  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3330  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3331  | 
show "Ex (homeomorphism u ((\<lambda>w. w ^ n) ` ball z d) (\<lambda>z. z ^ n))"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3332  | 
          if "u \<in> {ball z' d |z'. z' ^ n = z ^ n}" for u
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3333  | 
proof (rule invariance_of_domain_homeomorphism [of "u" "\<lambda>z. z^n"])  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3334  | 
show "open u"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3335  | 
using that by auto  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3336  | 
show "continuous_on u (\<lambda>z. z ^ n)"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3337  | 
by (intro continuous_intros)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3338  | 
show "inj_on (\<lambda>z. z ^ n) u"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3339  | 
using that by (auto simp: iff_x_eq_y inj_on_def)  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3340  | 
show "\<And>g. homeomorphism u ((\<lambda>z. z ^ n) ` u) (\<lambda>z. z ^ n) g \<Longrightarrow> Ex (homeomorphism u ((\<lambda>w. w ^ n) ` ball z d) (\<lambda>z. z ^ n))"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3341  | 
using im_eq that by clarify metis  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3342  | 
qed auto  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3343  | 
qed auto  | 
| 64287 | 3344  | 
qed  | 
3345  | 
qed  | 
|
3346  | 
show ?thesis  | 
|
3347  | 
using assms  | 
|
3348  | 
apply (simp add: covering_space_def zn1 zn2)  | 
|
3349  | 
apply (subst zn2 [symmetric])  | 
|
3350  | 
apply (simp add: openin_open_eq open_Compl)  | 
|
3351  | 
apply (blast intro: zn3)  | 
|
3352  | 
done  | 
|
3353  | 
qed  | 
|
3354  | 
||
3355  | 
corollary covering_space_square_punctured_plane:  | 
|
3356  | 
  "covering_space (- {0}) (\<lambda>z::complex. z^2) (- {0})"
 | 
|
3357  | 
by (simp add: covering_space_power_punctured_plane)  | 
|
3358  | 
||
3359  | 
||
3360  | 
proposition covering_space_exp_punctured_plane:  | 
|
3361  | 
  "covering_space UNIV (\<lambda>z::complex. exp z) (- {0})"
 | 
|
3362  | 
proof (simp add: covering_space_def, intro conjI ballI)  | 
|
3363  | 
show "continuous_on UNIV (\<lambda>z::complex. exp z)"  | 
|
3364  | 
by (rule continuous_on_exp [OF continuous_on_id])  | 
|
3365  | 
  show "range exp = - {0::complex}"
 | 
|
3366  | 
by auto (metis exp_Ln range_eqI)  | 
|
3367  | 
  show "\<exists>T. z \<in> T \<and> openin (subtopology euclidean (- {0})) T \<and>
 | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3368  | 
(\<exists>v. \<Union>v = exp -` T \<and> (\<forall>u\<in>v. open u) \<and> disjoint v \<and>  | 
| 64287 | 3369  | 
(\<forall>u\<in>v. \<exists>q. homeomorphism u T exp q))"  | 
3370  | 
        if "z \<in> - {0::complex}" for z
 | 
|
3371  | 
proof -  | 
|
3372  | 
have "z \<noteq> 0"  | 
|
3373  | 
using that by auto  | 
|
3374  | 
have inj_exp: "inj_on exp (ball (Ln z) 1)"  | 
|
3375  | 
apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]])  | 
|
3376  | 
using pi_ge_two by (simp add: ball_subset_ball_iff)  | 
|
| 64508 | 3377  | 
define \<V> where "\<V> \<equiv> range (\<lambda>n. (\<lambda>x. x + of_real (2 * of_int n * pi) * \<i>) ` (ball(Ln z) 1))"  | 
| 64287 | 3378  | 
show ?thesis  | 
3379  | 
proof (intro exI conjI)  | 
|
3380  | 
show "z \<in> exp ` (ball(Ln z) 1)"  | 
|
3381  | 
by (metis \<open>z \<noteq> 0\<close> centre_in_ball exp_Ln rev_image_eqI zero_less_one)  | 
|
3382  | 
      have "open (- {0::complex})"
 | 
|
3383  | 
by blast  | 
|
3384  | 
moreover have "inj_on exp (ball (Ln z) 1)"  | 
|
3385  | 
apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]])  | 
|
3386  | 
using pi_ge_two by (simp add: ball_subset_ball_iff)  | 
|
3387  | 
      ultimately show "openin (subtopology euclidean (- {0})) (exp ` ball (Ln z) 1)"
 | 
|
3388  | 
by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id])  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3389  | 
show "\<Union>\<V> = exp -` exp ` ball (Ln z) 1"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
3390  | 
by (force simp: \<V>_def Complex_Transcendental.exp_eq image_iff)  | 
| 64287 | 3391  | 
show "\<forall>V\<in>\<V>. open V"  | 
3392  | 
by (auto simp: \<V>_def inj_on_def continuous_intros invariance_of_domain)  | 
|
3393  | 
have xy: "2 \<le> cmod (2 * of_int x * of_real pi * \<i> - 2 * of_int y * of_real pi * \<i>)"  | 
|
3394  | 
if "x < y" for x y  | 
|
3395  | 
proof -  | 
|
3396  | 
have "1 \<le> abs (x - y)"  | 
|
3397  | 
using that by linarith  | 
|
3398  | 
then have "1 \<le> cmod (of_int x - of_int y) * 1"  | 
|
3399  | 
by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff)  | 
|
3400  | 
also have "... \<le> cmod (of_int x - of_int y) * of_real pi"  | 
|
3401  | 
apply (rule mult_left_mono)  | 
|
3402  | 
using pi_ge_two by auto  | 
|
3403  | 
also have "... \<le> cmod ((of_int x - of_int y) * of_real pi * \<i>)"  | 
|
3404  | 
by (simp add: norm_mult)  | 
|
3405  | 
also have "... \<le> cmod (of_int x * of_real pi * \<i> - of_int y * of_real pi * \<i>)"  | 
|
3406  | 
by (simp add: algebra_simps)  | 
|
3407  | 
finally have "1 \<le> cmod (of_int x * of_real pi * \<i> - of_int y * of_real pi * \<i>)" .  | 
|
3408  | 
then have "2 * 1 \<le> cmod (2 * (of_int x * of_real pi * \<i> - of_int y * of_real pi * \<i>))"  | 
|
3409  | 
by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral)  | 
|
3410  | 
then show ?thesis  | 
|
3411  | 
by (simp add: algebra_simps)  | 
|
3412  | 
qed  | 
|
3413  | 
show "disjoint \<V>"  | 
|
3414  | 
apply (clarsimp simp add: \<V>_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y]  | 
|
3415  | 
image_add_ball ball_eq_ball_iff)  | 
|
3416  | 
apply (rule disjoint_ballI)  | 
|
3417  | 
apply (auto simp: dist_norm neq_iff)  | 
|
3418  | 
by (metis norm_minus_commute xy)+  | 
|
3419  | 
show "\<forall>u\<in>\<V>. \<exists>q. homeomorphism u (exp ` ball (Ln z) 1) exp q"  | 
|
3420  | 
proof  | 
|
3421  | 
fix u  | 
|
3422  | 
assume "u \<in> \<V>"  | 
|
| 64508 | 3423  | 
then obtain n where n: "u = (\<lambda>x. x + of_real (2 * of_int n * pi) * \<i>) ` (ball(Ln z) 1)"  | 
| 64287 | 3424  | 
by (auto simp: \<V>_def)  | 
3425  | 
have "compact (cball (Ln z) 1)"  | 
|
3426  | 
by simp  | 
|
3427  | 
moreover have "continuous_on (cball (Ln z) 1) exp"  | 
|
3428  | 
by (rule continuous_on_exp [OF continuous_on_id])  | 
|
3429  | 
moreover have "inj_on exp (cball (Ln z) 1)"  | 
|
3430  | 
apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]])  | 
|
3431  | 
using pi_ge_two by (simp add: cball_subset_ball_iff)  | 
|
3432  | 
ultimately obtain \<gamma> where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \<gamma>"  | 
|
3433  | 
using homeomorphism_compact by blast  | 
|
3434  | 
have eq1: "exp ` u = exp ` ball (Ln z) 1"  | 
|
3435  | 
unfolding n  | 
|
3436  | 
apply (auto simp: algebra_simps)  | 
|
3437  | 
apply (rename_tac w)  | 
|
3438  | 
apply (rule_tac x = "w + \<i> * (of_int n * (of_real pi * 2))" in image_eqI)  | 
|
3439  | 
apply (auto simp: image_iff)  | 
|
3440  | 
done  | 
|
3441  | 
have \<gamma>exp: "\<gamma> (exp x) + 2 * of_int n * of_real pi * \<i> = x" if "x \<in> u" for x  | 
|
3442  | 
proof -  | 
|
3443  | 
have "exp x = exp (x - 2 * of_int n * of_real pi * \<i>)"  | 
|
3444  | 
by (simp add: exp_eq)  | 
|
3445  | 
then have "\<gamma> (exp x) = \<gamma> (exp (x - 2 * of_int n * of_real pi * \<i>))"  | 
|
3446  | 
by simp  | 
|
3447  | 
also have "... = x - 2 * of_int n * of_real pi * \<i>"  | 
|
3448  | 
apply (rule homeomorphism_apply1 [OF hom])  | 
|
3449  | 
using \<open>x \<in> u\<close> by (auto simp: n)  | 
|
3450  | 
finally show ?thesis  | 
|
3451  | 
by simp  | 
|
3452  | 
qed  | 
|
3453  | 
have exp2n: "exp (\<gamma> (exp x) + 2 * of_int n * complex_of_real pi * \<i>) = exp x"  | 
|
3454  | 
if "dist (Ln z) x < 1" for x  | 
|
3455  | 
using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom])  | 
|
3456  | 
have cont: "continuous_on (exp ` ball (Ln z) 1) (\<lambda>x. \<gamma> x + 2 * of_int n * complex_of_real pi * \<i>)"  | 
|
3457  | 
apply (intro continuous_intros)  | 
|
3458  | 
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom]])  | 
|
3459  | 
apply (force simp:)  | 
|
3460  | 
done  | 
|
3461  | 
show "\<exists>q. homeomorphism u (exp ` ball (Ln z) 1) exp q"  | 
|
| 64508 | 3462  | 
apply (rule_tac x="(\<lambda>x. x + of_real(2 * n * pi) * \<i>) \<circ> \<gamma>" in exI)  | 
| 64287 | 3463  | 
unfolding homeomorphism_def  | 
3464  | 
apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id])  | 
|
3465  | 
apply (auto simp: \<gamma>exp exp2n cont n)  | 
|
3466  | 
apply (simp add: homeomorphism_apply1 [OF hom])  | 
|
3467  | 
apply (simp add: image_comp [symmetric])  | 
|
3468  | 
using hom homeomorphism_apply1 apply (force simp: image_iff)  | 
|
3469  | 
done  | 
|
3470  | 
qed  | 
|
3471  | 
qed  | 
|
3472  | 
qed  | 
|
3473  | 
qed  | 
|
3474  | 
||
| 64845 | 3475  | 
|
3476  | 
subsection\<open>Hence the Borsukian results about mappings into circles\<close>  | 
|
3477  | 
||
3478  | 
lemma inessential_eq_continuous_logarithm:  | 
|
3479  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3480  | 
  shows "(\<exists>a. homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)) \<longleftrightarrow>
 | 
|
3481  | 
(\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x)))"  | 
|
3482  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
3483  | 
proof  | 
|
3484  | 
assume ?lhs thus ?rhs  | 
|
3485  | 
by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane)  | 
|
3486  | 
next  | 
|
3487  | 
assume ?rhs  | 
|
3488  | 
then obtain g where contg: "continuous_on S g" and f: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3489  | 
by metis  | 
|
3490  | 
  obtain a where "homotopic_with (\<lambda>h. True) S (- {of_real 0}) (exp \<circ> g) (\<lambda>x. a)"
 | 
|
3491  | 
proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV])  | 
|
3492  | 
show "continuous_on (UNIV::complex set) exp"  | 
|
3493  | 
by (intro continuous_intros)  | 
|
3494  | 
    show "range exp \<subseteq> - {0}"
 | 
|
3495  | 
by auto  | 
|
3496  | 
qed force  | 
|
3497  | 
thus ?lhs  | 
|
3498  | 
apply (rule_tac x=a in exI)  | 
|
3499  | 
by (simp add: f homotopic_with_eq)  | 
|
3500  | 
qed  | 
|
3501  | 
||
3502  | 
corollary inessential_imp_continuous_logarithm_circle:  | 
|
3503  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3504  | 
assumes "homotopic_with (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)"  | 
|
3505  | 
obtains g where "continuous_on S g" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3506  | 
proof -  | 
|
3507  | 
  have "homotopic_with (\<lambda>h. True) S (- {0}) f (\<lambda>t. a)"
 | 
|
3508  | 
using assms homotopic_with_subset_right by fastforce  | 
|
3509  | 
then show ?thesis  | 
|
3510  | 
by (metis inessential_eq_continuous_logarithm that)  | 
|
3511  | 
qed  | 
|
3512  | 
||
3513  | 
||
3514  | 
lemma inessential_eq_continuous_logarithm_circle:  | 
|
3515  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3516  | 
shows "(\<exists>a. homotopic_with (\<lambda>h. True) S (sphere 0 1) f (\<lambda>t. a)) \<longleftrightarrow>  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
3517  | 
(\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(\<i> * of_real(g x))))"  | 
| 64845 | 3518  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
3519  | 
proof  | 
|
3520  | 
assume L: ?lhs  | 
|
3521  | 
then obtain g where contg: "continuous_on S g" and g: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3522  | 
using inessential_imp_continuous_logarithm_circle by blast  | 
|
3523  | 
have "f ` S \<subseteq> sphere 0 1"  | 
|
3524  | 
by (metis L homotopic_with_imp_subset1)  | 
|
3525  | 
then have "\<And>x. x \<in> S \<Longrightarrow> Re (g x) = 0"  | 
|
3526  | 
using g by auto  | 
|
3527  | 
then show ?rhs  | 
|
3528  | 
apply (rule_tac x="Im \<circ> g" in exI)  | 
|
3529  | 
apply (intro conjI contg continuous_intros)  | 
|
3530  | 
apply (auto simp: Euler g)  | 
|
3531  | 
done  | 
|
3532  | 
next  | 
|
3533  | 
assume ?rhs  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
3534  | 
then obtain g where contg: "continuous_on S g" and g: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(\<i>* of_real(g x))"  | 
| 64845 | 3535  | 
by metis  | 
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
3536  | 
obtain a where "homotopic_with (\<lambda>h. True) S (sphere 0 1) ((exp \<circ> (\<lambda>z. \<i>*z)) \<circ> (of_real \<circ> g)) (\<lambda>x. a)"  | 
| 64845 | 3537  | 
proof (rule nullhomotopic_through_contractible)  | 
3538  | 
show "continuous_on S (complex_of_real \<circ> g)"  | 
|
3539  | 
by (intro conjI contg continuous_intros)  | 
|
3540  | 
show "(complex_of_real \<circ> g) ` S \<subseteq> \<real>"  | 
|
3541  | 
by auto  | 
|
| 67399 | 3542  | 
show "continuous_on \<real> (exp \<circ> ( * )\<i>)"  | 
| 64845 | 3543  | 
by (intro continuous_intros)  | 
| 67399 | 3544  | 
show "(exp \<circ> ( * )\<i>) ` \<real> \<subseteq> sphere 0 1"  | 
| 64845 | 3545  | 
by (auto simp: complex_is_Real_iff)  | 
3546  | 
qed (auto simp: convex_Reals convex_imp_contractible)  | 
|
| 67399 | 3547  | 
moreover have "\<And>x. x \<in> S \<Longrightarrow> (exp \<circ> ( * )\<i> \<circ> (complex_of_real \<circ> g)) x = f x"  | 
| 64845 | 3548  | 
by (simp add: g)  | 
3549  | 
ultimately show ?lhs  | 
|
3550  | 
apply (rule_tac x=a in exI)  | 
|
3551  | 
by (simp add: homotopic_with_eq)  | 
|
3552  | 
qed  | 
|
3553  | 
||
3554  | 
lemma homotopic_with_sphere_times:  | 
|
3555  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3556  | 
assumes hom: "homotopic_with (\<lambda>x. True) S (sphere 0 1) f g" and conth: "continuous_on S h"  | 
|
3557  | 
and hin: "\<And>x. x \<in> S \<Longrightarrow> h x \<in> sphere 0 1"  | 
|
| 64846 | 3558  | 
shows "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x * h x) (\<lambda>x. g x * h x)"  | 
| 64845 | 3559  | 
proof -  | 
3560  | 
  obtain k where contk: "continuous_on ({0..1::real} \<times> S) k"
 | 
|
3561  | 
             and kim: "k ` ({0..1} \<times> S) \<subseteq> sphere 0 1"
 | 
|
3562  | 
and k0: "\<And>x. k(0, x) = f x"  | 
|
3563  | 
and k1: "\<And>x. k(1, x) = g x"  | 
|
3564  | 
using hom by (auto simp: homotopic_with_def)  | 
|
3565  | 
show ?thesis  | 
|
3566  | 
apply (simp add: homotopic_with)  | 
|
3567  | 
apply (rule_tac x="\<lambda>z. k z*(h \<circ> snd)z" in exI)  | 
|
3568  | 
apply (intro conjI contk continuous_intros)  | 
|
3569  | 
apply (simp add: conth)  | 
|
3570  | 
using kim hin apply (force simp: norm_mult k0 k1)+  | 
|
3571  | 
done  | 
|
3572  | 
qed  | 
|
3573  | 
||
3574  | 
||
3575  | 
lemma homotopic_circlemaps_divide:  | 
|
3576  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3577  | 
shows "homotopic_with (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>  | 
|
3578  | 
continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>  | 
|
3579  | 
continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>  | 
|
3580  | 
(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c))"  | 
|
3581  | 
proof -  | 
|
3582  | 
have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"  | 
|
3583  | 
if "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c)" for c  | 
|
3584  | 
proof -  | 
|
3585  | 
    have "S = {} \<or> path_component (sphere 0 1) 1 c"
 | 
|
3586  | 
using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1]  | 
|
3587  | 
by (auto simp: path_connected_component)  | 
|
3588  | 
then have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. 1) (\<lambda>x. c)"  | 
|
3589  | 
by (metis homotopic_constant_maps)  | 
|
3590  | 
then show ?thesis  | 
|
3591  | 
using homotopic_with_symD homotopic_with_trans that by blast  | 
|
3592  | 
qed  | 
|
3593  | 
then have *: "(\<exists>c. homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. c)) \<longleftrightarrow>  | 
|
3594  | 
homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"  | 
|
3595  | 
by auto  | 
|
3596  | 
have "homotopic_with (\<lambda>x. True) S (sphere 0 1) f g \<longleftrightarrow>  | 
|
3597  | 
continuous_on S f \<and> f ` S \<subseteq> sphere 0 1 \<and>  | 
|
3598  | 
continuous_on S g \<and> g ` S \<subseteq> sphere 0 1 \<and>  | 
|
3599  | 
homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"  | 
|
3600  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
3601  | 
proof  | 
|
3602  | 
assume L: ?lhs  | 
|
3603  | 
have geq1 [simp]: "\<And>x. x \<in> S \<Longrightarrow> cmod (g x) = 1"  | 
|
3604  | 
using homotopic_with_imp_subset2 [OF L]  | 
|
3605  | 
by (simp add: image_subset_iff)  | 
|
3606  | 
have cont: "continuous_on S (inverse \<circ> g)"  | 
|
3607  | 
apply (rule continuous_intros)  | 
|
3608  | 
using homotopic_with_imp_continuous [OF L] apply blast  | 
|
3609  | 
apply (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse])  | 
|
3610  | 
apply (auto simp: continuous_on_id)  | 
|
3611  | 
done  | 
|
3612  | 
have "homotopic_with (\<lambda>x. True) S (sphere 0 1) (\<lambda>x. f x / g x) (\<lambda>x. 1)"  | 
|
3613  | 
using homotopic_with_sphere_times [OF L cont]  | 
|
3614  | 
apply (rule homotopic_with_eq)  | 
|
3615  | 
apply (auto simp: division_ring_class.divide_inverse norm_inverse)  | 
|
3616  | 
by (metis geq1 norm_zero right_inverse zero_neq_one)  | 
|
3617  | 
with L show ?rhs  | 
|
3618  | 
by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2)  | 
|
3619  | 
next  | 
|
3620  | 
assume ?rhs then show ?lhs  | 
|
3621  | 
by (force simp: elim: homotopic_with_eq dest: homotopic_with_sphere_times [where h=g])+  | 
|
3622  | 
qed  | 
|
3623  | 
then show ?thesis  | 
|
3624  | 
by (simp add: *)  | 
|
3625  | 
qed  | 
|
3626  | 
||
3627  | 
subsection\<open>Upper and lower hemicontinuous functions\<close>  | 
|
3628  | 
||
3629  | 
text\<open>And relation in the case of preimage map to open and closed maps, and fact that upper and lower  | 
|
3630  | 
hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the  | 
|
3631  | 
function gives a bounded and nonempty set).\<close>  | 
|
3632  | 
||
3633  | 
||
3634  | 
text\<open>Many similar proofs below.\<close>  | 
|
3635  | 
lemma upper_hemicontinuous:  | 
|
3636  | 
assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"  | 
|
3637  | 
shows "((\<forall>U. openin (subtopology euclidean T) U  | 
|
3638  | 
                 \<longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
 | 
|
3639  | 
(\<forall>U. closedin (subtopology euclidean T) U  | 
|
3640  | 
                 \<longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
 | 
|
3641  | 
(is "?lhs = ?rhs")  | 
|
3642  | 
proof (intro iffI allI impI)  | 
|
3643  | 
fix U  | 
|
3644  | 
assume * [rule_format]: ?lhs and "closedin (subtopology euclidean T) U"  | 
|
3645  | 
then have "openin (subtopology euclidean T) (T - U)"  | 
|
3646  | 
by (simp add: openin_diff)  | 
|
3647  | 
  then have "openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> T - U}"
 | 
|
3648  | 
using * [of "T-U"] by blast  | 
|
3649  | 
  moreover have "S - {x \<in> S. f x \<subseteq> T - U} = {x \<in> S. f x \<inter> U \<noteq> {}}"
 | 
|
3650  | 
using assms by blast  | 
|
3651  | 
  ultimately show "closedin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}"
 | 
|
3652  | 
by (simp add: openin_closedin_eq)  | 
|
3653  | 
next  | 
|
3654  | 
fix U  | 
|
3655  | 
assume * [rule_format]: ?rhs and "openin (subtopology euclidean T) U"  | 
|
3656  | 
then have "closedin (subtopology euclidean T) (T - U)"  | 
|
3657  | 
by (simp add: closedin_diff)  | 
|
3658  | 
  then have "closedin (subtopology euclidean S) {x \<in> S. f x \<inter> (T - U) \<noteq> {}}"
 | 
|
3659  | 
using * [of "T-U"] by blast  | 
|
3660  | 
  moreover have "{x \<in> S. f x \<inter> (T - U) \<noteq> {}} = S - {x \<in> S. f x \<subseteq> U}"
 | 
|
3661  | 
using assms by auto  | 
|
3662  | 
  ultimately show "openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
 | 
|
3663  | 
by (simp add: openin_closedin_eq)  | 
|
3664  | 
qed  | 
|
3665  | 
||
3666  | 
lemma lower_hemicontinuous:  | 
|
3667  | 
assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"  | 
|
3668  | 
shows "((\<forall>U. closedin (subtopology euclidean T) U  | 
|
3669  | 
                 \<longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}) \<longleftrightarrow>
 | 
|
3670  | 
(\<forall>U. openin (subtopology euclidean T) U  | 
|
3671  | 
                 \<longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}))"
 | 
|
3672  | 
(is "?lhs = ?rhs")  | 
|
3673  | 
proof (intro iffI allI impI)  | 
|
3674  | 
fix U  | 
|
3675  | 
assume * [rule_format]: ?lhs and "openin (subtopology euclidean T) U"  | 
|
3676  | 
then have "closedin (subtopology euclidean T) (T - U)"  | 
|
3677  | 
by (simp add: closedin_diff)  | 
|
3678  | 
  then have "closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> T-U}"
 | 
|
3679  | 
using * [of "T-U"] by blast  | 
|
3680  | 
  moreover have "{x \<in> S. f x \<subseteq> T-U} = S - {x \<in> S. f x \<inter> U \<noteq> {}}"
 | 
|
3681  | 
using assms by auto  | 
|
3682  | 
  ultimately show "openin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}"
 | 
|
3683  | 
by (simp add: openin_closedin_eq)  | 
|
3684  | 
next  | 
|
3685  | 
fix U  | 
|
3686  | 
assume * [rule_format]: ?rhs and "closedin (subtopology euclidean T) U"  | 
|
3687  | 
then have "openin (subtopology euclidean T) (T - U)"  | 
|
3688  | 
by (simp add: openin_diff)  | 
|
3689  | 
  then have "openin (subtopology euclidean S) {x \<in> S. f x \<inter> (T - U) \<noteq> {}}"
 | 
|
3690  | 
using * [of "T-U"] by blast  | 
|
3691  | 
  moreover have "S - {x \<in> S. f x \<inter> (T - U) \<noteq> {}} = {x \<in> S. f x \<subseteq> U}"
 | 
|
3692  | 
using assms by blast  | 
|
3693  | 
  ultimately show "closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
 | 
|
3694  | 
by (simp add: openin_closedin_eq)  | 
|
3695  | 
qed  | 
|
3696  | 
||
3697  | 
lemma open_map_iff_lower_hemicontinuous_preimage:  | 
|
3698  | 
assumes "f ` S \<subseteq> T"  | 
|
3699  | 
shows "((\<forall>U. openin (subtopology euclidean S) U  | 
|
3700  | 
\<longrightarrow> openin (subtopology euclidean T) (f ` U)) \<longleftrightarrow>  | 
|
3701  | 
(\<forall>U. closedin (subtopology euclidean S) U  | 
|
3702  | 
                 \<longrightarrow> closedin (subtopology euclidean T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}))"
 | 
|
3703  | 
(is "?lhs = ?rhs")  | 
|
3704  | 
proof (intro iffI allI impI)  | 
|
3705  | 
fix U  | 
|
3706  | 
assume * [rule_format]: ?lhs and "closedin (subtopology euclidean S) U"  | 
|
3707  | 
then have "openin (subtopology euclidean S) (S - U)"  | 
|
3708  | 
by (simp add: openin_diff)  | 
|
3709  | 
then have "openin (subtopology euclidean T) (f ` (S - U))"  | 
|
3710  | 
using * [of "S-U"] by blast  | 
|
3711  | 
  moreover have "T - (f ` (S - U)) = {y \<in> T. {x \<in> S. f x = y} \<subseteq> U}"
 | 
|
3712  | 
using assms by blast  | 
|
3713  | 
  ultimately show "closedin (subtopology euclidean T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> U}"
 | 
|
3714  | 
by (simp add: openin_closedin_eq)  | 
|
3715  | 
next  | 
|
3716  | 
fix U  | 
|
3717  | 
assume * [rule_format]: ?rhs and opeSU: "openin (subtopology euclidean S) U"  | 
|
3718  | 
then have "closedin (subtopology euclidean S) (S - U)"  | 
|
3719  | 
by (simp add: closedin_diff)  | 
|
3720  | 
  then have "closedin (subtopology euclidean T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
 | 
|
3721  | 
using * [of "S-U"] by blast  | 
|
3722  | 
  moreover have "{y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U} = T - (f ` U)"
 | 
|
3723  | 
using assms openin_imp_subset [OF opeSU] by auto  | 
|
3724  | 
ultimately show "openin (subtopology euclidean T) (f ` U)"  | 
|
3725  | 
using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq)  | 
|
3726  | 
qed  | 
|
3727  | 
||
3728  | 
lemma closed_map_iff_upper_hemicontinuous_preimage:  | 
|
3729  | 
assumes "f ` S \<subseteq> T"  | 
|
3730  | 
shows "((\<forall>U. closedin (subtopology euclidean S) U  | 
|
3731  | 
\<longrightarrow> closedin (subtopology euclidean T) (f ` U)) \<longleftrightarrow>  | 
|
3732  | 
(\<forall>U. openin (subtopology euclidean S) U  | 
|
3733  | 
                 \<longrightarrow> openin (subtopology euclidean T) {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}))"
 | 
|
3734  | 
(is "?lhs = ?rhs")  | 
|
3735  | 
proof (intro iffI allI impI)  | 
|
3736  | 
fix U  | 
|
3737  | 
assume * [rule_format]: ?lhs and opeSU: "openin (subtopology euclidean S) U"  | 
|
3738  | 
then have "closedin (subtopology euclidean S) (S - U)"  | 
|
3739  | 
by (simp add: closedin_diff)  | 
|
3740  | 
then have "closedin (subtopology euclidean T) (f ` (S - U))"  | 
|
3741  | 
using * [of "S-U"] by blast  | 
|
3742  | 
  moreover have "f ` (S - U) = T -  {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}"
 | 
|
3743  | 
using assms openin_imp_subset [OF opeSU] by auto  | 
|
3744  | 
  ultimately show "openin (subtopology euclidean T)  {y \<in> T. {x. x \<in> S \<and> f x = y} \<subseteq> U}"
 | 
|
3745  | 
using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq)  | 
|
3746  | 
next  | 
|
3747  | 
fix U  | 
|
3748  | 
assume * [rule_format]: ?rhs and cloSU: "closedin (subtopology euclidean S) U"  | 
|
3749  | 
then have "openin (subtopology euclidean S) (S - U)"  | 
|
3750  | 
by (simp add: openin_diff)  | 
|
3751  | 
  then have "openin (subtopology euclidean T) {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
 | 
|
3752  | 
using * [of "S-U"] by blast  | 
|
3753  | 
  moreover have "(f ` U) = T - {y \<in> T. {x \<in> S. f x = y} \<subseteq> S - U}"
 | 
|
3754  | 
using assms closedin_imp_subset [OF cloSU] by auto  | 
|
3755  | 
ultimately show "closedin (subtopology euclidean T) (f ` U)"  | 
|
3756  | 
by (simp add: openin_closedin_eq)  | 
|
3757  | 
qed  | 
|
3758  | 
||
3759  | 
proposition upper_lower_hemicontinuous_explicit:  | 
|
3760  | 
  fixes T :: "('b::{real_normed_vector,heine_borel}) set"
 | 
|
3761  | 
assumes fST: "\<And>x. x \<in> S \<Longrightarrow> f x \<subseteq> T"  | 
|
3762  | 
and ope: "\<And>U. openin (subtopology euclidean T) U  | 
|
3763  | 
                     \<Longrightarrow> openin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
 | 
|
3764  | 
and clo: "\<And>U. closedin (subtopology euclidean T) U  | 
|
3765  | 
                     \<Longrightarrow> closedin (subtopology euclidean S) {x \<in> S. f x \<subseteq> U}"
 | 
|
3766  | 
      and "x \<in> S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \<noteq> {}"
 | 
|
3767  | 
obtains d where "0 < d"  | 
|
3768  | 
"\<And>x'. \<lbrakk>x' \<in> S; dist x x' < d\<rbrakk>  | 
|
3769  | 
\<Longrightarrow> (\<forall>y \<in> f x. \<exists>y'. y' \<in> f x' \<and> dist y y' < e) \<and>  | 
|
3770  | 
(\<forall>y' \<in> f x'. \<exists>y. y \<in> f x \<and> dist y' y < e)"  | 
|
3771  | 
proof -  | 
|
3772  | 
  have "openin (subtopology euclidean T) (T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b}))"
 | 
|
3773  | 
by (auto simp: open_sums openin_open_Int)  | 
|
3774  | 
with ope have "openin (subtopology euclidean S)  | 
|
3775  | 
                    {u \<in> S. f u \<subseteq> T \<inter> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b})}" by blast
 | 
|
3776  | 
with \<open>0 < e\<close> \<open>x \<in> S\<close> obtain d1 where "d1 > 0" and  | 
|
3777  | 
         d1: "\<And>x'. \<lbrakk>x' \<in> S; dist x' x < d1\<rbrakk> \<Longrightarrow> f x' \<subseteq> T \<and> f x' \<subseteq> (\<Union>a \<in> f x. \<Union>b \<in> ball 0 e. {a + b})"
 | 
|
3778  | 
by (force simp: openin_euclidean_subtopology_iff dest: fST)  | 
|
3779  | 
have oo: "\<And>U. openin (subtopology euclidean T) U \<Longrightarrow>  | 
|
3780  | 
                 openin (subtopology euclidean S) {x \<in> S. f x \<inter> U \<noteq> {}}"
 | 
|
3781  | 
apply (rule lower_hemicontinuous [THEN iffD1, rule_format])  | 
|
3782  | 
using fST clo by auto  | 
|
3783  | 
have "compact (closure(f x))"  | 
|
3784  | 
by (simp add: bofx)  | 
|
3785  | 
moreover have "closure(f x) \<subseteq> (\<Union>a \<in> f x. ball a (e/2))"  | 
|
3786  | 
using \<open>0 < e\<close> by (force simp: closure_approachable simp del: divide_const_simps)  | 
|
3787  | 
ultimately obtain C where "C \<subseteq> f x" "finite C" "closure(f x) \<subseteq> (\<Union>a \<in> C. ball a (e/2))"  | 
|
3788  | 
apply (rule compactE, force)  | 
|
3789  | 
by (metis finite_subset_image)  | 
|
3790  | 
then have fx_cover: "f x \<subseteq> (\<Union>a \<in> C. ball a (e/2))"  | 
|
3791  | 
by (meson closure_subset order_trans)  | 
|
3792  | 
  with fx_ne have "C \<noteq> {}"
 | 
|
3793  | 
by blast  | 
|
3794  | 
  have xin: "x \<in> (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
 | 
|
3795  | 
using \<open>x \<in> S\<close> \<open>0 < e\<close> fST \<open>C \<subseteq> f x\<close> by force  | 
|
3796  | 
  have "openin (subtopology euclidean S) {x \<in> S. f x \<inter> (T \<inter> ball a (e/2)) \<noteq> {}}" for a
 | 
|
3797  | 
by (simp add: openin_open_Int oo)  | 
|
3798  | 
  then have "openin (subtopology euclidean S) (\<Inter>a \<in> C. {x \<in> S. f x \<inter> T \<inter> ball a (e/2) \<noteq> {}})"
 | 
|
3799  | 
    by (simp add: Int_assoc openin_INT2 [OF \<open>finite C\<close> \<open>C \<noteq> {}\<close>])
 | 
|
3800  | 
with xin obtain d2 where "d2>0"  | 
|
3801  | 
              and d2: "\<And>u v. \<lbrakk>u \<in> S; dist u x < d2; v \<in> C\<rbrakk> \<Longrightarrow> f u \<inter> T \<inter> ball v (e/2) \<noteq> {}"
 | 
|
| 
66955
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3802  | 
unfolding openin_euclidean_subtopology_iff using xin by fastforce  | 
| 64845 | 3803  | 
show ?thesis  | 
3804  | 
proof (intro that conjI ballI)  | 
|
3805  | 
show "0 < min d1 d2"  | 
|
3806  | 
using \<open>0 < d1\<close> \<open>0 < d2\<close> by linarith  | 
|
3807  | 
next  | 
|
3808  | 
fix x' y  | 
|
3809  | 
assume "x' \<in> S" "dist x x' < min d1 d2" "y \<in> f x"  | 
|
3810  | 
then have dd2: "dist x' x < d2"  | 
|
3811  | 
by (auto simp: dist_commute)  | 
|
3812  | 
obtain a where "a \<in> C" "y \<in> ball a (e/2)"  | 
|
3813  | 
using fx_cover \<open>y \<in> f x\<close> by auto  | 
|
3814  | 
then show "\<exists>y'. y' \<in> f x' \<and> dist y y' < e"  | 
|
3815  | 
using d2 [OF \<open>x' \<in> S\<close> dd2] dist_triangle_half_r by fastforce  | 
|
3816  | 
next  | 
|
3817  | 
fix x' y'  | 
|
3818  | 
assume "x' \<in> S" "dist x x' < min d1 d2" "y' \<in> f x'"  | 
|
3819  | 
then have "dist x' x < d1"  | 
|
3820  | 
by (auto simp: dist_commute)  | 
|
3821  | 
    then have "y' \<in> (\<Union>a\<in>f x. \<Union>b\<in>ball 0 e. {a + b})"
 | 
|
3822  | 
using d1 [OF \<open>x' \<in> S\<close>] \<open>y' \<in> f x'\<close> by force  | 
|
3823  | 
then show "\<exists>y. y \<in> f x \<and> dist y' y < e"  | 
|
3824  | 
apply auto  | 
|
3825  | 
by (metis add_diff_cancel_left' dist_norm)  | 
|
3826  | 
qed  | 
|
3827  | 
qed  | 
|
3828  | 
||
3829  | 
||
| 
66955
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3830  | 
subsection\<open>Complex logs exist on various "well-behaved" sets\<close>  | 
| 64845 | 3831  | 
|
3832  | 
lemma continuous_logarithm_on_contractible:  | 
|
3833  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3834  | 
assumes "continuous_on S f" "contractible S" "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
3835  | 
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3836  | 
proof -  | 
|
3837  | 
  obtain c where hom: "homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>x. c)"
 | 
|
3838  | 
using nullhomotopic_from_contractible assms  | 
|
3839  | 
by (metis imageE subset_Compl_singleton)  | 
|
3840  | 
then show ?thesis  | 
|
| 
66955
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3841  | 
by (metis inessential_eq_continuous_logarithm that)  | 
| 64845 | 3842  | 
qed  | 
3843  | 
||
3844  | 
lemma continuous_logarithm_on_simply_connected:  | 
|
3845  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3846  | 
assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S"  | 
|
3847  | 
and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
3848  | 
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3849  | 
using covering_space_lift [OF covering_space_exp_punctured_plane S contf]  | 
|
3850  | 
by (metis (full_types) f imageE subset_Compl_singleton)  | 
|
3851  | 
||
3852  | 
lemma continuous_logarithm_on_cball:  | 
|
3853  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3854  | 
assumes "continuous_on (cball a r) f" and "\<And>z. z \<in> cball a r \<Longrightarrow> f z \<noteq> 0"  | 
|
3855  | 
obtains h where "continuous_on (cball a r) h" "\<And>z. z \<in> cball a r \<Longrightarrow> f z = exp(h z)"  | 
|
3856  | 
using assms continuous_logarithm_on_contractible convex_imp_contractible by blast  | 
|
3857  | 
||
3858  | 
lemma continuous_logarithm_on_ball:  | 
|
3859  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3860  | 
assumes "continuous_on (ball a r) f" and "\<And>z. z \<in> ball a r \<Longrightarrow> f z \<noteq> 0"  | 
|
3861  | 
obtains h where "continuous_on (ball a r) h" "\<And>z. z \<in> ball a r \<Longrightarrow> f z = exp(h z)"  | 
|
3862  | 
using assms continuous_logarithm_on_contractible convex_imp_contractible by blast  | 
|
3863  | 
||
3864  | 
lemma continuous_sqrt_on_contractible:  | 
|
3865  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3866  | 
assumes "continuous_on S f" "contractible S"  | 
|
3867  | 
and "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
3868  | 
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = (g x) ^ 2"  | 
|
3869  | 
proof -  | 
|
3870  | 
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3871  | 
using continuous_logarithm_on_contractible [OF assms] by blast  | 
|
3872  | 
show ?thesis  | 
|
3873  | 
proof  | 
|
3874  | 
show "continuous_on S (\<lambda>z. exp (g z / 2))"  | 
|
3875  | 
by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto  | 
|
3876  | 
show "\<And>x. x \<in> S \<Longrightarrow> f x = (exp (g x / 2))\<^sup>2"  | 
|
3877  | 
by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)  | 
|
3878  | 
qed  | 
|
3879  | 
qed  | 
|
3880  | 
||
3881  | 
lemma continuous_sqrt_on_simply_connected:  | 
|
3882  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> complex"  | 
|
3883  | 
assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S"  | 
|
3884  | 
and f: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
3885  | 
obtains g where "continuous_on S g" "\<And>x. x \<in> S \<Longrightarrow> f x = (g x) ^ 2"  | 
|
3886  | 
proof -  | 
|
3887  | 
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
3888  | 
using continuous_logarithm_on_simply_connected [OF assms] by blast  | 
|
3889  | 
show ?thesis  | 
|
3890  | 
proof  | 
|
3891  | 
show "continuous_on S (\<lambda>z. exp (g z / 2))"  | 
|
3892  | 
by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto  | 
|
3893  | 
show "\<And>x. x \<in> S \<Longrightarrow> f x = (exp (g x / 2))\<^sup>2"  | 
|
3894  | 
by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)  | 
|
3895  | 
qed  | 
|
3896  | 
qed  | 
|
3897  | 
||
3898  | 
||
| 
66955
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3899  | 
subsection\<open>Another simple case where sphere maps are nullhomotopic.\<close>  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3900  | 
|
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3901  | 
lemma inessential_spheremap_2_aux:  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3902  | 
fixes f :: "'a::euclidean_space \<Rightarrow> complex"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3903  | 
  assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" 
 | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3904  | 
and fim: "f `(sphere a r) \<subseteq> (sphere 0 1)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3905  | 
obtains c where "homotopic_with (\<lambda>z. True) (sphere a r) (sphere 0 1) f (\<lambda>x. c)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3906  | 
proof -  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3907  | 
obtain g where contg: "continuous_on (sphere a r) g"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3908  | 
and feq: "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = exp(g x)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3909  | 
proof (rule continuous_logarithm_on_simply_connected [OF contf])  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3910  | 
show "simply_connected (sphere a r)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3911  | 
using 2 by (simp add: simply_connected_sphere_eq)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3912  | 
show "locally path_connected (sphere a r)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3913  | 
by (simp add: locally_path_connected_sphere)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3914  | 
show "\<And>z. z \<in> sphere a r \<Longrightarrow> f z \<noteq> 0"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3915  | 
using fim by force  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3916  | 
qed auto  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3917  | 
have "\<exists>g. continuous_on (sphere a r) g \<and> (\<forall>x\<in>sphere a r. f x = exp (\<i> * complex_of_real (g x)))"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3918  | 
proof (intro exI conjI)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3919  | 
show "continuous_on (sphere a r) (Im \<circ> g)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3920  | 
by (intro contg continuous_intros continuous_on_compose)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3921  | 
show "\<forall>x\<in>sphere a r. f x = exp (\<i> * complex_of_real ((Im \<circ> g) x))"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3922  | 
using exp_eq_polar feq fim norm_exp_eq_Re by auto  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3923  | 
qed  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3924  | 
with inessential_eq_continuous_logarithm_circle that show ?thesis  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3925  | 
by metis  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3926  | 
qed  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3927  | 
|
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3928  | 
lemma inessential_spheremap_2:  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3929  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3930  | 
  assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" 
 | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3931  | 
and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \<subseteq> (sphere b s)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3932  | 
obtains c where "homotopic_with (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. c)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3933  | 
proof (cases "s \<le> 0")  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3934  | 
case True  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3935  | 
then show ?thesis  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3936  | 
using contf contractible_sphere fim nullhomotopic_into_contractible that by blast  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3937  | 
next  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3938  | 
case False  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3939  | 
then have "sphere b s homeomorphic sphere (0::complex) 1"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3940  | 
using assms by (simp add: homeomorphic_spheres_gen)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3941  | 
then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3942  | 
by (auto simp: homeomorphic_def)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3943  | 
then have conth: "continuous_on (sphere b s) h"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3944  | 
and contk: "continuous_on (sphere 0 1) k"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3945  | 
and him: "h ` sphere b s \<subseteq> sphere 0 1"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3946  | 
and kim: "k ` sphere 0 1 \<subseteq> sphere b s"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3947  | 
by (simp_all add: homeomorphism_def)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3948  | 
obtain c where "homotopic_with (\<lambda>z. True) (sphere a r) (sphere 0 1) (h \<circ> f) (\<lambda>x. c)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3949  | 
proof (rule inessential_spheremap_2_aux [OF a2])  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3950  | 
show "continuous_on (sphere a r) (h \<circ> f)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3951  | 
by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3952  | 
show "(h \<circ> f) ` sphere a r \<subseteq> sphere 0 1"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3953  | 
using fim him by force  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3954  | 
qed auto  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3955  | 
then have "homotopic_with (\<lambda>f. True) (sphere a r) (sphere b s) (k \<circ> (h \<circ> f)) (k \<circ> (\<lambda>x. c))"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3956  | 
by (rule homotopic_compose_continuous_left [OF _ contk kim])  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3957  | 
then have "homotopic_with (\<lambda>z. True) (sphere a r) (sphere b s) f (\<lambda>x. k c)"  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3958  | 
apply (rule homotopic_with_eq, auto)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3959  | 
by (metis fim hk homeomorphism_def image_subset_iff mem_sphere)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3960  | 
then show ?thesis  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3961  | 
by (metis that)  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3962  | 
qed  | 
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3963  | 
|
| 
 
289f390c4e57
A few more topological results. And made some slow proofs faster
 
paulson <lp15@cam.ac.uk> 
parents: 
66941 
diff
changeset
 | 
3964  | 
|
| 64845 | 3965  | 
subsection\<open>Holomorphic logarithms and square roots.\<close>  | 
3966  | 
||
3967  | 
lemma contractible_imp_holomorphic_log:  | 
|
3968  | 
assumes holf: "f holomorphic_on S"  | 
|
3969  | 
and S: "contractible S"  | 
|
3970  | 
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
3971  | 
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"  | 
|
3972  | 
proof -  | 
|
3973  | 
have contf: "continuous_on S f"  | 
|
3974  | 
by (simp add: holf holomorphic_on_imp_continuous_on)  | 
|
3975  | 
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp (g x)"  | 
|
3976  | 
by (metis continuous_logarithm_on_contractible [OF contf S fnz])  | 
|
3977  | 
have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \<in> S" for z  | 
|
3978  | 
proof -  | 
|
3979  | 
obtain f' where f': "((\<lambda>y. (f y - f z) / (y - z)) \<longlongrightarrow> f') (at z within S)"  | 
|
3980  | 
using \<open>f field_differentiable at z within S\<close> by (auto simp: field_differentiable_def DERIV_iff2)  | 
|
3981  | 
then have ee: "((\<lambda>x. (exp(g x) - exp(g z)) / (x - z)) \<longlongrightarrow> f') (at z within S)"  | 
|
3982  | 
by (simp add: feq \<open>z \<in> S\<close> Lim_transform_within [OF _ zero_less_one])  | 
|
3983  | 
have "(((\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<circ> g) \<longlongrightarrow> exp (g z))  | 
|
3984  | 
(at z within S)"  | 
|
3985  | 
proof (rule tendsto_compose_at)  | 
|
3986  | 
show "(g \<longlongrightarrow> g z) (at z within S)"  | 
|
3987  | 
using contg continuous_on \<open>z \<in> S\<close> by blast  | 
|
3988  | 
show "(\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<midarrow>g z\<rightarrow> exp (g z)"  | 
|
3989  | 
apply (subst Lim_at_zero)  | 
|
3990  | 
apply (simp add: DERIV_D cong: if_cong Lim_cong_within)  | 
|
3991  | 
done  | 
|
3992  | 
qed auto  | 
|
3993  | 
then have dd: "((\<lambda>x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \<longlongrightarrow> exp(g z)) (at z within S)"  | 
|
3994  | 
by (simp add: o_def)  | 
|
3995  | 
have "continuous (at z within S) g"  | 
|
3996  | 
using contg continuous_on_eq_continuous_within \<open>z \<in> S\<close> by blast  | 
|
3997  | 
then have "(\<forall>\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)"  | 
|
3998  | 
by (simp add: continuous_within tendsto_iff)  | 
|
3999  | 
then have "\<forall>\<^sub>F x in at z within S. exp (g x) = exp (g z) \<longrightarrow> g x \<noteq> g z \<longrightarrow> x = z"  | 
|
4000  | 
apply (rule eventually_mono)  | 
|
4001  | 
apply (auto simp: exp_eq dist_norm norm_mult)  | 
|
4002  | 
done  | 
|
4003  | 
then have "((\<lambda>y. (g y - g z) / (y - z)) \<longlongrightarrow> f' / exp (g z)) (at z within S)"  | 
|
4004  | 
by (auto intro!: Lim_transform_eventually [OF _ tendsto_divide [OF ee dd]])  | 
|
4005  | 
then show ?thesis  | 
|
4006  | 
by (auto simp: field_differentiable_def DERIV_iff2)  | 
|
4007  | 
qed  | 
|
4008  | 
then have "g holomorphic_on S"  | 
|
4009  | 
using holf holomorphic_on_def by auto  | 
|
4010  | 
then show ?thesis  | 
|
4011  | 
using feq that by auto  | 
|
4012  | 
qed  | 
|
4013  | 
||
4014  | 
(*Identical proofs*)  | 
|
4015  | 
lemma simply_connected_imp_holomorphic_log:  | 
|
4016  | 
assumes holf: "f holomorphic_on S"  | 
|
4017  | 
and S: "simply_connected S" "locally path_connected S"  | 
|
4018  | 
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
4019  | 
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"  | 
|
4020  | 
proof -  | 
|
4021  | 
have contf: "continuous_on S f"  | 
|
4022  | 
by (simp add: holf holomorphic_on_imp_continuous_on)  | 
|
4023  | 
obtain g where contg: "continuous_on S g" and feq: "\<And>x. x \<in> S \<Longrightarrow> f x = exp (g x)"  | 
|
4024  | 
by (metis continuous_logarithm_on_simply_connected [OF contf S fnz])  | 
|
4025  | 
have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \<in> S" for z  | 
|
4026  | 
proof -  | 
|
4027  | 
obtain f' where f': "((\<lambda>y. (f y - f z) / (y - z)) \<longlongrightarrow> f') (at z within S)"  | 
|
4028  | 
using \<open>f field_differentiable at z within S\<close> by (auto simp: field_differentiable_def DERIV_iff2)  | 
|
4029  | 
then have ee: "((\<lambda>x. (exp(g x) - exp(g z)) / (x - z)) \<longlongrightarrow> f') (at z within S)"  | 
|
4030  | 
by (simp add: feq \<open>z \<in> S\<close> Lim_transform_within [OF _ zero_less_one])  | 
|
4031  | 
have "(((\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<circ> g) \<longlongrightarrow> exp (g z))  | 
|
4032  | 
(at z within S)"  | 
|
4033  | 
proof (rule tendsto_compose_at)  | 
|
4034  | 
show "(g \<longlongrightarrow> g z) (at z within S)"  | 
|
4035  | 
using contg continuous_on \<open>z \<in> S\<close> by blast  | 
|
4036  | 
show "(\<lambda>y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \<midarrow>g z\<rightarrow> exp (g z)"  | 
|
4037  | 
apply (subst Lim_at_zero)  | 
|
4038  | 
apply (simp add: DERIV_D cong: if_cong Lim_cong_within)  | 
|
4039  | 
done  | 
|
4040  | 
qed auto  | 
|
4041  | 
then have dd: "((\<lambda>x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \<longlongrightarrow> exp(g z)) (at z within S)"  | 
|
4042  | 
by (simp add: o_def)  | 
|
4043  | 
have "continuous (at z within S) g"  | 
|
4044  | 
using contg continuous_on_eq_continuous_within \<open>z \<in> S\<close> by blast  | 
|
4045  | 
then have "(\<forall>\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)"  | 
|
4046  | 
by (simp add: continuous_within tendsto_iff)  | 
|
4047  | 
then have "\<forall>\<^sub>F x in at z within S. exp (g x) = exp (g z) \<longrightarrow> g x \<noteq> g z \<longrightarrow> x = z"  | 
|
4048  | 
apply (rule eventually_mono)  | 
|
4049  | 
apply (auto simp: exp_eq dist_norm norm_mult)  | 
|
4050  | 
done  | 
|
4051  | 
then have "((\<lambda>y. (g y - g z) / (y - z)) \<longlongrightarrow> f' / exp (g z)) (at z within S)"  | 
|
4052  | 
by (auto intro!: Lim_transform_eventually [OF _ tendsto_divide [OF ee dd]])  | 
|
4053  | 
then show ?thesis  | 
|
4054  | 
by (auto simp: field_differentiable_def DERIV_iff2)  | 
|
4055  | 
qed  | 
|
4056  | 
then have "g holomorphic_on S"  | 
|
4057  | 
using holf holomorphic_on_def by auto  | 
|
4058  | 
then show ?thesis  | 
|
4059  | 
using feq that by auto  | 
|
4060  | 
qed  | 
|
4061  | 
||
4062  | 
||
4063  | 
lemma contractible_imp_holomorphic_sqrt:  | 
|
4064  | 
assumes holf: "f holomorphic_on S"  | 
|
4065  | 
and S: "contractible S"  | 
|
4066  | 
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
4067  | 
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = g z ^ 2"  | 
|
4068  | 
proof -  | 
|
4069  | 
obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"  | 
|
4070  | 
using contractible_imp_holomorphic_log [OF assms] by blast  | 
|
4071  | 
show ?thesis  | 
|
4072  | 
proof  | 
|
4073  | 
show "exp \<circ> (\<lambda>z. z / 2) \<circ> g holomorphic_on S"  | 
|
4074  | 
by (intro holomorphic_on_compose holg holomorphic_intros) auto  | 
|
4075  | 
show "\<And>z. z \<in> S \<Longrightarrow> f z = ((exp \<circ> (\<lambda>z. z / 2) \<circ> g) z)\<^sup>2"  | 
|
4076  | 
apply (auto simp: feq)  | 
|
4077  | 
by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral)  | 
|
4078  | 
qed  | 
|
4079  | 
qed  | 
|
4080  | 
||
4081  | 
lemma simply_connected_imp_holomorphic_sqrt:  | 
|
4082  | 
assumes holf: "f holomorphic_on S"  | 
|
4083  | 
and S: "simply_connected S" "locally path_connected S"  | 
|
4084  | 
and fnz: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0"  | 
|
4085  | 
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = g z ^ 2"  | 
|
4086  | 
proof -  | 
|
4087  | 
obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = exp(g z)"  | 
|
4088  | 
using simply_connected_imp_holomorphic_log [OF assms] by blast  | 
|
4089  | 
show ?thesis  | 
|
4090  | 
proof  | 
|
4091  | 
show "exp \<circ> (\<lambda>z. z / 2) \<circ> g holomorphic_on S"  | 
|
4092  | 
by (intro holomorphic_on_compose holg holomorphic_intros) auto  | 
|
4093  | 
show "\<And>z. z \<in> S \<Longrightarrow> f z = ((exp \<circ> (\<lambda>z. z / 2) \<circ> g) z)\<^sup>2"  | 
|
4094  | 
apply (auto simp: feq)  | 
|
4095  | 
by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral)  | 
|
4096  | 
qed  | 
|
4097  | 
qed  | 
|
4098  | 
||
4099  | 
text\<open> Related theorems about holomorphic inverse cosines.\<close>  | 
|
4100  | 
||
4101  | 
lemma contractible_imp_holomorphic_arccos:  | 
|
4102  | 
assumes holf: "f holomorphic_on S" and S: "contractible S"  | 
|
4103  | 
and non1: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"  | 
|
4104  | 
obtains g where "g holomorphic_on S" "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"  | 
|
4105  | 
proof -  | 
|
4106  | 
have hol1f: "(\<lambda>z. 1 - f z ^ 2) holomorphic_on S"  | 
|
4107  | 
by (intro holomorphic_intros holf)  | 
|
4108  | 
obtain g where holg: "g holomorphic_on S" and eq: "\<And>z. z \<in> S \<Longrightarrow> 1 - (f z)\<^sup>2 = (g z)\<^sup>2"  | 
|
4109  | 
using contractible_imp_holomorphic_sqrt [OF hol1f S]  | 
|
4110  | 
by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff)  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4111  | 
have holfg: "(\<lambda>z. f z + \<i>*g z) holomorphic_on S"  | 
| 64845 | 4112  | 
by (intro holf holg holomorphic_intros)  | 
4113  | 
have "\<And>z. z \<in> S \<Longrightarrow> f z + \<i>*g z \<noteq> 0"  | 
|
4114  | 
by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff)  | 
|
4115  | 
then obtain h where holh: "h holomorphic_on S" and fgeq: "\<And>z. z \<in> S \<Longrightarrow> f z + \<i>*g z = exp (h z)"  | 
|
4116  | 
using contractible_imp_holomorphic_log [OF holfg S] by metis  | 
|
4117  | 
show ?thesis  | 
|
4118  | 
proof  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4119  | 
show "(\<lambda>z. -\<i>*h z) holomorphic_on S"  | 
| 64845 | 4120  | 
by (intro holh holomorphic_intros)  | 
4121  | 
show "f z = cos (- \<i>*h z)" if "z \<in> S" for z  | 
|
4122  | 
proof -  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4123  | 
have "(f z + \<i>*g z)*(f z - \<i>*g z) = 1"  | 
| 64845 | 4124  | 
using that eq by (auto simp: algebra_simps power2_eq_square)  | 
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4125  | 
then have "f z - \<i>*g z = inverse (f z + \<i>*g z)"  | 
| 64845 | 4126  | 
using inverse_unique by force  | 
4127  | 
also have "... = exp (- h z)"  | 
|
4128  | 
by (simp add: exp_minus fgeq that)  | 
|
4129  | 
finally have "f z = exp (- h z) + \<i>*g z"  | 
|
4130  | 
by (simp add: diff_eq_eq)  | 
|
4131  | 
then show ?thesis  | 
|
4132  | 
apply (simp add: cos_exp_eq)  | 
|
4133  | 
by (metis fgeq add.assoc mult_2_right that)  | 
|
4134  | 
qed  | 
|
4135  | 
qed  | 
|
4136  | 
qed  | 
|
4137  | 
||
4138  | 
||
4139  | 
lemma contractible_imp_holomorphic_arccos_bounded:  | 
|
4140  | 
assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \<in> S"  | 
|
4141  | 
and non1: "\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 1 \<and> f z \<noteq> -1"  | 
|
4142  | 
obtains g where "g holomorphic_on S" "norm(g a) \<le> pi + norm(f a)" "\<And>z. z \<in> S \<Longrightarrow> f z = cos(g z)"  | 
|
4143  | 
proof -  | 
|
4144  | 
obtain g where holg: "g holomorphic_on S" and feq: "\<And>z. z \<in> S \<Longrightarrow> f z = cos (g z)"  | 
|
4145  | 
using contractible_imp_holomorphic_arccos [OF holf S non1] by blast  | 
|
4146  | 
obtain b where "cos b = f a" "norm b \<le> pi + norm (f a)"  | 
|
4147  | 
using cos_Arccos norm_Arccos_bounded by blast  | 
|
4148  | 
then have "cos b = cos (g a)"  | 
|
4149  | 
by (simp add: \<open>a \<in> S\<close> feq)  | 
|
4150  | 
then consider n where "n \<in> \<int>" "b = g a + of_real(2*n*pi)" | n where "n \<in> \<int>" "b = -g a + of_real(2*n*pi)"  | 
|
4151  | 
by (auto simp: complex_cos_eq)  | 
|
4152  | 
then show ?thesis  | 
|
4153  | 
proof cases  | 
|
4154  | 
case 1  | 
|
4155  | 
show ?thesis  | 
|
4156  | 
proof  | 
|
4157  | 
show "(\<lambda>z. g z + of_real(2*n*pi)) holomorphic_on S"  | 
|
4158  | 
by (intro holomorphic_intros holg)  | 
|
4159  | 
show "cmod (g a + of_real(2*n*pi)) \<le> pi + cmod (f a)"  | 
|
4160  | 
using "1" \<open>cmod b \<le> pi + cmod (f a)\<close> by blast  | 
|
4161  | 
show "\<And>z. z \<in> S \<Longrightarrow> f z = cos (g z + complex_of_real (2*n*pi))"  | 
|
4162  | 
by (metis \<open>n \<in> \<int>\<close> complex_cos_eq feq)  | 
|
4163  | 
qed  | 
|
4164  | 
next  | 
|
4165  | 
case 2  | 
|
4166  | 
show ?thesis  | 
|
4167  | 
proof  | 
|
4168  | 
show "(\<lambda>z. -g z + of_real(2*n*pi)) holomorphic_on S"  | 
|
4169  | 
by (intro holomorphic_intros holg)  | 
|
4170  | 
show "cmod (-g a + of_real(2*n*pi)) \<le> pi + cmod (f a)"  | 
|
4171  | 
using "2" \<open>cmod b \<le> pi + cmod (f a)\<close> by blast  | 
|
4172  | 
show "\<And>z. z \<in> S \<Longrightarrow> f z = cos (-g z + complex_of_real (2*n*pi))"  | 
|
4173  | 
by (metis \<open>n \<in> \<int>\<close> complex_cos_eq feq)  | 
|
4174  | 
qed  | 
|
4175  | 
qed  | 
|
4176  | 
qed  | 
|
4177  | 
||
4178  | 
||
4179  | 
subsection\<open>The "Borsukian" property of sets\<close>  | 
|
4180  | 
||
| 64847 | 4181  | 
text\<open>This doesn't have a standard name. Kuratowski uses ``contractible with respect to $[S^1]$''  | 
4182  | 
while Whyburn uses ``property b''. It's closely related to unicoherence.\<close>  | 
|
| 64845 | 4183  | 
|
4184  | 
definition Borsukian where  | 
|
4185  | 
"Borsukian S \<equiv>  | 
|
4186  | 
        \<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
 | 
|
4187  | 
            \<longrightarrow> (\<exists>a. homotopic_with (\<lambda>h. True) S (- {0}) f (\<lambda>x. a))"
 | 
|
4188  | 
||
4189  | 
lemma Borsukian_retraction_gen:  | 
|
4190  | 
assumes "Borsukian S" "continuous_on S h" "h ` S = T"  | 
|
4191  | 
"continuous_on T k" "k ` T \<subseteq> S" "\<And>y. y \<in> T \<Longrightarrow> h(k y) = y"  | 
|
4192  | 
shows "Borsukian T"  | 
|
4193  | 
proof -  | 
|
4194  | 
interpret R: Retracts S h T k  | 
|
4195  | 
using assms by (simp add: Retracts.intro)  | 
|
4196  | 
show ?thesis  | 
|
4197  | 
using assms  | 
|
4198  | 
apply (simp add: Borsukian_def, clarify)  | 
|
4199  | 
    apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto)
 | 
|
4200  | 
done  | 
|
4201  | 
qed  | 
|
4202  | 
||
4203  | 
lemma retract_of_Borsukian: "\<lbrakk>Borsukian T; S retract_of T\<rbrakk> \<Longrightarrow> Borsukian S"  | 
|
4204  | 
apply (auto simp: retract_of_def retraction_def)  | 
|
4205  | 
apply (erule (1) Borsukian_retraction_gen)  | 
|
4206  | 
apply (meson retraction retraction_def)  | 
|
4207  | 
apply (auto simp: continuous_on_id)  | 
|
4208  | 
done  | 
|
4209  | 
||
4210  | 
lemma homeomorphic_Borsukian: "\<lbrakk>Borsukian S; S homeomorphic T\<rbrakk> \<Longrightarrow> Borsukian T"  | 
|
4211  | 
using Borsukian_retraction_gen order_refl  | 
|
4212  | 
by (fastforce simp add: homeomorphism_def homeomorphic_def)  | 
|
4213  | 
||
4214  | 
lemma homeomorphic_Borsukian_eq:  | 
|
4215  | 
"S homeomorphic T \<Longrightarrow> Borsukian S \<longleftrightarrow> Borsukian T"  | 
|
4216  | 
by (meson homeomorphic_Borsukian homeomorphic_sym)  | 
|
4217  | 
||
4218  | 
lemma Borsukian_translation:  | 
|
4219  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4220  | 
shows "Borsukian (image (\<lambda>x. a + x) S) \<longleftrightarrow> Borsukian S"  | 
|
4221  | 
apply (rule homeomorphic_Borsukian_eq)  | 
|
4222  | 
using homeomorphic_translation homeomorphic_sym by blast  | 
|
4223  | 
||
4224  | 
lemma Borsukian_injective_linear_image:  | 
|
4225  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
4226  | 
assumes "linear f" "inj f"  | 
|
4227  | 
shows "Borsukian(f ` S) \<longleftrightarrow> Borsukian S"  | 
|
4228  | 
apply (rule homeomorphic_Borsukian_eq)  | 
|
4229  | 
using assms homeomorphic_sym linear_homeomorphic_image by blast  | 
|
4230  | 
||
4231  | 
lemma homotopy_eqv_Borsukianness:  | 
|
4232  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4233  | 
and T :: "'b::real_normed_vector set"  | 
|
4234  | 
assumes "S homotopy_eqv T"  | 
|
4235  | 
shows "(Borsukian S \<longleftrightarrow> Borsukian T)"  | 
|
4236  | 
by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null)  | 
|
4237  | 
||
4238  | 
lemma Borsukian_alt:  | 
|
4239  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4240  | 
shows  | 
|
4241  | 
"Borsukian S \<longleftrightarrow>  | 
|
4242  | 
        (\<forall>f g. continuous_on S f \<and> f ` S \<subseteq> -{0} \<and>
 | 
|
4243  | 
               continuous_on S g \<and> g ` S \<subseteq> -{0}
 | 
|
4244  | 
               \<longrightarrow> homotopic_with (\<lambda>h. True) S (- {0::complex}) f g)"
 | 
|
4245  | 
unfolding Borsukian_def homotopic_triviality  | 
|
4246  | 
by (simp add: path_connected_punctured_universe)  | 
|
4247  | 
||
4248  | 
lemma Borsukian_continuous_logarithm:  | 
|
4249  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4250  | 
shows "Borsukian S \<longleftrightarrow>  | 
|
4251  | 
            (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> (- {0::complex})
 | 
|
4252  | 
\<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))))"  | 
|
4253  | 
by (simp add: Borsukian_def inessential_eq_continuous_logarithm)  | 
|
4254  | 
||
4255  | 
lemma Borsukian_continuous_logarithm_circle:  | 
|
4256  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4257  | 
shows "Borsukian S \<longleftrightarrow>  | 
|
4258  | 
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1  | 
|
4259  | 
\<longrightarrow> (\<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))))"  | 
|
4260  | 
(is "?lhs = ?rhs")  | 
|
4261  | 
proof  | 
|
4262  | 
assume ?lhs then show ?rhs  | 
|
4263  | 
by (force simp: Borsukian_continuous_logarithm)  | 
|
4264  | 
next  | 
|
4265  | 
assume RHS [rule_format]: ?rhs  | 
|
4266  | 
show ?lhs  | 
|
4267  | 
proof (clarsimp simp: Borsukian_continuous_logarithm)  | 
|
4268  | 
fix f :: "'a \<Rightarrow> complex"  | 
|
4269  | 
assume contf: "continuous_on S f" and 0: "0 \<notin> f ` S"  | 
|
4270  | 
then have "continuous_on S (\<lambda>x. f x / complex_of_real (cmod (f x)))"  | 
|
4271  | 
by (intro continuous_intros) auto  | 
|
4272  | 
moreover have "(\<lambda>x. f x / complex_of_real (cmod (f x))) ` S \<subseteq> sphere 0 1"  | 
|
4273  | 
using 0 by (auto simp: norm_divide)  | 
|
4274  | 
ultimately obtain g where contg: "continuous_on S g"  | 
|
4275  | 
and fg: "\<forall>x \<in> S. f x / complex_of_real (cmod (f x)) = exp(g x)"  | 
|
4276  | 
using RHS [of "\<lambda>x. f x / of_real(norm(f x))"] by auto  | 
|
4277  | 
show "\<exists>g. continuous_on S g \<and> (\<forall>x\<in>S. f x = exp (g x))"  | 
|
4278  | 
proof (intro exI ballI conjI)  | 
|
4279  | 
show "continuous_on S (\<lambda>x. (Ln \<circ> of_real \<circ> norm \<circ> f)x + g x)"  | 
|
4280  | 
by (intro continuous_intros contf contg conjI) (use "0" in auto)  | 
|
4281  | 
show "f x = exp ((Ln \<circ> complex_of_real \<circ> cmod \<circ> f) x + g x)" if "x \<in> S" for x  | 
|
4282  | 
using 0 that  | 
|
4283  | 
apply (clarsimp simp: exp_add)  | 
|
4284  | 
apply (subst exp_Ln, force)  | 
|
4285  | 
by (metis eq_divide_eq exp_not_eq_zero fg mult.commute)  | 
|
4286  | 
qed  | 
|
4287  | 
qed  | 
|
4288  | 
qed  | 
|
4289  | 
||
4290  | 
||
4291  | 
lemma Borsukian_continuous_logarithm_circle_real:  | 
|
4292  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4293  | 
shows "Borsukian S \<longleftrightarrow>  | 
|
4294  | 
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4295  | 
\<longrightarrow> (\<exists>g. continuous_on S (complex_of_real \<circ> g) \<and> (\<forall>x \<in> S. f x = exp(\<i> * of_real(g x)))))"  | 
| 64845 | 4296  | 
(is "?lhs = ?rhs")  | 
4297  | 
proof  | 
|
4298  | 
assume LHS: ?lhs  | 
|
4299  | 
show ?rhs  | 
|
4300  | 
proof (clarify)  | 
|
4301  | 
fix f :: "'a \<Rightarrow> complex"  | 
|
4302  | 
assume "continuous_on S f" and f01: "f ` S \<subseteq> sphere 0 1"  | 
|
4303  | 
then obtain g where contg: "continuous_on S g" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
4304  | 
using LHS by (auto simp: Borsukian_continuous_logarithm_circle)  | 
|
4305  | 
then have "\<forall>x\<in>S. f x = exp (\<i> * complex_of_real ((Im \<circ> g) x))"  | 
|
4306  | 
using f01 apply (simp add: image_iff subset_iff)  | 
|
4307  | 
by (metis cis_conv_exp exp_eq_polar mult.left_neutral norm_exp_eq_Re of_real_1)  | 
|
4308  | 
then show "\<exists>g. continuous_on S (complex_of_real \<circ> g) \<and> (\<forall>x\<in>S. f x = exp (\<i> * complex_of_real (g x)))"  | 
|
4309  | 
by (rule_tac x="Im \<circ> g" in exI) (force intro: continuous_intros contg)  | 
|
4310  | 
qed  | 
|
4311  | 
next  | 
|
4312  | 
assume RHS [rule_format]: ?rhs  | 
|
4313  | 
show ?lhs  | 
|
4314  | 
proof (clarsimp simp: Borsukian_continuous_logarithm_circle)  | 
|
4315  | 
fix f :: "'a \<Rightarrow> complex"  | 
|
4316  | 
assume "continuous_on S f" and f01: "f ` S \<subseteq> sphere 0 1"  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4317  | 
then obtain g where contg: "continuous_on S (complex_of_real \<circ> g)" and "\<And>x. x \<in> S \<Longrightarrow> f x = exp(\<i> * of_real(g x))"  | 
| 64845 | 4318  | 
by (metis RHS)  | 
4319  | 
then show "\<exists>g. continuous_on S g \<and> (\<forall>x\<in>S. f x = exp (g x))"  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4320  | 
by (rule_tac x="\<lambda>x. \<i>* of_real(g x)" in exI) (auto simp: continuous_intros contg)  | 
| 64845 | 4321  | 
qed  | 
4322  | 
qed  | 
|
4323  | 
||
4324  | 
lemma Borsukian_circle:  | 
|
4325  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4326  | 
shows "Borsukian S \<longleftrightarrow>  | 
|
4327  | 
(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> sphere (0::complex) 1  | 
|
4328  | 
\<longrightarrow> (\<exists>a. homotopic_with (\<lambda>h. True) S (sphere (0::complex) 1) f (\<lambda>x. a)))"  | 
|
4329  | 
by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real)  | 
|
4330  | 
||
4331  | 
lemma contractible_imp_Borsukian: "contractible S \<Longrightarrow> Borsukian S"  | 
|
4332  | 
by (meson Borsukian_def nullhomotopic_from_contractible)  | 
|
4333  | 
||
4334  | 
lemma simply_connected_imp_Borsukian:  | 
|
4335  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4336  | 
shows "\<lbrakk>simply_connected S; locally path_connected S\<rbrakk> \<Longrightarrow> Borsukian S"  | 
|
4337  | 
apply (simp add: Borsukian_continuous_logarithm)  | 
|
4338  | 
by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff)  | 
|
4339  | 
||
4340  | 
lemma starlike_imp_Borsukian:  | 
|
4341  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4342  | 
shows "starlike S \<Longrightarrow> Borsukian S"  | 
|
4343  | 
by (simp add: contractible_imp_Borsukian starlike_imp_contractible)  | 
|
4344  | 
||
4345  | 
lemma Borsukian_empty: "Borsukian {}"
 | 
|
4346  | 
by (auto simp: contractible_imp_Borsukian)  | 
|
4347  | 
||
4348  | 
lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)"  | 
|
4349  | 
by (auto simp: contractible_imp_Borsukian)  | 
|
4350  | 
||
4351  | 
lemma convex_imp_Borsukian:  | 
|
4352  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4353  | 
shows "convex S \<Longrightarrow> Borsukian S"  | 
|
4354  | 
by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible)  | 
|
4355  | 
||
4356  | 
lemma Borsukian_sphere:  | 
|
4357  | 
fixes a :: "'a::euclidean_space"  | 
|
4358  | 
  shows "3 \<le> DIM('a) \<Longrightarrow> Borsukian (sphere a r)"
 | 
|
4359  | 
apply (rule simply_connected_imp_Borsukian)  | 
|
4360  | 
using simply_connected_sphere apply blast  | 
|
4361  | 
using ENR_imp_locally_path_connected ENR_sphere by blast  | 
|
4362  | 
||
4363  | 
lemma Borsukian_open_Un:  | 
|
4364  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4365  | 
assumes opeS: "openin (subtopology euclidean (S \<union> T)) S"  | 
|
4366  | 
and opeT: "openin (subtopology euclidean (S \<union> T)) T"  | 
|
4367  | 
and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"  | 
|
4368  | 
shows "Borsukian(S \<union> T)"  | 
|
4369  | 
proof (clarsimp simp add: Borsukian_continuous_logarithm)  | 
|
4370  | 
fix f :: "'a \<Rightarrow> complex"  | 
|
4371  | 
assume contf: "continuous_on (S \<union> T) f" and 0: "0 \<notin> f ` (S \<union> T)"  | 
|
4372  | 
then have contfS: "continuous_on S f" and contfT: "continuous_on T f"  | 
|
4373  | 
using continuous_on_subset by auto  | 
|
4374  | 
  have "\<lbrakk>continuous_on S f; f ` S \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))"
 | 
|
4375  | 
using BS by (auto simp: Borsukian_continuous_logarithm)  | 
|
4376  | 
then obtain g where contg: "continuous_on S g" and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
4377  | 
using "0" contfS by blast  | 
|
4378  | 
  have "\<lbrakk>continuous_on T f; f ` T \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on T g \<and> (\<forall>x \<in> T. f x = exp(g x))"
 | 
|
4379  | 
using BT by (auto simp: Borsukian_continuous_logarithm)  | 
|
4380  | 
then obtain h where conth: "continuous_on T h" and fh: "\<And>x. x \<in> T \<Longrightarrow> f x = exp(h x)"  | 
|
4381  | 
using "0" contfT by blast  | 
|
4382  | 
show "\<exists>g. continuous_on (S \<union> T) g \<and> (\<forall>x\<in>S \<union> T. f x = exp (g x))"  | 
|
4383  | 
  proof (cases "S \<inter> T = {}")
 | 
|
4384  | 
case True  | 
|
4385  | 
show ?thesis  | 
|
4386  | 
proof (intro exI conjI)  | 
|
4387  | 
show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else h x)"  | 
|
4388  | 
apply (rule continuous_on_cases_local_open [OF opeS opeT contg conth])  | 
|
4389  | 
using True by blast  | 
|
4390  | 
show "\<forall>x\<in>S \<union> T. f x = exp (if x \<in> S then g x else h x)"  | 
|
4391  | 
using fg fh by auto  | 
|
4392  | 
qed  | 
|
4393  | 
next  | 
|
4394  | 
case False  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4395  | 
have "(\<lambda>x. g x - h x) constant_on S \<inter> T"  | 
| 64845 | 4396  | 
proof (rule continuous_discrete_range_constant [OF ST])  | 
4397  | 
show "continuous_on (S \<inter> T) (\<lambda>x. g x - h x)"  | 
|
4398  | 
apply (intro continuous_intros)  | 
|
4399  | 
apply (meson contg continuous_on_subset inf_le1)  | 
|
4400  | 
by (meson conth continuous_on_subset inf_sup_ord(2))  | 
|
4401  | 
show "\<exists>e>0. \<forall>y. y \<in> S \<inter> T \<and> g y - h y \<noteq> g x - h x \<longrightarrow> e \<le> cmod (g y - h y - (g x - h x))"  | 
|
4402  | 
if "x \<in> S \<inter> T" for x  | 
|
4403  | 
proof -  | 
|
4404  | 
have "g y - g x = h y - h x"  | 
|
4405  | 
if "y \<in> S" "y \<in> T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y  | 
|
4406  | 
proof (rule exp_complex_eqI)  | 
|
4407  | 
have "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> \<le> cmod (g y - g x - (h y - h x))"  | 
|
4408  | 
by (metis abs_Im_le_cmod minus_complex.simps(2))  | 
|
4409  | 
then show "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> < 2 * pi"  | 
|
4410  | 
using that by linarith  | 
|
4411  | 
have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)"  | 
|
4412  | 
using fg fh that \<open>x \<in> S \<inter> T\<close> by fastforce+  | 
|
4413  | 
then show "exp (g y - g x) = exp (h y - h x)"  | 
|
4414  | 
by (simp add: exp_diff)  | 
|
4415  | 
qed  | 
|
4416  | 
then show ?thesis  | 
|
4417  | 
by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps)  | 
|
4418  | 
qed  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4419  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4420  | 
then obtain a where a: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x - h x = a"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4421  | 
by (auto simp: constant_on_def)  | 
| 64845 | 4422  | 
with False have "exp a = 1"  | 
4423  | 
by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh)  | 
|
4424  | 
with a show ?thesis  | 
|
4425  | 
apply (rule_tac x="\<lambda>x. if x \<in> S then g x else a + h x" in exI)  | 
|
4426  | 
apply (intro continuous_on_cases_local_open opeS opeT contg conth continuous_intros conjI)  | 
|
4427  | 
apply (auto simp: algebra_simps fg fh exp_add)  | 
|
| 
65037
 
2cf841ff23be
some new material, also recasting some theorems using “obtains”
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
4428  | 
done  | 
| 64845 | 4429  | 
qed  | 
4430  | 
qed  | 
|
4431  | 
||
| 64911 | 4432  | 
text\<open>The proof is a duplicate of that of \<open>Borsukian_open_Un\<close>.\<close>  | 
| 64845 | 4433  | 
lemma Borsukian_closed_Un:  | 
4434  | 
fixes S :: "'a::real_normed_vector set"  | 
|
4435  | 
assumes cloS: "closedin (subtopology euclidean (S \<union> T)) S"  | 
|
4436  | 
and cloT: "closedin (subtopology euclidean (S \<union> T)) T"  | 
|
4437  | 
and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \<inter> T)"  | 
|
4438  | 
shows "Borsukian(S \<union> T)"  | 
|
4439  | 
proof (clarsimp simp add: Borsukian_continuous_logarithm)  | 
|
4440  | 
fix f :: "'a \<Rightarrow> complex"  | 
|
4441  | 
assume contf: "continuous_on (S \<union> T) f" and 0: "0 \<notin> f ` (S \<union> T)"  | 
|
4442  | 
then have contfS: "continuous_on S f" and contfT: "continuous_on T f"  | 
|
4443  | 
using continuous_on_subset by auto  | 
|
4444  | 
  have "\<lbrakk>continuous_on S f; f ` S \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on S g \<and> (\<forall>x \<in> S. f x = exp(g x))"
 | 
|
4445  | 
using BS by (auto simp: Borsukian_continuous_logarithm)  | 
|
4446  | 
then obtain g where contg: "continuous_on S g" and fg: "\<And>x. x \<in> S \<Longrightarrow> f x = exp(g x)"  | 
|
4447  | 
using "0" contfS by blast  | 
|
4448  | 
  have "\<lbrakk>continuous_on T f; f ` T \<subseteq> -{0}\<rbrakk> \<Longrightarrow> \<exists>g. continuous_on T g \<and> (\<forall>x \<in> T. f x = exp(g x))"
 | 
|
4449  | 
using BT by (auto simp: Borsukian_continuous_logarithm)  | 
|
4450  | 
then obtain h where conth: "continuous_on T h" and fh: "\<And>x. x \<in> T \<Longrightarrow> f x = exp(h x)"  | 
|
4451  | 
using "0" contfT by blast  | 
|
4452  | 
show "\<exists>g. continuous_on (S \<union> T) g \<and> (\<forall>x\<in>S \<union> T. f x = exp (g x))"  | 
|
4453  | 
  proof (cases "S \<inter> T = {}")
 | 
|
4454  | 
case True  | 
|
4455  | 
show ?thesis  | 
|
4456  | 
proof (intro exI conjI)  | 
|
4457  | 
show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then g x else h x)"  | 
|
4458  | 
apply (rule continuous_on_cases_local [OF cloS cloT contg conth])  | 
|
4459  | 
using True by blast  | 
|
4460  | 
show "\<forall>x\<in>S \<union> T. f x = exp (if x \<in> S then g x else h x)"  | 
|
4461  | 
using fg fh by auto  | 
|
4462  | 
qed  | 
|
4463  | 
next  | 
|
4464  | 
case False  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4465  | 
have "(\<lambda>x. g x - h x) constant_on S \<inter> T"  | 
| 64845 | 4466  | 
proof (rule continuous_discrete_range_constant [OF ST])  | 
4467  | 
show "continuous_on (S \<inter> T) (\<lambda>x. g x - h x)"  | 
|
4468  | 
apply (intro continuous_intros)  | 
|
4469  | 
apply (meson contg continuous_on_subset inf_le1)  | 
|
4470  | 
by (meson conth continuous_on_subset inf_sup_ord(2))  | 
|
4471  | 
show "\<exists>e>0. \<forall>y. y \<in> S \<inter> T \<and> g y - h y \<noteq> g x - h x \<longrightarrow> e \<le> cmod (g y - h y - (g x - h x))"  | 
|
4472  | 
if "x \<in> S \<inter> T" for x  | 
|
4473  | 
proof -  | 
|
4474  | 
have "g y - g x = h y - h x"  | 
|
4475  | 
if "y \<in> S" "y \<in> T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y  | 
|
4476  | 
proof (rule exp_complex_eqI)  | 
|
4477  | 
have "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> \<le> cmod (g y - g x - (h y - h x))"  | 
|
4478  | 
by (metis abs_Im_le_cmod minus_complex.simps(2))  | 
|
4479  | 
then show "\<bar>Im (g y - g x) - Im (h y - h x)\<bar> < 2 * pi"  | 
|
4480  | 
using that by linarith  | 
|
4481  | 
have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)"  | 
|
4482  | 
using fg fh that \<open>x \<in> S \<inter> T\<close> by fastforce+  | 
|
4483  | 
then show "exp (g y - g x) = exp (h y - h x)"  | 
|
4484  | 
by (simp add: exp_diff)  | 
|
4485  | 
qed  | 
|
4486  | 
then show ?thesis  | 
|
4487  | 
by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps)  | 
|
4488  | 
qed  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4489  | 
qed  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4490  | 
then obtain a where a: "\<And>x. x \<in> S \<inter> T \<Longrightarrow> g x - h x = a"  | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4491  | 
by (auto simp: constant_on_def)  | 
| 64845 | 4492  | 
with False have "exp a = 1"  | 
4493  | 
by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh)  | 
|
4494  | 
with a show ?thesis  | 
|
4495  | 
apply (rule_tac x="\<lambda>x. if x \<in> S then g x else a + h x" in exI)  | 
|
4496  | 
apply (intro continuous_on_cases_local cloS cloT contg conth continuous_intros conjI)  | 
|
4497  | 
apply (auto simp: algebra_simps fg fh exp_add)  | 
|
| 
65037
 
2cf841ff23be
some new material, also recasting some theorems using “obtains”
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
4498  | 
done  | 
| 64845 | 4499  | 
qed  | 
4500  | 
qed  | 
|
4501  | 
||
4502  | 
lemma Borsukian_separation_compact:  | 
|
4503  | 
fixes S :: "complex set"  | 
|
4504  | 
assumes "compact S"  | 
|
4505  | 
shows "Borsukian S \<longleftrightarrow> connected(- S)"  | 
|
4506  | 
by (simp add: Borsuk_separation_theorem Borsukian_circle assms)  | 
|
4507  | 
||
4508  | 
lemma Borsukian_monotone_image_compact:  | 
|
4509  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
4510  | 
assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T"  | 
|
4511  | 
      and "compact S" and conn: "\<And>y. y \<in> T \<Longrightarrow> connected {x. x \<in> S \<and> f x = y}"
 | 
|
4512  | 
shows "Borsukian T"  | 
|
4513  | 
proof (clarsimp simp add: Borsukian_continuous_logarithm)  | 
|
4514  | 
fix g :: "'b \<Rightarrow> complex"  | 
|
4515  | 
assume contg: "continuous_on T g" and 0: "0 \<notin> g ` T"  | 
|
4516  | 
have "continuous_on S (g \<circ> f)"  | 
|
4517  | 
using contf contg continuous_on_compose fim by blast  | 
|
4518  | 
  moreover have "(g \<circ> f) ` S \<subseteq> -{0}"
 | 
|
4519  | 
using fim 0 by auto  | 
|
4520  | 
ultimately obtain h where conth: "continuous_on S h" and gfh: "\<And>x. x \<in> S \<Longrightarrow> (g \<circ> f) x = exp(h x)"  | 
|
4521  | 
using \<open>Borsukian S\<close> by (auto simp: Borsukian_continuous_logarithm)  | 
|
4522  | 
have "\<And>y. \<exists>x. y \<in> T \<longrightarrow> x \<in> S \<and> f x = y"  | 
|
4523  | 
using fim by auto  | 
|
4524  | 
then obtain f' where f': "\<And>y. y \<in> T \<longrightarrow> f' y \<in> S \<and> f (f' y) = y"  | 
|
4525  | 
by metis  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4526  | 
  have *: "(\<lambda>x. h x - h(f' y)) constant_on {x. x \<in> S \<and> f x = y}" if "y \<in> T" for y
 | 
| 
65037
 
2cf841ff23be
some new material, also recasting some theorems using “obtains”
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
4527  | 
proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\<lambda>x. h x - h (f' y)"], simp_all add: algebra_simps)  | 
| 
 
2cf841ff23be
some new material, also recasting some theorems using “obtains”
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
4528  | 
    show "continuous_on {x \<in> S. f x = y} (\<lambda>x. h x - h (f' y))"
 | 
| 
 
2cf841ff23be
some new material, also recasting some theorems using “obtains”
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
4529  | 
by (intro continuous_intros continuous_on_subset [OF conth]) auto  | 
| 64845 | 4530  | 
show "\<exists>e>0. \<forall>u. u \<in> S \<and> f u = y \<and> h u \<noteq> h x \<longrightarrow> e \<le> cmod (h u - h x)"  | 
4531  | 
if x: "x \<in> S \<and> f x = y" for x  | 
|
4532  | 
proof -  | 
|
4533  | 
have "h u = h x" if "u \<in> S" "f u = y" "cmod (h u - h x) < 2 * pi" for u  | 
|
4534  | 
proof (rule exp_complex_eqI)  | 
|
4535  | 
have "\<bar>Im (h u) - Im (h x)\<bar> \<le> cmod (h u - h x)"  | 
|
4536  | 
by (metis abs_Im_le_cmod minus_complex.simps(2))  | 
|
4537  | 
then show "\<bar>Im (h u) - Im (h x)\<bar> < 2 * pi"  | 
|
4538  | 
using that by linarith  | 
|
4539  | 
show "exp (h u) = exp (h x)"  | 
|
4540  | 
by (simp add: gfh [symmetric] x that)  | 
|
4541  | 
qed  | 
|
4542  | 
then show ?thesis  | 
|
4543  | 
by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps)  | 
|
4544  | 
qed  | 
|
| 
65037
 
2cf841ff23be
some new material, also recasting some theorems using “obtains”
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
4545  | 
qed  | 
| 64845 | 4546  | 
have "h x = h (f' (f x))" if "x \<in> S" for x  | 
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4547  | 
using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq)  | 
| 64845 | 4548  | 
moreover have "\<And>x. x \<in> T \<Longrightarrow> \<exists>u. u \<in> S \<and> x = f u \<and> h (f' x) = h u"  | 
4549  | 
using f' by fastforce  | 
|
4550  | 
ultimately  | 
|
4551  | 
have eq: "((\<lambda>x. (x, (h \<circ> f') x)) ` T) =  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4552  | 
            {p. \<exists>x. x \<in> S \<and> (x, p) \<in> (S \<times> UNIV) \<inter> ((\<lambda>z. snd z - ((f \<circ> fst) z, (h \<circ> fst) z)) -` {0})}"
 | 
| 64845 | 4553  | 
using fim by (auto simp: image_iff)  | 
4554  | 
show "\<exists>h. continuous_on T h \<and> (\<forall>x\<in>T. g x = exp (h x))"  | 
|
4555  | 
proof (intro exI conjI)  | 
|
4556  | 
show "continuous_on T (h \<circ> f')"  | 
|
4557  | 
proof (rule continuous_from_closed_graph [of "h ` S"])  | 
|
4558  | 
show "compact (h ` S)"  | 
|
4559  | 
by (simp add: \<open>compact S\<close> compact_continuous_image conth)  | 
|
4560  | 
show "(h \<circ> f') ` T \<subseteq> h ` S"  | 
|
4561  | 
by (auto simp: f')  | 
|
4562  | 
show "closed ((\<lambda>x. (x, (h \<circ> f') x)) ` T)"  | 
|
4563  | 
apply (subst eq)  | 
|
4564  | 
apply (intro closed_compact_projection [OF \<open>compact S\<close>] continuous_closed_preimage  | 
|
4565  | 
continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth])  | 
|
4566  | 
apply (auto simp: \<open>compact S\<close> closed_Times compact_imp_closed)  | 
|
4567  | 
done  | 
|
4568  | 
qed  | 
|
4569  | 
qed (use f' gfh in fastforce)  | 
|
4570  | 
qed  | 
|
4571  | 
||
4572  | 
||
4573  | 
lemma Borsukian_open_map_image_compact:  | 
|
4574  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
4575  | 
assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S"  | 
|
4576  | 
and ope: "\<And>U. openin (subtopology euclidean S) U  | 
|
4577  | 
\<Longrightarrow> openin (subtopology euclidean T) (f ` U)"  | 
|
4578  | 
shows "Borsukian T"  | 
|
4579  | 
proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real)  | 
|
4580  | 
fix g :: "'b \<Rightarrow> complex"  | 
|
4581  | 
assume contg: "continuous_on T g" and gim: "g ` T \<subseteq> sphere 0 1"  | 
|
4582  | 
have "continuous_on S (g \<circ> f)"  | 
|
4583  | 
using contf contg continuous_on_compose fim by blast  | 
|
4584  | 
moreover have "(g \<circ> f) ` S \<subseteq> sphere 0 1"  | 
|
4585  | 
using fim gim by auto  | 
|
4586  | 
ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \<circ> h)"  | 
|
| 
65064
 
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name.  Also replaced some legacy def commands
 
paulson <lp15@cam.ac.uk> 
parents: 
65037 
diff
changeset
 | 
4587  | 
and gfh: "\<And>x. x \<in> S \<Longrightarrow> (g \<circ> f) x = exp(\<i> * of_real(h x))"  | 
| 64845 | 4588  | 
using \<open>Borsukian S\<close> Borsukian_continuous_logarithm_circle_real by metis  | 
4589  | 
then have conth: "continuous_on S h"  | 
|
4590  | 
by simp  | 
|
4591  | 
have "\<exists>x. x \<in> S \<and> f x = y \<and> (\<forall>x' \<in> S. f x' = y \<longrightarrow> h x \<le> h x')" if "y \<in> T" for y  | 
|
4592  | 
proof -  | 
|
4593  | 
    have 1: "compact (h ` {x \<in> S. f x = y})"
 | 
|
4594  | 
proof (rule compact_continuous_image)  | 
|
4595  | 
      show "continuous_on {x \<in> S. f x = y} h"
 | 
|
4596  | 
by (rule continuous_on_subset [OF conth]) auto  | 
|
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4597  | 
      have "compact (S \<inter> f -` {y})"
 | 
| 64845 | 4598  | 
by (rule proper_map_from_compact [OF contf _ \<open>compact S\<close>, of T]) (simp_all add: fim that)  | 
| 
66884
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4599  | 
      then show "compact {x \<in> S. f x = y}" 
 | 
| 
 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 
paulson <lp15@cam.ac.uk> 
parents: 
66827 
diff
changeset
 | 
4600  | 
by (auto simp: vimage_def Int_def)  | 
| 64845 | 4601  | 
qed  | 
4602  | 
    have 2: "h ` {x \<in> S. f x = y} \<noteq> {}"
 | 
|
4603  | 
using fim that by auto  | 
|
4604  | 
    have "\<exists>s \<in> h ` {x \<in> S. f x = y}. \<forall>t \<in> h ` {x \<in> S. f x = y}. s \<le> t"
 | 
|
4605  | 
using compact_attains_inf [OF 1 2] by blast  | 
|
4606  | 
then show ?thesis by auto  | 
|
4607  | 
qed  | 
|
4608  | 
then obtain k where kTS: "\<And>y. y \<in> T \<Longrightarrow> k y \<in> S"  | 
|
4609  | 
and fk: "\<And>y. y \<in> T \<Longrightarrow> f (k y) = y "  | 
|
4610  | 
and hle: "\<And>x' y. \<lbrakk>y \<in> T; x' \<in> S; f x' = y\<rbrakk> \<Longrightarrow> h (k y) \<le> h x'"  | 
|
4611  | 
by metis  | 
|
4612  | 
have "continuous_on T (h \<circ> k)"  | 
|
4613  | 
proof (clarsimp simp add: continuous_on_iff)  | 
|
4614  | 
fix y and e::real  | 
|
4615  | 
assume "y \<in> T" "0 < e"  | 
|
4616  | 
moreover have "uniformly_continuous_on S (complex_of_real \<circ> h)"  | 
|
4617  | 
using \<open>compact S\<close> cont_cxh compact_uniformly_continuous by blast  | 
|
4618  | 
ultimately obtain d where "0 < d"  | 
|
4619  | 
and d: "\<And>x x'. \<lbrakk>x\<in>S; x'\<in>S; dist x' x < d\<rbrakk> \<Longrightarrow> dist (h x') (h x) < e"  | 
|
4620  | 
by (force simp: uniformly_continuous_on_def)  | 
|
4621  | 
obtain \<delta> where "0 < \<delta>" and \<delta>:  | 
|
4622  | 
"\<And>x'. \<lbrakk>x' \<in> T; dist y x' < \<delta>\<rbrakk>  | 
|
4623  | 
               \<Longrightarrow> (\<forall>v \<in> {z \<in> S. f z = y}. \<exists>v'. v' \<in> {z \<in> S. f z = x'} \<and> dist v v' < d) \<and>
 | 
|
4624  | 
                   (\<forall>v' \<in> {z \<in> S. f z = x'}. \<exists>v. v \<in> {z \<in> S. f z = y} \<and> dist v' v < d)"
 | 
|
4625  | 
    proof (rule upper_lower_hemicontinuous_explicit [of T "\<lambda>y. {z \<in> S. f z = y}" S])
 | 
|
4626  | 
show "\<And>U. openin (subtopology euclidean S) U  | 
|
4627  | 
                 \<Longrightarrow> openin (subtopology euclidean T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
 | 
|
4628  | 
using continuous_imp_closed_map closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]]  | 
|
4629  | 
by (simp add: continuous_imp_closed_map \<open>compact S\<close> contf fim)  | 
|
4630  | 
show "\<And>U. closedin (subtopology euclidean S) U \<Longrightarrow>  | 
|
4631  | 
                 closedin (subtopology euclidean T) {x \<in> T. {z \<in> S. f z = x} \<subseteq> U}"
 | 
|
4632  | 
using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]]  | 
|
4633  | 
by meson  | 
|
4634  | 
      show "bounded {z \<in> S. f z = y}"
 | 
|
4635  | 
by (metis (no_types, lifting) compact_imp_bounded [OF \<open>compact S\<close>] bounded_subset mem_Collect_eq subsetI)  | 
|
4636  | 
qed (use \<open>y \<in> T\<close> \<open>0 < d\<close> fk kTS in \<open>force+\<close>)  | 
|
4637  | 
have "dist (h (k y')) (h (k y)) < e" if "y' \<in> T" "dist y y' < \<delta>" for y'  | 
|
4638  | 
proof -  | 
|
4639  | 
have k1: "k y \<in> S" "f (k y) = y" and k2: "k y' \<in> S" "f (k y') = y'"  | 
|
4640  | 
by (auto simp: \<open>y \<in> T\<close> \<open>y' \<in> T\<close> kTS fk)  | 
|
4641  | 
      have 1: "\<And>v. \<lbrakk>v \<in> S; f v = y\<rbrakk> \<Longrightarrow> \<exists>v'. v' \<in> {z \<in> S. f z = y'} \<and> dist v v' < d"
 | 
|
4642  | 
       and 2: "\<And>v'. \<lbrakk>v' \<in> S; f v' = y'\<rbrakk> \<Longrightarrow> \<exists>v. v \<in> {z \<in> S. f z = y} \<and> dist v' v < d"
 | 
|
4643  | 
using \<delta> [OF that] by auto  | 
|
4644  | 
then obtain w' w where "w' \<in> S" "f w' = y'" "dist (k y) w' < d"  | 
|
4645  | 
and "w \<in> S" "f w = y" "dist (k y') w < d"  | 
|
4646  | 
using 1 [OF k1] 2 [OF k2] by auto  | 
|
4647  | 
then show ?thesis  | 
|
4648  | 
using d [of w "k y'"] d [of w' "k y"] k1 k2 \<open>y' \<in> T\<close> \<open>y \<in> T\<close> hle  | 
|
4649  | 
by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps)  | 
|
4650  | 
qed  | 
|
4651  | 
then show "\<exists>d>0. \<forall>x'\<in>T. dist x' y < d \<longrightarrow> dist (h (k x')) (h (k y)) < e"  | 
|
4652  | 
using \<open>0 < \<delta>\<close> by (auto simp: dist_commute)  | 
|
4653  | 
qed  | 
|
4654  | 
then show "\<exists>h. continuous_on T h \<and> (\<forall>x\<in>T. g x = exp (\<i> * complex_of_real (h x)))"  | 
|
4655  | 
using fk gfh kTS by force  | 
|
4656  | 
qed  | 
|
4657  | 
||
| 
66939
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4658  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4659  | 
text\<open>If two points are separated by a closed set, there's a minimal one.\<close>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4660  | 
proposition closed_irreducible_separator:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4661  | 
fixes a :: "'a::real_normed_vector"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4662  | 
assumes "closed S" and ab: "\<not> connected_component (- S) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4663  | 
  obtains T where "T \<subseteq> S" "closed T" "T \<noteq> {}" "\<not> connected_component (- T) a b"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4664  | 
"\<And>U. U \<subset> T \<Longrightarrow> connected_component (- U) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4665  | 
proof (cases "a \<in> S \<or> b \<in> S")  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4666  | 
case True  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4667  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4668  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4669  | 
assume *: "a \<in> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4670  | 
show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4671  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4672  | 
      show "{a} \<subseteq> S"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4673  | 
using * by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4674  | 
      show "\<not> connected_component (- {a}) a b"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4675  | 
using connected_component_in by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4676  | 
      show "\<And>U. U \<subset> {a} \<Longrightarrow> connected_component (- U) a b"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4677  | 
by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4678  | 
qed auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4679  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4680  | 
assume *: "b \<in> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4681  | 
show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4682  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4683  | 
      show "{b} \<subseteq> S"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4684  | 
using * by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4685  | 
      show "\<not> connected_component (- {b}) a b"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4686  | 
using connected_component_in by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4687  | 
      show "\<And>U. U \<subset> {b} \<Longrightarrow> connected_component (- U) a b"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4688  | 
by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4689  | 
qed auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4690  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4691  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4692  | 
case False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4693  | 
define A where "A \<equiv> connected_component_set (- S) a"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4694  | 
define B where "B \<equiv> connected_component_set (- (closure A)) b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4695  | 
have "a \<in> A"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4696  | 
using False A_def by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4697  | 
have "b \<in> B"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4698  | 
unfolding A_def B_def closure_Un_frontier  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4699  | 
using ab False \<open>closed S\<close> frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4700  | 
have "frontier B \<subseteq> frontier (connected_component_set (- closure A) b)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4701  | 
using B_def by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4702  | 
also have frsub: "... \<subseteq> frontier A"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4703  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4704  | 
have "\<And>A. closure (- closure (- A)) \<subseteq> closure A"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4705  | 
by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4706  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4707  | 
by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4708  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4709  | 
finally have frBA: "frontier B \<subseteq> frontier A" .  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4710  | 
show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4711  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4712  | 
show "frontier B \<subseteq> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4713  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4714  | 
have "frontier S \<subseteq> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4715  | 
by (simp add: \<open>closed S\<close> frontier_subset_closed)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4716  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4717  | 
using frsub frontier_complement frontier_of_connected_component_subset  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4718  | 
unfolding A_def B_def by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4719  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4720  | 
show "closed (frontier B)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4721  | 
by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4722  | 
show "\<not> connected_component (- frontier B) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4723  | 
unfolding connected_component_def  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4724  | 
proof clarify  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4725  | 
fix T  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4726  | 
assume "connected T" and TB: "T \<subseteq> - frontier B" and "a \<in> T" and "b \<in> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4727  | 
have "a \<notin> B"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4728  | 
by (metis A_def B_def ComplD \<open>a \<in> A\<close> assms(1) closed_open connected_component_subset in_closure_connected_component set_mp)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4729  | 
      have "T \<inter> B \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4730  | 
using \<open>b \<in> B\<close> \<open>b \<in> T\<close> by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4731  | 
      moreover have "T - B \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4732  | 
using \<open>a \<notin> B\<close> \<open>a \<in> T\<close> by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4733  | 
ultimately show "False"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4734  | 
using connected_Int_frontier [of T B] TB \<open>connected T\<close> by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4735  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4736  | 
    moreover have "connected_component (- frontier B) a b" if "frontier B = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4737  | 
apply (simp add: that)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4738  | 
using connected_component_eq_UNIV by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4739  | 
    ultimately show "frontier B \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4740  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4741  | 
show "connected_component (- U) a b" if "U \<subset> frontier B" for U  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4742  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4743  | 
obtain p where Usub: "U \<subseteq> frontier B" and p: "p \<in> frontier B" "p \<notin> U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4744  | 
using \<open>U \<subset> frontier B\<close> by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4745  | 
show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4746  | 
unfolding connected_component_def  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4747  | 
proof (intro exI conjI)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4748  | 
have "connected ((insert p A) \<union> (insert p B))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4749  | 
proof (rule connected_Un)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4750  | 
show "connected (insert p A)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4751  | 
by (metis A_def IntD1 frBA \<open>p \<in> frontier B\<close> closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4752  | 
show "connected (insert p B)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4753  | 
by (metis B_def IntD1 \<open>p \<in> frontier B\<close> closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4754  | 
qed blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4755  | 
then show "connected (insert p (B \<union> A))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4756  | 
by (simp add: sup.commute)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4757  | 
have "A \<subseteq> - U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4758  | 
using A_def Usub \<open>frontier B \<subseteq> S\<close> connected_component_subset by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4759  | 
moreover have "B \<subseteq> - U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4760  | 
using B_def Usub connected_component_subset frBA frontier_closures by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4761  | 
ultimately show "insert p (B \<union> A) \<subseteq> - U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4762  | 
using p by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4763  | 
qed (auto simp: \<open>a \<in> A\<close> \<open>b \<in> B\<close>)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4764  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4765  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4766  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4767  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4768  | 
lemma frontier_minimal_separating_closed_pointwise:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4769  | 
fixes S :: "'a::real_normed_vector set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4770  | 
assumes S: "closed S" "a \<notin> S" and nconn: "\<not> connected_component (- S) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4771  | 
and conn: "\<And>T. \<lbrakk>closed T; T \<subset> S\<rbrakk> \<Longrightarrow> connected_component (- T) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4772  | 
shows "frontier(connected_component_set (- S) a) = S" (is "?F = S")  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4773  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4774  | 
have "?F \<subseteq> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4775  | 
by (simp add: S componentsI frontier_of_components_closed_complement)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4776  | 
moreover have False if "?F \<subset> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4777  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4778  | 
have "connected_component (- ?F) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4779  | 
by (simp add: conn that)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4780  | 
then obtain T where "connected T" "T \<subseteq> -?F" "a \<in> T" "b \<in> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4781  | 
by (auto simp: connected_component_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4782  | 
    moreover have "T \<inter> ?F \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4783  | 
proof (rule connected_Int_frontier [OF \<open>connected T\<close>])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4784  | 
      show "T \<inter> connected_component_set (- S) a \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4785  | 
using \<open>a \<notin> S\<close> \<open>a \<in> T\<close> by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4786  | 
      show "T - connected_component_set (- S) a \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4787  | 
using \<open>b \<in> T\<close> nconn by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4788  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4789  | 
ultimately show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4790  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4791  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4792  | 
ultimately show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4793  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4794  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4795  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4796  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4797  | 
subsection\<open>Unicoherence (closed)\<close>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4798  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4799  | 
definition unicoherent where  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4800  | 
"unicoherent U \<equiv>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4801  | 
\<forall>S T. connected S \<and> connected T \<and> S \<union> T = U \<and>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4802  | 
closedin (subtopology euclidean U) S \<and> closedin (subtopology euclidean U) T  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4803  | 
\<longrightarrow> connected (S \<inter> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4804  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4805  | 
lemma unicoherentI [intro?]:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4806  | 
assumes "\<And>S T. \<lbrakk>connected S; connected T; U = S \<union> T; closedin (subtopology euclidean U) S; closedin (subtopology euclidean U) T\<rbrakk>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4807  | 
\<Longrightarrow> connected (S \<inter> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4808  | 
shows "unicoherent U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4809  | 
using assms unfolding unicoherent_def by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4810  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4811  | 
lemma unicoherentD:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4812  | 
assumes "unicoherent U" "connected S" "connected T" "U = S \<union> T" "closedin (subtopology euclidean U) S" "closedin (subtopology euclidean U) T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4813  | 
shows "connected (S \<inter> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4814  | 
using assms unfolding unicoherent_def by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4815  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4816  | 
lemma homeomorphic_unicoherent:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4817  | 
assumes ST: "S homeomorphic T" and S: "unicoherent S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4818  | 
shows "unicoherent T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4819  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4820  | 
obtain f g where gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4821  | 
and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4822  | 
using ST by (auto simp: homeomorphic_def homeomorphism_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4823  | 
show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4824  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4825  | 
fix U V  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4826  | 
assume "connected U" "connected V" and T: "T = U \<union> V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4827  | 
and cloU: "closedin (subtopology euclidean T) U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4828  | 
and cloV: "closedin (subtopology euclidean T) V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4829  | 
have "f ` (g ` U \<inter> g ` V) \<subseteq> U" "f ` (g ` U \<inter> g ` V) \<subseteq> V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4830  | 
using gf fim T by auto (metis UnCI image_iff)+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4831  | 
moreover have "U \<inter> V \<subseteq> f ` (g ` U \<inter> g ` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4832  | 
using gf fim by (force simp: image_iff T)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4833  | 
ultimately have "U \<inter> V = f ` (g ` U \<inter> g ` V)" by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4834  | 
moreover have "connected (f ` (g ` U \<inter> g ` V))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4835  | 
proof (rule connected_continuous_image)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4836  | 
show "continuous_on (g ` U \<inter> g ` V) f"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4837  | 
apply (rule continuous_on_subset [OF contf])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4838  | 
using T fim gfim by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4839  | 
show "connected (g ` U \<inter> g ` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4840  | 
proof (intro conjI unicoherentD [OF S])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4841  | 
show "connected (g ` U)" "connected (g ` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4842  | 
using \<open>connected U\<close> cloU \<open>connected V\<close> cloV  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4843  | 
by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4844  | 
show "S = g ` U \<union> g ` V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4845  | 
using T fim gfim by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4846  | 
have hom: "homeomorphism T S g f"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4847  | 
by (simp add: contf contg fim gf gfim homeomorphism_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4848  | 
have "closedin (subtopology euclidean T) U" "closedin (subtopology euclidean T) V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4849  | 
by (simp_all add: cloU cloV)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4850  | 
then show "closedin (subtopology euclidean S) (g ` U)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4851  | 
"closedin (subtopology euclidean S) (g ` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4852  | 
by (blast intro: homeomorphism_imp_closed_map [OF hom])+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4853  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4854  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4855  | 
ultimately show "connected (U \<inter> V)" by metis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4856  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4857  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4858  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4859  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4860  | 
lemma homeomorphic_unicoherent_eq:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4861  | 
"S homeomorphic T \<Longrightarrow> (unicoherent S \<longleftrightarrow> unicoherent T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4862  | 
by (meson homeomorphic_sym homeomorphic_unicoherent)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4863  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4864  | 
lemma unicoherent_translation:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4865  | 
fixes S :: "'a::real_normed_vector set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4866  | 
shows  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4867  | 
"unicoherent (image (\<lambda>x. a + x) S) \<longleftrightarrow> unicoherent S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4868  | 
using homeomorphic_translation homeomorphic_unicoherent_eq by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4869  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4870  | 
lemma unicoherent_injective_linear_image:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4871  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4872  | 
assumes "linear f" "inj f"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4873  | 
shows "(unicoherent(f ` S) \<longleftrightarrow> unicoherent S)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4874  | 
using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4875  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4876  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4877  | 
lemma Borsukian_imp_unicoherent:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4878  | 
fixes U :: "'a::euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4879  | 
assumes "Borsukian U" shows "unicoherent U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4880  | 
unfolding unicoherent_def  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4881  | 
proof clarify  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4882  | 
fix S T  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4883  | 
assume "connected S" "connected T" "U = S \<union> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4884  | 
and cloS: "closedin (subtopology euclidean (S \<union> T)) S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4885  | 
and cloT: "closedin (subtopology euclidean (S \<union> T)) T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4886  | 
show "connected (S \<inter> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4887  | 
unfolding connected_closedin_eq  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4888  | 
proof clarify  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4889  | 
fix V W  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4890  | 
assume "closedin (subtopology euclidean (S \<inter> T)) V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4891  | 
and "closedin (subtopology euclidean (S \<inter> T)) W"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4892  | 
       and VW: "V \<union> W = S \<inter> T" "V \<inter> W = {}" and "V \<noteq> {}" "W \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4893  | 
then have cloV: "closedin (subtopology euclidean U) V" and cloW: "closedin (subtopology euclidean U) W"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4894  | 
using \<open>U = S \<union> T\<close> cloS cloT closedin_trans by blast+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4895  | 
obtain q where contq: "continuous_on U q"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4896  | 
         and q01: "\<And>x. x \<in> U \<Longrightarrow> q x \<in> {0..1::real}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4897  | 
and qV: "\<And>x. x \<in> V \<Longrightarrow> q x = 0" and qW: "\<And>x. x \<in> W \<Longrightarrow> q x = 1"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4898  | 
      by (rule Urysohn_local [OF cloV cloW \<open>V \<inter> W = {}\<close>, of 0 1])
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4899  | 
(fastforce simp: closed_segment_eq_real_ivl)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4900  | 
let ?h = "\<lambda>x. if x \<in> S then exp(pi * \<i> * q x) else 1 / exp(pi * \<i> * q x)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4901  | 
have eqST: "exp(pi * \<i> * q x) = 1 / exp(pi * \<i> * q x)" if "x \<in> S \<inter> T" for x  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4902  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4903  | 
have "x \<in> V \<union> W"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4904  | 
using that \<open>V \<union> W = S \<inter> T\<close> by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4905  | 
with qV qW show ?thesis by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4906  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4907  | 
obtain g where contg: "continuous_on U g"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4908  | 
and circle: "g ` U \<subseteq> sphere 0 1"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4909  | 
and S: "\<And>x. x \<in> S \<Longrightarrow> g x = exp(pi * \<i> * q x)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4910  | 
and T: "\<And>x. x \<in> T \<Longrightarrow> g x = 1 / exp(pi * \<i> * q x)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4911  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4912  | 
show "continuous_on U ?h"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4913  | 
unfolding \<open>U = S \<union> T\<close>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4914  | 
proof (rule continuous_on_cases_local [OF cloS cloT])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4915  | 
show "continuous_on S (\<lambda>x. exp (pi * \<i> * q x))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4916  | 
apply (intro continuous_intros)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4917  | 
using \<open>U = S \<union> T\<close> continuous_on_subset contq by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4918  | 
show "continuous_on T (\<lambda>x. 1 / exp (pi * \<i> * q x))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4919  | 
apply (intro continuous_intros)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4920  | 
using \<open>U = S \<union> T\<close> continuous_on_subset contq by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4921  | 
qed (use eqST in auto)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4922  | 
qed (use eqST in \<open>auto simp: norm_divide\<close>)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4923  | 
then obtain h where conth: "continuous_on U h" and heq: "\<And>x. x \<in> U \<Longrightarrow> g x = exp (h x)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4924  | 
by (metis Borsukian_continuous_logarithm_circle assms)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4925  | 
obtain v w where "v \<in> V" "w \<in> W"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4926  | 
      using \<open>V \<noteq> {}\<close> \<open>W \<noteq> {}\<close> by blast
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4927  | 
then have vw: "v \<in> S \<inter> T" "w \<in> S \<inter> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4928  | 
using VW by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4929  | 
have iff: "2 * pi \<le> cmod (2 * of_int m * of_real pi * \<i> - 2 * of_int n * of_real pi * \<i>)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4930  | 
\<longleftrightarrow> 1 \<le> abs (m - n)" for m n  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4931  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4932  | 
have "2 * pi \<le> cmod (2 * of_int m * of_real pi * \<i> - 2 * of_int n * of_real pi * \<i>)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4933  | 
\<longleftrightarrow> 2 * pi \<le> cmod ((2 * pi * \<i>) * (of_int m - of_int n))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4934  | 
by (simp add: algebra_simps)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4935  | 
also have "... \<longleftrightarrow> 2 * pi \<le> 2 * pi * cmod (of_int m - of_int n)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4936  | 
by (simp add: norm_mult)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4937  | 
also have "... \<longleftrightarrow> 1 \<le> abs (m - n)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4938  | 
by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4939  | 
finally show ?thesis .  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4940  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4941  | 
have *: "\<exists>n::int. h x - (pi * \<i> * q x) = (of_int(2*n) * pi) * \<i>" if "x \<in> S" for x  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4942  | 
using that S \<open>U = S \<union> T\<close> heq exp_eq [symmetric] by (simp add: algebra_simps)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4943  | 
moreover have "(\<lambda>x. h x - (pi * \<i> * q x)) constant_on S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4944  | 
proof (rule continuous_discrete_range_constant [OF \<open>connected S\<close>])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4945  | 
have "continuous_on S h" "continuous_on S q"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4946  | 
using \<open>U = S \<union> T\<close> continuous_on_subset conth contq by blast+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4947  | 
then show "continuous_on S (\<lambda>x. h x - (pi * \<i> * q x))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4948  | 
by (intro continuous_intros)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4949  | 
have "2*pi \<le> cmod (h y - (pi * \<i> * q y) - (h x - (pi * \<i> * q x)))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4950  | 
if "x \<in> S" "y \<in> S" and ne: "h y - (pi * \<i> * q y) \<noteq> h x - (pi * \<i> * q x)" for x y  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4951  | 
using * [OF \<open>x \<in> S\<close>] * [OF \<open>y \<in> S\<close>] ne by (auto simp: iff)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4952  | 
then show "\<And>x. x \<in> S \<Longrightarrow>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4953  | 
\<exists>e>0. \<forall>y. y \<in> S \<and> h y - (pi * \<i> * q y) \<noteq> h x - (pi * \<i> * q x) \<longrightarrow>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4954  | 
e \<le> cmod (h y - (pi * \<i> * q y) - (h x - (pi * \<i> * q x)))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4955  | 
by (rule_tac x="2*pi" in exI) auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4956  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4957  | 
ultimately  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4958  | 
obtain m where m: "\<And>x. x \<in> S \<Longrightarrow> h x - (pi * \<i> * q x) = (of_int(2*m) * pi) * \<i>"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4959  | 
using vw by (force simp: constant_on_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4960  | 
have *: "\<exists>n::int. h x = - (pi * \<i> * q x) + (of_int(2*n) * pi) * \<i>" if "x \<in> T" for x  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4961  | 
unfolding exp_eq [symmetric]  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4962  | 
using that T \<open>U = S \<union> T\<close> by (simp add: exp_minus field_simps heq [symmetric])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4963  | 
moreover have "(\<lambda>x. h x + (pi * \<i> * q x)) constant_on T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4964  | 
proof (rule continuous_discrete_range_constant [OF \<open>connected T\<close>])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4965  | 
have "continuous_on T h" "continuous_on T q"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4966  | 
using \<open>U = S \<union> T\<close> continuous_on_subset conth contq by blast+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4967  | 
then show "continuous_on T (\<lambda>x. h x + (pi * \<i> * q x))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4968  | 
by (intro continuous_intros)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4969  | 
have "2*pi \<le> cmod (h y + (pi * \<i> * q y) - (h x + (pi * \<i> * q x)))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4970  | 
if "x \<in> T" "y \<in> T" and ne: "h y + (pi * \<i> * q y) \<noteq> h x + (pi * \<i> * q x)" for x y  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4971  | 
using * [OF \<open>x \<in> T\<close>] * [OF \<open>y \<in> T\<close>] ne by (auto simp: iff)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4972  | 
then show "\<And>x. x \<in> T \<Longrightarrow>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4973  | 
\<exists>e>0. \<forall>y. y \<in> T \<and> h y + (pi * \<i> * q y) \<noteq> h x + (pi * \<i> * q x) \<longrightarrow>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4974  | 
e \<le> cmod (h y + (pi * \<i> * q y) - (h x + (pi * \<i> * q x)))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4975  | 
by (rule_tac x="2*pi" in exI) auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4976  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4977  | 
ultimately  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4978  | 
obtain n where n: "\<And>x. x \<in> T \<Longrightarrow> h x + (pi * \<i> * q x) = (of_int(2*n) * pi) * \<i>"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4979  | 
using vw by (force simp: constant_on_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4980  | 
show "False"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4981  | 
using m [of v] m [of w] n [of v] n [of w] vw  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4982  | 
by (auto simp: algebra_simps \<open>v \<in> V\<close> \<open>w \<in> W\<close> qV qW)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4983  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4984  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4985  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4986  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4987  | 
corollary contractible_imp_unicoherent:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4988  | 
fixes U :: "'a::euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4989  | 
assumes "contractible U" shows "unicoherent U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4990  | 
by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4991  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4992  | 
corollary convex_imp_unicoherent:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4993  | 
fixes U :: "'a::euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4994  | 
assumes "convex U" shows "unicoherent U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4995  | 
by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4996  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4997  | 
text\<open>If the type class constraint can be relaxed, I don't know how!\<close>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4998  | 
corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
4999  | 
by (simp add: convex_imp_unicoherent)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5000  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5001  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5002  | 
lemma unicoherent_monotone_image_compact:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5003  | 
fixes T :: "'b :: t2_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5004  | 
assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5005  | 
  and conn: "\<And>y. y \<in> T \<Longrightarrow> connected (S \<inter> f -` {y})"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5006  | 
shows "unicoherent T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5007  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5008  | 
fix U V  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5009  | 
assume UV: "connected U" "connected V" "T = U \<union> V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5010  | 
and cloU: "closedin (subtopology euclidean T) U"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5011  | 
and cloV: "closedin (subtopology euclidean T) V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5012  | 
moreover have "compact T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5013  | 
using \<open>compact S\<close> compact_continuous_image contf fim by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5014  | 
ultimately have "closed U" "closed V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5015  | 
by (auto simp: closedin_closed_eq compact_imp_closed)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5016  | 
let ?SUV = "(S \<inter> f -` U) \<inter> (S \<inter> f -` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5017  | 
have UV_eq: "f ` ?SUV = U \<inter> V"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5018  | 
using \<open>T = U \<union> V\<close> fim by force+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5019  | 
have "connected (f ` ?SUV)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5020  | 
proof (rule connected_continuous_image)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5021  | 
show "continuous_on ?SUV f"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5022  | 
by (meson contf continuous_on_subset inf_le1)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5023  | 
show "connected ?SUV"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5024  | 
proof (rule unicoherentD [OF \<open>unicoherent S\<close>, of "S \<inter> f -` U" "S \<inter> f -` V"])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5025  | 
have "\<And>C. closedin (subtopology euclidean S) C \<Longrightarrow> closedin (subtopology euclidean T) (f ` C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5026  | 
by (metis \<open>compact S\<close> closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5027  | 
then show "connected (S \<inter> f -` U)" "connected (S \<inter> f -` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5028  | 
using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5029  | 
show "S = (S \<inter> f -` U) \<union> (S \<inter> f -` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5030  | 
using UV fim by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5031  | 
show "closedin (subtopology euclidean S) (S \<inter> f -` U)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5032  | 
"closedin (subtopology euclidean S) (S \<inter> f -` V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5033  | 
by (auto simp: continuous_on_imp_closedin cloU cloV contf fim)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5034  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5035  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5036  | 
with UV_eq show "connected (U \<inter> V)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5037  | 
by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5038  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5039  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5040  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5041  | 
subsection\<open>Several common variants of unicoherence\<close>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5042  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5043  | 
lemma connected_frontier_simple:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5044  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5045  | 
assumes "connected S" "connected(- S)" shows "connected(frontier S)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5046  | 
unfolding frontier_closures  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5047  | 
apply (rule unicoherentD [OF unicoherent_UNIV])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5048  | 
apply (simp_all add: assms connected_imp_connected_closure)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5049  | 
by (simp add: closure_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5050  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5051  | 
lemma connected_frontier_component_complement:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5052  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5053  | 
assumes "connected S" and C: "C \<in> components(- S)" shows "connected(frontier C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5054  | 
apply (rule connected_frontier_simple)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5055  | 
using C in_components_connected apply blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5056  | 
by (metis Compl_eq_Diff_UNIV connected_UNIV assms top_greatest component_complement_connected)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5057  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5058  | 
lemma connected_frontier_disjoint:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5059  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5060  | 
assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \<subseteq> frontier T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5061  | 
shows "connected(frontier S)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5062  | 
proof (cases "S = UNIV")  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5063  | 
case True then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5064  | 
by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5065  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5066  | 
case False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5067  | 
  then have "-S \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5068  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5069  | 
then obtain C where C: "C \<in> components(- S)" and "T \<subseteq> C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5070  | 
by (metis ComplI disjnt_iff subsetI exists_component_superset \<open>disjnt S T\<close> \<open>connected T\<close>)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5071  | 
moreover have "frontier S = frontier C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5072  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5073  | 
have "frontier C \<subseteq> frontier S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5074  | 
using C frontier_complement frontier_of_components_subset by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5075  | 
moreover have "x \<in> frontier C" if "x \<in> frontier S" for x  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5076  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5077  | 
have "x \<in> closure C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5078  | 
using that unfolding frontier_def  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5079  | 
by (metis (no_types) Diff_eq ST \<open>T \<subseteq> C\<close> closure_mono contra_subsetD frontier_def le_inf_iff that)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5080  | 
moreover have "x \<notin> interior C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5081  | 
using that unfolding frontier_def  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5082  | 
by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5083  | 
ultimately show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5084  | 
by (auto simp: frontier_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5085  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5086  | 
ultimately show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5087  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5088  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5089  | 
ultimately show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5090  | 
using \<open>connected S\<close> connected_frontier_component_complement by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5091  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5092  | 
|
| 
66941
 
c67bb79a0ceb
More topological results overlooked last time
 
paulson <lp15@cam.ac.uk> 
parents: 
66939 
diff
changeset
 | 
5093  | 
|
| 
 
c67bb79a0ceb
More topological results overlooked last time
 
paulson <lp15@cam.ac.uk> 
parents: 
66939 
diff
changeset
 | 
5094  | 
subsection\<open>Some separation results\<close>  | 
| 
 
c67bb79a0ceb
More topological results overlooked last time
 
paulson <lp15@cam.ac.uk> 
parents: 
66939 
diff
changeset
 | 
5095  | 
|
| 
66939
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5096  | 
lemma separation_by_component_closed_pointwise:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5097  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5098  | 
assumes "closed S" "\<not> connected_component (- S) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5099  | 
obtains C where "C \<in> components S" "\<not> connected_component(- C) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5100  | 
proof (cases "a \<in> S \<or> b \<in> S")  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5101  | 
case True  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5102  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5103  | 
using connected_component_in componentsI that by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5104  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5105  | 
case False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5106  | 
  obtain T where "T \<subseteq> S" "closed T" "T \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5107  | 
and nab: "\<not> connected_component (- T) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5108  | 
and conn: "\<And>U. U \<subset> T \<Longrightarrow> connected_component (- U) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5109  | 
using closed_irreducible_separator [OF assms] by metis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5110  | 
moreover have "connected T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5111  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5112  | 
have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5113  | 
using frontier_minimal_separating_closed_pointwise  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5114  | 
by (metis False \<open>T \<subseteq> S\<close> \<open>closed T\<close> connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5115  | 
have "connected (frontier (connected_component_set (- T) a))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5116  | 
proof (rule connected_frontier_disjoint)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5117  | 
show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5118  | 
unfolding disjnt_iff  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5119  | 
by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5120  | 
show "frontier (connected_component_set (- T) a) \<subseteq> frontier (connected_component_set (- T) b)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5121  | 
by (simp add: ab)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5122  | 
qed auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5123  | 
with ab \<open>closed T\<close> show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5124  | 
by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5125  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5126  | 
ultimately obtain C where "C \<in> components S" "T \<subseteq> C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5127  | 
using exists_component_superset [of T S] by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5128  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5129  | 
by (meson Compl_anti_mono connected_component_of_subset nab that)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5130  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5131  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5132  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5133  | 
lemma separation_by_component_closed:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5134  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5135  | 
assumes "closed S" "\<not> connected(- S)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5136  | 
obtains C where "C \<in> components S" "\<not> connected(- C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5137  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5138  | 
obtain x y where "closed S" "x \<notin> S" "y \<notin> S" and "\<not> connected_component (- S) x y"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5139  | 
using assms by (auto simp: connected_iff_connected_component)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5140  | 
then obtain C where "C \<in> components S" "\<not> connected_component(- C) x y"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5141  | 
using separation_by_component_closed_pointwise by metis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5142  | 
then show "thesis"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5143  | 
apply (clarify elim!: componentsE)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5144  | 
by (metis Compl_iff \<open>C \<in> components S\<close> \<open>x \<notin> S\<close> \<open>y \<notin> S\<close> connected_component_eq connected_component_eq_eq connected_iff_connected_component that)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5145  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5146  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5147  | 
lemma separation_by_Un_closed_pointwise:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5148  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5149  | 
  assumes ST: "closed S" "closed T" "S \<inter> T = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5150  | 
and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5151  | 
shows "connected_component (- (S \<union> T)) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5152  | 
proof (rule ccontr)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5153  | 
have "a \<notin> S" "b \<notin> S" "a \<notin> T" "b \<notin> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5154  | 
using conS conT connected_component_in by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5155  | 
assume "\<not> connected_component (- (S \<union> T)) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5156  | 
then obtain C where "C \<in> components (S \<union> T)" and C: "\<not> connected_component(- C) a b"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5157  | 
using separation_by_component_closed_pointwise assms by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5158  | 
then have "C \<subseteq> S \<or> C \<subseteq> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5159  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5160  | 
have "connected C" "C \<subseteq> S \<union> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5161  | 
using \<open>C \<in> components (S \<union> T)\<close> in_components_subset by (blast elim: componentsE)+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5162  | 
    moreover then have "C \<inter> T = {} \<or> C \<inter> S = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5163  | 
by (metis Int_empty_right ST inf.commute connected_closed)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5164  | 
ultimately show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5165  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5166  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5167  | 
then show False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5168  | 
by (meson Compl_anti_mono C conS conT connected_component_of_subset)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5169  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5170  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5171  | 
lemma separation_by_Un_closed:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5172  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5173  | 
  assumes ST: "closed S" "closed T" "S \<inter> T = {}" and conS: "connected(- S)" and conT: "connected(- T)"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5174  | 
shows "connected(- (S \<union> T))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5175  | 
using assms separation_by_Un_closed_pointwise  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5176  | 
by (fastforce simp add: connected_iff_connected_component)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5177  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5178  | 
lemma open_unicoherent_UNIV:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5179  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5180  | 
assumes "open S" "open T" "connected S" "connected T" "S \<union> T = UNIV"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5181  | 
shows "connected(S \<inter> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5182  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5183  | 
have "connected(- (-S \<union> -T))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5184  | 
by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5185  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5186  | 
by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5187  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5188  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5189  | 
lemma separation_by_component_open_aux:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5190  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5191  | 
  assumes ST: "closed S" "closed T" "S \<inter> T = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5192  | 
      and "S \<noteq> {}" "T \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5193  | 
  obtains C where "C \<in> components(-(S \<union> T))" "C \<noteq> {}" "frontier C \<inter> S \<noteq> {}" "frontier C \<inter> T \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5194  | 
proof (rule ccontr)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5195  | 
  let ?S = "S \<union> \<Union>{C \<in> components(- (S \<union> T)). frontier C \<subseteq> S}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5196  | 
  let ?T = "T \<union> \<Union>{C \<in> components(- (S \<union> T)). frontier C \<subseteq> T}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5197  | 
assume "~ thesis"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5198  | 
  with that have *: "frontier C \<inter> S = {} \<or> frontier C \<inter> T = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5199  | 
            if C: "C \<in> components (- (S \<union> T))" "C \<noteq> {}" for C
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5200  | 
using C by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5201  | 
  have "\<exists>A B::'a set. closed A \<and> closed B \<and> UNIV \<subseteq> A \<union> B \<and> A \<inter> B = {} \<and> A \<noteq> {} \<and> B \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5202  | 
proof (intro exI conjI)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5203  | 
    have "frontier (\<Union>{C \<in> components (- S \<inter> - T). frontier C \<subseteq> S}) \<subseteq> S"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5204  | 
apply (rule subset_trans [OF frontier_Union_subset_closure])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5205  | 
by (metis (no_types, lifting) SUP_least \<open>closed S\<close> closure_minimal mem_Collect_eq)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5206  | 
then have "frontier ?S \<subseteq> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5207  | 
by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5208  | 
then show "closed ?S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5209  | 
using frontier_subset_eq by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5210  | 
    have "frontier (\<Union>{C \<in> components (- S \<inter> - T). frontier C \<subseteq> T}) \<subseteq> T"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5211  | 
apply (rule subset_trans [OF frontier_Union_subset_closure])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5212  | 
by (metis (no_types, lifting) SUP_least \<open>closed T\<close> closure_minimal mem_Collect_eq)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5213  | 
then have "frontier ?T \<subseteq> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5214  | 
by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5215  | 
then show "closed ?T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5216  | 
using frontier_subset_eq by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5217  | 
have "UNIV \<subseteq> (S \<union> T) \<union> \<Union>(components(- (S \<union> T)))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5218  | 
using Union_components by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5219  | 
also have "... \<subseteq> ?S \<union> ?T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5220  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5221  | 
have "C \<in> components (-(S \<union> T)) \<and> frontier C \<subseteq> S \<or>  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5222  | 
C \<in> components (-(S \<union> T)) \<and> frontier C \<subseteq> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5223  | 
        if "C \<in> components (- (S \<union> T))" "C \<noteq> {}" for C
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5224  | 
using * [OF that] that  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5225  | 
by clarify (metis (no_types, lifting) UnE \<open>closed S\<close> \<open>closed T\<close> closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5226  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5227  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5228  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5229  | 
finally show "UNIV \<subseteq> ?S \<union> ?T" .  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5230  | 
    have "\<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> S} \<union>
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5231  | 
          \<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> T} \<subseteq> - (S \<union> T)"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5232  | 
using in_components_subset by fastforce  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5233  | 
    moreover have "\<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> S} \<inter>
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5234  | 
                   \<Union>{C \<in> components (- (S \<union> T)). frontier C \<subseteq> T} = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5235  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5236  | 
      have "C \<inter> C' = {}" if "C \<in> components (- (S \<union> T))" "frontier C \<subseteq> S"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5237  | 
"C' \<in> components (- (S \<union> T))" "frontier C' \<subseteq> T" for C C'  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5238  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5239  | 
have NUN: "- S \<inter> - T \<noteq> UNIV"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5240  | 
          using \<open>T \<noteq> {}\<close> by blast
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5241  | 
have "C \<noteq> C'"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5242  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5243  | 
assume "C = C'"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5244  | 
with that have "frontier C' \<subseteq> S \<inter> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5245  | 
by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5246  | 
          also have "... = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5247  | 
            using \<open>S \<inter> T = {}\<close> by blast
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5248  | 
          finally have "C' = {} \<or> C' = UNIV"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5249  | 
using frontier_eq_empty by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5250  | 
then show False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5251  | 
using \<open>C = C'\<close> NUN that by (force simp: dest: in_components_nonempty in_components_subset)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5252  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5253  | 
with that show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5254  | 
by (simp add: components_nonoverlap [of _ "-(S \<union> T)"])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5255  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5256  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5257  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5258  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5259  | 
    ultimately show "?S \<inter> ?T = {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5260  | 
using ST by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5261  | 
    show "?S \<noteq> {}" "?T \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5262  | 
      using \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> by blast+
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5263  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5264  | 
then show False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5265  | 
by (metis Compl_disjoint Convex_Euclidean_Space.connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5266  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5267  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5268  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5269  | 
lemma separation_by_component_open:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5270  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5271  | 
assumes "open S" and non: "\<not> connected(- S)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5272  | 
obtains C where "C \<in> components S" "\<not> connected(- C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5273  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5274  | 
obtain T U  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5275  | 
    where "closed T" "closed U" and TU: "T \<union> U = - S" "T \<inter> U = {}" "T \<noteq> {}" "U \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5276  | 
using assms by (auto simp: connected_closed_set closed_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5277  | 
  then obtain C where C: "C \<in> components(-(T \<union> U))" "C \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5278  | 
          and "frontier C \<inter> T \<noteq> {}" "frontier C \<inter> U \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5279  | 
    using separation_by_component_open_aux [OF \<open>closed T\<close> \<open>closed U\<close> \<open>T \<inter> U = {}\<close>] by force
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5280  | 
show "thesis"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5281  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5282  | 
show "C \<in> components S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5283  | 
using C(1) TU(1) by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5284  | 
show "\<not> connected (- C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5285  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5286  | 
assume "connected (- C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5287  | 
then have "connected (frontier C)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5288  | 
using connected_frontier_simple [of C] \<open>C \<in> components S\<close> in_components_connected by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5289  | 
then show False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5290  | 
unfolding connected_closed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5291  | 
        by (metis C(1) TU(2) \<open>closed T\<close> \<open>closed U\<close> \<open>frontier C \<inter> T \<noteq> {}\<close> \<open>frontier C \<inter> U \<noteq> {}\<close> closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute)
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5292  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5293  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5294  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5295  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5296  | 
lemma separation_by_Un_open:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5297  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5298  | 
  assumes "open S" "open T" "S \<inter> T = {}" and cS: "connected(-S)" and cT: "connected(-T)"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5299  | 
shows "connected(- (S \<union> T))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5300  | 
using assms unicoherent_UNIV unfolding unicoherent_def by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5301  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5302  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5303  | 
lemma nonseparation_by_component_eq:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5304  | 
fixes S :: "'a :: euclidean_space set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5305  | 
assumes "open S \<or> closed S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5306  | 
shows "((\<forall>C \<in> components S. connected(-C)) \<longleftrightarrow> connected(- S))" (is "?lhs = ?rhs")  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5307  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5308  | 
assume ?lhs with assms show ?rhs  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5309  | 
by (meson separation_by_component_closed separation_by_component_open)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5310  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5311  | 
assume ?rhs with assms show ?lhs  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5312  | 
using component_complement_connected by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5313  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5314  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5315  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5316  | 
text\<open>Another interesting equivalent of an inessential mapping into C-{0}\<close>
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5317  | 
proposition inessential_eq_extensible:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5318  | 
fixes f :: "'a::euclidean_space \<Rightarrow> complex"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5319  | 
assumes "closed S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5320  | 
  shows "(\<exists>a. homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)) \<longleftrightarrow>
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5321  | 
(\<exists>g. continuous_on UNIV g \<and> (\<forall>x \<in> S. g x = f x) \<and> (\<forall>x. g x \<noteq> 0))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5322  | 
(is "?lhs = ?rhs")  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5323  | 
proof  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5324  | 
assume ?lhs  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5325  | 
  then obtain a where a: "homotopic_with (\<lambda>h. True) S (-{0}) f (\<lambda>t. a)" ..
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5326  | 
show ?rhs  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5327  | 
  proof (cases "S = {}")
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5328  | 
case True  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5329  | 
with a show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5330  | 
using continuous_on_const by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5331  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5332  | 
case False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5333  | 
    have anr: "ANR (-{0::complex})"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5334  | 
by (simp add: ANR_delete open_Compl open_imp_ANR)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5335  | 
    obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \<subseteq> -{0}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5336  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5337  | 
proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5338  | 
show "closedin (subtopology euclidean UNIV) S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5339  | 
using assms by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5340  | 
      show "range (\<lambda>t. a) \<subseteq> - {0}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5341  | 
using a homotopic_with_imp_subset2 False by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5342  | 
qed (use anr that in \<open>force+\<close>)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5343  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5344  | 
by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5345  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5346  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5347  | 
assume ?rhs  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5348  | 
then obtain g where contg: "continuous_on UNIV g"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5349  | 
and gf: "\<And>x. x \<in> S \<Longrightarrow> g x = f x" and non0: "\<And>x. g x \<noteq> 0"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5350  | 
by metis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5351  | 
obtain h k::"'a\<Rightarrow>'a" where hk: "homeomorphism (ball 0 1) UNIV h k"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5352  | 
using homeomorphic_ball01_UNIV homeomorphic_def by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5353  | 
then have "continuous_on (ball 0 1) (g \<circ> h)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5354  | 
by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5355  | 
then obtain j where contj: "continuous_on (ball 0 1) j"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5356  | 
and j: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> exp(j z) = (g \<circ> h) z"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5357  | 
by (metis (mono_tags, hide_lams) continuous_logarithm_on_ball comp_apply non0)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5358  | 
have [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (k x) = x"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5359  | 
using hk homeomorphism_apply2 by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5360  | 
have "\<exists>\<zeta>. continuous_on S \<zeta>\<and> (\<forall>x\<in>S. f x = exp (\<zeta> x))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5361  | 
proof (intro exI conjI ballI)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5362  | 
show "continuous_on S (j \<circ> k)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5363  | 
proof (rule continuous_on_compose)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5364  | 
show "continuous_on S k"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5365  | 
by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5366  | 
show "continuous_on (k ` S) j"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5367  | 
apply (rule continuous_on_subset [OF contj])  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5368  | 
using homeomorphism_image2 [OF hk] continuous_on_subset [OF contj] by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5369  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5370  | 
show "f x = exp ((j \<circ> k) x)" if "x \<in> S" for x  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5371  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5372  | 
have "f x = (g \<circ> h) (k x)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5373  | 
by (simp add: gf that)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5374  | 
also have "... = exp (j (k x))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5375  | 
by (metis rangeI homeomorphism_image2 [OF hk] j)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5376  | 
finally show ?thesis by simp  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5377  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5378  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5379  | 
then show ?lhs  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5380  | 
by (simp add: inessential_eq_continuous_logarithm)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5381  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5382  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5383  | 
lemma inessential_on_clopen_Union:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5384  | 
fixes \<F> :: "'a::euclidean_space set set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5385  | 
assumes T: "path_connected T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5386  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5387  | 
and "\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5388  | 
and hom: "\<And>S. S \<in> \<F> \<Longrightarrow> \<exists>a. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. a)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5389  | 
obtains a where "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5390  | 
proof (cases "\<Union>\<F> = {}")
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5391  | 
case True  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5392  | 
with that show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5393  | 
by force  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5394  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5395  | 
case False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5396  | 
  then obtain C where "C \<in> \<F>" "C \<noteq> {}"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5397  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5398  | 
then obtain a where clo: "closedin (subtopology euclidean (\<Union>\<F>)) C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5399  | 
and ope: "openin (subtopology euclidean (\<Union>\<F>)) C"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5400  | 
and "homotopic_with (\<lambda>x. True) C T f (\<lambda>x. a)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5401  | 
using assms by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5402  | 
  with \<open>C \<noteq> {}\<close> have "f ` C \<subseteq> T" "a \<in> T"
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5403  | 
using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5404  | 
have "homotopic_with (\<lambda>x. True) (\<Union>\<F>) T f (\<lambda>x. a)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5405  | 
proof (rule homotopic_on_clopen_Union)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5406  | 
show "\<And>S. S \<in> \<F> \<Longrightarrow> closedin (subtopology euclidean (\<Union>\<F>)) S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5407  | 
"\<And>S. S \<in> \<F> \<Longrightarrow> openin (subtopology euclidean (\<Union>\<F>)) S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5408  | 
by (simp_all add: assms)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5409  | 
show "homotopic_with (\<lambda>x. True) S T f (\<lambda>x. a)" if "S \<in> \<F>" for S  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5410  | 
    proof (cases "S = {}")
 | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5411  | 
case True  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5412  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5413  | 
by auto  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5414  | 
next  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5415  | 
case False  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5416  | 
then obtain b where "b \<in> S"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5417  | 
by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5418  | 
obtain c where c: "homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5419  | 
using \<open>S \<in> \<F>\<close> hom by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5420  | 
then have "c \<in> T"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5421  | 
using \<open>b \<in> S\<close> homotopic_with_imp_subset2 by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5422  | 
then have "homotopic_with (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. c)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5423  | 
using T \<open>a \<in> T\<close> homotopic_constant_maps path_connected_component by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5424  | 
then show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5425  | 
using c homotopic_with_symD homotopic_with_trans by blast  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5426  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5427  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5428  | 
then show ?thesis ..  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5429  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5430  | 
|
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5431  | 
lemma Janiszewski_dual:  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5432  | 
fixes S :: "complex set"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5433  | 
assumes  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5434  | 
"compact S" "compact T" "connected S" "connected T" "connected(- (S \<union> T))"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5435  | 
shows "connected(S \<inter> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5436  | 
proof -  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5437  | 
have ST: "compact (S \<union> T)"  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5438  | 
by (simp add: assms compact_Un)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5439  | 
with Borsukian_imp_unicoherent [of "S \<union> T"] ST assms  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5440  | 
show ?thesis  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5441  | 
by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def)  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5442  | 
qed  | 
| 
 
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
 
paulson <lp15@cam.ac.uk> 
parents: 
66884 
diff
changeset
 | 
5443  | 
|
| 
64006
 
0de4736dad8b
new theorems including the theory FurtherTopology
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
5444  | 
end  |