| author | wenzelm |
| Tue, 18 Oct 2016 15:09:52 +0200 | |
| changeset 64301 | 8053c882839f |
| parent 64122 | 74fde524799e |
| child 64287 | d85d88722745 |
| 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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
5 |
theory "FurtherTopology" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
6 |
imports Equivalence_Lebesgue_Henstock_Integration Weierstrass_Theorems Polytope |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
7 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
8 |
begin |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
9 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
10 |
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
|
11 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
12 |
lemma spheremap_lemma1: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
13 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
14 |
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
|
15 |
and "S \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
16 |
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
|
17 |
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
|
18 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
19 |
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
|
20 |
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
|
21 |
using subspace_mul \<open>subspace S\<close> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
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
|
23 |
using \<open>subspace S\<close> subspace_mul by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
24 |
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
|
25 |
by (rule differentiable_on_subset [OF diff_f]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
26 |
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
|
27 |
have gdiff: "g differentiable_on S - {0}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
28 |
unfolding g_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
29 |
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
|
30 |
have geq: "g ` (S - {0}) = T - {0}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
31 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
32 |
have "g ` (S - {0}) \<subseteq> T"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
33 |
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
|
34 |
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
|
35 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
36 |
moreover have "g ` (S - {0}) \<subseteq> UNIV - {0}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
37 |
proof (clarsimp simp: g_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
38 |
fix y |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
39 |
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
|
40 |
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
|
41 |
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
|
42 |
then show "y = 0" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
43 |
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
|
44 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
45 |
ultimately show "g ` (S - {0}) \<subseteq> T - {0}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
46 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
47 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
48 |
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
|
49 |
using fim by (simp add: image_subset_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
50 |
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
|
51 |
if "x \<in> T" "x \<noteq> 0" for x |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
52 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
53 |
have "x /\<^sub>R norm x \<in> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
54 |
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
|
55 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
56 |
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
|
57 |
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
|
58 |
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
|
59 |
using \<open>subspace S\<close> subspace_mul apply force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
60 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
61 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
62 |
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
|
63 |
by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
64 |
then show "T - {0} \<subseteq> g ` (S - {0})"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
65 |
by (simp add: g_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
66 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
67 |
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
|
68 |
have "subspace T'" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
69 |
by (simp add: subspace_orthogonal_to_vectors T'_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
70 |
have dim_eq: "dim T' + dim T = DIM('a)"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
71 |
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
|
72 |
by (simp add: dim_UNIV T'_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
73 |
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
|
74 |
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
|
75 |
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
|
76 |
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
|
77 |
and eq: "p1 x + p2 x = x" for x |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
78 |
by metis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
79 |
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
|
80 |
using span_eq \<open>subspace T\<close> by blast+ |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
81 |
then have p2: "\<And>z. p2 z \<in> T'" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
82 |
by (simp add: T'_def orthogonal_commute) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
83 |
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
|
84 |
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
|
85 |
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
|
86 |
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
|
87 |
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
|
88 |
using T'_def by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
89 |
qed (auto simp: span_superset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
90 |
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
|
91 |
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
|
92 |
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
|
93 |
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
|
94 |
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
|
95 |
by (simp add: add.assoc add.left_commute eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
96 |
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
|
97 |
using T'_def by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
98 |
qed (auto simp: p1span p2 span_superset subspace_add) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
99 |
ultimately have "linear p1" "linear p2" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
100 |
by unfold_locales auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
101 |
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
|
102 |
apply (rule differentiable_on_compose [where f=g]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
103 |
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
|
104 |
apply (rule differentiable_on_subset [OF gdiff]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
105 |
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
|
106 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
107 |
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
|
108 |
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
|
109 |
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
|
110 |
by (blast intro: dim_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
111 |
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
|
112 |
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
|
113 |
by (simp add: algebra_simps) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
114 |
also have "... < DIM('a)"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
115 |
using dimST dim_eq by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
116 |
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
|
117 |
by (rule negligible_lowdim) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
118 |
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
|
119 |
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
|
120 |
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
|
121 |
proof (rule negligible_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
122 |
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
|
123 |
\<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
|
124 |
{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
|
125 |
apply (rule_tac x="s + t'" in image_eqI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
126 |
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
|
127 |
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
|
128 |
\<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
|
129 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
130 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
131 |
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
|
132 |
proof clarsimp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
133 |
fix z assume "z \<notin> T'" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
134 |
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
|
135 |
apply (rule_tac x="p1 z" in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
136 |
apply (rule_tac x="p2 z" in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
137 |
apply (simp add: p1 eq p2 geq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
138 |
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
|
139 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
140 |
ultimately have "negligible (-T')" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
141 |
using negligible_subset by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
142 |
moreover have "negligible T'" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
143 |
using negligible_lowdim |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
144 |
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
|
145 |
ultimately have "negligible (-T' \<union> T')" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
146 |
by (metis negligible_Un_eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
147 |
then show False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
148 |
using negligible_Un_eq non_negligible_UNIV by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
149 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
152 |
lemma spheremap_lemma2: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
153 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'a::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
154 |
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
|
155 |
and "S \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
156 |
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
|
157 |
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
|
158 |
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
|
159 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
160 |
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
|
161 |
using fim by (simp add: image_subset_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
162 |
have "compact (sphere 0 1 \<inter> S)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
163 |
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
|
164 |
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
|
165 |
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
|
166 |
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
|
167 |
using fim apply auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
168 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
169 |
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
|
170 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
171 |
have "norm (f x) = 1" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
172 |
using fim that by (simp add: image_subset_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
173 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
174 |
using g12 [OF that] by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
175 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
176 |
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
|
177 |
by (metis pfg differentiable_on_polynomial_function) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
178 |
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
|
179 |
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
|
180 |
unfolding h_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
181 |
using gnz [of x] |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
182 |
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
|
183 |
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
|
184 |
unfolding h_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
185 |
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
|
186 |
using gnz apply auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
187 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
188 |
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
|
189 |
proof (rule homotopic_with_linear [OF contf]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
190 |
show "continuous_on (sphere 0 1 \<inter> S) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
191 |
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
|
192 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
193 |
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
|
194 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
195 |
have "f x \<in> sphere 0 1" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
196 |
using fim that by (simp add: image_subset_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
197 |
moreover have "norm(f x - g x) < 1/2" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
198 |
apply (rule g12) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
199 |
using that by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
200 |
ultimately show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
201 |
by (auto simp: norm_minus_commute dest: segment_bound) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
202 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
203 |
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
|
204 |
apply (simp add: subset_Diff_insert non0fg) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
205 |
apply (simp add: segment_convex_hull) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
206 |
apply (rule hull_minimal) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
207 |
using fim image_eqI gim apply force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
208 |
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
|
209 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
210 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
211 |
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
|
212 |
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
|
213 |
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
|
214 |
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
|
215 |
by (auto simp: between_mem_segment midpoint_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
216 |
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
|
217 |
using differentiable_imp_continuous_on diffh by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
218 |
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
|
219 |
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
|
220 |
apply (simp add: subset_Diff_insert non0hd) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
221 |
apply (simp add: segment_convex_hull) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
222 |
apply (rule hull_minimal) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
223 |
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
|
224 |
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
|
225 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
226 |
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
|
227 |
by (intro continuous_intros) auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
228 |
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
|
229 |
by (fastforce simp: assms(2) subspace_mul) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
230 |
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
|
231 |
apply (rule_tac c="-d" in that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
232 |
apply (rule homotopic_with_eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
233 |
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
|
234 |
using d apply (auto simp: h_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
235 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
236 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
237 |
apply (rule_tac x=c in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
238 |
apply (rule homotopic_with_trans [OF _ homhc]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
239 |
apply (rule homotopic_with_eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
240 |
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
|
241 |
apply (auto simp: h_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
242 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
243 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
246 |
lemma spheremap_lemma3: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
247 |
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
|
248 |
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
|
249 |
"(rel_frontier S) homeomorphic (sphere 0 1 \<inter> T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
250 |
proof (cases "S = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
251 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
252 |
with \<open>subspace U\<close> subspace_0 show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
253 |
by (rule_tac T = "{0}" in that) auto
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
254 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
255 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
256 |
then obtain a where "a \<in> S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
257 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
258 |
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
|
259 |
by (metis hull_inc aff_dim_eq_dim) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
260 |
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
|
261 |
by linarith |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
262 |
with choose_subspace_of_subspace |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
263 |
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
|
264 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
265 |
proof (rule that [OF \<open>subspace T\<close>]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
266 |
show "T \<subseteq> U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
267 |
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
|
268 |
show "aff_dim T = aff_dim S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
269 |
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
|
270 |
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
|
271 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
272 |
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
|
273 |
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
|
274 |
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
|
275 |
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
|
276 |
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
|
277 |
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
|
278 |
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
|
279 |
apply (simp add: bounded_Int) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
280 |
apply (rule affS_eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
281 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
282 |
also have "... = frontier (ball 0 1) \<inter> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
283 |
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
|
284 |
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
|
285 |
using \<open>subspace T\<close> subspace_0 by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
286 |
also have "... = sphere 0 1 \<inter> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
287 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
288 |
finally show ?thesis . |
|
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 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
294 |
proposition inessential_spheremap_lowdim_gen: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
295 |
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
296 |
assumes "convex S" "bounded S" "convex T" "bounded T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
297 |
and affST: "aff_dim S < aff_dim T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
298 |
and contf: "continuous_on (rel_frontier S) f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
299 |
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
|
300 |
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
|
301 |
proof (cases "S = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
302 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
303 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
304 |
by (simp add: that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
305 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
306 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
307 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
308 |
proof (cases "T = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
309 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
310 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
311 |
using fim that by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
312 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
313 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
314 |
obtain T':: "'a set" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
315 |
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
|
316 |
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
|
317 |
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
|
318 |
apply (simp add: dim_UNIV aff_dim_le_DIM) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
319 |
using \<open>T \<noteq> {}\<close> by blast
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
320 |
with homeomorphic_imp_homotopy_eqv |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
321 |
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
|
322 |
using homotopy_eqv_sym by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
323 |
have "aff_dim S \<le> int (dim T')" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
324 |
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
|
325 |
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
|
326 |
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
|
327 |
and affS': "aff_dim S' = aff_dim S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
328 |
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
|
329 |
by metis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
330 |
with homeomorphic_imp_homotopy_eqv |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
331 |
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
|
332 |
using homotopy_eqv_sym by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
333 |
have dimST': "dim S' < dim T'" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
334 |
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
|
335 |
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
|
336 |
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
|
337 |
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
|
338 |
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
|
339 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
340 |
with that show ?thesis by blast |
|
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 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
343 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
344 |
lemma inessential_spheremap_lowdim: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
345 |
fixes f :: "'M::euclidean_space \<Rightarrow> 'a::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
346 |
assumes |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
347 |
"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
|
348 |
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
|
349 |
proof (cases "s \<le> 0") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
350 |
case True then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
351 |
by (meson nullhomotopic_into_contractible f contractible_sphere that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
352 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
353 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
354 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
355 |
proof (cases "r \<le> 0") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
356 |
case True then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
357 |
by (meson f nullhomotopic_from_contractible contractible_sphere that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
358 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
359 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
360 |
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
|
361 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
362 |
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
|
363 |
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
|
364 |
apply (simp_all add: f aff_dim_cball) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
365 |
using that by blast |
|
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 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
371 |
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
|
372 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
373 |
lemma extending_maps_Union_aux: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
374 |
assumes fin: "finite \<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
375 |
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
376 |
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
|
377 |
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
|
378 |
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
|
379 |
using assms |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
380 |
proof (induction \<F>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
381 |
case empty show ?case by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
382 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
383 |
case (insert S \<F>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
384 |
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
|
385 |
by (meson insertI1) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
386 |
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
|
387 |
using insert by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
388 |
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
|
389 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
390 |
have "T \<inter> S \<subseteq> K \<or> S = T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
391 |
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
|
392 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
393 |
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
|
394 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
395 |
show ?case |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
396 |
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
|
397 |
apply (intro conjI continuous_on_cases) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
398 |
apply (simp_all add: insert closed_Union contf contg) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
399 |
using fim gim feq geq |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
400 |
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
|
401 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
402 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
403 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
404 |
lemma extending_maps_Union: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
405 |
assumes fin: "finite \<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
406 |
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
|
407 |
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
408 |
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
|
409 |
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
|
410 |
apply (simp add: Union_maximal_sets [OF fin, symmetric]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
411 |
apply (rule extending_maps_Union_aux) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
412 |
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
|
413 |
by (metis K psubsetI) |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
416 |
lemma extend_map_lemma: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
417 |
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
|
418 |
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
|
419 |
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
|
420 |
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
|
421 |
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
|
422 |
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
|
423 |
proof (cases "\<F> - \<G> = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
424 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
425 |
then have "\<Union>\<F> \<subseteq> \<Union>\<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
426 |
by (simp add: Union_mono) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
427 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
428 |
apply (rule_tac g=f in that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
429 |
using contf continuous_on_subset apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
430 |
using fim apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
431 |
by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
432 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
433 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
434 |
then have "0 \<le> aff_dim T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
435 |
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
|
436 |
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
|
437 |
by (metis nonneg_eq_int) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
438 |
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
|
439 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
440 |
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
|
441 |
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
|
442 |
(\<forall>x \<in> \<Union>\<G>. g x = f x)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
443 |
if "i \<le> aff_dim T" for i::nat |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
444 |
using that |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
445 |
proof (induction i) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
446 |
case 0 then show ?case |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
447 |
apply (simp add: Union_empty_eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
448 |
apply (rule_tac x=f in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
449 |
apply (intro conjI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
450 |
using contf continuous_on_subset apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
451 |
using fim apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
452 |
by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
453 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
454 |
case (Suc p) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
455 |
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
|
456 |
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
|
457 |
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
|
458 |
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
|
459 |
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
|
460 |
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
|
461 |
\<subseteq> rel_frontier T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
462 |
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
|
463 |
using Suc.IH [OF ple] by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
464 |
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
|
465 |
have extendh: "\<exists>g. continuous_on D g \<and> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
466 |
g ` D \<subseteq> rel_frontier T \<and> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
467 |
(\<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
|
468 |
if D: "D \<in> \<G> \<union> ?Faces" for D |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
469 |
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
|
470 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
471 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
472 |
apply (rule_tac x=h in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
473 |
apply (intro conjI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
474 |
apply (blast intro: continuous_on_subset [OF conth]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
475 |
using him apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
476 |
by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
477 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
478 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
479 |
note notDsub = False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
480 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
481 |
proof (cases "\<exists>a. D = {a}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
482 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
483 |
then obtain a where "D = {a}" by auto
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
484 |
with notDsub t show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
485 |
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
|
486 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
487 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
488 |
have "D \<noteq> {}" using notDsub by auto
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
489 |
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
|
490 |
using notDsub by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
491 |
then have "D \<notin> \<G>" by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
492 |
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
|
493 |
using Dnotin that by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
494 |
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
|
495 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
496 |
then have "bounded D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
497 |
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
|
498 |
then have [simp]: "\<not> affine D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
499 |
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
|
500 |
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
|
501 |
apply clarify |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
502 |
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
|
503 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
504 |
moreover have "polyhedron D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
505 |
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
|
506 |
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
|
507 |
by (simp add: rel_frontier_of_polyhedron Union_mono) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
508 |
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
|
509 |
using \<open>C \<in> \<F>\<close> him by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
510 |
have "convex D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
511 |
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
|
512 |
have affD_lessT: "aff_dim D < aff_dim T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
513 |
using Suc.prems affD by linarith |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
514 |
have contDh: "continuous_on (rel_frontier D) h" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
515 |
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
|
516 |
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
|
517 |
(\<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
|
518 |
(\<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
|
519 |
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
|
520 |
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
|
521 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
522 |
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
|
523 |
(rel_frontier T) h (\<lambda>x. c))" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
524 |
by (metis inessential_spheremap_lowdim_gen |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
525 |
[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
|
526 |
then obtain g where contg: "continuous_on UNIV g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
527 |
and gim: "range g \<subseteq> rel_frontier T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
528 |
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
|
529 |
by (metis *) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
530 |
have "D \<inter> E \<subseteq> rel_frontier D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
531 |
if "E \<in> \<G> \<union> {D. Bex \<F> (op face_of D) \<and> aff_dim D < int p}" for E
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
532 |
proof (rule face_of_subset_rel_frontier) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
533 |
show "D \<inter> E face_of D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
534 |
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
|
535 |
apply auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
536 |
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
|
537 |
using face_of_Int_subface apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
538 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
539 |
show "D \<inter> E \<noteq> D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
540 |
using that notDsub by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
541 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
542 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
543 |
apply (rule_tac x=g in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
544 |
apply (intro conjI ballI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
545 |
using continuous_on_subset contg apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
546 |
using gim apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
547 |
using gh by fastforce |
|
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 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
550 |
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
|
551 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
552 |
have "finite \<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
553 |
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
|
554 |
then have fin: "finite (\<G> \<union> ?Faces)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
555 |
apply simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
556 |
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
|
557 |
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
|
558 |
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
|
559 |
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
|
560 |
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
|
561 |
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
|
562 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
563 |
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
|
564 |
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
|
565 |
apply (rule face_of_Int_subface [OF _ _ XY]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
566 |
apply (auto simp: face DE) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
567 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
568 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
569 |
using that |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
570 |
apply auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
571 |
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
|
572 |
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
|
573 |
apply (fastforce dest: face_of_aff_dim_lt) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
574 |
by (meson face_of_trans ff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
575 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
576 |
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
|
577 |
"g ` \<Union>(\<G> \<union> ?Faces) \<subseteq> rel_frontier T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
578 |
"(\<forall>x \<in> \<Union>(\<G> \<union> ?Faces) \<inter> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
579 |
\<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
|
580 |
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
|
581 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
582 |
then show ?case |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
583 |
apply (simp add: intle local.heq [symmetric], blast) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
584 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
585 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
586 |
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
|
587 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
588 |
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
|
589 |
apply (rule Union_subsetI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
590 |
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
|
591 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
592 |
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
|
593 |
apply (rule Union_mono) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
594 |
using face apply (fastforce simp: aff i) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
595 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
596 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
597 |
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
|
598 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
599 |
using extendf [of i] unfolding eq by (metis that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
600 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
601 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
602 |
lemma extend_map_lemma_cofinite0: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
603 |
assumes "finite \<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
604 |
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
|
605 |
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
|
606 |
and "\<And>S. S \<in> \<F> \<Longrightarrow> closed S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
607 |
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
|
608 |
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
|
609 |
\<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
|
610 |
using assms |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
611 |
proof induction |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
612 |
case empty then show ?case |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
613 |
by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
614 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
615 |
case (insert X \<F>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
616 |
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
|
617 |
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
|
618 |
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
|
619 |
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
|
620 |
by (simp_all add: pairwise_insert) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
621 |
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
|
622 |
and contg: "continuous_on (\<Union>\<F> - C) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
623 |
and gim: "g ` (\<Union>\<F> - C) \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
624 |
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
|
625 |
using insert.IH [OF pwF \<F> clo] by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
626 |
obtain a f where "a \<notin> U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
627 |
and contf: "continuous_on (X - {a}) f"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
628 |
and fim: "f ` (X - {a}) \<subseteq> T"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
629 |
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
|
630 |
using insert.prems by (meson insertI1) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
631 |
show ?case |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
632 |
proof (intro exI conjI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
633 |
show "finite (insert a C)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
634 |
by (simp add: C) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
635 |
show "disjnt (insert a C) U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
636 |
using C \<open>a \<notin> U\<close> by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
637 |
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
|
638 |
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
|
639 |
have "closed (\<Union>\<F>)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
640 |
using clo insert.hyps by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
641 |
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
|
642 |
apply (rule continuous_on_cases_local) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
643 |
apply (simp_all add: closedin_closed) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
644 |
using \<open>closed X\<close> apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
645 |
using \<open>closed (\<Union>\<F>)\<close> apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
646 |
using contf apply (force simp: elim: continuous_on_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
647 |
using contg apply (force simp: elim: continuous_on_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
648 |
using fh gh insert.hyps pwX by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
649 |
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
|
650 |
by (blast intro: continuous_on_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
651 |
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
|
652 |
using gh by (auto simp: fh) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
653 |
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
|
654 |
using fim gim by auto force |
|
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 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
659 |
lemma extend_map_lemma_cofinite1: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
660 |
assumes "finite \<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
661 |
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
|
662 |
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
|
663 |
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
|
664 |
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
|
665 |
"g ` (\<Union>\<F> - C) \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
666 |
"\<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
|
667 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
668 |
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
|
669 |
have [simp]: "\<Union>?\<F> = \<Union>\<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
670 |
by (simp add: Union_maximal_sets assms) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
671 |
have fin: "finite ?\<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
672 |
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
|
673 |
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
|
674 |
by (simp add: pairwise_def) (metis K psubsetI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
675 |
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
|
676 |
by (simp add: \<open>finite \<F>\<close> card_mono) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
677 |
moreover |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
678 |
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
|
679 |
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
|
680 |
\<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
|
681 |
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
|
682 |
apply (fastforce intro!: clo \<F>)+ |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
683 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
684 |
ultimately show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
685 |
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
|
686 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
689 |
lemma extend_map_lemma_cofinite: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
690 |
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
|
691 |
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
|
692 |
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
|
693 |
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
|
694 |
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
|
695 |
obtains C g where |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
696 |
"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
|
697 |
"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
|
698 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
699 |
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
|
700 |
have "finite \<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
701 |
using assms finite_subset by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
702 |
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
|
703 |
apply (rule finite_Union) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
704 |
apply (simp add: \<open>finite \<F>\<close>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
705 |
using finite_polytope_faces poly by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
706 |
ultimately have "finite \<H>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
707 |
apply (simp add: \<H>_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
708 |
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
|
709 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
710 |
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
|
711 |
unfolding \<H>_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
712 |
apply (elim UnE bexE CollectE DiffE) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
713 |
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
|
714 |
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
|
715 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
716 |
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
|
717 |
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
|
718 |
using \<open>finite \<H>\<close> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
719 |
unfolding \<H>_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
720 |
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
|
721 |
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
|
722 |
using * apply (auto simp: \<H>_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
723 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
724 |
have "bounded (\<Union>\<G>)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
725 |
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
|
726 |
then have "\<Union>\<G> \<noteq> UNIV" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
727 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
728 |
then obtain a where a: "a \<notin> \<Union>\<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
729 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
730 |
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
|
731 |
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
|
732 |
if "D \<in> \<F>" for D |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
733 |
proof (cases "D \<subseteq> \<Union>\<H>") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
734 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
735 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
736 |
apply (rule_tac x=a in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
737 |
apply (rule_tac x=h in exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
738 |
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
|
739 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
740 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
741 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
742 |
note D_not_subset = False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
743 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
744 |
proof (cases "D \<in> \<G>") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
745 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
746 |
with D_not_subset show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
747 |
by (auto simp: \<H>_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
748 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
749 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
750 |
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
|
751 |
by (simp add: \<open>D \<in> \<F>\<close> aff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
752 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
753 |
proof (cases "rel_interior D = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
754 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
755 |
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
|
756 |
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
|
757 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
758 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
759 |
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
|
760 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
761 |
have "polyhedron D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
762 |
by (simp add: poly polytope_imp_polyhedron that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
763 |
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
|
764 |
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
|
765 |
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
|
766 |
and contr: "continuous_on (affine hull D - {b}) r"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
767 |
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
|
768 |
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
|
769 |
by (auto simp: retract_of_def retraction_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
770 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
771 |
proof (intro exI conjI ballI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
772 |
show "b \<notin> \<Union>\<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
773 |
proof clarify |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
774 |
fix E |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
775 |
assume "b \<in> E" "E \<in> \<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
776 |
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
|
777 |
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
|
778 |
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
|
779 |
D_not_subset rel_frontier_def \<H>_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
780 |
show False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
781 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
782 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
783 |
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
|
784 |
by (simp add: Diff_mono hull_subset image_mono) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
785 |
also have "... \<subseteq> rel_frontier D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
786 |
by (rule rim) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
787 |
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
|
788 |
using affD |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
789 |
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
|
790 |
also have "... \<subseteq> \<Union>(\<H>)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
791 |
using D_not_subset \<H>_def that by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
792 |
finally have rsub: "r ` (D - {b}) \<subseteq> \<Union>(\<H>)" .
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
793 |
show "continuous_on (D - {b}) (h \<circ> r)"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
794 |
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
|
795 |
apply (rule continuous_on_subset [OF contr]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
796 |
apply (simp add: Diff_mono hull_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
797 |
apply (rule continuous_on_subset [OF conth rsub]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
798 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
799 |
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
|
800 |
using brelD him rsub by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
801 |
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
|
802 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
803 |
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
|
804 |
| 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
|
805 |
using x by (auto simp: \<H>_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
806 |
then have xrel: "x \<in> rel_frontier D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
807 |
proof cases |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
808 |
case 1 show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
809 |
proof (rule face_of_subset_rel_frontier [THEN subsetD]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
810 |
show "D \<inter> A face_of D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
811 |
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
|
812 |
show "D \<inter> A \<noteq> D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
813 |
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
|
814 |
qed (auto simp: 1) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
815 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
816 |
case 2 show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
817 |
proof (rule face_of_subset_rel_frontier [THEN subsetD]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
818 |
show "D \<inter> A face_of D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
819 |
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
|
820 |
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
|
821 |
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
|
822 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
823 |
show "D \<inter> A \<noteq> D" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
824 |
using "2" D_not_subset \<H>_def by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
825 |
qed (auto simp: 2) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
826 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
827 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
828 |
by (simp add: rid xrel) |
|
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 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
834 |
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
|
835 |
by (simp add: poly polytope_imp_closed) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
836 |
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
|
837 |
"g ` (\<Union>\<F> - C) \<subseteq> rel_frontier T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
838 |
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
|
839 |
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
|
840 |
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
|
841 |
proof (cases "X \<in> \<G>") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
842 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
843 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
844 |
by (auto simp: \<H>_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
845 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
846 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
847 |
have "X \<inter> Y \<noteq> X" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
848 |
using \<open>\<not> X \<subseteq> Y\<close> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
849 |
with XY |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
850 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
851 |
by (clarsimp simp: \<H>_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
852 |
(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
|
853 |
not_le poly polytope_imp_convex) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
854 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
855 |
qed (blast)+ |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
856 |
with \<open>\<G> \<subseteq> \<F>\<close> show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
857 |
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
|
858 |
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
|
859 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
860 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
861 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
862 |
text\<open>The next two proofs are similar\<close> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
863 |
theorem extend_map_cell_complex_to_sphere: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
864 |
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
|
865 |
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
|
866 |
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
|
867 |
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
|
868 |
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
|
869 |
obtains g where "continuous_on (\<Union>\<F>) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
870 |
"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
|
871 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
872 |
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
|
873 |
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
|
874 |
have "compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
875 |
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
|
876 |
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
|
877 |
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
|
878 |
obtain \<G> where "finite \<G>" "\<Union>\<G> = \<Union>\<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
879 |
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
|
880 |
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
|
881 |
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
|
882 |
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
|
883 |
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
|
884 |
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
|
885 |
by (simp add: aff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
886 |
qed auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
887 |
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
|
888 |
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
|
889 |
show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
890 |
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
|
891 |
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
|
892 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
893 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
894 |
show "continuous_on (\<Union>\<F>) h" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
895 |
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
|
896 |
show "h ` \<Union>\<F> \<subseteq> rel_frontier T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
897 |
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
|
898 |
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
|
899 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
900 |
have "x \<in> \<Union>\<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
901 |
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
|
902 |
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
|
903 |
then have "diameter X < d" "bounded X" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
904 |
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
|
905 |
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
|
906 |
by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
907 |
have "h x = g x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
908 |
apply (rule hg) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
909 |
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
|
910 |
also have "... = f x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
911 |
by (simp add: gf that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
912 |
finally show "h x = f x" . |
|
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 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
918 |
theorem extend_map_cell_complex_to_sphere_cofinite: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
919 |
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
|
920 |
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
|
921 |
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
|
922 |
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
|
923 |
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
|
924 |
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
|
925 |
"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
|
926 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
927 |
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
|
928 |
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
|
929 |
have "compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
930 |
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
|
931 |
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
|
932 |
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
|
933 |
obtain \<G> where "finite \<G>" "\<Union>\<G> = \<Union>\<F>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
934 |
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
|
935 |
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
|
936 |
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
|
937 |
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
|
938 |
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
|
939 |
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
|
940 |
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
|
941 |
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
|
942 |
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
|
943 |
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
|
944 |
show "continuous_on (\<Union>(\<G> \<inter> Pow V)) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
945 |
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
|
946 |
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
|
947 |
using gim by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
948 |
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
|
949 |
have Ssub: "S \<subseteq> \<Union>(\<G> \<inter> Pow V)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
950 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
951 |
fix x |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
952 |
assume "x \<in> S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
953 |
then have "x \<in> \<Union>\<G>" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
954 |
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
|
955 |
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
|
956 |
then have "diameter X < d" "bounded X" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
957 |
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
|
958 |
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
|
959 |
by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
960 |
then show "x \<in> \<Union>(\<G> \<inter> Pow V)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
961 |
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
|
962 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
963 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
964 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
965 |
show "continuous_on (\<Union>\<F>-C) h" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
966 |
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
|
967 |
show "h ` (\<Union>\<F> - C) \<subseteq> rel_frontier T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
968 |
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
|
969 |
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
|
970 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
971 |
have "h x = g x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
972 |
apply (rule hg) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
973 |
using Ssub that by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
974 |
also have "... = f x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
975 |
by (simp add: gf that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
976 |
finally show "h x = f x" . |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
977 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
978 |
show "disjnt C S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
979 |
using dis Ssub by (meson disjnt_iff subset_eq) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
980 |
qed (intro \<open>finite C\<close>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
981 |
qed |
|
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 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
985 |
subsection\<open> Special cases and corollaries involving spheres.\<close> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
986 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
987 |
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
|
988 |
by (auto simp: disjnt_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
989 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
990 |
proposition extend_map_affine_to_sphere_cofinite_simple: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
991 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
992 |
assumes "compact S" "convex U" "bounded U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
993 |
and aff: "aff_dim T \<le> aff_dim U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
994 |
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
|
995 |
and fim: "f ` S \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
996 |
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
|
997 |
"g ` (T - K) \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
998 |
"\<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
|
999 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1000 |
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
|
1001 |
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
|
1002 |
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
|
1003 |
proof (cases "S = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1004 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1005 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1006 |
proof (cases "rel_frontier U = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1007 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1008 |
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
|
1009 |
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
|
1010 |
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
|
1011 |
then obtain a where "T \<subseteq> {a}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1012 |
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
|
1013 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1014 |
using \<open>S = {}\<close> fim
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1015 |
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
|
1016 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1017 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1018 |
then obtain a where "a \<in> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1019 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1020 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1021 |
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
|
1022 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1023 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1024 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1025 |
have "bounded S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1026 |
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
|
1027 |
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
|
1028 |
using bounded_subset_cbox_symmetric by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1029 |
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
|
1030 |
have "cbox (-b) b \<subseteq> bbox" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1031 |
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
|
1032 |
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
|
1033 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1034 |
then have Ssub: "S \<subseteq> \<Union>{bbox \<inter> T}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1035 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1036 |
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
|
1037 |
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
|
1038 |
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
|
1039 |
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
|
1040 |
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
|
1041 |
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
|
1042 |
proof (rule extend_map_cell_complex_to_sphere_cofinite |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1043 |
[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
|
1044 |
show "closed S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1045 |
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
|
1046 |
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
|
1047 |
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
|
1048 |
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
|
1049 |
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
|
1050 |
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
|
1051 |
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
|
1052 |
qed auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1053 |
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
|
1054 |
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
|
1055 |
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
|
1056 |
show "infinite {1/2..1::real}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1057 |
by (simp add: infinite_Icc) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1058 |
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
|
1059 |
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
|
1060 |
then show "disjoint_family_on fro {1/2..1}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1061 |
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
|
1062 |
qed auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1063 |
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
|
1064 |
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
|
1065 |
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
|
1066 |
have clo_cbT: "closed (cbox (- c) c \<inter> T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1067 |
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
|
1068 |
have cpT_ne: "cbox (- c) c \<inter> T \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1069 |
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
|
1070 |
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
|
1071 |
proof (cases "x \<in> cbox (-c) c") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1072 |
case True with that show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1073 |
by (simp add: closest_point_self) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1074 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1075 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1076 |
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
|
1077 |
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
|
1078 |
have "convex T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1079 |
by (meson \<open>affine T\<close> affine_imp_convex) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1080 |
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
|
1081 |
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
|
1082 |
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
|
1083 |
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
|
1084 |
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
|
1085 |
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
|
1086 |
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
|
1087 |
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
|
1088 |
apply (auto simp: fro_def c_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1089 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1090 |
ultimately show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1091 |
using dd by (force simp: disjnt_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1092 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1093 |
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
|
1094 |
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
|
1095 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1096 |
proof (intro conjI ballI exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1097 |
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
|
1098 |
apply (rule continuous_on_closest_point) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1099 |
using \<open>S \<noteq> {}\<close> cbsub(2) b that
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1100 |
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
|
1101 |
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
|
1102 |
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
|
1103 |
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
|
1104 |
by (metis image_comp image_mono cpt_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1105 |
also have "... \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1106 |
by (rule gim) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1107 |
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
|
1108 |
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
|
1109 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1110 |
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
|
1111 |
unfolding o_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1112 |
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
|
1113 |
also have "... = f x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1114 |
by (simp add: that gf) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1115 |
finally show ?thesis . |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1116 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1117 |
qed (auto simp: K) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1118 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1119 |
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
|
1120 |
and contg: "continuous_on (affine hull T - K) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1121 |
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
|
1122 |
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
|
1123 |
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
|
1124 |
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
|
1125 |
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
|
1126 |
and contg: "continuous_on (T - K) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1127 |
and gim: "g ` (T - K) \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1128 |
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
|
1129 |
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
|
1130 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1131 |
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
|
1132 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1133 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1134 |
subsection\<open>Extending maps to spheres\<close> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1135 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1136 |
(*Up to extend_map_affine_to_sphere_cofinite_gen*) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1137 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1138 |
lemma closedin_closed_subset: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1139 |
"\<lbrakk>closedin (subtopology euclidean U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1140 |
\<Longrightarrow> closedin (subtopology euclidean T) S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1141 |
by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1142 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1143 |
lemma extend_map_affine_to_sphere1: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1144 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::topological_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1145 |
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
|
1146 |
and fim: "f ` (U - K) \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1147 |
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
|
1148 |
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
|
1149 |
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
|
1150 |
proof (cases "K = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1151 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1152 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1153 |
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
|
1154 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1155 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1156 |
have "S \<subseteq> U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1157 |
using clo closedin_limpt by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1158 |
then have "(U - S) \<inter> K \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1159 |
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
|
1160 |
then have "\<Union>(components (U - S)) \<inter> K \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1161 |
using Union_components by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1162 |
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
|
1163 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1164 |
have "convex U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1165 |
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
|
1166 |
then have "locally connected U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1167 |
by (rule convex_imp_locally_connected) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1168 |
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
|
1169 |
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
|
1170 |
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
|
1171 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1172 |
have "C \<subseteq> U-S" "C \<inter> L \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1173 |
by (simp_all add: in_components_subset comps that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1174 |
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
|
1175 |
have opeUC: "openin (subtopology euclidean U) C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1176 |
proof (rule openin_trans) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1177 |
show "openin (subtopology euclidean (U-S)) C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1178 |
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
|
1179 |
show "openin (subtopology euclidean U) (U - S)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1180 |
by (simp add: clo openin_diff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1181 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1182 |
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
|
1183 |
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
|
1184 |
then have "ball a d \<inter> U \<subseteq> C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1185 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1186 |
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
|
1187 |
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
|
1188 |
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
|
1189 |
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
|
1190 |
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
|
1191 |
show "openin (subtopology euclidean C) (ball a d \<inter> U)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1192 |
by (metis Topology_Euclidean_Space.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) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1193 |
show "openin (subtopology euclidean (affine hull C)) C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1194 |
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
|
1195 |
show "ball a d \<inter> U \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1196 |
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
|
1197 |
show "finite (C \<inter> K)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1198 |
by (simp add: \<open>finite K\<close>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1199 |
show "S \<union> C \<subseteq> affine hull C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1200 |
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
|
1201 |
show "connected C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1202 |
by (metis C in_components_connected) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1203 |
qed auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1204 |
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
|
1205 |
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
|
1206 |
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
|
1207 |
apply (rule rel_frontier_retract_of_punctured_affine_hull) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1208 |
apply (auto simp: \<open>convex U\<close> convex_Int) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1209 |
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
|
1210 |
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
|
1211 |
apply (rule convex_affine_rel_frontier_Int) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1212 |
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
|
1213 |
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
|
1214 |
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
|
1215 |
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
|
1216 |
by metis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1217 |
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
|
1218 |
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
|
1219 |
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
|
1220 |
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
|
1221 |
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
|
1222 |
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
|
1223 |
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
|
1224 |
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
|
1225 |
using \<open>0 < d\<close> by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1226 |
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
|
1227 |
proof clarify |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1228 |
fix y assume "y \<in> S \<union> (C - {a})"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1229 |
then have "y \<in> U - {a}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1230 |
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
|
1231 |
then have "r y \<in> sphere a d" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1232 |
using rim by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1233 |
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
|
1234 |
apply (simp add: j_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1235 |
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
|
1236 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1237 |
have contj: "continuous_on (U - {a}) j"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1238 |
unfolding j_def Uaeq |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1239 |
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
|
1240 |
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
|
1241 |
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
|
1242 |
using affine_closed \<open>affine U\<close> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1243 |
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
|
1244 |
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
|
1245 |
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
|
1246 |
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
|
1247 |
by (force intro: continuous_on_subset [OF contr]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1248 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1249 |
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
|
1250 |
using fim by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1251 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1252 |
proof (intro conjI exI) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1253 |
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
|
1254 |
proof (intro continuous_on_compose) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1255 |
show "continuous_on (S \<union> (C - {a})) j"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1256 |
apply (rule continuous_on_subset [OF contj]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1257 |
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
|
1258 |
show "continuous_on (j ` (S \<union> (C - {a}))) k"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1259 |
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
|
1260 |
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
|
1261 |
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
|
1262 |
proof (clarify intro!: continuous_on_subset [OF contf]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1263 |
fix y assume "y \<in> S \<union> (C - {a})"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1264 |
have ky: "k y \<in> S \<union> C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1265 |
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
|
1266 |
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
|
1267 |
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
|
1268 |
show "k (j y) \<in> U - K" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1269 |
apply safe |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1270 |
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
|
1271 |
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
|
1272 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1273 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1274 |
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
|
1275 |
apply (simp add: kj) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1276 |
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
|
1277 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1278 |
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
|
1279 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1280 |
have rx: "r x \<in> sphere a d" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1281 |
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
|
1282 |
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
|
1283 |
using jim that by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1284 |
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
|
1285 |
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
|
1286 |
then have "k (j x) \<in> C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1287 |
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
|
1288 |
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
|
1289 |
with jj \<open>C \<subseteq> U\<close> show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1290 |
apply safe |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1291 |
using ST j_def apply fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1292 |
apply (auto simp: not_less intro!: fT) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1293 |
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
|
1294 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1295 |
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
|
1296 |
by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1297 |
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
|
1298 |
qed (auto simp: a) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1299 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1300 |
then obtain a h where |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1301 |
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
|
1302 |
\<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
|
1303 |
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
|
1304 |
using that by metis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1305 |
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
|
1306 |
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
|
1307 |
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
|
1308 |
have "C0 \<in> F" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1309 |
by (auto simp: F_def C0) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1310 |
have "finite F" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1311 |
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
|
1312 |
show "inj_on (\<lambda>C. C \<inter> K) F" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1313 |
unfolding F_def inj_on_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1314 |
using components_nonoverlap by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1315 |
show "finite ((\<lambda>C. C \<inter> K) ` F)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1316 |
unfolding F_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1317 |
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
|
1318 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1319 |
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
|
1320 |
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
|
1321 |
\<Longrightarrow> g x = h i x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1322 |
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
|
1323 |
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
|
1324 |
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
|
1325 |
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
|
1326 |
if "C \<in> F" for C |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1327 |
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
|
1328 |
show "closedin (subtopology euclidean U) (S \<union> C)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1329 |
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
|
1330 |
using F_def that by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1331 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1332 |
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
|
1333 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1334 |
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
|
1335 |
using \<open>x \<in> C'\<close> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1336 |
with that show "x = a C'" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1337 |
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
|
1338 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1339 |
then show "S \<union> UF \<subseteq> U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1340 |
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
|
1341 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1342 |
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
|
1343 |
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
|
1344 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1345 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1346 |
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
|
1347 |
using ah F_def that by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1348 |
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
|
1349 |
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
|
1350 |
\<Longrightarrow> h i x = h j x" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1351 |
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
|
1352 |
qed blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1353 |
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
|
1354 |
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
|
1355 |
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
|
1356 |
proof (rule closedin_closed_subset [OF _ SU']) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1357 |
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
|
1358 |
unfolding F_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1359 |
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
|
1360 |
show "closedin (subtopology euclidean U) (U - UF)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1361 |
unfolding UF_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1362 |
by (force intro: openin_delete *) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1363 |
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
|
1364 |
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
|
1365 |
apply (metis Diff_iff UnionI Union_components) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1366 |
apply (metis DiffD1 UnionI Union_components) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1367 |
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
|
1368 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1369 |
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
|
1370 |
proof (rule closedin_closed_subset [OF _ SU']) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1371 |
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
|
1372 |
apply (rule closedin_Union) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1373 |
apply (simp add: \<open>finite F\<close>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1374 |
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
|
1375 |
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
|
1376 |
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
|
1377 |
using C0 apply blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1378 |
by (metis components_nonoverlap disjnt_def disjnt_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1379 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1380 |
have SUG: "S \<union> \<Union>G \<subseteq> U - K" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1381 |
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
|
1382 |
by (meson Diff_iff subsetD in_components_subset) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1383 |
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
|
1384 |
by (rule continuous_on_subset [OF contf]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1385 |
have contg': "continuous_on (S \<union> UF) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1386 |
apply (rule continuous_on_subset [OF contg]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1387 |
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
|
1388 |
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
|
1389 |
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
|
1390 |
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
|
1391 |
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
|
1392 |
using components_eq by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1393 |
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
|
1394 |
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
|
1395 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1396 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1397 |
have UF: "\<Union>F - L \<subseteq> UF" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1398 |
unfolding F_def UF_def using ah by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1399 |
have "U - S - L = \<Union>(components (U - S)) - L" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1400 |
by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1401 |
also have "... = \<Union>F \<union> \<Union>G - L" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1402 |
unfolding F_def G_def by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1403 |
also have "... \<subseteq> UF \<union> \<Union>G" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1404 |
using UF by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1405 |
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
|
1406 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1407 |
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
|
1408 |
by (rule continuous_on_subset [OF cont]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1409 |
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
|
1410 |
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
|
1411 |
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
|
1412 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1413 |
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
|
1414 |
proof (subst gh) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1415 |
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
|
1416 |
using that by (auto simp: UF_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1417 |
show "h C x \<in> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1418 |
using ah that by (fastforce simp add: F_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1419 |
qed (rule that) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1420 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1421 |
by (force simp: UF_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1422 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1423 |
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
|
1424 |
using image_mono order_trans by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1425 |
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
|
1426 |
using fim SUG by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1427 |
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
|
1428 |
by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1429 |
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
|
1430 |
by (simp add: F_def G_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1431 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1432 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1433 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1434 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1435 |
lemma extend_map_affine_to_sphere2: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1436 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1437 |
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
|
1438 |
and affTU: "aff_dim T \<le> aff_dim U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1439 |
and contf: "continuous_on S f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1440 |
and fim: "f ` S \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1441 |
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
|
1442 |
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
|
1443 |
"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
|
1444 |
"\<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
|
1445 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1446 |
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
|
1447 |
and contg: "continuous_on (T - K) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1448 |
and gim: "g ` (T - K) \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1449 |
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
|
1450 |
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
|
1451 |
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
|
1452 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1453 |
have "x \<in> T-S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1454 |
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
|
1455 |
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
|
1456 |
by (metis UnionE Union_components) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1457 |
with ovlap [of C] show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1458 |
by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1459 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1460 |
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
|
1461 |
by metis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1462 |
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
|
1463 |
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
|
1464 |
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
|
1465 |
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
|
1466 |
show cloTS: "closedin (subtopology euclidean T) S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1467 |
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
|
1468 |
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
|
1469 |
using \<xi> components_eq by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1470 |
qed (use K in auto) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1471 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1472 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1473 |
show *: "\<xi> ` K \<subseteq> L" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1474 |
using \<xi> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1475 |
show "finite (\<xi> ` K)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1476 |
by (simp add: K) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1477 |
show "\<xi> ` K \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1478 |
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
|
1479 |
show "continuous_on (T - \<xi> ` K) h" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1480 |
by (rule conth) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1481 |
show "disjnt (\<xi> ` K) S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1482 |
using K |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1483 |
apply (auto simp: disjnt_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1484 |
by (metis \<xi> DiffD2 UnionI Union_components) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1485 |
qed (simp_all add: him hg gf) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1486 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1487 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1488 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1489 |
proposition extend_map_affine_to_sphere_cofinite_gen: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1490 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1491 |
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
|
1492 |
and aff: "aff_dim T \<le> aff_dim U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1493 |
and contf: "continuous_on S f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1494 |
and fim: "f ` S \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1495 |
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
|
1496 |
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
|
1497 |
"g ` (T - K) \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1498 |
"\<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
|
1499 |
proof (cases "S = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1500 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1501 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1502 |
proof (cases "rel_frontier U = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1503 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1504 |
with aff have "aff_dim T \<le> 0" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1505 |
apply (simp add: rel_frontier_eq_empty) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1506 |
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
|
1507 |
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
|
1508 |
by linarith |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1509 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1510 |
proof cases |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1511 |
assume "aff_dim T = -1" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1512 |
then have "T = {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1513 |
by (simp add: aff_dim_empty) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1514 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1515 |
by (rule_tac K="{}" in that) auto
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1516 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1517 |
assume "aff_dim T = 0" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1518 |
then obtain a where "T = {a}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1519 |
using aff_dim_eq_0 by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1520 |
then have "a \<in> L" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1521 |
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
|
1522 |
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
|
1523 |
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
|
1524 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1525 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1526 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1527 |
then obtain y where "y \<in> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1528 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1529 |
with \<open>S = {}\<close> show ?thesis
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1530 |
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
|
1531 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1532 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1533 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1534 |
have "bounded S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1535 |
by (simp add: assms compact_imp_bounded) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1536 |
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
|
1537 |
using bounded_subset_cbox_symmetric by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1538 |
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
|
1539 |
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
|
1540 |
and contg: "continuous_on (T - K) g" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1541 |
and gim: "g ` (T - K) \<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1542 |
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
|
1543 |
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
|
1544 |
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
|
1545 |
proof (cases "bounded C") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1546 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1547 |
with dis that show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1548 |
unfolding LU_def by fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1549 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1550 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1551 |
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
|
1552 |
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
|
1553 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1554 |
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
|
1555 |
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
|
1556 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1557 |
qed blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1558 |
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
|
1559 |
"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
|
1560 |
"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
|
1561 |
using that by (auto simp: mem_box algebra_simps) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1562 |
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
|
1563 |
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
|
1564 |
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
|
1565 |
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
|
1566 |
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
|
1567 |
by (auto simp: \<open>finite K\<close>) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1568 |
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
|
1569 |
have cbsub: "cbox (-b) b \<subseteq> box (-c) c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1570 |
"cbox (-b) b \<subseteq> cbox (-c) c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1571 |
"cbox (-c) c \<subseteq> cbox (-(b+One)) (b+One)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1572 |
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
|
1573 |
have clo_cT: "closed (cbox (- c) c \<inter> T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1574 |
using affine_closed \<open>affine T\<close> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1575 |
have cT_ne: "cbox (- c) c \<inter> T \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1576 |
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
|
1577 |
have S_sub_cc: "S \<subseteq> cbox (- c) c" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1578 |
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
|
1579 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1580 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1581 |
show "finite (K \<inter> cbox (-(b+One)) (b+One))" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1582 |
using \<open>finite K\<close> by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1583 |
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
|
1584 |
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
|
1585 |
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
|
1586 |
using \<open>K \<subseteq> T\<close> by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1587 |
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
|
1588 |
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
|
1589 |
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
|
1590 |
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
|
1591 |
proof (cases "x \<in> cbox (- c) c") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1592 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1593 |
with \<open>x \<in> T\<close> show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1594 |
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
|
1595 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1596 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1597 |
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
|
1598 |
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
|
1599 |
have "T \<inter> interior (cbox (- c) c) \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1600 |
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
|
1601 |
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
|
1602 |
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
|
1603 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1604 |
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
|
1605 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1606 |
have "interior (cbox (- c) c) \<inter> T \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1607 |
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
|
1608 |
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
|
1609 |
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
|
1610 |
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
|
1611 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1612 |
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
|
1613 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1614 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1615 |
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
|
1616 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1617 |
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
|
1618 |
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
|
1619 |
by (metis ddis disjnt_iff) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1620 |
then show False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1621 |
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
|
1622 |
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
|
1623 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1624 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1625 |
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
|
1626 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1627 |
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
|
1628 |
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
|
1629 |
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
|
1630 |
using cloTK by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1631 |
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
|
1632 |
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
|
1633 |
apply (rule gim [THEN subsetD]) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1634 |
using that cloTK by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1635 |
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
|
1636 |
\<subseteq> rel_frontier U" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1637 |
by force |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1638 |
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
|
1639 |
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
|
1640 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1641 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1642 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1643 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1644 |
corollary extend_map_affine_to_sphere_cofinite: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1645 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1646 |
assumes SUT: "compact S" "affine T" "S \<subseteq> T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1647 |
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
|
1648 |
and contf: "continuous_on S f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1649 |
and fim: "f ` S \<subseteq> sphere a r" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1650 |
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
|
1651 |
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
|
1652 |
"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
|
1653 |
proof (cases "r = 0") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1654 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1655 |
with fim show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1656 |
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
|
1657 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1658 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1659 |
with assms have "0 < r" by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1660 |
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
|
1661 |
by (simp add: aff aff_dim_cball) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1662 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1663 |
apply (rule extend_map_affine_to_sphere_cofinite_gen |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1664 |
[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
|
1665 |
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
|
1666 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1667 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1668 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1669 |
corollary extend_map_UNIV_to_sphere_cofinite: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1670 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1671 |
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
|
1672 |
and SUT: "compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1673 |
and contf: "continuous_on S f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1674 |
and fim: "f ` S \<subseteq> sphere a r" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1675 |
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
|
1676 |
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
|
1677 |
"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
|
1678 |
apply (rule extend_map_affine_to_sphere_cofinite |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1679 |
[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
|
1680 |
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
|
1681 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1682 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1683 |
corollary extend_map_UNIV_to_sphere_no_bounded_component: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1684 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1685 |
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
|
1686 |
and SUT: "compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1687 |
and contf: "continuous_on S f" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1688 |
and fim: "f ` S \<subseteq> sphere a r" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1689 |
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
|
1690 |
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
|
1691 |
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
|
1692 |
apply (auto simp: that dest: dis) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1693 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1694 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1695 |
theorem Borsuk_separation_theorem_gen: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1696 |
fixes S :: "'a::euclidean_space set" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1697 |
assumes "compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1698 |
shows "(\<forall>c \<in> components(- S). ~bounded c) \<longleftrightarrow> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1699 |
(\<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
|
1700 |
\<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
|
1701 |
(is "?lhs = ?rhs") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1702 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1703 |
assume L [rule_format]: ?lhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1704 |
show ?rhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1705 |
proof clarify |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1706 |
fix f :: "'a \<Rightarrow> 'a" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1707 |
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
|
1708 |
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
|
1709 |
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
|
1710 |
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
|
1711 |
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
|
1712 |
using nullhomotopic_from_contractible [OF contg gim] |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1713 |
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
|
1714 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1715 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1716 |
assume R [rule_format]: ?rhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1717 |
show ?lhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1718 |
unfolding components_def |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1719 |
proof clarify |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1720 |
fix a |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1721 |
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
|
1722 |
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
|
1723 |
apply (intro continuous_intros) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1724 |
using \<open>a \<notin> S\<close> by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1725 |
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
|
1726 |
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
|
1727 |
show False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1728 |
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
|
1729 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1730 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1731 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1732 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1733 |
corollary Borsuk_separation_theorem: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1734 |
fixes S :: "'a::euclidean_space set" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1735 |
assumes "compact S" and 2: "2 \<le> DIM('a)"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1736 |
shows "connected(- S) \<longleftrightarrow> |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1737 |
(\<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
|
1738 |
\<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
|
1739 |
(is "?lhs = ?rhs") |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1740 |
proof |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1741 |
assume L: ?lhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1742 |
show ?rhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1743 |
proof (cases "S = {}")
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1744 |
case True |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1745 |
then show ?thesis by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1746 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1747 |
case False |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1748 |
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
|
1749 |
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
|
1750 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1751 |
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
|
1752 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1753 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1754 |
assume R: ?rhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1755 |
then show ?lhs |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1756 |
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
|
1757 |
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
|
1758 |
using connected_component_in apply fastforce |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1759 |
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
|
1760 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1761 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1762 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1763 |
lemma homotopy_eqv_separation: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1764 |
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
|
1765 |
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
|
1766 |
shows "connected(- S) \<longleftrightarrow> connected(- T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1767 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1768 |
consider "DIM('a) = 1" | "2 \<le> DIM('a)"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1769 |
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
|
1770 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1771 |
proof cases |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1772 |
case 1 |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1773 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1774 |
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
|
1775 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1776 |
case 2 |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1777 |
with assms show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1778 |
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
|
1779 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1780 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1781 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1782 |
lemma Jordan_Brouwer_separation: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1783 |
fixes S :: "'a::euclidean_space set" and a::'a |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1784 |
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
|
1785 |
shows "\<not> connected(- S)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1786 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1787 |
have "- sphere a r \<inter> ball a r \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1788 |
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
|
1789 |
moreover |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1790 |
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
|
1791 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1792 |
have "- cball a r \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1793 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1794 |
have "frontier (cball a r) \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1795 |
using \<open>0 < r\<close> by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1796 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1797 |
by (metis frontier_complement frontier_empty) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1798 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1799 |
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
|
1800 |
by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1801 |
moreover |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1802 |
have "connected (- S) = connected (- sphere a r)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1803 |
proof (rule homotopy_eqv_separation) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1804 |
show "S homotopy_eqv sphere a r" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1805 |
using hom homeomorphic_imp_homotopy_eqv by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1806 |
show "compact (sphere a r)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1807 |
by simp |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1808 |
then show " compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1809 |
using hom homeomorphic_compactness by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1810 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1811 |
ultimately show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1812 |
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
|
1813 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1814 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1815 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1816 |
lemma Jordan_Brouwer_frontier: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1817 |
fixes S :: "'a::euclidean_space set" and a::'a |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1818 |
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
|
1819 |
shows "frontier T = S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1820 |
proof (cases r rule: linorder_cases) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1821 |
assume "r < 0" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1822 |
with S T show ?thesis by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1823 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1824 |
assume "r = 0" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1825 |
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
|
1826 |
by (auto simp: homeomorphic_finite [of "{a}" S])
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1827 |
have "components (- {b}) = { -{b}}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1828 |
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
|
1829 |
with T show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1830 |
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
|
1831 |
next |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1832 |
assume "r > 0" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1833 |
have "compact S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1834 |
using homeomorphic_compactness compact_sphere S by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1835 |
show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1836 |
proof (rule frontier_minimal_separating_closed) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1837 |
show "closed S" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1838 |
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
|
1839 |
show "\<not> connected (- S)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1840 |
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
|
1841 |
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
|
1842 |
using S by (auto simp: homeomorphic_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1843 |
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
|
1844 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1845 |
have "f ` T \<subseteq> sphere a r" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1846 |
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
|
1847 |
moreover have "f ` T \<noteq> sphere a r" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1848 |
using \<open>T \<subset> S\<close> hom |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1849 |
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
|
1850 |
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
|
1851 |
then have "connected (- f ` T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1852 |
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
|
1853 |
moreover have "compact T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1854 |
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
|
1855 |
moreover then have "compact (f ` T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1856 |
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
|
1857 |
moreover have "T homotopy_eqv f ` T" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1858 |
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
|
1859 |
ultimately show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1860 |
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
|
1861 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1862 |
qed (rule T) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1863 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1864 |
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1865 |
lemma Jordan_Brouwer_nonseparation: |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1866 |
fixes S :: "'a::euclidean_space set" and a::'a |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1867 |
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
|
1868 |
shows "connected(- T)" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1869 |
proof - |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1870 |
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
|
1871 |
proof (rule connected_intermediate_closure) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1872 |
show "connected C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1873 |
using in_components_connected that by auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1874 |
have "S = frontier C" |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1875 |
using "2" Jordan_Brouwer_frontier S that by blast |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1876 |
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
|
1877 |
by (auto simp: frontier_def) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1878 |
qed auto |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1879 |
have "components(- S) \<noteq> {}"
|
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1880 |
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
|
1881 |
components_eq_empty homeomorphic_compactness) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1882 |
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
|
1883 |
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
|
1884 |
then show ?thesis |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1885 |
apply (rule ssubst) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1886 |
apply (rule connected_Union) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1887 |
using \<open>T \<subset> S\<close> apply (auto simp: *) |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1888 |
done |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1889 |
qed |
|
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1890 |
|
| 64122 | 1891 |
subsection\<open> Invariance of domain and corollaries\<close> |
1892 |
||
1893 |
lemma invariance_of_domain_ball: |
|
1894 |
fixes f :: "'a \<Rightarrow> 'a::euclidean_space" |
|
1895 |
assumes contf: "continuous_on (cball a r) f" and "0 < r" |
|
1896 |
and inj: "inj_on f (cball a r)" |
|
1897 |
shows "open(f ` ball a r)" |
|
1898 |
proof (cases "DIM('a) = 1")
|
|
1899 |
case True |
|
1900 |
obtain h::"'a\<Rightarrow>real" and k |
|
1901 |
where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" |
|
1902 |
"\<And>x. norm(h x) = norm x" "\<And>x. norm(k x) = norm x" |
|
1903 |
"\<And>x. k(h x) = x" "\<And>x. h(k x) = x" |
|
1904 |
apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real]) |
|
1905 |
using True |
|
1906 |
apply force |
|
1907 |
by (metis UNIV_I UNIV_eq_I imageI) |
|
1908 |
have cont: "continuous_on S h" "continuous_on T k" for S T |
|
1909 |
by (simp_all add: \<open>linear h\<close> \<open>linear k\<close> linear_continuous_on linear_linear) |
|
1910 |
have "continuous_on (h ` cball a r) (h \<circ> f \<circ> k)" |
|
1911 |
apply (intro continuous_on_compose cont continuous_on_subset [OF contf]) |
|
1912 |
apply (auto simp: \<open>\<And>x. k (h x) = x\<close>) |
|
1913 |
done |
|
1914 |
moreover have "is_interval (h ` cball a r)" |
|
1915 |
by (simp add: is_interval_connected_1 \<open>linear h\<close> linear_continuous_on linear_linear connected_continuous_image) |
|
1916 |
moreover have "inj_on (h \<circ> f \<circ> k) (h ` cball a r)" |
|
1917 |
using inj by (simp add: inj_on_def) (metis \<open>\<And>x. k (h x) = x\<close>) |
|
1918 |
ultimately have *: "\<And>T. \<lbrakk>open T; T \<subseteq> h ` cball a r\<rbrakk> \<Longrightarrow> open ((h \<circ> f \<circ> k) ` T)" |
|
1919 |
using injective_eq_1d_open_map_UNIV by blast |
|
1920 |
have "open ((h \<circ> f \<circ> k) ` (h ` ball a r))" |
|
1921 |
by (rule *) (auto simp: \<open>linear h\<close> \<open>range h = UNIV\<close> open_surjective_linear_image) |
|
1922 |
then have "open ((h \<circ> f) ` ball a r)" |
|
1923 |
by (simp add: image_comp \<open>\<And>x. k (h x) = x\<close> cong: image_cong) |
|
1924 |
then show ?thesis |
|
1925 |
apply (simp add: image_comp [symmetric]) |
|
1926 |
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) |
|
1927 |
done |
|
1928 |
next |
|
1929 |
case False |
|
1930 |
then have 2: "DIM('a) \<ge> 2"
|
|
1931 |
by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq) |
|
1932 |
have fimsub: "f ` ball a r \<subseteq> - f ` sphere a r" |
|
1933 |
using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl) |
|
1934 |
have hom: "f ` sphere a r homeomorphic sphere a r" |
|
1935 |
by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball) |
|
1936 |
then have nconn: "\<not> connected (- f ` sphere a r)" |
|
1937 |
by (rule Jordan_Brouwer_separation) (auto simp: \<open>0 < r\<close>) |
|
1938 |
obtain C where C: "C \<in> components (- f ` sphere a r)" and "bounded C" |
|
1939 |
apply (rule cobounded_has_bounded_component [OF _ nconn]) |
|
1940 |
apply (simp_all add: 2) |
|
1941 |
by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball) |
|
1942 |
moreover have "f ` (ball a r) = C" |
|
1943 |
proof |
|
1944 |
have "C \<noteq> {}"
|
|
1945 |
by (rule in_components_nonempty [OF C]) |
|
1946 |
show "C \<subseteq> f ` ball a r" |
|
1947 |
proof (rule ccontr) |
|
1948 |
assume nonsub: "\<not> C \<subseteq> f ` ball a r" |
|
1949 |
have "- f ` cball a r \<subseteq> C" |
|
1950 |
proof (rule components_maximal [OF C]) |
|
1951 |
have "f ` cball a r homeomorphic cball a r" |
|
1952 |
using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast |
|
1953 |
then show "connected (- f ` cball a r)" |
|
1954 |
by (auto intro: connected_complement_homeomorphic_convex_compact 2) |
|
1955 |
show "- f ` cball a r \<subseteq> - f ` sphere a r" |
|
1956 |
by auto |
|
1957 |
then show "C \<inter> - f ` cball a r \<noteq> {}"
|
|
1958 |
using \<open>C \<noteq> {}\<close> in_components_subset [OF C] nonsub
|
|
1959 |
using image_iff by fastforce |
|
1960 |
qed |
|
1961 |
then have "bounded (- f ` cball a r)" |
|
1962 |
using bounded_subset \<open>bounded C\<close> by auto |
|
1963 |
then have "\<not> bounded (f ` cball a r)" |
|
1964 |
using cobounded_imp_unbounded by blast |
|
1965 |
then show "False" |
|
1966 |
using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast |
|
1967 |
qed |
|
1968 |
with \<open>C \<noteq> {}\<close> have "C \<inter> f ` ball a r \<noteq> {}"
|
|
1969 |
by (simp add: inf.absorb_iff1) |
|
1970 |
then show "f ` ball a r \<subseteq> C" |
|
1971 |
by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset) |
|
1972 |
qed |
|
1973 |
moreover have "open (- f ` sphere a r)" |
|
1974 |
using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast |
|
1975 |
ultimately show ?thesis |
|
1976 |
using open_components by blast |
|
1977 |
qed |
|
1978 |
||
1979 |
||
1980 |
text\<open>Proved by L. E. J. Brouwer (1912)\<close> |
|
1981 |
theorem invariance_of_domain: |
|
1982 |
fixes f :: "'a \<Rightarrow> 'a::euclidean_space" |
|
1983 |
assumes "continuous_on S f" "open S" "inj_on f S" |
|
1984 |
shows "open(f ` S)" |
|
1985 |
unfolding open_subopen [of "f`S"] |
|
1986 |
proof clarify |
|
1987 |
fix a |
|
1988 |
assume "a \<in> S" |
|
1989 |
obtain \<delta> where "\<delta> > 0" and \<delta>: "cball a \<delta> \<subseteq> S" |
|
1990 |
using \<open>open S\<close> \<open>a \<in> S\<close> open_contains_cball_eq by blast |
|
1991 |
show "\<exists>T. open T \<and> f a \<in> T \<and> T \<subseteq> f ` S" |
|
1992 |
proof (intro exI conjI) |
|
1993 |
show "open (f ` (ball a \<delta>))" |
|
1994 |
by (meson \<delta> \<open>0 < \<delta>\<close> assms continuous_on_subset inj_on_subset invariance_of_domain_ball) |
|
1995 |
show "f a \<in> f ` ball a \<delta>" |
|
1996 |
by (simp add: \<open>0 < \<delta>\<close>) |
|
1997 |
show "f ` ball a \<delta> \<subseteq> f ` S" |
|
1998 |
using \<delta> ball_subset_cball by blast |
|
1999 |
qed |
|
2000 |
qed |
|
2001 |
||
2002 |
lemma inv_of_domain_ss0: |
|
2003 |
fixes f :: "'a \<Rightarrow> 'a::euclidean_space" |
|
2004 |
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S" |
|
2005 |
and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)"
|
|
2006 |
and ope: "openin (subtopology euclidean S) U" |
|
2007 |
shows "openin (subtopology euclidean S) (f ` U)" |
|
2008 |
proof - |
|
2009 |
have "U \<subseteq> S" |
|
2010 |
using ope openin_imp_subset by blast |
|
2011 |
have "(UNIV::'b set) homeomorphic S" |
|
2012 |
by (simp add: \<open>subspace S\<close> dimS dim_UNIV homeomorphic_subspaces) |
|
2013 |
then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" |
|
2014 |
using homeomorphic_def by blast |
|
2015 |
have homkh: "homeomorphism S (k ` S) k h" |
|
2016 |
using homhk homeomorphism_image2 homeomorphism_sym by fastforce |
|
2017 |
have "open ((k \<circ> f \<circ> h) ` k ` U)" |
|
2018 |
proof (rule invariance_of_domain) |
|
2019 |
show "continuous_on (k ` U) (k \<circ> f \<circ> h)" |
|
2020 |
proof (intro continuous_intros) |
|
2021 |
show "continuous_on (k ` U) h" |
|
2022 |
by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest) |
|
2023 |
show "continuous_on (h ` k ` U) f" |
|
2024 |
apply (rule continuous_on_subset [OF contf], clarify) |
|
2025 |
apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD) |
|
2026 |
done |
|
2027 |
show "continuous_on (f ` h ` k ` U) k" |
|
2028 |
apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) |
|
2029 |
using fim homhk homeomorphism_apply2 ope openin_subset by fastforce |
|
2030 |
qed |
|
2031 |
have ope_iff: "\<And>T. open T \<longleftrightarrow> openin (subtopology euclidean (k ` S)) T" |
|
2032 |
using homhk homeomorphism_image2 open_openin by fastforce |
|
2033 |
show "open (k ` U)" |
|
2034 |
by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) |
|
2035 |
show "inj_on (k \<circ> f \<circ> h) (k ` U)" |
|
2036 |
apply (clarsimp simp: inj_on_def) |
|
2037 |
by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \<open>U \<subseteq> S\<close>) |
|
2038 |
qed |
|
2039 |
moreover |
|
2040 |
have eq: "f ` U = h ` (k \<circ> f \<circ> h \<circ> k) ` U" |
|
2041 |
apply (auto simp: image_comp [symmetric]) |
|
2042 |
apply (metis homkh \<open>U \<subseteq> S\<close> fim homeomorphism_image2 homeomorphism_of_subsets homhk imageI subset_UNIV) |
|
2043 |
by (metis \<open>U \<subseteq> S\<close> subsetD fim homeomorphism_def homhk image_eqI) |
|
2044 |
ultimately show ?thesis |
|
2045 |
by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV) |
|
2046 |
qed |
|
2047 |
||
2048 |
lemma inv_of_domain_ss1: |
|
2049 |
fixes f :: "'a \<Rightarrow> 'a::euclidean_space" |
|
2050 |
assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \<subseteq> S" |
|
2051 |
and "subspace S" |
|
2052 |
and ope: "openin (subtopology euclidean S) U" |
|
2053 |
shows "openin (subtopology euclidean S) (f ` U)" |
|
2054 |
proof - |
|
2055 |
define S' where "S' \<equiv> {y. \<forall>x \<in> S. orthogonal x y}"
|
|
2056 |
have "subspace S'" |
|
2057 |
by (simp add: S'_def subspace_orthogonal_to_vectors) |
|
2058 |
define g where "g \<equiv> \<lambda>z::'a*'a. ((f \<circ> fst)z, snd z)" |
|
2059 |
have "openin (subtopology euclidean (S \<times> S')) (g ` (U \<times> S'))" |
|
2060 |
proof (rule inv_of_domain_ss0) |
|
2061 |
show "continuous_on (U \<times> S') g" |
|
2062 |
apply (simp add: g_def) |
|
2063 |
apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto) |
|
2064 |
done |
|
2065 |
show "g ` (U \<times> S') \<subseteq> S \<times> S'" |
|
2066 |
using fim by (auto simp: g_def) |
|
2067 |
show "inj_on g (U \<times> S')" |
|
2068 |
using injf by (auto simp: g_def inj_on_def) |
|
2069 |
show "subspace (S \<times> S')" |
|
2070 |
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> subspace_Times) |
|
2071 |
show "openin (subtopology euclidean (S \<times> S')) (U \<times> S')" |
|
2072 |
by (simp add: openin_Times [OF ope]) |
|
2073 |
have "dim (S \<times> S') = dim S + dim S'" |
|
2074 |
by (simp add: \<open>subspace S'\<close> \<open>subspace S\<close> dim_Times) |
|
2075 |
also have "... = DIM('a)"
|
|
2076 |
using dim_subspace_orthogonal_to_vectors [OF \<open>subspace S\<close> subspace_UNIV] |
|
2077 |
by (simp add: add.commute S'_def) |
|
2078 |
finally show "dim (S \<times> S') = DIM('a)" .
|
|
2079 |
qed |
|
2080 |
moreover have "g ` (U \<times> S') = f ` U \<times> S'" |
|
2081 |
by (auto simp: g_def image_iff) |
|
2082 |
moreover have "0 \<in> S'" |
|
2083 |
using \<open>subspace S'\<close> subspace_affine by blast |
|
2084 |
ultimately show ?thesis |
|
2085 |
by (auto simp: openin_Times_eq) |
|
2086 |
qed |
|
2087 |
||
2088 |
||
2089 |
corollary invariance_of_domain_subspaces: |
|
2090 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2091 |
assumes ope: "openin (subtopology euclidean U) S" |
|
2092 |
and "subspace U" "subspace V" and VU: "dim V \<le> dim U" |
|
2093 |
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V" |
|
2094 |
and injf: "inj_on f S" |
|
2095 |
shows "openin (subtopology euclidean V) (f ` S)" |
|
2096 |
proof - |
|
2097 |
obtain V' where "subspace V'" "V' \<subseteq> U" "dim V' = dim V" |
|
2098 |
using choose_subspace_of_subspace [OF VU] |
|
2099 |
by (metis span_eq \<open>subspace U\<close>) |
|
2100 |
then have "V homeomorphic V'" |
|
2101 |
by (simp add: \<open>subspace V\<close> homeomorphic_subspaces) |
|
2102 |
then obtain h k where homhk: "homeomorphism V V' h k" |
|
2103 |
using homeomorphic_def by blast |
|
2104 |
have eq: "f ` S = k ` (h \<circ> f) ` S" |
|
2105 |
proof - |
|
2106 |
have "k ` h ` f ` S = f ` S" |
|
2107 |
by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl) |
|
2108 |
then show ?thesis |
|
2109 |
by (simp add: image_comp) |
|
2110 |
qed |
|
2111 |
show ?thesis |
|
2112 |
unfolding eq |
|
2113 |
proof (rule homeomorphism_imp_open_map) |
|
2114 |
show homkh: "homeomorphism V' V k h" |
|
2115 |
by (simp add: homeomorphism_symD homhk) |
|
2116 |
have hfV': "(h \<circ> f) ` S \<subseteq> V'" |
|
2117 |
using fim homeomorphism_image1 homhk by fastforce |
|
2118 |
moreover have "openin (subtopology euclidean U) ((h \<circ> f) ` S)" |
|
2119 |
proof (rule inv_of_domain_ss1) |
|
2120 |
show "continuous_on S (h \<circ> f)" |
|
2121 |
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) |
|
2122 |
show "inj_on (h \<circ> f) S" |
|
2123 |
apply (clarsimp simp: inj_on_def) |
|
2124 |
by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf) |
|
2125 |
show "(h \<circ> f) ` S \<subseteq> U" |
|
2126 |
using \<open>V' \<subseteq> U\<close> hfV' by auto |
|
2127 |
qed (auto simp: assms) |
|
2128 |
ultimately show "openin (subtopology euclidean V') ((h \<circ> f) ` S)" |
|
2129 |
using openin_subset_trans \<open>V' \<subseteq> U\<close> by force |
|
2130 |
qed |
|
2131 |
qed |
|
2132 |
||
2133 |
corollary invariance_of_dimension_subspaces: |
|
2134 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2135 |
assumes ope: "openin (subtopology euclidean U) S" |
|
2136 |
and "subspace U" "subspace V" |
|
2137 |
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V" |
|
2138 |
and injf: "inj_on f S" and "S \<noteq> {}"
|
|
2139 |
shows "dim U \<le> dim V" |
|
2140 |
proof - |
|
2141 |
have "False" if "dim V < dim U" |
|
2142 |
proof - |
|
2143 |
obtain T where "subspace T" "T \<subseteq> U" "dim T = dim V" |
|
2144 |
using choose_subspace_of_subspace [of "dim V" U] |
|
2145 |
by (metis span_eq \<open>subspace U\<close> \<open>dim V < dim U\<close> linear not_le) |
|
2146 |
then have "V homeomorphic T" |
|
2147 |
by (simp add: \<open>subspace V\<close> homeomorphic_subspaces) |
|
2148 |
then obtain h k where homhk: "homeomorphism V T h k" |
|
2149 |
using homeomorphic_def by blast |
|
2150 |
have "continuous_on S (h \<circ> f)" |
|
2151 |
by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) |
|
2152 |
moreover have "(h \<circ> f) ` S \<subseteq> U" |
|
2153 |
using \<open>T \<subseteq> U\<close> fim homeomorphism_image1 homhk by fastforce |
|
2154 |
moreover have "inj_on (h \<circ> f) S" |
|
2155 |
apply (clarsimp simp: inj_on_def) |
|
2156 |
by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) |
|
2157 |
ultimately have ope_hf: "openin (subtopology euclidean U) ((h \<circ> f) ` S)" |
|
2158 |
using invariance_of_domain_subspaces [OF ope \<open>subspace U\<close> \<open>subspace U\<close>] by auto |
|
2159 |
have "(h \<circ> f) ` S \<subseteq> T" |
|
2160 |
using fim homeomorphism_image1 homhk by fastforce |
|
2161 |
then show ?thesis |
|
2162 |
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)
|
|
2163 |
qed |
|
2164 |
then show ?thesis |
|
2165 |
using not_less by blast |
|
2166 |
qed |
|
2167 |
||
2168 |
corollary invariance_of_domain_affine_sets: |
|
2169 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2170 |
assumes ope: "openin (subtopology euclidean U) S" |
|
2171 |
and aff: "affine U" "affine V" "aff_dim V \<le> aff_dim U" |
|
2172 |
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V" |
|
2173 |
and injf: "inj_on f S" |
|
2174 |
shows "openin (subtopology euclidean V) (f ` S)" |
|
2175 |
proof (cases "S = {}")
|
|
2176 |
case True |
|
2177 |
then show ?thesis by auto |
|
2178 |
next |
|
2179 |
case False |
|
2180 |
obtain a b where "a \<in> S" "a \<in> U" "b \<in> V" |
|
2181 |
using False fim ope openin_contains_cball by fastforce |
|
2182 |
have "openin (subtopology euclidean (op + (- b) ` V)) ((op + (- b) \<circ> f \<circ> op + a) ` op + (- a) ` S)" |
|
2183 |
proof (rule invariance_of_domain_subspaces) |
|
2184 |
show "openin (subtopology euclidean (op + (- a) ` U)) (op + (- a) ` S)" |
|
2185 |
by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) |
|
2186 |
show "subspace (op + (- a) ` U)" |
|
2187 |
by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace \<open>affine U\<close>) |
|
2188 |
show "subspace (op + (- b) ` V)" |
|
2189 |
by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace \<open>affine V\<close>) |
|
2190 |
show "dim (op + (- b) ` V) \<le> dim (op + (- a) ` U)" |
|
2191 |
by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) |
|
2192 |
show "continuous_on (op + (- a) ` S) (op + (- b) \<circ> f \<circ> op + a)" |
|
2193 |
by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) |
|
2194 |
show "(op + (- b) \<circ> f \<circ> op + a) ` op + (- a) ` S \<subseteq> op + (- b) ` V" |
|
2195 |
using fim by auto |
|
2196 |
show "inj_on (op + (- b) \<circ> f \<circ> op + a) (op + (- a) ` S)" |
|
2197 |
by (auto simp: inj_on_def) (meson inj_onD injf) |
|
2198 |
qed |
|
2199 |
then show ?thesis |
|
2200 |
by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) |
|
2201 |
qed |
|
2202 |
||
2203 |
corollary invariance_of_dimension_affine_sets: |
|
2204 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2205 |
assumes ope: "openin (subtopology euclidean U) S" |
|
2206 |
and aff: "affine U" "affine V" |
|
2207 |
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V" |
|
2208 |
and injf: "inj_on f S" and "S \<noteq> {}"
|
|
2209 |
shows "aff_dim U \<le> aff_dim V" |
|
2210 |
proof - |
|
2211 |
obtain a b where "a \<in> S" "a \<in> U" "b \<in> V" |
|
2212 |
using \<open>S \<noteq> {}\<close> fim ope openin_contains_cball by fastforce
|
|
2213 |
have "dim (op + (- a) ` U) \<le> dim (op + (- b) ` V)" |
|
2214 |
proof (rule invariance_of_dimension_subspaces) |
|
2215 |
show "openin (subtopology euclidean (op + (- a) ` U)) (op + (- a) ` S)" |
|
2216 |
by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) |
|
2217 |
show "subspace (op + (- a) ` U)" |
|
2218 |
by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace \<open>affine U\<close>) |
|
2219 |
show "subspace (op + (- b) ` V)" |
|
2220 |
by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace \<open>affine V\<close>) |
|
2221 |
show "continuous_on (op + (- a) ` S) (op + (- b) \<circ> f \<circ> op + a)" |
|
2222 |
by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) |
|
2223 |
show "(op + (- b) \<circ> f \<circ> op + a) ` op + (- a) ` S \<subseteq> op + (- b) ` V" |
|
2224 |
using fim by auto |
|
2225 |
show "inj_on (op + (- b) \<circ> f \<circ> op + a) (op + (- a) ` S)" |
|
2226 |
by (auto simp: inj_on_def) (meson inj_onD injf) |
|
2227 |
qed (use \<open>S \<noteq> {}\<close> in auto)
|
|
2228 |
then show ?thesis |
|
2229 |
by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) |
|
2230 |
qed |
|
2231 |
||
2232 |
corollary invariance_of_dimension: |
|
2233 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2234 |
assumes contf: "continuous_on S f" and "open S" |
|
2235 |
and injf: "inj_on f S" and "S \<noteq> {}"
|
|
2236 |
shows "DIM('a) \<le> DIM('b)"
|
|
2237 |
using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms |
|
2238 |
by auto |
|
2239 |
||
2240 |
||
2241 |
corollary continuous_injective_image_subspace_dim_le: |
|
2242 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2243 |
assumes "subspace S" "subspace T" |
|
2244 |
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> T" |
|
2245 |
and injf: "inj_on f S" |
|
2246 |
shows "dim S \<le> dim T" |
|
2247 |
apply (rule invariance_of_dimension_subspaces [of S S _ f]) |
|
2248 |
using assms by (auto simp: subspace_affine) |
|
2249 |
||
2250 |
lemma invariance_of_dimension_convex_domain: |
|
2251 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2252 |
assumes "convex S" |
|
2253 |
and contf: "continuous_on S f" and fim: "f ` S \<subseteq> affine hull T" |
|
2254 |
and injf: "inj_on f S" |
|
2255 |
shows "aff_dim S \<le> aff_dim T" |
|
2256 |
proof (cases "S = {}")
|
|
2257 |
case True |
|
2258 |
then show ?thesis by (simp add: aff_dim_geq) |
|
2259 |
next |
|
2260 |
case False |
|
2261 |
have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)" |
|
2262 |
proof (rule invariance_of_dimension_affine_sets) |
|
2263 |
show "openin (subtopology euclidean (affine hull S)) (rel_interior S)" |
|
2264 |
by (simp add: openin_rel_interior) |
|
2265 |
show "continuous_on (rel_interior S) f" |
|
2266 |
using contf continuous_on_subset rel_interior_subset by blast |
|
2267 |
show "f ` rel_interior S \<subseteq> affine hull T" |
|
2268 |
using fim rel_interior_subset by blast |
|
2269 |
show "inj_on f (rel_interior S)" |
|
2270 |
using inj_on_subset injf rel_interior_subset by blast |
|
2271 |
show "rel_interior S \<noteq> {}"
|
|
2272 |
by (simp add: False \<open>convex S\<close> rel_interior_eq_empty) |
|
2273 |
qed auto |
|
2274 |
then show ?thesis |
|
2275 |
by simp |
|
2276 |
qed |
|
2277 |
||
2278 |
||
2279 |
lemma homeomorphic_convex_sets_le: |
|
2280 |
assumes "convex S" "S homeomorphic T" |
|
2281 |
shows "aff_dim S \<le> aff_dim T" |
|
2282 |
proof - |
|
2283 |
obtain h k where homhk: "homeomorphism S T h k" |
|
2284 |
using homeomorphic_def assms by blast |
|
2285 |
show ?thesis |
|
2286 |
proof (rule invariance_of_dimension_convex_domain [OF \<open>convex S\<close>]) |
|
2287 |
show "continuous_on S h" |
|
2288 |
using homeomorphism_def homhk by blast |
|
2289 |
show "h ` S \<subseteq> affine hull T" |
|
2290 |
by (metis homeomorphism_def homhk hull_subset) |
|
2291 |
show "inj_on h S" |
|
2292 |
by (meson homeomorphism_apply1 homhk inj_on_inverseI) |
|
2293 |
qed |
|
2294 |
qed |
|
2295 |
||
2296 |
lemma homeomorphic_convex_sets: |
|
2297 |
assumes "convex S" "convex T" "S homeomorphic T" |
|
2298 |
shows "aff_dim S = aff_dim T" |
|
2299 |
by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) |
|
2300 |
||
2301 |
lemma homeomorphic_convex_compact_sets_eq: |
|
2302 |
assumes "convex S" "compact S" "convex T" "compact T" |
|
2303 |
shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T" |
|
2304 |
by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) |
|
2305 |
||
2306 |
lemma invariance_of_domain_gen: |
|
2307 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2308 |
assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
|
|
2309 |
shows "open(f ` S)" |
|
2310 |
using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto |
|
2311 |
||
2312 |
lemma injective_into_1d_imp_open_map_UNIV: |
|
2313 |
fixes f :: "'a::euclidean_space \<Rightarrow> real" |
|
2314 |
assumes "open T" "continuous_on S f" "inj_on f S" "T \<subseteq> S" |
|
2315 |
shows "open (f ` T)" |
|
2316 |
apply (rule invariance_of_domain_gen [OF \<open>open T\<close>]) |
|
2317 |
using assms apply (auto simp: elim: continuous_on_subset subset_inj_on) |
|
2318 |
done |
|
2319 |
||
2320 |
lemma continuous_on_inverse_open: |
|
2321 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2322 |
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
|
|
2323 |
shows "continuous_on (f ` S) g" |
|
2324 |
proof (clarsimp simp add: continuous_openin_preimage_eq) |
|
2325 |
fix T :: "'a set" |
|
2326 |
assume "open T" |
|
2327 |
have eq: "{x. x \<in> f ` S \<and> g x \<in> T} = f ` (S \<inter> T)"
|
|
2328 |
by (auto simp: gf) |
|
2329 |
show "openin (subtopology euclidean (f ` S)) {x \<in> f ` S. g x \<in> T}"
|
|
2330 |
apply (subst eq) |
|
2331 |
apply (rule open_openin_trans) |
|
2332 |
apply (rule invariance_of_domain_gen) |
|
2333 |
using assms |
|
2334 |
apply auto |
|
2335 |
using inj_on_inverseI apply auto[1] |
|
2336 |
by (metis \<open>open T\<close> continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq) |
|
2337 |
qed |
|
2338 |
||
2339 |
lemma invariance_of_domain_homeomorphism: |
|
2340 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2341 |
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
|
|
2342 |
obtains g where "homeomorphism S (f ` S) f g" |
|
2343 |
proof |
|
2344 |
show "homeomorphism S (f ` S) f (inv_into S f)" |
|
2345 |
by (simp add: assms continuous_on_inverse_open homeomorphism_def) |
|
2346 |
qed |
|
2347 |
||
2348 |
corollary invariance_of_domain_homeomorphic: |
|
2349 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
2350 |
assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
|
|
2351 |
shows "S homeomorphic (f ` S)" |
|
2352 |
using invariance_of_domain_homeomorphism [OF assms] |
|
2353 |
by (meson homeomorphic_def) |
|
2354 |
||
|
64006
0de4736dad8b
new theorems including the theory FurtherTopology
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2355 |
end |