| author | wenzelm |
| Mon, 05 May 2014 16:29:09 +0200 | |
| changeset 56866 | c4512e94e15c |
| parent 56742 | 678a52e676b6 |
| child 57275 | 0ddb5b755cdc |
| permissions | -rw-r--r-- |
|
33714
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
1 |
(* title: HOL/Library/Topology_Euclidian_Space.thy |
| 33175 | 2 |
Author: Amine Chaieb, University of Cambridge |
3 |
Author: Robert Himmelmann, TU Muenchen |
|
| 44075 | 4 |
Author: Brian Huffman, Portland State University |
| 33175 | 5 |
*) |
6 |
||
7 |
header {* Elementary topology in Euclidean space. *}
|
|
8 |
||
9 |
theory Topology_Euclidean_Space |
|
| 50087 | 10 |
imports |
| 50938 | 11 |
Complex_Main |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
12 |
"~~/src/HOL/Library/Countable_Set" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
13 |
"~~/src/HOL/Library/FuncSet" |
| 50938 | 14 |
Linear_Algebra |
| 50087 | 15 |
Norm_Arith |
16 |
begin |
|
17 |
||
| 50972 | 18 |
lemma dist_0_norm: |
19 |
fixes x :: "'a::real_normed_vector" |
|
20 |
shows "dist 0 x = norm x" |
|
21 |
unfolding dist_norm by simp |
|
22 |
||
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
23 |
lemma dist_double: "dist x y < d / 2 \<Longrightarrow> dist x z < d / 2 \<Longrightarrow> dist y z < d" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
24 |
using dist_triangle[of y z x] by (simp add: dist_commute) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
25 |
|
| 50972 | 26 |
(* LEGACY *) |
| 53640 | 27 |
lemma lim_subseq: "subseq r \<Longrightarrow> s ----> l \<Longrightarrow> (s \<circ> r) ----> l" |
| 50972 | 28 |
by (rule LIMSEQ_subseq_LIMSEQ) |
29 |
||
| 53282 | 30 |
lemma countable_PiE: |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
31 |
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (PiE I F)" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
32 |
by (induct I arbitrary: F rule: finite_induct) (auto simp: PiE_insert_eq) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
33 |
|
|
51481
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
34 |
lemma Lim_within_open: |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
35 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
36 |
shows "a \<in> S \<Longrightarrow> open S \<Longrightarrow> (f ---> l)(at a within S) \<longleftrightarrow> (f ---> l)(at a)" |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
37 |
by (fact tendsto_within_open) |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
38 |
|
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
39 |
lemma continuous_on_union: |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
40 |
"closed s \<Longrightarrow> closed t \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on t f \<Longrightarrow> continuous_on (s \<union> t) f" |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
41 |
by (fact continuous_on_closed_Un) |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
42 |
|
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
43 |
lemma continuous_on_cases: |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
44 |
"closed s \<Longrightarrow> closed t \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on t g \<Longrightarrow> |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
45 |
\<forall>x. (x\<in>s \<and> \<not> P x) \<or> (x \<in> t \<and> P x) \<longrightarrow> f x = g x \<Longrightarrow> |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
46 |
continuous_on (s \<union> t) (\<lambda>x. if P x then f x else g x)" |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
47 |
by (rule continuous_on_If) auto |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
48 |
|
| 53255 | 49 |
|
| 50087 | 50 |
subsection {* Topological Basis *}
|
51 |
||
52 |
context topological_space |
|
53 |
begin |
|
54 |
||
| 53291 | 55 |
definition "topological_basis B \<longleftrightarrow> |
56 |
(\<forall>b\<in>B. open b) \<and> (\<forall>x. open x \<longrightarrow> (\<exists>B'. B' \<subseteq> B \<and> \<Union>B' = x))" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
57 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
58 |
lemma topological_basis: |
| 53291 | 59 |
"topological_basis B \<longleftrightarrow> (\<forall>x. open x \<longleftrightarrow> (\<exists>B'. B' \<subseteq> B \<and> \<Union>B' = x))" |
| 50998 | 60 |
unfolding topological_basis_def |
61 |
apply safe |
|
62 |
apply fastforce |
|
63 |
apply fastforce |
|
64 |
apply (erule_tac x="x" in allE) |
|
65 |
apply simp |
|
66 |
apply (rule_tac x="{x}" in exI)
|
|
67 |
apply auto |
|
68 |
done |
|
69 |
||
| 50087 | 70 |
lemma topological_basis_iff: |
71 |
assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'" |
|
72 |
shows "topological_basis B \<longleftrightarrow> (\<forall>O'. open O' \<longrightarrow> (\<forall>x\<in>O'. \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'))" |
|
73 |
(is "_ \<longleftrightarrow> ?rhs") |
|
74 |
proof safe |
|
75 |
fix O' and x::'a |
|
76 |
assume H: "topological_basis B" "open O'" "x \<in> O'" |
|
| 53282 | 77 |
then have "(\<exists>B'\<subseteq>B. \<Union>B' = O')" by (simp add: topological_basis_def) |
| 50087 | 78 |
then obtain B' where "B' \<subseteq> B" "O' = \<Union>B'" by auto |
| 53282 | 79 |
then show "\<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'" using H by auto |
| 50087 | 80 |
next |
81 |
assume H: ?rhs |
|
| 53282 | 82 |
show "topological_basis B" |
83 |
using assms unfolding topological_basis_def |
|
| 50087 | 84 |
proof safe |
| 53640 | 85 |
fix O' :: "'a set" |
| 53282 | 86 |
assume "open O'" |
| 50087 | 87 |
with H obtain f where "\<forall>x\<in>O'. f x \<in> B \<and> x \<in> f x \<and> f x \<subseteq> O'" |
88 |
by (force intro: bchoice simp: Bex_def) |
|
| 53282 | 89 |
then show "\<exists>B'\<subseteq>B. \<Union>B' = O'" |
| 50087 | 90 |
by (auto intro: exI[where x="{f x |x. x \<in> O'}"])
|
91 |
qed |
|
92 |
qed |
|
93 |
||
94 |
lemma topological_basisI: |
|
95 |
assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'" |
|
| 53282 | 96 |
and "\<And>O' x. open O' \<Longrightarrow> x \<in> O' \<Longrightarrow> \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'" |
| 50087 | 97 |
shows "topological_basis B" |
98 |
using assms by (subst topological_basis_iff) auto |
|
99 |
||
100 |
lemma topological_basisE: |
|
101 |
fixes O' |
|
102 |
assumes "topological_basis B" |
|
| 53282 | 103 |
and "open O'" |
104 |
and "x \<in> O'" |
|
| 50087 | 105 |
obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'" |
106 |
proof atomize_elim |
|
| 53282 | 107 |
from assms have "\<And>B'. B'\<in>B \<Longrightarrow> open B'" |
108 |
by (simp add: topological_basis_def) |
|
| 50087 | 109 |
with topological_basis_iff assms |
| 53282 | 110 |
show "\<exists>B'. B' \<in> B \<and> x \<in> B' \<and> B' \<subseteq> O'" |
111 |
using assms by (simp add: Bex_def) |
|
| 50087 | 112 |
qed |
113 |
||
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
114 |
lemma topological_basis_open: |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
115 |
assumes "topological_basis B" |
| 53282 | 116 |
and "X \<in> B" |
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
117 |
shows "open X" |
| 53282 | 118 |
using assms by (simp add: topological_basis_def) |
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
119 |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
120 |
lemma topological_basis_imp_subbasis: |
| 53255 | 121 |
assumes B: "topological_basis B" |
122 |
shows "open = generate_topology B" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
123 |
proof (intro ext iffI) |
| 53255 | 124 |
fix S :: "'a set" |
125 |
assume "open S" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
126 |
with B obtain B' where "B' \<subseteq> B" "S = \<Union>B'" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
127 |
unfolding topological_basis_def by blast |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
128 |
then show "generate_topology B S" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
129 |
by (auto intro: generate_topology.intros dest: topological_basis_open) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
130 |
next |
| 53255 | 131 |
fix S :: "'a set" |
132 |
assume "generate_topology B S" |
|
133 |
then show "open S" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
134 |
by induct (auto dest: topological_basis_open[OF B]) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
135 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
136 |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
137 |
lemma basis_dense: |
| 53640 | 138 |
fixes B :: "'a set set" |
139 |
and f :: "'a set \<Rightarrow> 'a" |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
140 |
assumes "topological_basis B" |
| 53255 | 141 |
and choosefrom_basis: "\<And>B'. B' \<noteq> {} \<Longrightarrow> f B' \<in> B'"
|
| 55522 | 142 |
shows "\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>B' \<in> B. f B' \<in> X)"
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
143 |
proof (intro allI impI) |
| 53640 | 144 |
fix X :: "'a set" |
145 |
assume "open X" and "X \<noteq> {}"
|
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
146 |
from topological_basisE[OF `topological_basis B` `open X` choosefrom_basis[OF `X \<noteq> {}`]]
|
| 55522 | 147 |
obtain B' where "B' \<in> B" "f X \<in> B'" "B' \<subseteq> X" . |
| 53255 | 148 |
then show "\<exists>B'\<in>B. f B' \<in> X" |
149 |
by (auto intro!: choosefrom_basis) |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
150 |
qed |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
151 |
|
| 50087 | 152 |
end |
153 |
||
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
154 |
lemma topological_basis_prod: |
| 53255 | 155 |
assumes A: "topological_basis A" |
156 |
and B: "topological_basis B" |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
157 |
shows "topological_basis ((\<lambda>(a, b). a \<times> b) ` (A \<times> B))" |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
158 |
unfolding topological_basis_def |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
159 |
proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric]) |
| 53255 | 160 |
fix S :: "('a \<times> 'b) set"
|
161 |
assume "open S" |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
162 |
then show "\<exists>X\<subseteq>A \<times> B. (\<Union>(a,b)\<in>X. a \<times> b) = S" |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
163 |
proof (safe intro!: exI[of _ "{x\<in>A \<times> B. fst x \<times> snd x \<subseteq> S}"])
|
| 53255 | 164 |
fix x y |
165 |
assume "(x, y) \<in> S" |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
166 |
from open_prod_elim[OF `open S` this] |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
167 |
obtain a b where a: "open a""x \<in> a" and b: "open b" "y \<in> b" and "a \<times> b \<subseteq> S" |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
168 |
by (metis mem_Sigma_iff) |
| 55522 | 169 |
moreover |
170 |
from A a obtain A0 where "A0 \<in> A" "x \<in> A0" "A0 \<subseteq> a" |
|
171 |
by (rule topological_basisE) |
|
172 |
moreover |
|
173 |
from B b obtain B0 where "B0 \<in> B" "y \<in> B0" "B0 \<subseteq> b" |
|
174 |
by (rule topological_basisE) |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
175 |
ultimately show "(x, y) \<in> (\<Union>(a, b)\<in>{X \<in> A \<times> B. fst X \<times> snd X \<subseteq> S}. a \<times> b)"
|
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
176 |
by (intro UN_I[of "(A0, B0)"]) auto |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
177 |
qed auto |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
178 |
qed (metis A B topological_basis_open open_Times) |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
179 |
|
| 53255 | 180 |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
181 |
subsection {* Countable Basis *}
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
182 |
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
183 |
locale countable_basis = |
| 53640 | 184 |
fixes B :: "'a::topological_space set set" |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
185 |
assumes is_basis: "topological_basis B" |
| 53282 | 186 |
and countable_basis: "countable B" |
| 33175 | 187 |
begin |
188 |
||
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
189 |
lemma open_countable_basis_ex: |
| 50087 | 190 |
assumes "open X" |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
191 |
shows "\<exists>B' \<subseteq> B. X = Union B'" |
| 53255 | 192 |
using assms countable_basis is_basis |
193 |
unfolding topological_basis_def by blast |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
194 |
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
195 |
lemma open_countable_basisE: |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
196 |
assumes "open X" |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
197 |
obtains B' where "B' \<subseteq> B" "X = Union B'" |
| 53255 | 198 |
using assms open_countable_basis_ex |
199 |
by (atomize_elim) simp |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
200 |
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
201 |
lemma countable_dense_exists: |
| 53291 | 202 |
"\<exists>D::'a set. countable D \<and> (\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>d \<in> D. d \<in> X))"
|
| 50087 | 203 |
proof - |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
204 |
let ?f = "(\<lambda>B'. SOME x. x \<in> B')" |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
205 |
have "countable (?f ` B)" using countable_basis by simp |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
206 |
with basis_dense[OF is_basis, of ?f] show ?thesis |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
207 |
by (intro exI[where x="?f ` B"]) (metis (mono_tags) all_not_in_conv imageI someI) |
| 50087 | 208 |
qed |
209 |
||
210 |
lemma countable_dense_setE: |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
211 |
obtains D :: "'a set" |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
212 |
where "countable D" "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d \<in> D. d \<in> X"
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
213 |
using countable_dense_exists by blast |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
214 |
|
| 50087 | 215 |
end |
216 |
||
| 50883 | 217 |
lemma (in first_countable_topology) first_countable_basisE: |
218 |
obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a" |
|
219 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)" |
|
220 |
using first_countable_basis[of x] |
|
| 51473 | 221 |
apply atomize_elim |
222 |
apply (elim exE) |
|
223 |
apply (rule_tac x="range A" in exI) |
|
224 |
apply auto |
|
225 |
done |
|
| 50883 | 226 |
|
| 51105 | 227 |
lemma (in first_countable_topology) first_countable_basis_Int_stableE: |
228 |
obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a" |
|
229 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)" |
|
230 |
"\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A" |
|
231 |
proof atomize_elim |
|
| 55522 | 232 |
obtain A' where A': |
233 |
"countable A'" |
|
234 |
"\<And>a. a \<in> A' \<Longrightarrow> x \<in> a" |
|
235 |
"\<And>a. a \<in> A' \<Longrightarrow> open a" |
|
236 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A'. a \<subseteq> S" |
|
237 |
by (rule first_countable_basisE) blast |
|
| 51105 | 238 |
def A \<equiv> "(\<lambda>N. \<Inter>((\<lambda>n. from_nat_into A' n) ` N)) ` (Collect finite::nat set set)" |
| 53255 | 239 |
then show "\<exists>A. countable A \<and> (\<forall>a. a \<in> A \<longrightarrow> x \<in> a) \<and> (\<forall>a. a \<in> A \<longrightarrow> open a) \<and> |
| 51105 | 240 |
(\<forall>S. open S \<longrightarrow> x \<in> S \<longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)) \<and> (\<forall>a b. a \<in> A \<longrightarrow> b \<in> A \<longrightarrow> a \<inter> b \<in> A)" |
241 |
proof (safe intro!: exI[where x=A]) |
|
| 53255 | 242 |
show "countable A" |
243 |
unfolding A_def by (intro countable_image countable_Collect_finite) |
|
244 |
fix a |
|
245 |
assume "a \<in> A" |
|
246 |
then show "x \<in> a" "open a" |
|
247 |
using A'(4)[OF open_UNIV] by (auto simp: A_def intro: A' from_nat_into) |
|
| 51105 | 248 |
next |
|
52141
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
haftmann
parents:
51773
diff
changeset
|
249 |
let ?int = "\<lambda>N. \<Inter>(from_nat_into A' ` N)" |
| 53255 | 250 |
fix a b |
251 |
assume "a \<in> A" "b \<in> A" |
|
252 |
then obtain N M where "a = ?int N" "b = ?int M" "finite (N \<union> M)" |
|
253 |
by (auto simp: A_def) |
|
254 |
then show "a \<inter> b \<in> A" |
|
255 |
by (auto simp: A_def intro!: image_eqI[where x="N \<union> M"]) |
|
| 51105 | 256 |
next |
| 53255 | 257 |
fix S |
258 |
assume "open S" "x \<in> S" |
|
259 |
then obtain a where a: "a\<in>A'" "a \<subseteq> S" using A' by blast |
|
260 |
then show "\<exists>a\<in>A. a \<subseteq> S" using a A' |
|
| 51105 | 261 |
by (intro bexI[where x=a]) (auto simp: A_def intro: image_eqI[where x="{to_nat_on A' a}"])
|
262 |
qed |
|
263 |
qed |
|
264 |
||
| 51473 | 265 |
lemma (in topological_space) first_countableI: |
| 53255 | 266 |
assumes "countable A" |
267 |
and 1: "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a" |
|
268 |
and 2: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A. a \<subseteq> S" |
|
| 51473 | 269 |
shows "\<exists>A::nat \<Rightarrow> 'a set. (\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))" |
270 |
proof (safe intro!: exI[of _ "from_nat_into A"]) |
|
| 53255 | 271 |
fix i |
| 51473 | 272 |
have "A \<noteq> {}" using 2[of UNIV] by auto
|
| 53255 | 273 |
show "x \<in> from_nat_into A i" "open (from_nat_into A i)" |
274 |
using range_from_nat_into_subset[OF `A \<noteq> {}`] 1 by auto
|
|
275 |
next |
|
276 |
fix S |
|
277 |
assume "open S" "x\<in>S" from 2[OF this] |
|
278 |
show "\<exists>i. from_nat_into A i \<subseteq> S" |
|
279 |
using subset_range_from_nat_into[OF `countable A`] by auto |
|
| 51473 | 280 |
qed |
| 51350 | 281 |
|
| 50883 | 282 |
instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology |
283 |
proof |
|
284 |
fix x :: "'a \<times> 'b" |
|
| 55522 | 285 |
obtain A where A: |
286 |
"countable A" |
|
287 |
"\<And>a. a \<in> A \<Longrightarrow> fst x \<in> a" |
|
288 |
"\<And>a. a \<in> A \<Longrightarrow> open a" |
|
289 |
"\<And>S. open S \<Longrightarrow> fst x \<in> S \<Longrightarrow> \<exists>a\<in>A. a \<subseteq> S" |
|
290 |
by (rule first_countable_basisE[of "fst x"]) blast |
|
291 |
obtain B where B: |
|
292 |
"countable B" |
|
293 |
"\<And>a. a \<in> B \<Longrightarrow> snd x \<in> a" |
|
294 |
"\<And>a. a \<in> B \<Longrightarrow> open a" |
|
295 |
"\<And>S. open S \<Longrightarrow> snd x \<in> S \<Longrightarrow> \<exists>a\<in>B. a \<subseteq> S" |
|
296 |
by (rule first_countable_basisE[of "snd x"]) blast |
|
| 53282 | 297 |
show "\<exists>A::nat \<Rightarrow> ('a \<times> 'b) set.
|
298 |
(\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))" |
|
| 51473 | 299 |
proof (rule first_countableI[of "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"], safe) |
| 53255 | 300 |
fix a b |
301 |
assume x: "a \<in> A" "b \<in> B" |
|
| 53640 | 302 |
with A(2, 3)[of a] B(2, 3)[of b] show "x \<in> a \<times> b" and "open (a \<times> b)" |
303 |
unfolding mem_Times_iff |
|
304 |
by (auto intro: open_Times) |
|
| 50883 | 305 |
next |
| 53255 | 306 |
fix S |
307 |
assume "open S" "x \<in> S" |
|
| 55522 | 308 |
then obtain a' b' where a'b': "open a'" "open b'" "x \<in> a' \<times> b'" "a' \<times> b' \<subseteq> S" |
309 |
by (rule open_prod_elim) |
|
310 |
moreover |
|
311 |
from a'b' A(4)[of a'] B(4)[of b'] |
|
312 |
obtain a b where "a \<in> A" "a \<subseteq> a'" "b \<in> B" "b \<subseteq> b'" |
|
313 |
by auto |
|
314 |
ultimately |
|
315 |
show "\<exists>a\<in>(\<lambda>(a, b). a \<times> b) ` (A \<times> B). a \<subseteq> S" |
|
| 50883 | 316 |
by (auto intro!: bexI[of _ "a \<times> b"] bexI[of _ a] bexI[of _ b]) |
317 |
qed (simp add: A B) |
|
318 |
qed |
|
319 |
||
|
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50526
diff
changeset
|
320 |
class second_countable_topology = topological_space + |
| 53282 | 321 |
assumes ex_countable_subbasis: |
322 |
"\<exists>B::'a::topological_space set set. countable B \<and> open = generate_topology B" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
323 |
begin |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
324 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
325 |
lemma ex_countable_basis: "\<exists>B::'a set set. countable B \<and> topological_basis B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
326 |
proof - |
| 53255 | 327 |
from ex_countable_subbasis obtain B where B: "countable B" "open = generate_topology B" |
328 |
by blast |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
329 |
let ?B = "Inter ` {b. finite b \<and> b \<subseteq> B }"
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
330 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
331 |
show ?thesis |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
332 |
proof (intro exI conjI) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
333 |
show "countable ?B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
334 |
by (intro countable_image countable_Collect_finite_subset B) |
| 53255 | 335 |
{
|
336 |
fix S |
|
337 |
assume "open S" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
338 |
then have "\<exists>B'\<subseteq>{b. finite b \<and> b \<subseteq> B}. (\<Union>b\<in>B'. \<Inter>b) = S"
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
339 |
unfolding B |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
340 |
proof induct |
| 53255 | 341 |
case UNIV |
342 |
show ?case by (intro exI[of _ "{{}}"]) simp
|
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
343 |
next |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
344 |
case (Int a b) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
345 |
then obtain x y where x: "a = UNION x Inter" "\<And>i. i \<in> x \<Longrightarrow> finite i \<and> i \<subseteq> B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
346 |
and y: "b = UNION y Inter" "\<And>i. i \<in> y \<Longrightarrow> finite i \<and> i \<subseteq> B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
347 |
by blast |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
348 |
show ?case |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
349 |
unfolding x y Int_UN_distrib2 |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
350 |
by (intro exI[of _ "{i \<union> j| i j. i \<in> x \<and> j \<in> y}"]) (auto dest: x(2) y(2))
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
351 |
next |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
352 |
case (UN K) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
353 |
then have "\<forall>k\<in>K. \<exists>B'\<subseteq>{b. finite b \<and> b \<subseteq> B}. UNION B' Inter = k" by auto
|
| 55522 | 354 |
then obtain k where |
355 |
"\<forall>ka\<in>K. k ka \<subseteq> {b. finite b \<and> b \<subseteq> B} \<and> UNION (k ka) Inter = ka"
|
|
356 |
unfolding bchoice_iff .. |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
357 |
then show "\<exists>B'\<subseteq>{b. finite b \<and> b \<subseteq> B}. UNION B' Inter = \<Union>K"
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
358 |
by (intro exI[of _ "UNION K k"]) auto |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
359 |
next |
| 53255 | 360 |
case (Basis S) |
361 |
then show ?case |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
362 |
by (intro exI[of _ "{{S}}"]) auto
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
363 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
364 |
then have "(\<exists>B'\<subseteq>Inter ` {b. finite b \<and> b \<subseteq> B}. \<Union>B' = S)"
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
365 |
unfolding subset_image_iff by blast } |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
366 |
then show "topological_basis ?B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
367 |
unfolding topological_space_class.topological_basis_def |
| 53282 | 368 |
by (safe intro!: topological_space_class.open_Inter) |
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
369 |
(simp_all add: B generate_topology.Basis subset_eq) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
370 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
371 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
372 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
373 |
end |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
374 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
375 |
sublocale second_countable_topology < |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
376 |
countable_basis "SOME B. countable B \<and> topological_basis B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
377 |
using someI_ex[OF ex_countable_basis] |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
378 |
by unfold_locales safe |
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
379 |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
380 |
instance prod :: (second_countable_topology, second_countable_topology) second_countable_topology |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
381 |
proof |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
382 |
obtain A :: "'a set set" where "countable A" "topological_basis A" |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
383 |
using ex_countable_basis by auto |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
384 |
moreover |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
385 |
obtain B :: "'b set set" where "countable B" "topological_basis B" |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
386 |
using ex_countable_basis by auto |
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
387 |
ultimately show "\<exists>B::('a \<times> 'b) set set. countable B \<and> open = generate_topology B"
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
388 |
by (auto intro!: exI[of _ "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"] topological_basis_prod |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
389 |
topological_basis_imp_subbasis) |
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
390 |
qed |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
391 |
|
| 50883 | 392 |
instance second_countable_topology \<subseteq> first_countable_topology |
393 |
proof |
|
394 |
fix x :: 'a |
|
395 |
def B \<equiv> "SOME B::'a set set. countable B \<and> topological_basis B" |
|
396 |
then have B: "countable B" "topological_basis B" |
|
397 |
using countable_basis is_basis |
|
398 |
by (auto simp: countable_basis is_basis) |
|
| 53282 | 399 |
then show "\<exists>A::nat \<Rightarrow> 'a set. |
400 |
(\<forall>i. x \<in> A i \<and> open (A i)) \<and> (\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. A i \<subseteq> S))" |
|
| 51473 | 401 |
by (intro first_countableI[of "{b\<in>B. x \<in> b}"])
|
402 |
(fastforce simp: topological_space_class.topological_basis_def)+ |
|
| 50883 | 403 |
qed |
404 |
||
| 53255 | 405 |
|
| 50087 | 406 |
subsection {* Polish spaces *}
|
407 |
||
408 |
text {* Textbooks define Polish spaces as completely metrizable.
|
|
409 |
We assume the topology to be complete for a given metric. *} |
|
410 |
||
|
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50526
diff
changeset
|
411 |
class polish_space = complete_space + second_countable_topology |
| 50087 | 412 |
|
| 44517 | 413 |
subsection {* General notion of a topology as a value *}
|
| 33175 | 414 |
|
| 53255 | 415 |
definition "istopology L \<longleftrightarrow> |
416 |
L {} \<and> (\<forall>S T. L S \<longrightarrow> L T \<longrightarrow> L (S \<inter> T)) \<and> (\<forall>K. Ball K L \<longrightarrow> L (\<Union> K))"
|
|
417 |
||
| 49834 | 418 |
typedef 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
|
| 33175 | 419 |
morphisms "openin" "topology" |
420 |
unfolding istopology_def by blast |
|
421 |
||
422 |
lemma istopology_open_in[intro]: "istopology(openin U)" |
|
423 |
using openin[of U] by blast |
|
424 |
||
425 |
lemma topology_inverse': "istopology U \<Longrightarrow> openin (topology U) = U" |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
426 |
using topology_inverse[unfolded mem_Collect_eq] . |
| 33175 | 427 |
|
428 |
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U" |
|
429 |
using topology_inverse[of U] istopology_open_in[of "topology U"] by auto |
|
430 |
||
431 |
lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)" |
|
| 53255 | 432 |
proof |
433 |
assume "T1 = T2" |
|
434 |
then show "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp |
|
435 |
next |
|
436 |
assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" |
|
437 |
then have "openin T1 = openin T2" by (simp add: fun_eq_iff) |
|
438 |
then have "topology (openin T1) = topology (openin T2)" by simp |
|
439 |
then show "T1 = T2" unfolding openin_inverse . |
|
| 33175 | 440 |
qed |
441 |
||
442 |
text{* Infer the "universe" from union of all sets in the topology. *}
|
|
443 |
||
| 53640 | 444 |
definition "topspace T = \<Union>{S. openin T S}"
|
| 33175 | 445 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
446 |
subsubsection {* Main properties of open sets *}
|
| 33175 | 447 |
|
448 |
lemma openin_clauses: |
|
449 |
fixes U :: "'a topology" |
|
| 53282 | 450 |
shows |
451 |
"openin U {}"
|
|
452 |
"\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)" |
|
453 |
"\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)" |
|
454 |
using openin[of U] unfolding istopology_def mem_Collect_eq by fast+ |
|
| 33175 | 455 |
|
456 |
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U" |
|
457 |
unfolding topspace_def by blast |
|
| 53255 | 458 |
|
459 |
lemma openin_empty[simp]: "openin U {}"
|
|
460 |
by (simp add: openin_clauses) |
|
| 33175 | 461 |
|
462 |
lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)" |
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
463 |
using openin_clauses by simp |
|
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
464 |
|
|
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
465 |
lemma openin_Union[intro]: "(\<forall>S \<in>K. openin U S) \<Longrightarrow> openin U (\<Union> K)" |
|
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
466 |
using openin_clauses by simp |
| 33175 | 467 |
|
468 |
lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)" |
|
469 |
using openin_Union[of "{S,T}" U] by auto
|
|
470 |
||
| 53255 | 471 |
lemma openin_topspace[intro, simp]: "openin U (topspace U)" |
472 |
by (simp add: openin_Union topspace_def) |
|
| 33175 | 473 |
|
| 49711 | 474 |
lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)" |
475 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
| 36584 | 476 |
proof |
| 49711 | 477 |
assume ?lhs |
478 |
then show ?rhs by auto |
|
| 36584 | 479 |
next |
480 |
assume H: ?rhs |
|
481 |
let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
|
|
482 |
have "openin U ?t" by (simp add: openin_Union) |
|
483 |
also have "?t = S" using H by auto |
|
484 |
finally show "openin U S" . |
|
| 33175 | 485 |
qed |
486 |
||
| 49711 | 487 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
488 |
subsubsection {* Closed sets *}
|
| 33175 | 489 |
|
490 |
definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)" |
|
491 |
||
| 53255 | 492 |
lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" |
493 |
by (metis closedin_def) |
|
494 |
||
495 |
lemma closedin_empty[simp]: "closedin U {}"
|
|
496 |
by (simp add: closedin_def) |
|
497 |
||
498 |
lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" |
|
499 |
by (simp add: closedin_def) |
|
500 |
||
| 33175 | 501 |
lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)" |
502 |
by (auto simp add: Diff_Un closedin_def) |
|
503 |
||
| 53255 | 504 |
lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union> {A - s|s. s\<in>S}"
|
505 |
by auto |
|
506 |
||
507 |
lemma closedin_Inter[intro]: |
|
508 |
assumes Ke: "K \<noteq> {}"
|
|
509 |
and Kc: "\<forall>S \<in>K. closedin U S" |
|
510 |
shows "closedin U (\<Inter> K)" |
|
511 |
using Ke Kc unfolding closedin_def Diff_Inter by auto |
|
| 33175 | 512 |
|
513 |
lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)" |
|
514 |
using closedin_Inter[of "{S,T}" U] by auto
|
|
515 |
||
| 53255 | 516 |
lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B" |
517 |
by blast |
|
518 |
||
| 33175 | 519 |
lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)" |
520 |
apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2) |
|
521 |
apply (metis openin_subset subset_eq) |
|
522 |
done |
|
523 |
||
| 53255 | 524 |
lemma openin_closedin: "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))" |
| 33175 | 525 |
by (simp add: openin_closedin_eq) |
526 |
||
| 53255 | 527 |
lemma openin_diff[intro]: |
528 |
assumes oS: "openin U S" |
|
529 |
and cT: "closedin U T" |
|
530 |
shows "openin U (S - T)" |
|
531 |
proof - |
|
| 33175 | 532 |
have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S] oS cT |
533 |
by (auto simp add: topspace_def openin_subset) |
|
| 53282 | 534 |
then show ?thesis using oS cT |
535 |
by (auto simp add: closedin_def) |
|
| 33175 | 536 |
qed |
537 |
||
| 53255 | 538 |
lemma closedin_diff[intro]: |
539 |
assumes oS: "closedin U S" |
|
540 |
and cT: "openin U T" |
|
541 |
shows "closedin U (S - T)" |
|
542 |
proof - |
|
543 |
have "S - T = S \<inter> (topspace U - T)" |
|
| 53282 | 544 |
using closedin_subset[of U S] oS cT by (auto simp add: topspace_def) |
| 53255 | 545 |
then show ?thesis |
546 |
using oS cT by (auto simp add: openin_closedin_eq) |
|
547 |
qed |
|
548 |
||
| 33175 | 549 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
550 |
subsubsection {* Subspace topology *}
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
551 |
|
|
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
552 |
definition "subtopology U V = topology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)" |
|
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
553 |
|
|
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
554 |
lemma istopology_subtopology: "istopology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)" |
|
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
555 |
(is "istopology ?L") |
| 53255 | 556 |
proof - |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
557 |
have "?L {}" by blast
|
| 53255 | 558 |
{
|
559 |
fix A B |
|
560 |
assume A: "?L A" and B: "?L B" |
|
561 |
from A B obtain Sa and Sb where Sa: "openin U Sa" "A = Sa \<inter> V" and Sb: "openin U Sb" "B = Sb \<inter> V" |
|
562 |
by blast |
|
563 |
have "A \<inter> B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)" |
|
564 |
using Sa Sb by blast+ |
|
565 |
then have "?L (A \<inter> B)" by blast |
|
566 |
} |
|
| 33175 | 567 |
moreover |
| 53255 | 568 |
{
|
| 53282 | 569 |
fix K |
570 |
assume K: "K \<subseteq> Collect ?L" |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
571 |
have th0: "Collect ?L = (\<lambda>S. S \<inter> V) ` Collect (openin U)" |
| 55775 | 572 |
by blast |
| 33175 | 573 |
from K[unfolded th0 subset_image_iff] |
| 53255 | 574 |
obtain Sk where Sk: "Sk \<subseteq> Collect (openin U)" "K = (\<lambda>S. S \<inter> V) ` Sk" |
575 |
by blast |
|
576 |
have "\<Union>K = (\<Union>Sk) \<inter> V" |
|
577 |
using Sk by auto |
|
578 |
moreover have "openin U (\<Union> Sk)" |
|
579 |
using Sk by (auto simp add: subset_eq) |
|
580 |
ultimately have "?L (\<Union>K)" by blast |
|
581 |
} |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
582 |
ultimately show ?thesis |
|
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
583 |
unfolding subset_eq mem_Collect_eq istopology_def by blast |
| 33175 | 584 |
qed |
585 |
||
| 53255 | 586 |
lemma openin_subtopology: "openin (subtopology U V) S \<longleftrightarrow> (\<exists>T. openin U T \<and> S = T \<inter> V)" |
| 33175 | 587 |
unfolding subtopology_def topology_inverse'[OF istopology_subtopology] |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
588 |
by auto |
| 33175 | 589 |
|
| 53255 | 590 |
lemma topspace_subtopology: "topspace (subtopology U V) = topspace U \<inter> V" |
| 33175 | 591 |
by (auto simp add: topspace_def openin_subtopology) |
592 |
||
| 53255 | 593 |
lemma closedin_subtopology: "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)" |
| 33175 | 594 |
unfolding closedin_def topspace_subtopology |
| 55775 | 595 |
by (auto simp add: openin_subtopology) |
| 33175 | 596 |
|
597 |
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U" |
|
598 |
unfolding openin_subtopology |
|
| 55775 | 599 |
by auto (metis IntD1 in_mono openin_subset) |
| 49711 | 600 |
|
601 |
lemma subtopology_superset: |
|
602 |
assumes UV: "topspace U \<subseteq> V" |
|
| 33175 | 603 |
shows "subtopology U V = U" |
| 53255 | 604 |
proof - |
605 |
{
|
|
606 |
fix S |
|
607 |
{
|
|
608 |
fix T |
|
609 |
assume T: "openin U T" "S = T \<inter> V" |
|
610 |
from T openin_subset[OF T(1)] UV have eq: "S = T" |
|
611 |
by blast |
|
612 |
have "openin U S" |
|
613 |
unfolding eq using T by blast |
|
614 |
} |
|
| 33175 | 615 |
moreover |
| 53255 | 616 |
{
|
617 |
assume S: "openin U S" |
|
618 |
then have "\<exists>T. openin U T \<and> S = T \<inter> V" |
|
619 |
using openin_subset[OF S] UV by auto |
|
620 |
} |
|
621 |
ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" |
|
622 |
by blast |
|
623 |
} |
|
624 |
then show ?thesis |
|
625 |
unfolding topology_eq openin_subtopology by blast |
|
| 33175 | 626 |
qed |
627 |
||
628 |
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" |
|
629 |
by (simp add: subtopology_superset) |
|
630 |
||
631 |
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" |
|
632 |
by (simp add: subtopology_superset) |
|
633 |
||
| 53255 | 634 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
635 |
subsubsection {* The standard Euclidean topology *}
|
| 33175 | 636 |
|
| 53255 | 637 |
definition euclidean :: "'a::topological_space topology" |
638 |
where "euclidean = topology open" |
|
| 33175 | 639 |
|
640 |
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S" |
|
641 |
unfolding euclidean_def |
|
642 |
apply (rule cong[where x=S and y=S]) |
|
643 |
apply (rule topology_inverse[symmetric]) |
|
644 |
apply (auto simp add: istopology_def) |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
645 |
done |
| 33175 | 646 |
|
647 |
lemma topspace_euclidean: "topspace euclidean = UNIV" |
|
648 |
apply (simp add: topspace_def) |
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
649 |
apply (rule set_eqI) |
| 53255 | 650 |
apply (auto simp add: open_openin[symmetric]) |
651 |
done |
|
| 33175 | 652 |
|
653 |
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S" |
|
654 |
by (simp add: topspace_euclidean topspace_subtopology) |
|
655 |
||
656 |
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S" |
|
657 |
by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV) |
|
658 |
||
659 |
lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)" |
|
660 |
by (simp add: open_openin openin_subopen[symmetric]) |
|
661 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
662 |
text {* Basic "localization" results are handy for connectedness. *}
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
663 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
664 |
lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
665 |
by (auto simp add: openin_subtopology open_openin[symmetric]) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
666 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
667 |
lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
668 |
by (auto simp add: openin_open) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
669 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
670 |
lemma open_openin_trans[trans]: |
| 53255 | 671 |
"open S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> openin (subtopology euclidean S) T" |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
672 |
by (metis Int_absorb1 openin_open_Int) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
673 |
|
| 53255 | 674 |
lemma open_subset: "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S" |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
675 |
by (auto simp add: openin_open) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
676 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
677 |
lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
678 |
by (simp add: closedin_subtopology closed_closedin Int_ac) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
679 |
|
| 53291 | 680 |
lemma closedin_closed_Int: "closed S \<Longrightarrow> closedin (subtopology euclidean U) (U \<inter> S)" |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
681 |
by (metis closedin_closed) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
682 |
|
| 53282 | 683 |
lemma closed_closedin_trans: |
684 |
"closed S \<Longrightarrow> closed T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> closedin (subtopology euclidean S) T" |
|
| 55775 | 685 |
by (metis closedin_closed inf.absorb2) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
686 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
687 |
lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
688 |
by (auto simp add: closedin_closed) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
689 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
690 |
lemma openin_euclidean_subtopology_iff: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
691 |
fixes S U :: "'a::metric_space set" |
| 53255 | 692 |
shows "openin (subtopology euclidean U) S \<longleftrightarrow> |
693 |
S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" |
|
694 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
695 |
proof |
| 53255 | 696 |
assume ?lhs |
| 53282 | 697 |
then show ?rhs |
698 |
unfolding openin_open open_dist by blast |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
699 |
next |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
700 |
def T \<equiv> "{x. \<exists>a\<in>S. \<exists>d>0. (\<forall>y\<in>U. dist y a < d \<longrightarrow> y \<in> S) \<and> dist x a < d}"
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
701 |
have 1: "\<forall>x\<in>T. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> T" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
702 |
unfolding T_def |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
703 |
apply clarsimp |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
704 |
apply (rule_tac x="d - dist x a" in exI) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
705 |
apply (clarsimp simp add: less_diff_eq) |
| 55775 | 706 |
by (metis dist_commute dist_triangle_lt) |
| 53282 | 707 |
assume ?rhs then have 2: "S = U \<inter> T" |
| 55775 | 708 |
unfolding T_def |
709 |
by auto (metis dist_self) |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
710 |
from 1 2 show ?lhs |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
711 |
unfolding openin_open open_dist by fast |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
712 |
qed |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
713 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
714 |
text {* These "transitivity" results are handy too *}
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
715 |
|
| 53255 | 716 |
lemma openin_trans[trans]: |
717 |
"openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T \<Longrightarrow> |
|
718 |
openin (subtopology euclidean U) S" |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
719 |
unfolding open_openin openin_open by blast |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
720 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
721 |
lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
722 |
by (auto simp add: openin_open intro: openin_trans) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
723 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
724 |
lemma closedin_trans[trans]: |
| 53255 | 725 |
"closedin (subtopology euclidean T) S \<Longrightarrow> closedin (subtopology euclidean U) T \<Longrightarrow> |
726 |
closedin (subtopology euclidean U) S" |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
727 |
by (auto simp add: closedin_closed closed_closedin closed_Inter Int_assoc) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
728 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
729 |
lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
730 |
by (auto simp add: closedin_closed intro: closedin_trans) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
731 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
732 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
733 |
subsection {* Open and closed balls *}
|
| 33175 | 734 |
|
| 53255 | 735 |
definition ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" |
736 |
where "ball x e = {y. dist x y < e}"
|
|
737 |
||
738 |
definition cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" |
|
739 |
where "cball x e = {y. dist x y \<le> e}"
|
|
| 33175 | 740 |
|
|
45776
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
741 |
lemma mem_ball [simp]: "y \<in> ball x e \<longleftrightarrow> dist x y < e" |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
742 |
by (simp add: ball_def) |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
743 |
|
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
744 |
lemma mem_cball [simp]: "y \<in> cball x e \<longleftrightarrow> dist x y \<le> e" |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
745 |
by (simp add: cball_def) |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
746 |
|
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
747 |
lemma mem_ball_0: |
| 33175 | 748 |
fixes x :: "'a::real_normed_vector" |
749 |
shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e" |
|
750 |
by (simp add: dist_norm) |
|
751 |
||
|
45776
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
752 |
lemma mem_cball_0: |
| 33175 | 753 |
fixes x :: "'a::real_normed_vector" |
754 |
shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e" |
|
755 |
by (simp add: dist_norm) |
|
756 |
||
|
45776
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
757 |
lemma centre_in_ball: "x \<in> ball x e \<longleftrightarrow> 0 < e" |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
758 |
by simp |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
759 |
|
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
760 |
lemma centre_in_cball: "x \<in> cball x e \<longleftrightarrow> 0 \<le> e" |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
761 |
by simp |
|
714100f5fda4
remove mem_(c)ball_0 and centre_in_(c)ball from simpset, as rules mem_(c)ball always match instead
huffman
parents:
45548
diff
changeset
|
762 |
|
| 53255 | 763 |
lemma ball_subset_cball[simp,intro]: "ball x e \<subseteq> cball x e" |
764 |
by (simp add: subset_eq) |
|
765 |
||
| 53282 | 766 |
lemma subset_ball[intro]: "d \<le> e \<Longrightarrow> ball x d \<subseteq> ball x e" |
| 53255 | 767 |
by (simp add: subset_eq) |
768 |
||
| 53282 | 769 |
lemma subset_cball[intro]: "d \<le> e \<Longrightarrow> cball x d \<subseteq> cball x e" |
| 53255 | 770 |
by (simp add: subset_eq) |
771 |
||
| 33175 | 772 |
lemma ball_max_Un: "ball a (max r s) = ball a r \<union> ball a s" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
773 |
by (simp add: set_eq_iff) arith |
| 33175 | 774 |
|
775 |
lemma ball_min_Int: "ball a (min r s) = ball a r \<inter> ball a s" |
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
776 |
by (simp add: set_eq_iff) |
| 33175 | 777 |
|
| 53255 | 778 |
lemma diff_less_iff: |
779 |
"(a::real) - b > 0 \<longleftrightarrow> a > b" |
|
| 33175 | 780 |
"(a::real) - b < 0 \<longleftrightarrow> a < b" |
| 53255 | 781 |
"a - b < c \<longleftrightarrow> a < c + b" "a - b > c \<longleftrightarrow> a > c + b" |
782 |
by arith+ |
|
783 |
||
784 |
lemma diff_le_iff: |
|
785 |
"(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b" |
|
786 |
"(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b" |
|
787 |
"a - b \<le> c \<longleftrightarrow> a \<le> c + b" |
|
788 |
"a - b \<ge> c \<longleftrightarrow> a \<ge> c + b" |
|
789 |
by arith+ |
|
| 33175 | 790 |
|
| 54070 | 791 |
lemma open_ball [intro, simp]: "open (ball x e)" |
792 |
proof - |
|
793 |
have "open (dist x -` {..<e})"
|
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
794 |
by (intro open_vimage open_lessThan continuous_intros) |
| 54070 | 795 |
also have "dist x -` {..<e} = ball x e"
|
796 |
by auto |
|
797 |
finally show ?thesis . |
|
798 |
qed |
|
| 33175 | 799 |
|
800 |
lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)" |
|
801 |
unfolding open_dist subset_eq mem_ball Ball_def dist_commute .. |
|
802 |
||
|
33714
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
803 |
lemma openE[elim?]: |
| 53282 | 804 |
assumes "open S" "x\<in>S" |
|
33714
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
805 |
obtains e where "e>0" "ball x e \<subseteq> S" |
|
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
806 |
using assms unfolding open_contains_ball by auto |
|
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
807 |
|
| 33175 | 808 |
lemma open_contains_ball_eq: "open S \<Longrightarrow> \<forall>x. x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)" |
809 |
by (metis open_contains_ball subset_eq centre_in_ball) |
|
810 |
||
811 |
lemma ball_eq_empty[simp]: "ball x e = {} \<longleftrightarrow> e \<le> 0"
|
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
812 |
unfolding mem_ball set_eq_iff |
| 33175 | 813 |
apply (simp add: not_less) |
| 52624 | 814 |
apply (metis zero_le_dist order_trans dist_self) |
815 |
done |
|
| 33175 | 816 |
|
| 53291 | 817 |
lemma ball_empty[intro]: "e \<le> 0 \<Longrightarrow> ball x e = {}" by simp
|
| 33175 | 818 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
819 |
lemma euclidean_dist_l2: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
820 |
fixes x y :: "'a :: euclidean_space" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
821 |
shows "dist x y = setL2 (\<lambda>i. dist (x \<bullet> i) (y \<bullet> i)) Basis" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
822 |
unfolding dist_norm norm_eq_sqrt_inner setL2_def |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
823 |
by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
824 |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
825 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
826 |
subsection {* Boxes *}
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
827 |
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
828 |
definition (in euclidean_space) eucl_less (infix "<e" 50) |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
829 |
where "eucl_less a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i)" |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
830 |
|
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
831 |
definition box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
|
| 56188 | 832 |
definition "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}"
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
833 |
|
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
834 |
lemma box_def: "box a b = {x. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
|
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
835 |
and in_box_eucl_less: "x \<in> box a b \<longleftrightarrow> a <e x \<and> x <e b" |
| 56188 | 836 |
and mem_box: "x \<in> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i)" |
837 |
"x \<in> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)" |
|
838 |
by (auto simp: box_eucl_less eucl_less_def cbox_def) |
|
839 |
||
840 |
lemma mem_box_real[simp]: |
|
841 |
"(x::real) \<in> box a b \<longleftrightarrow> a < x \<and> x < b" |
|
842 |
"(x::real) \<in> cbox a b \<longleftrightarrow> a \<le> x \<and> x \<le> b" |
|
843 |
by (auto simp: mem_box) |
|
844 |
||
845 |
lemma box_real[simp]: |
|
846 |
fixes a b:: real |
|
847 |
shows "box a b = {a <..< b}" "cbox a b = {a .. b}"
|
|
848 |
by auto |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
849 |
|
| 50087 | 850 |
lemma rational_boxes: |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
851 |
fixes x :: "'a\<Colon>euclidean_space" |
| 53291 | 852 |
assumes "e > 0" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
853 |
shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat> ) \<and> x \<in> box a b \<and> box a b \<subseteq> ball x e" |
| 50087 | 854 |
proof - |
855 |
def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
|
|
| 53291 | 856 |
then have e: "e' > 0" |
| 56541 | 857 |
using assms by (auto simp: DIM_positive) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
858 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i") |
| 50087 | 859 |
proof |
| 53255 | 860 |
fix i |
861 |
from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e |
|
862 |
show "?th i" by auto |
|
| 50087 | 863 |
qed |
| 55522 | 864 |
from choice[OF this] obtain a where |
865 |
a: "\<forall>xa. a xa \<in> \<rat> \<and> a xa < x \<bullet> xa \<and> x \<bullet> xa - a xa < e'" .. |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
866 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i") |
| 50087 | 867 |
proof |
| 53255 | 868 |
fix i |
869 |
from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e |
|
870 |
show "?th i" by auto |
|
| 50087 | 871 |
qed |
| 55522 | 872 |
from choice[OF this] obtain b where |
873 |
b: "\<forall>xa. b xa \<in> \<rat> \<and> x \<bullet> xa < b xa \<and> b xa - x \<bullet> xa < e'" .. |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
874 |
let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
875 |
show ?thesis |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
876 |
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) |
| 53255 | 877 |
fix y :: 'a |
878 |
assume *: "y \<in> box ?a ?b" |
|
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
879 |
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" |
| 50087 | 880 |
unfolding setL2_def[symmetric] by (rule euclidean_dist_l2) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
881 |
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
|
| 50087 | 882 |
proof (rule real_sqrt_less_mono, rule setsum_strict_mono) |
| 53255 | 883 |
fix i :: "'a" |
884 |
assume i: "i \<in> Basis" |
|
885 |
have "a i < y\<bullet>i \<and> y\<bullet>i < b i" |
|
886 |
using * i by (auto simp: box_def) |
|
887 |
moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" |
|
888 |
using a by auto |
|
889 |
moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" |
|
890 |
using b by auto |
|
891 |
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" |
|
892 |
by auto |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
893 |
then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
|
| 50087 | 894 |
unfolding e'_def by (auto simp: dist_real_def) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
895 |
then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
|
| 50087 | 896 |
by (rule power_strict_mono) auto |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
897 |
then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
|
| 50087 | 898 |
by (simp add: power_divide) |
899 |
qed auto |
|
| 53255 | 900 |
also have "\<dots> = e" |
901 |
using `0 < e` by (simp add: real_eq_of_nat) |
|
902 |
finally show "y \<in> ball x e" |
|
903 |
by (auto simp: ball_def) |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
904 |
qed (insert a b, auto simp: box_def) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
905 |
qed |
| 51103 | 906 |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
907 |
lemma open_UNION_box: |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
908 |
fixes M :: "'a\<Colon>euclidean_space set" |
| 53282 | 909 |
assumes "open M" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
910 |
defines "a' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
911 |
defines "b' \<equiv> \<lambda>f :: 'a \<Rightarrow> real \<times> real. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
912 |
defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (a' f) (b' f) \<subseteq> M}"
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
913 |
shows "M = (\<Union>f\<in>I. box (a' f) (b' f))" |
| 52624 | 914 |
proof - |
915 |
{
|
|
916 |
fix x assume "x \<in> M" |
|
917 |
obtain e where e: "e > 0" "ball x e \<subseteq> M" |
|
918 |
using openE[OF `open M` `x \<in> M`] by auto |
|
| 53282 | 919 |
moreover obtain a b where ab: |
920 |
"x \<in> box a b" |
|
921 |
"\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" |
|
922 |
"\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>" |
|
923 |
"box a b \<subseteq> ball x e" |
|
| 52624 | 924 |
using rational_boxes[OF e(1)] by metis |
925 |
ultimately have "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))" |
|
926 |
by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) |
|
927 |
(auto simp: euclidean_representation I_def a'_def b'_def) |
|
928 |
} |
|
929 |
then show ?thesis by (auto simp: I_def) |
|
930 |
qed |
|
931 |
||
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
932 |
lemma box_eq_empty: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
933 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
934 |
shows "(box a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
935 |
and "(cbox a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
936 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
937 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
938 |
fix i x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
939 |
assume i: "i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
940 |
then have "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
941 |
unfolding mem_box by (auto simp: box_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
942 |
then have "a\<bullet>i < b\<bullet>i" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
943 |
then have False using as by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
944 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
945 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
946 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
947 |
assume as: "\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
948 |
let ?x = "(1/2) *\<^sub>R (a + b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
949 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
950 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
951 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
952 |
have "a\<bullet>i < b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
953 |
using as[THEN bspec[where x=i]] i by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
954 |
then have "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
955 |
by (auto simp: inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
956 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
957 |
then have "box a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
958 |
using mem_box(1)[of "?x" a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
959 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
960 |
ultimately show ?th1 by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
961 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
962 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
963 |
fix i x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
964 |
assume i: "i \<in> Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
965 |
then have "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
966 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
967 |
then have "a\<bullet>i \<le> b\<bullet>i" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
968 |
then have False using as by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
969 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
970 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
971 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
972 |
assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
973 |
let ?x = "(1/2) *\<^sub>R (a + b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
974 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
975 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
976 |
assume i:"i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
977 |
have "a\<bullet>i \<le> b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
978 |
using as[THEN bspec[where x=i]] i by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
979 |
then have "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
980 |
by (auto simp: inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
981 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
982 |
then have "cbox a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
983 |
using mem_box(2)[of "?x" a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
984 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
985 |
ultimately show ?th2 by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
986 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
987 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
988 |
lemma box_ne_empty: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
989 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
990 |
shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i)"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
991 |
and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
992 |
unfolding box_eq_empty[of a b] by fastforce+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
993 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
994 |
lemma |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
995 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
996 |
shows cbox_sing: "cbox a a = {a}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
997 |
and box_sing: "box a a = {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
998 |
unfolding set_eq_iff mem_box eq_iff [symmetric] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
999 |
by (auto intro!: euclidean_eqI[where 'a='a]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1000 |
(metis all_not_in_conv nonempty_Basis) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1001 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1002 |
lemma subset_box_imp: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1003 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1004 |
shows "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1005 |
and "(\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1006 |
and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1007 |
and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1008 |
unfolding subset_eq[unfolded Ball_def] unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1009 |
by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1010 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1011 |
lemma box_subset_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1012 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1013 |
shows "box a b \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1014 |
unfolding subset_eq [unfolded Ball_def] mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1015 |
by (fast intro: less_imp_le) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1016 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1017 |
lemma subset_box: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1018 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1019 |
shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th1) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1020 |
and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i)" (is ?th2) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1021 |
and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th3) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1022 |
and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) --> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th4) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1023 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1024 |
show ?th1 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1025 |
unfolding subset_eq and Ball_def and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1026 |
by (auto intro: order_trans) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1027 |
show ?th2 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1028 |
unfolding subset_eq and Ball_def and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1029 |
by (auto intro: le_less_trans less_le_trans order_trans less_imp_le) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1030 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1031 |
assume as: "box c d \<subseteq> cbox a b" "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1032 |
then have "box c d \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1033 |
unfolding box_eq_empty by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1034 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1035 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1036 |
(** TODO combine the following two parts as done in the HOL_light version. **) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1037 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1038 |
let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((min (a\<bullet>j) (d\<bullet>j))+c\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1039 |
assume as2: "a\<bullet>i > c\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1040 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1041 |
fix j :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1042 |
assume j: "j \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1043 |
then have "c \<bullet> j < ?x \<bullet> j \<and> ?x \<bullet> j < d \<bullet> j" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1044 |
apply (cases "j = i") |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1045 |
using as(2)[THEN bspec[where x=j]] i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1046 |
apply (auto simp add: as2) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1047 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1048 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1049 |
then have "?x\<in>box c d" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1050 |
using i unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1051 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1052 |
have "?x \<notin> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1053 |
unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1054 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1055 |
apply (rule_tac x=i in bexI) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1056 |
using as(2)[THEN bspec[where x=i]] and as2 i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1057 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1058 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1059 |
ultimately have False using as by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1060 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1061 |
then have "a\<bullet>i \<le> c\<bullet>i" by (rule ccontr) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1062 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1063 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1064 |
let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((max (b\<bullet>j) (c\<bullet>j))+d\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1065 |
assume as2: "b\<bullet>i < d\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1066 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1067 |
fix j :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1068 |
assume "j\<in>Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1069 |
then have "d \<bullet> j > ?x \<bullet> j \<and> ?x \<bullet> j > c \<bullet> j" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1070 |
apply (cases "j = i") |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1071 |
using as(2)[THEN bspec[where x=j]] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1072 |
apply (auto simp add: as2) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1073 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1074 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1075 |
then have "?x\<in>box c d" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1076 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1077 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1078 |
have "?x\<notin>cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1079 |
unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1080 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1081 |
apply (rule_tac x=i in bexI) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1082 |
using as(2)[THEN bspec[where x=i]] and as2 using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1083 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1084 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1085 |
ultimately have False using as by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1086 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1087 |
then have "b\<bullet>i \<ge> d\<bullet>i" by (rule ccontr) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1088 |
ultimately |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1089 |
have "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1090 |
} note part1 = this |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1091 |
show ?th3 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1092 |
unfolding subset_eq and Ball_def and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1093 |
apply (rule, rule, rule, rule) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1094 |
apply (rule part1) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1095 |
unfolding subset_eq and Ball_def and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1096 |
prefer 4 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1097 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1098 |
apply (erule_tac x=xa in allE, erule_tac x=xa in allE, fastforce)+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1099 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1100 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1101 |
assume as: "box c d \<subseteq> box a b" "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1102 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1103 |
assume i:"i\<in>Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1104 |
from as(1) have "box c d \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1105 |
using box_subset_cbox[of a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1106 |
then have "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1107 |
using part1 and as(2) using i by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1108 |
} note * = this |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1109 |
show ?th4 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1110 |
unfolding subset_eq and Ball_def and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1111 |
apply (rule, rule, rule, rule) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1112 |
apply (rule *) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1113 |
unfolding subset_eq and Ball_def and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1114 |
prefer 4 |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1115 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1116 |
apply (erule_tac x=xa in allE, simp)+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1117 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1118 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1119 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1120 |
lemma inter_interval: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1121 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1122 |
shows "cbox a b \<inter> cbox c d = |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1123 |
cbox (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1124 |
unfolding set_eq_iff and Int_iff and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1125 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1126 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1127 |
lemma disjoint_interval: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1128 |
fixes a::"'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1129 |
shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i < c\<bullet>i \<or> d\<bullet>i < a\<bullet>i))" (is ?th1)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1130 |
and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th2)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1131 |
and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th3)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1132 |
and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th4)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1133 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1134 |
let ?z = "(\<Sum>i\<in>Basis. (((max (a\<bullet>i) (c\<bullet>i)) + (min (b\<bullet>i) (d\<bullet>i))) / 2) *\<^sub>R i)::'a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1135 |
have **: "\<And>P Q. (\<And>i :: 'a. i \<in> Basis \<Longrightarrow> Q ?z i \<Longrightarrow> P i) \<Longrightarrow> |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1136 |
(\<And>i x :: 'a. i \<in> Basis \<Longrightarrow> P i \<Longrightarrow> Q x i) \<Longrightarrow> (\<forall>x. \<exists>i\<in>Basis. Q x i) \<longleftrightarrow> (\<exists>i\<in>Basis. P i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1137 |
by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1138 |
note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1139 |
show ?th1 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1140 |
show ?th2 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1141 |
show ?th3 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1142 |
show ?th4 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1143 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1144 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1145 |
text {* Intervals in general, including infinite and mixtures of open and closed. *}
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1146 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1147 |
definition "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1148 |
(\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1149 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1150 |
lemma is_interval_cbox: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1151 |
and is_interval_box: "is_interval (box a b)" (is ?th2) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1152 |
unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1153 |
by (meson order_trans le_less_trans less_le_trans less_trans)+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1154 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1155 |
lemma is_interval_empty: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1156 |
"is_interval {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1157 |
unfolding is_interval_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1158 |
by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1159 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1160 |
lemma is_interval_univ: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1161 |
"is_interval UNIV" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1162 |
unfolding is_interval_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1163 |
by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1164 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1165 |
lemma mem_is_intervalI: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1166 |
assumes "is_interval s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1167 |
assumes "a \<in> s" "b \<in> s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1168 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i \<or> b \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> a \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1169 |
shows "x \<in> s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1170 |
by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1171 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1172 |
lemma interval_subst: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1173 |
fixes S::"'a::euclidean_space set" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1174 |
assumes "is_interval S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1175 |
assumes "x \<in> S" "y j \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1176 |
assumes "j \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1177 |
shows "(\<Sum>i\<in>Basis. (if i = j then y i \<bullet> i else x \<bullet> i) *\<^sub>R i) \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1178 |
by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1179 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1180 |
lemma mem_box_componentwiseI: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1181 |
fixes S::"'a::euclidean_space set" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1182 |
assumes "is_interval S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1183 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<in> ((\<lambda>x. x \<bullet> i) ` S)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1184 |
shows "x \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1185 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1186 |
from assms have "\<forall>i \<in> Basis. \<exists>s \<in> S. x \<bullet> i = s \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1187 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1188 |
with finite_Basis obtain s and bs::"'a list" where |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1189 |
s: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i = s i \<bullet> i" "\<And>i. i \<in> Basis \<Longrightarrow> s i \<in> S" and |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1190 |
bs: "set bs = Basis" "distinct bs" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1191 |
by (metis finite_distinct_list) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1192 |
from nonempty_Basis s obtain j where j: "j \<in> Basis" "s j \<in> S" by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1193 |
def y \<equiv> "rec_list |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1194 |
(s j) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1195 |
(\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1196 |
have "x = (\<Sum>i\<in>Basis. (if i \<in> set bs then s i \<bullet> i else s j \<bullet> i) *\<^sub>R i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1197 |
using bs by (auto simp add: s(1)[symmetric] euclidean_representation) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1198 |
also have [symmetric]: "y bs = \<dots>" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1199 |
using bs(2) bs(1)[THEN equalityD1] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1200 |
by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1201 |
also have "y bs \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1202 |
using bs(1)[THEN equalityD1] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1203 |
apply (induct bs) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1204 |
apply (auto simp: y_def j) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1205 |
apply (rule interval_subst[OF assms(1)]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1206 |
apply (auto simp: s) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1207 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1208 |
finally show ?thesis . |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1209 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1210 |
|
| 33175 | 1211 |
|
1212 |
subsection{* Connectedness *}
|
|
1213 |
||
1214 |
lemma connected_local: |
|
| 53255 | 1215 |
"connected S \<longleftrightarrow> |
1216 |
\<not> (\<exists>e1 e2. |
|
1217 |
openin (subtopology euclidean S) e1 \<and> |
|
1218 |
openin (subtopology euclidean S) e2 \<and> |
|
1219 |
S \<subseteq> e1 \<union> e2 \<and> |
|
1220 |
e1 \<inter> e2 = {} \<and>
|
|
1221 |
e1 \<noteq> {} \<and>
|
|
1222 |
e2 \<noteq> {})"
|
|
| 53282 | 1223 |
unfolding connected_def openin_open |
| 55775 | 1224 |
by blast |
| 33175 | 1225 |
|
| 34105 | 1226 |
lemma exists_diff: |
1227 |
fixes P :: "'a set \<Rightarrow> bool" |
|
1228 |
shows "(\<exists>S. P(- S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs") |
|
| 53255 | 1229 |
proof - |
1230 |
{
|
|
1231 |
assume "?lhs" |
|
1232 |
then have ?rhs by blast |
|
1233 |
} |
|
| 33175 | 1234 |
moreover |
| 53255 | 1235 |
{
|
1236 |
fix S |
|
1237 |
assume H: "P S" |
|
| 34105 | 1238 |
have "S = - (- S)" by auto |
| 53255 | 1239 |
with H have "P (- (- S))" by metis |
1240 |
} |
|
| 33175 | 1241 |
ultimately show ?thesis by metis |
1242 |
qed |
|
1243 |
||
1244 |
lemma connected_clopen: "connected S \<longleftrightarrow> |
|
| 53255 | 1245 |
(\<forall>T. openin (subtopology euclidean S) T \<and> |
1246 |
closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
|
|
1247 |
proof - |
|
1248 |
have "\<not> connected S \<longleftrightarrow> |
|
1249 |
(\<exists>e1 e2. open e1 \<and> open (- e2) \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
|
|
| 33175 | 1250 |
unfolding connected_def openin_open closedin_closed |
| 55775 | 1251 |
by (metis double_complement) |
| 53282 | 1252 |
then have th0: "connected S \<longleftrightarrow> |
| 53255 | 1253 |
\<not> (\<exists>e2 e1. closed e2 \<and> open e1 \<and> S \<subseteq> e1 \<union> (- e2) \<and> e1 \<inter> (- e2) \<inter> S = {} \<and> e1 \<inter> S \<noteq> {} \<and> (- e2) \<inter> S \<noteq> {})"
|
| 52624 | 1254 |
(is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)") |
1255 |
apply (simp add: closed_def) |
|
1256 |
apply metis |
|
1257 |
done |
|
| 33175 | 1258 |
have th1: "?rhs \<longleftrightarrow> \<not> (\<exists>t' t. closed t'\<and>t = S\<inter>t' \<and> t\<noteq>{} \<and> t\<noteq>S \<and> (\<exists>t'. open t' \<and> t = S \<inter> t'))"
|
1259 |
(is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)") |
|
1260 |
unfolding connected_def openin_open closedin_closed by auto |
|
| 53255 | 1261 |
{
|
1262 |
fix e2 |
|
1263 |
{
|
|
1264 |
fix e1 |
|
| 53282 | 1265 |
have "?P e2 e1 \<longleftrightarrow> (\<exists>t. closed e2 \<and> t = S\<inter>e2 \<and> open e1 \<and> t = S\<inter>e1 \<and> t\<noteq>{} \<and> t \<noteq> S)"
|
| 53255 | 1266 |
by auto |
1267 |
} |
|
1268 |
then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" |
|
1269 |
by metis |
|
1270 |
} |
|
1271 |
then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)" |
|
1272 |
by blast |
|
1273 |
then show ?thesis |
|
1274 |
unfolding th0 th1 by simp |
|
| 33175 | 1275 |
qed |
1276 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1277 |
|
| 33175 | 1278 |
subsection{* Limit points *}
|
1279 |
||
| 53282 | 1280 |
definition (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60) |
| 53255 | 1281 |
where "x islimpt S \<longleftrightarrow> (\<forall>T. x\<in>T \<longrightarrow> open T \<longrightarrow> (\<exists>y\<in>S. y\<in>T \<and> y\<noteq>x))" |
| 33175 | 1282 |
|
1283 |
lemma islimptI: |
|
1284 |
assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x" |
|
1285 |
shows "x islimpt S" |
|
1286 |
using assms unfolding islimpt_def by auto |
|
1287 |
||
1288 |
lemma islimptE: |
|
1289 |
assumes "x islimpt S" and "x \<in> T" and "open T" |
|
1290 |
obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x" |
|
1291 |
using assms unfolding islimpt_def by auto |
|
1292 |
||
| 44584 | 1293 |
lemma islimpt_iff_eventually: "x islimpt S \<longleftrightarrow> \<not> eventually (\<lambda>y. y \<notin> S) (at x)" |
1294 |
unfolding islimpt_def eventually_at_topological by auto |
|
1295 |
||
| 53255 | 1296 |
lemma islimpt_subset: "x islimpt S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> x islimpt T" |
| 44584 | 1297 |
unfolding islimpt_def by fast |
| 33175 | 1298 |
|
1299 |
lemma islimpt_approachable: |
|
1300 |
fixes x :: "'a::metric_space" |
|
1301 |
shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e)" |
|
| 44584 | 1302 |
unfolding islimpt_iff_eventually eventually_at by fast |
| 33175 | 1303 |
|
1304 |
lemma islimpt_approachable_le: |
|
1305 |
fixes x :: "'a::metric_space" |
|
| 53640 | 1306 |
shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x \<le> e)" |
| 33175 | 1307 |
unfolding islimpt_approachable |
| 44584 | 1308 |
using approachable_lt_le [where f="\<lambda>y. dist y x" and P="\<lambda>y. y \<notin> S \<or> y = x", |
1309 |
THEN arg_cong [where f=Not]] |
|
1310 |
by (simp add: Bex_def conj_commute conj_left_commute) |
|
| 33175 | 1311 |
|
| 44571 | 1312 |
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
|
1313 |
unfolding islimpt_def by (safe, fast, case_tac "T = {x}", fast, fast)
|
|
1314 |
||
| 51351 | 1315 |
lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})"
|
1316 |
unfolding islimpt_def by blast |
|
1317 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1318 |
text {* A perfect space has no isolated points. *}
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1319 |
|
| 44571 | 1320 |
lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV" |
1321 |
unfolding islimpt_UNIV_iff by (rule not_open_singleton) |
|
| 33175 | 1322 |
|
1323 |
lemma perfect_choose_dist: |
|
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1324 |
fixes x :: "'a::{perfect_space, metric_space}"
|
| 33175 | 1325 |
shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r" |
| 53255 | 1326 |
using islimpt_UNIV [of x] |
1327 |
by (simp add: islimpt_approachable) |
|
| 33175 | 1328 |
|
1329 |
lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)" |
|
1330 |
unfolding closed_def |
|
1331 |
apply (subst open_subopen) |
|
| 34105 | 1332 |
apply (simp add: islimpt_def subset_eq) |
| 52624 | 1333 |
apply (metis ComplE ComplI) |
1334 |
done |
|
| 33175 | 1335 |
|
1336 |
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
|
|
1337 |
unfolding islimpt_def by auto |
|
1338 |
||
1339 |
lemma finite_set_avoid: |
|
1340 |
fixes a :: "'a::metric_space" |
|
| 53255 | 1341 |
assumes fS: "finite S" |
| 53640 | 1342 |
shows "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x" |
| 53255 | 1343 |
proof (induct rule: finite_induct[OF fS]) |
1344 |
case 1 |
|
1345 |
then show ?case by (auto intro: zero_less_one) |
|
| 33175 | 1346 |
next |
1347 |
case (2 x F) |
|
| 53255 | 1348 |
from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x" |
1349 |
by blast |
|
1350 |
show ?case |
|
1351 |
proof (cases "x = a") |
|
1352 |
case True |
|
1353 |
then show ?thesis using d by auto |
|
1354 |
next |
|
1355 |
case False |
|
| 33175 | 1356 |
let ?d = "min d (dist a x)" |
| 53255 | 1357 |
have dp: "?d > 0" |
1358 |
using False d(1) using dist_nz by auto |
|
1359 |
from d have d': "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> ?d \<le> dist a x" |
|
1360 |
by auto |
|
1361 |
with dp False show ?thesis |
|
1362 |
by (auto intro!: exI[where x="?d"]) |
|
1363 |
qed |
|
| 33175 | 1364 |
qed |
1365 |
||
1366 |
lemma islimpt_Un: "x islimpt (S \<union> T) \<longleftrightarrow> x islimpt S \<or> x islimpt T" |
|
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
1367 |
by (simp add: islimpt_iff_eventually eventually_conj_iff) |
| 33175 | 1368 |
|
1369 |
lemma discrete_imp_closed: |
|
1370 |
fixes S :: "'a::metric_space set" |
|
| 53255 | 1371 |
assumes e: "0 < e" |
1372 |
and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x" |
|
| 33175 | 1373 |
shows "closed S" |
| 53255 | 1374 |
proof - |
1375 |
{
|
|
1376 |
fix x |
|
1377 |
assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" |
|
| 33175 | 1378 |
from e have e2: "e/2 > 0" by arith |
| 53282 | 1379 |
from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y \<noteq> x" "dist y x < e/2" |
| 53255 | 1380 |
by blast |
| 33175 | 1381 |
let ?m = "min (e/2) (dist x y) " |
| 53255 | 1382 |
from e2 y(2) have mp: "?m > 0" |
| 53291 | 1383 |
by (simp add: dist_nz[symmetric]) |
| 53282 | 1384 |
from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z \<noteq> x" "dist z x < ?m" |
| 53255 | 1385 |
by blast |
| 33175 | 1386 |
have th: "dist z y < e" using z y |
1387 |
by (intro dist_triangle_lt [where z=x], simp) |
|
1388 |
from d[rule_format, OF y(1) z(1) th] y z |
|
1389 |
have False by (auto simp add: dist_commute)} |
|
| 53255 | 1390 |
then show ?thesis |
1391 |
by (metis islimpt_approachable closed_limpt [where 'a='a]) |
|
| 33175 | 1392 |
qed |
1393 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1394 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1395 |
subsection {* Interior of a Set *}
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1396 |
|
| 44519 | 1397 |
definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
|
1398 |
||
1399 |
lemma interiorI [intro?]: |
|
1400 |
assumes "open T" and "x \<in> T" and "T \<subseteq> S" |
|
1401 |
shows "x \<in> interior S" |
|
1402 |
using assms unfolding interior_def by fast |
|
1403 |
||
1404 |
lemma interiorE [elim?]: |
|
1405 |
assumes "x \<in> interior S" |
|
1406 |
obtains T where "open T" and "x \<in> T" and "T \<subseteq> S" |
|
1407 |
using assms unfolding interior_def by fast |
|
1408 |
||
1409 |
lemma open_interior [simp, intro]: "open (interior S)" |
|
1410 |
by (simp add: interior_def open_Union) |
|
1411 |
||
1412 |
lemma interior_subset: "interior S \<subseteq> S" |
|
1413 |
by (auto simp add: interior_def) |
|
1414 |
||
1415 |
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S" |
|
1416 |
by (auto simp add: interior_def) |
|
1417 |
||
1418 |
lemma interior_open: "open S \<Longrightarrow> interior S = S" |
|
1419 |
by (intro equalityI interior_subset interior_maximal subset_refl) |
|
| 33175 | 1420 |
|
1421 |
lemma interior_eq: "interior S = S \<longleftrightarrow> open S" |
|
| 44519 | 1422 |
by (metis open_interior interior_open) |
1423 |
||
1424 |
lemma open_subset_interior: "open S \<Longrightarrow> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T" |
|
| 33175 | 1425 |
by (metis interior_maximal interior_subset subset_trans) |
1426 |
||
| 44519 | 1427 |
lemma interior_empty [simp]: "interior {} = {}"
|
1428 |
using open_empty by (rule interior_open) |
|
1429 |
||
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1430 |
lemma interior_UNIV [simp]: "interior UNIV = UNIV" |
|
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1431 |
using open_UNIV by (rule interior_open) |
|
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1432 |
|
| 44519 | 1433 |
lemma interior_interior [simp]: "interior (interior S) = interior S" |
1434 |
using open_interior by (rule interior_open) |
|
1435 |
||
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1436 |
lemma interior_mono: "S \<subseteq> T \<Longrightarrow> interior S \<subseteq> interior T" |
|
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1437 |
by (auto simp add: interior_def) |
| 44519 | 1438 |
|
1439 |
lemma interior_unique: |
|
1440 |
assumes "T \<subseteq> S" and "open T" |
|
1441 |
assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T" |
|
1442 |
shows "interior S = T" |
|
1443 |
by (intro equalityI assms interior_subset open_interior interior_maximal) |
|
1444 |
||
1445 |
lemma interior_inter [simp]: "interior (S \<inter> T) = interior S \<inter> interior T" |
|
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1446 |
by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1 |
| 44519 | 1447 |
Int_lower2 interior_maximal interior_subset open_Int open_interior) |
1448 |
||
1449 |
lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)" |
|
1450 |
using open_contains_ball_eq [where S="interior S"] |
|
1451 |
by (simp add: open_subset_interior) |
|
| 33175 | 1452 |
|
1453 |
lemma interior_limit_point [intro]: |
|
1454 |
fixes x :: "'a::perfect_space" |
|
| 53255 | 1455 |
assumes x: "x \<in> interior S" |
1456 |
shows "x islimpt S" |
|
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1457 |
using x islimpt_UNIV [of x] |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1458 |
unfolding interior_def islimpt_def |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1459 |
apply (clarsimp, rename_tac T T') |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1460 |
apply (drule_tac x="T \<inter> T'" in spec) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1461 |
apply (auto simp add: open_Int) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1462 |
done |
| 33175 | 1463 |
|
1464 |
lemma interior_closed_Un_empty_interior: |
|
| 53255 | 1465 |
assumes cS: "closed S" |
1466 |
and iT: "interior T = {}"
|
|
| 44519 | 1467 |
shows "interior (S \<union> T) = interior S" |
| 33175 | 1468 |
proof |
| 44519 | 1469 |
show "interior S \<subseteq> interior (S \<union> T)" |
| 53255 | 1470 |
by (rule interior_mono) (rule Un_upper1) |
| 33175 | 1471 |
show "interior (S \<union> T) \<subseteq> interior S" |
1472 |
proof |
|
| 53255 | 1473 |
fix x |
1474 |
assume "x \<in> interior (S \<union> T)" |
|
| 44519 | 1475 |
then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T" .. |
| 33175 | 1476 |
show "x \<in> interior S" |
1477 |
proof (rule ccontr) |
|
1478 |
assume "x \<notin> interior S" |
|
1479 |
with `x \<in> R` `open R` obtain y where "y \<in> R - S" |
|
| 44519 | 1480 |
unfolding interior_def by fast |
| 53282 | 1481 |
from `open R` `closed S` have "open (R - S)" |
1482 |
by (rule open_Diff) |
|
1483 |
from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T" |
|
1484 |
by fast |
|
1485 |
from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}` show False
|
|
1486 |
unfolding interior_def by fast |
|
| 33175 | 1487 |
qed |
1488 |
qed |
|
1489 |
qed |
|
1490 |
||
| 44365 | 1491 |
lemma interior_Times: "interior (A \<times> B) = interior A \<times> interior B" |
1492 |
proof (rule interior_unique) |
|
1493 |
show "interior A \<times> interior B \<subseteq> A \<times> B" |
|
1494 |
by (intro Sigma_mono interior_subset) |
|
1495 |
show "open (interior A \<times> interior B)" |
|
1496 |
by (intro open_Times open_interior) |
|
| 53255 | 1497 |
fix T |
1498 |
assume "T \<subseteq> A \<times> B" and "open T" |
|
1499 |
then show "T \<subseteq> interior A \<times> interior B" |
|
| 53282 | 1500 |
proof safe |
| 53255 | 1501 |
fix x y |
1502 |
assume "(x, y) \<in> T" |
|
| 44519 | 1503 |
then obtain C D where "open C" "open D" "C \<times> D \<subseteq> T" "x \<in> C" "y \<in> D" |
1504 |
using `open T` unfolding open_prod_def by fast |
|
| 53255 | 1505 |
then have "open C" "open D" "C \<subseteq> A" "D \<subseteq> B" "x \<in> C" "y \<in> D" |
| 44519 | 1506 |
using `T \<subseteq> A \<times> B` by auto |
| 53255 | 1507 |
then show "x \<in> interior A" and "y \<in> interior B" |
| 44519 | 1508 |
by (auto intro: interiorI) |
1509 |
qed |
|
| 44365 | 1510 |
qed |
1511 |
||
| 33175 | 1512 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1513 |
subsection {* Closure of a Set *}
|
| 33175 | 1514 |
|
1515 |
definition "closure S = S \<union> {x | x. x islimpt S}"
|
|
1516 |
||
| 44518 | 1517 |
lemma interior_closure: "interior S = - (closure (- S))" |
1518 |
unfolding interior_def closure_def islimpt_def by auto |
|
1519 |
||
| 34105 | 1520 |
lemma closure_interior: "closure S = - interior (- S)" |
| 44518 | 1521 |
unfolding interior_closure by simp |
| 33175 | 1522 |
|
1523 |
lemma closed_closure[simp, intro]: "closed (closure S)" |
|
| 44518 | 1524 |
unfolding closure_interior by (simp add: closed_Compl) |
1525 |
||
1526 |
lemma closure_subset: "S \<subseteq> closure S" |
|
1527 |
unfolding closure_def by simp |
|
| 33175 | 1528 |
|
1529 |
lemma closure_hull: "closure S = closed hull S" |
|
| 44519 | 1530 |
unfolding hull_def closure_interior interior_def by auto |
| 33175 | 1531 |
|
1532 |
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S" |
|
| 44519 | 1533 |
unfolding closure_hull using closed_Inter by (rule hull_eq) |
1534 |
||
1535 |
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S" |
|
1536 |
unfolding closure_eq . |
|
1537 |
||
1538 |
lemma closure_closure [simp]: "closure (closure S) = closure S" |
|
| 44518 | 1539 |
unfolding closure_hull by (rule hull_hull) |
| 33175 | 1540 |
|
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1541 |
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T" |
| 44518 | 1542 |
unfolding closure_hull by (rule hull_mono) |
| 33175 | 1543 |
|
| 44519 | 1544 |
lemma closure_minimal: "S \<subseteq> T \<Longrightarrow> closed T \<Longrightarrow> closure S \<subseteq> T" |
| 44518 | 1545 |
unfolding closure_hull by (rule hull_minimal) |
| 33175 | 1546 |
|
| 44519 | 1547 |
lemma closure_unique: |
| 53255 | 1548 |
assumes "S \<subseteq> T" |
1549 |
and "closed T" |
|
1550 |
and "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'" |
|
| 44519 | 1551 |
shows "closure S = T" |
1552 |
using assms unfolding closure_hull by (rule hull_unique) |
|
1553 |
||
1554 |
lemma closure_empty [simp]: "closure {} = {}"
|
|
| 44518 | 1555 |
using closed_empty by (rule closure_closed) |
| 33175 | 1556 |
|
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
1557 |
lemma closure_UNIV [simp]: "closure UNIV = UNIV" |
| 44518 | 1558 |
using closed_UNIV by (rule closure_closed) |
1559 |
||
1560 |
lemma closure_union [simp]: "closure (S \<union> T) = closure S \<union> closure T" |
|
1561 |
unfolding closure_interior by simp |
|
| 33175 | 1562 |
|
1563 |
lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
|
|
1564 |
using closure_empty closure_subset[of S] |
|
1565 |
by blast |
|
1566 |
||
1567 |
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S" |
|
1568 |
using closure_eq[of S] closure_subset[of S] |
|
1569 |
by simp |
|
1570 |
||
1571 |
lemma open_inter_closure_eq_empty: |
|
1572 |
"open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
|
|
| 34105 | 1573 |
using open_subset_interior[of S "- T"] |
1574 |
using interior_subset[of "- T"] |
|
| 33175 | 1575 |
unfolding closure_interior |
1576 |
by auto |
|
1577 |
||
1578 |
lemma open_inter_closure_subset: |
|
1579 |
"open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)" |
|
1580 |
proof |
|
1581 |
fix x |
|
1582 |
assume as: "open S" "x \<in> S \<inter> closure T" |
|
| 53255 | 1583 |
{
|
| 53282 | 1584 |
assume *: "x islimpt T" |
| 33175 | 1585 |
have "x islimpt (S \<inter> T)" |
1586 |
proof (rule islimptI) |
|
1587 |
fix A |
|
1588 |
assume "x \<in> A" "open A" |
|
1589 |
with as have "x \<in> A \<inter> S" "open (A \<inter> S)" |
|
1590 |
by (simp_all add: open_Int) |
|
1591 |
with * obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x" |
|
1592 |
by (rule islimptE) |
|
| 53255 | 1593 |
then have "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x" |
| 33175 | 1594 |
by simp_all |
| 53255 | 1595 |
then show "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" .. |
| 33175 | 1596 |
qed |
1597 |
} |
|
1598 |
then show "x \<in> closure (S \<inter> T)" using as |
|
1599 |
unfolding closure_def |
|
1600 |
by blast |
|
1601 |
qed |
|
1602 |
||
| 44519 | 1603 |
lemma closure_complement: "closure (- S) = - interior S" |
| 44518 | 1604 |
unfolding closure_interior by simp |
| 33175 | 1605 |
|
| 44519 | 1606 |
lemma interior_complement: "interior (- S) = - closure S" |
| 44518 | 1607 |
unfolding closure_interior by simp |
| 33175 | 1608 |
|
| 44365 | 1609 |
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B" |
| 44519 | 1610 |
proof (rule closure_unique) |
| 44365 | 1611 |
show "A \<times> B \<subseteq> closure A \<times> closure B" |
1612 |
by (intro Sigma_mono closure_subset) |
|
1613 |
show "closed (closure A \<times> closure B)" |
|
1614 |
by (intro closed_Times closed_closure) |
|
| 53282 | 1615 |
fix T |
1616 |
assume "A \<times> B \<subseteq> T" and "closed T" |
|
1617 |
then show "closure A \<times> closure B \<subseteq> T" |
|
| 44365 | 1618 |
apply (simp add: closed_def open_prod_def, clarify) |
1619 |
apply (rule ccontr) |
|
1620 |
apply (drule_tac x="(a, b)" in bspec, simp, clarify, rename_tac C D) |
|
1621 |
apply (simp add: closure_interior interior_def) |
|
1622 |
apply (drule_tac x=C in spec) |
|
1623 |
apply (drule_tac x=D in spec) |
|
1624 |
apply auto |
|
1625 |
done |
|
1626 |
qed |
|
1627 |
||
| 51351 | 1628 |
lemma islimpt_in_closure: "(x islimpt S) = (x:closure(S-{x}))"
|
1629 |
unfolding closure_def using islimpt_punctured by blast |
|
1630 |
||
1631 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1632 |
subsection {* Frontier (aka boundary) *}
|
| 33175 | 1633 |
|
1634 |
definition "frontier S = closure S - interior S" |
|
1635 |
||
| 53255 | 1636 |
lemma frontier_closed: "closed (frontier S)" |
| 33175 | 1637 |
by (simp add: frontier_def closed_Diff) |
1638 |
||
| 34105 | 1639 |
lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(- S))" |
| 33175 | 1640 |
by (auto simp add: frontier_def interior_closure) |
1641 |
||
1642 |
lemma frontier_straddle: |
|
1643 |
fixes a :: "'a::metric_space" |
|
| 44909 | 1644 |
shows "a \<in> frontier S \<longleftrightarrow> (\<forall>e>0. (\<exists>x\<in>S. dist a x < e) \<and> (\<exists>x. x \<notin> S \<and> dist a x < e))" |
1645 |
unfolding frontier_def closure_interior |
|
1646 |
by (auto simp add: mem_interior subset_eq ball_def) |
|
| 33175 | 1647 |
|
1648 |
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S" |
|
1649 |
by (metis frontier_def closure_closed Diff_subset) |
|
1650 |
||
| 34964 | 1651 |
lemma frontier_empty[simp]: "frontier {} = {}"
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
1652 |
by (simp add: frontier_def) |
| 33175 | 1653 |
|
1654 |
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S" |
|
1655 |
proof- |
|
| 53255 | 1656 |
{
|
1657 |
assume "frontier S \<subseteq> S" |
|
1658 |
then have "closure S \<subseteq> S" |
|
1659 |
using interior_subset unfolding frontier_def by auto |
|
1660 |
then have "closed S" |
|
1661 |
using closure_subset_eq by auto |
|
| 33175 | 1662 |
} |
| 53255 | 1663 |
then show ?thesis using frontier_subset_closed[of S] .. |
| 33175 | 1664 |
qed |
1665 |
||
| 34105 | 1666 |
lemma frontier_complement: "frontier(- S) = frontier S" |
| 33175 | 1667 |
by (auto simp add: frontier_def closure_complement interior_complement) |
1668 |
||
1669 |
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
|
|
| 34105 | 1670 |
using frontier_complement frontier_subset_eq[of "- S"] |
1671 |
unfolding open_closed by auto |
|
| 33175 | 1672 |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44076
diff
changeset
|
1673 |
subsection {* Filters and the ``eventually true'' quantifier *}
|
|
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44076
diff
changeset
|
1674 |
|
| 52624 | 1675 |
definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter" |
1676 |
(infixr "indirection" 70) |
|
1677 |
where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
|
|
| 33175 | 1678 |
|
| 36437 | 1679 |
text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
|
| 33175 | 1680 |
|
| 52624 | 1681 |
lemma trivial_limit_within: "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S" |
| 33175 | 1682 |
proof |
1683 |
assume "trivial_limit (at a within S)" |
|
| 53255 | 1684 |
then show "\<not> a islimpt S" |
| 33175 | 1685 |
unfolding trivial_limit_def |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1686 |
unfolding eventually_at_topological |
| 33175 | 1687 |
unfolding islimpt_def |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
1688 |
apply (clarsimp simp add: set_eq_iff) |
| 33175 | 1689 |
apply (rename_tac T, rule_tac x=T in exI) |
|
36358
246493d61204
define nets directly as filters, instead of as filter bases
huffman
parents:
36336
diff
changeset
|
1690 |
apply (clarsimp, drule_tac x=y in bspec, simp_all) |
| 33175 | 1691 |
done |
1692 |
next |
|
1693 |
assume "\<not> a islimpt S" |
|
| 53255 | 1694 |
then show "trivial_limit (at a within S)" |
| 55775 | 1695 |
unfolding trivial_limit_def eventually_at_topological islimpt_def |
1696 |
by metis |
|
| 33175 | 1697 |
qed |
1698 |
||
1699 |
lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV" |
|
| 45031 | 1700 |
using trivial_limit_within [of a UNIV] by simp |
| 33175 | 1701 |
|
1702 |
lemma trivial_limit_at: |
|
1703 |
fixes a :: "'a::perfect_space" |
|
1704 |
shows "\<not> trivial_limit (at a)" |
|
| 44571 | 1705 |
by (rule at_neq_bot) |
| 33175 | 1706 |
|
1707 |
lemma trivial_limit_at_infinity: |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44076
diff
changeset
|
1708 |
"\<not> trivial_limit (at_infinity :: ('a::{real_normed_vector,perfect_space}) filter)"
|
|
36358
246493d61204
define nets directly as filters, instead of as filter bases
huffman
parents:
36336
diff
changeset
|
1709 |
unfolding trivial_limit_def eventually_at_infinity |
|
246493d61204
define nets directly as filters, instead of as filter bases
huffman
parents:
36336
diff
changeset
|
1710 |
apply clarsimp |
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1711 |
apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1712 |
apply (rule_tac x="scaleR (b / norm x) x" in exI, simp) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1713 |
apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def]) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
1714 |
apply (drule_tac x=UNIV in spec, simp) |
| 33175 | 1715 |
done |
1716 |
||
| 53640 | 1717 |
lemma not_trivial_limit_within: "\<not> trivial_limit (at x within S) = (x \<in> closure (S - {x}))"
|
1718 |
using islimpt_in_closure |
|
1719 |
by (metis trivial_limit_within) |
|
| 51351 | 1720 |
|
| 36437 | 1721 |
text {* Some property holds "sufficiently close" to the limit point. *}
|
| 33175 | 1722 |
|
|
51530
609914f0934a
rename eventually_at / _within, to distinguish them from the lemmas in the HOL image
hoelzl
parents:
51518
diff
changeset
|
1723 |
lemma eventually_at2: |
| 33175 | 1724 |
"eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)" |
| 53255 | 1725 |
unfolding eventually_at dist_nz by auto |
1726 |
||
1727 |
lemma eventually_happens: "eventually P net \<Longrightarrow> trivial_limit net \<or> (\<exists>x. P x)" |
|
|
36358
246493d61204
define nets directly as filters, instead of as filter bases
huffman
parents:
36336
diff
changeset
|
1728 |
unfolding trivial_limit_def |
|
246493d61204
define nets directly as filters, instead of as filter bases
huffman
parents:
36336
diff
changeset
|
1729 |
by (auto elim: eventually_rev_mp) |
| 33175 | 1730 |
|
1731 |
lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net" |
|
| 45031 | 1732 |
by simp |
| 33175 | 1733 |
|
1734 |
lemma trivial_limit_eq: "trivial_limit net \<longleftrightarrow> (\<forall>P. eventually P net)" |
|
|
44342
8321948340ea
redefine constant 'trivial_limit' as an abbreviation
huffman
parents:
44286
diff
changeset
|
1735 |
by (simp add: filter_eq_iff) |
| 33175 | 1736 |
|
1737 |
text{* Combining theorems for "eventually" *}
|
|
1738 |
||
1739 |
lemma eventually_rev_mono: |
|
1740 |
"eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net" |
|
| 53255 | 1741 |
using eventually_mono [of P Q] by fast |
| 33175 | 1742 |
|
| 53282 | 1743 |
lemma not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> \<not> trivial_limit net \<Longrightarrow> \<not> eventually (\<lambda>x. P x) net" |
| 33175 | 1744 |
by (simp add: eventually_False) |
1745 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1746 |
|
| 36437 | 1747 |
subsection {* Limits *}
|
| 33175 | 1748 |
|
1749 |
lemma Lim: |
|
| 53255 | 1750 |
"(f ---> l) net \<longleftrightarrow> |
| 33175 | 1751 |
trivial_limit net \<or> |
1752 |
(\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)" |
|
1753 |
unfolding tendsto_iff trivial_limit_eq by auto |
|
1754 |
||
1755 |
text{* Show that they yield usual definitions in the various cases. *}
|
|
1756 |
||
1757 |
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow> |
|
| 53640 | 1758 |
(\<forall>e>0. \<exists>d>0. \<forall>x\<in>S. 0 < dist x a \<and> dist x a \<le> d \<longrightarrow> dist (f x) l < e)" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1759 |
by (auto simp add: tendsto_iff eventually_at_le dist_nz) |
| 33175 | 1760 |
|
1761 |
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow> |
|
| 53640 | 1762 |
(\<forall>e >0. \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1763 |
by (auto simp add: tendsto_iff eventually_at dist_nz) |
| 33175 | 1764 |
|
1765 |
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow> |
|
| 53640 | 1766 |
(\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)" |
|
51530
609914f0934a
rename eventually_at / _within, to distinguish them from the lemmas in the HOL image
hoelzl
parents:
51518
diff
changeset
|
1767 |
by (auto simp add: tendsto_iff eventually_at2) |
| 33175 | 1768 |
|
1769 |
lemma Lim_at_infinity: |
|
| 53640 | 1770 |
"(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x \<ge> b \<longrightarrow> dist (f x) l < e)" |
| 33175 | 1771 |
by (auto simp add: tendsto_iff eventually_at_infinity) |
1772 |
||
1773 |
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net" |
|
1774 |
by (rule topological_tendstoI, auto elim: eventually_rev_mono) |
|
1775 |
||
1776 |
text{* The expected monotonicity property. *}
|
|
1777 |
||
| 53255 | 1778 |
lemma Lim_Un: |
1779 |
assumes "(f ---> l) (at x within S)" "(f ---> l) (at x within T)" |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1780 |
shows "(f ---> l) (at x within (S \<union> T))" |
| 53860 | 1781 |
using assms unfolding at_within_union by (rule filterlim_sup) |
| 33175 | 1782 |
|
1783 |
lemma Lim_Un_univ: |
|
| 53282 | 1784 |
"(f ---> l) (at x within S) \<Longrightarrow> (f ---> l) (at x within T) \<Longrightarrow> |
| 53255 | 1785 |
S \<union> T = UNIV \<Longrightarrow> (f ---> l) (at x)" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1786 |
by (metis Lim_Un) |
| 33175 | 1787 |
|
1788 |
text{* Interrelations between restricted and unrestricted limits. *}
|
|
1789 |
||
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1790 |
lemma Lim_at_within: (* FIXME: rename *) |
|
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1791 |
"(f ---> l) (at x) \<Longrightarrow> (f ---> l) (at x within S)" |
|
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1792 |
by (metis order_refl filterlim_mono subset_UNIV at_le) |
| 33175 | 1793 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1794 |
lemma eventually_within_interior: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1795 |
assumes "x \<in> interior S" |
| 53255 | 1796 |
shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)" |
1797 |
(is "?lhs = ?rhs") |
|
1798 |
proof |
|
| 44519 | 1799 |
from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S" .. |
| 53255 | 1800 |
{
|
1801 |
assume "?lhs" |
|
| 53640 | 1802 |
then obtain A where "open A" and "x \<in> A" and "\<forall>y\<in>A. y \<noteq> x \<longrightarrow> y \<in> S \<longrightarrow> P y" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1803 |
unfolding eventually_at_topological |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1804 |
by auto |
| 53640 | 1805 |
with T have "open (A \<inter> T)" and "x \<in> A \<inter> T" and "\<forall>y \<in> A \<inter> T. y \<noteq> x \<longrightarrow> P y" |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1806 |
by auto |
| 53255 | 1807 |
then show "?rhs" |
| 51471 | 1808 |
unfolding eventually_at_topological by auto |
| 53255 | 1809 |
next |
1810 |
assume "?rhs" |
|
1811 |
then show "?lhs" |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1812 |
by (auto elim: eventually_elim1 simp: eventually_at_filter) |
| 52624 | 1813 |
} |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1814 |
qed |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1815 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1816 |
lemma at_within_interior: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1817 |
"x \<in> interior S \<Longrightarrow> at x within S = at x" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1818 |
unfolding filter_eq_iff by (intro allI eventually_within_interior) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1819 |
|
| 43338 | 1820 |
lemma Lim_within_LIMSEQ: |
| 53862 | 1821 |
fixes a :: "'a::first_countable_topology" |
| 43338 | 1822 |
assumes "\<forall>S. (\<forall>n. S n \<noteq> a \<and> S n \<in> T) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L" |
1823 |
shows "(X ---> L) (at a within T)" |
|
| 44584 | 1824 |
using assms unfolding tendsto_def [where l=L] |
1825 |
by (simp add: sequentially_imp_eventually_within) |
|
| 43338 | 1826 |
|
1827 |
lemma Lim_right_bound: |
|
| 51773 | 1828 |
fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder, no_top} \<Rightarrow>
|
1829 |
'b::{linorder_topology, conditionally_complete_linorder}"
|
|
| 43338 | 1830 |
assumes mono: "\<And>a b. a \<in> I \<Longrightarrow> b \<in> I \<Longrightarrow> x < a \<Longrightarrow> a \<le> b \<Longrightarrow> f a \<le> f b" |
| 53255 | 1831 |
and bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a" |
| 43338 | 1832 |
shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
|
| 53640 | 1833 |
proof (cases "{x<..} \<inter> I = {}")
|
1834 |
case True |
|
| 53859 | 1835 |
then show ?thesis by simp |
| 43338 | 1836 |
next |
| 53640 | 1837 |
case False |
| 43338 | 1838 |
show ?thesis |
|
51518
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
hoelzl
parents:
51481
diff
changeset
|
1839 |
proof (rule order_tendstoI) |
| 53282 | 1840 |
fix a |
1841 |
assume a: "a < Inf (f ` ({x<..} \<inter> I))"
|
|
| 53255 | 1842 |
{
|
1843 |
fix y |
|
1844 |
assume "y \<in> {x<..} \<inter> I"
|
|
| 53640 | 1845 |
with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
|
| 56166 | 1846 |
by (auto intro!: cInf_lower bdd_belowI2 simp del: Inf_image_eq) |
| 53255 | 1847 |
with a have "a < f y" |
1848 |
by (blast intro: less_le_trans) |
|
1849 |
} |
|
|
51518
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
hoelzl
parents:
51481
diff
changeset
|
1850 |
then show "eventually (\<lambda>x. a < f x) (at x within ({x<..} \<inter> I))"
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1851 |
by (auto simp: eventually_at_filter intro: exI[of _ 1] zero_less_one) |
|
51518
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
hoelzl
parents:
51481
diff
changeset
|
1852 |
next |
| 53255 | 1853 |
fix a |
1854 |
assume "Inf (f ` ({x<..} \<inter> I)) < a"
|
|
| 53640 | 1855 |
from cInf_lessD[OF _ this] False obtain y where y: "x < y" "y \<in> I" "f y < a" |
| 53255 | 1856 |
by auto |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1857 |
then have "eventually (\<lambda>x. x \<in> I \<longrightarrow> f x < a) (at_right x)" |
|
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1858 |
unfolding eventually_at_right by (metis less_imp_le le_less_trans mono) |
|
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1859 |
then show "eventually (\<lambda>x. f x < a) (at x within ({x<..} \<inter> I))"
|
|
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1860 |
unfolding eventually_at_filter by eventually_elim simp |
| 43338 | 1861 |
qed |
1862 |
qed |
|
1863 |
||
| 33175 | 1864 |
text{* Another limit point characterization. *}
|
1865 |
||
1866 |
lemma islimpt_sequential: |
|
| 50883 | 1867 |
fixes x :: "'a::first_countable_topology" |
1868 |
shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S - {x}) \<and> (f ---> x) sequentially)"
|
|
| 33175 | 1869 |
(is "?lhs = ?rhs") |
1870 |
proof |
|
1871 |
assume ?lhs |
|
| 55522 | 1872 |
from countable_basis_at_decseq[of x] obtain A where A: |
1873 |
"\<And>i. open (A i)" |
|
1874 |
"\<And>i. x \<in> A i" |
|
1875 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
|
1876 |
by blast |
|
| 50883 | 1877 |
def f \<equiv> "\<lambda>n. SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y" |
| 53255 | 1878 |
{
|
1879 |
fix n |
|
| 50883 | 1880 |
from `?lhs` have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y" |
1881 |
unfolding islimpt_def using A(1,2)[of n] by auto |
|
1882 |
then have "f n \<in> S \<and> f n \<in> A n \<and> x \<noteq> f n" |
|
1883 |
unfolding f_def by (rule someI_ex) |
|
| 53255 | 1884 |
then have "f n \<in> S" "f n \<in> A n" "x \<noteq> f n" by auto |
1885 |
} |
|
| 50883 | 1886 |
then have "\<forall>n. f n \<in> S - {x}" by auto
|
1887 |
moreover have "(\<lambda>n. f n) ----> x" |
|
1888 |
proof (rule topological_tendstoI) |
|
| 53255 | 1889 |
fix S |
1890 |
assume "open S" "x \<in> S" |
|
| 50883 | 1891 |
from A(3)[OF this] `\<And>n. f n \<in> A n` |
| 53255 | 1892 |
show "eventually (\<lambda>x. f x \<in> S) sequentially" |
1893 |
by (auto elim!: eventually_elim1) |
|
| 44584 | 1894 |
qed |
1895 |
ultimately show ?rhs by fast |
|
| 33175 | 1896 |
next |
1897 |
assume ?rhs |
|
| 53255 | 1898 |
then obtain f :: "nat \<Rightarrow> 'a" where f: "\<And>n. f n \<in> S - {x}" and lim: "f ----> x"
|
1899 |
by auto |
|
| 50883 | 1900 |
show ?lhs |
1901 |
unfolding islimpt_def |
|
1902 |
proof safe |
|
| 53255 | 1903 |
fix T |
1904 |
assume "open T" "x \<in> T" |
|
| 50883 | 1905 |
from lim[THEN topological_tendstoD, OF this] f |
1906 |
show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> x" |
|
1907 |
unfolding eventually_sequentially by auto |
|
1908 |
qed |
|
| 33175 | 1909 |
qed |
1910 |
||
1911 |
lemma Lim_null: |
|
1912 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 44125 | 1913 |
shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net" |
| 33175 | 1914 |
by (simp add: Lim dist_norm) |
1915 |
||
1916 |
lemma Lim_null_comparison: |
|
1917 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
1918 |
assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net" |
|
1919 |
shows "(f ---> 0) net" |
|
| 53282 | 1920 |
using assms(2) |
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
1921 |
proof (rule metric_tendsto_imp_tendsto) |
|
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
1922 |
show "eventually (\<lambda>x. dist (f x) 0 \<le> dist (g x) 0) net" |
| 53255 | 1923 |
using assms(1) by (rule eventually_elim1) (simp add: dist_norm) |
| 33175 | 1924 |
qed |
1925 |
||
1926 |
lemma Lim_transform_bound: |
|
1927 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 53255 | 1928 |
and g :: "'a \<Rightarrow> 'c::real_normed_vector" |
| 53640 | 1929 |
assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) net" |
| 53255 | 1930 |
and "(g ---> 0) net" |
| 33175 | 1931 |
shows "(f ---> 0) net" |
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
1932 |
using assms(1) tendsto_norm_zero [OF assms(2)] |
|
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
1933 |
by (rule Lim_null_comparison) |
| 33175 | 1934 |
|
1935 |
text{* Deducing things about the limit from the elements. *}
|
|
1936 |
||
1937 |
lemma Lim_in_closed_set: |
|
| 53255 | 1938 |
assumes "closed S" |
1939 |
and "eventually (\<lambda>x. f(x) \<in> S) net" |
|
| 53640 | 1940 |
and "\<not> trivial_limit net" "(f ---> l) net" |
| 33175 | 1941 |
shows "l \<in> S" |
1942 |
proof (rule ccontr) |
|
1943 |
assume "l \<notin> S" |
|
1944 |
with `closed S` have "open (- S)" "l \<in> - S" |
|
1945 |
by (simp_all add: open_Compl) |
|
1946 |
with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net" |
|
1947 |
by (rule topological_tendstoD) |
|
1948 |
with assms(2) have "eventually (\<lambda>x. False) net" |
|
1949 |
by (rule eventually_elim2) simp |
|
1950 |
with assms(3) show "False" |
|
1951 |
by (simp add: eventually_False) |
|
1952 |
qed |
|
1953 |
||
1954 |
text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
|
|
1955 |
||
1956 |
lemma Lim_dist_ubound: |
|
| 53255 | 1957 |
assumes "\<not>(trivial_limit net)" |
1958 |
and "(f ---> l) net" |
|
| 53640 | 1959 |
and "eventually (\<lambda>x. dist a (f x) \<le> e) net" |
1960 |
shows "dist a l \<le> e" |
|
| 56290 | 1961 |
using assms by (fast intro: tendsto_le tendsto_intros) |
| 33175 | 1962 |
|
1963 |
lemma Lim_norm_ubound: |
|
1964 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 53255 | 1965 |
assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) \<le> e) net" |
1966 |
shows "norm(l) \<le> e" |
|
| 56290 | 1967 |
using assms by (fast intro: tendsto_le tendsto_intros) |
| 33175 | 1968 |
|
1969 |
lemma Lim_norm_lbound: |
|
1970 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 53640 | 1971 |
assumes "\<not> trivial_limit net" |
1972 |
and "(f ---> l) net" |
|
1973 |
and "eventually (\<lambda>x. e \<le> norm (f x)) net" |
|
| 33175 | 1974 |
shows "e \<le> norm l" |
| 56290 | 1975 |
using assms by (fast intro: tendsto_le tendsto_intros) |
| 33175 | 1976 |
|
1977 |
text{* Limit under bilinear function *}
|
|
1978 |
||
1979 |
lemma Lim_bilinear: |
|
| 53282 | 1980 |
assumes "(f ---> l) net" |
1981 |
and "(g ---> m) net" |
|
1982 |
and "bounded_bilinear h" |
|
| 33175 | 1983 |
shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net" |
| 52624 | 1984 |
using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net` |
1985 |
by (rule bounded_bilinear.tendsto) |
|
| 33175 | 1986 |
|
1987 |
text{* These are special for limits out of the same vector space. *}
|
|
1988 |
||
1989 |
lemma Lim_within_id: "(id ---> a) (at a within s)" |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
1990 |
unfolding id_def by (rule tendsto_ident_at) |
| 33175 | 1991 |
|
1992 |
lemma Lim_at_id: "(id ---> a) (at a)" |
|
| 45031 | 1993 |
unfolding id_def by (rule tendsto_ident_at) |
| 33175 | 1994 |
|
1995 |
lemma Lim_at_zero: |
|
1996 |
fixes a :: "'a::real_normed_vector" |
|
| 53291 | 1997 |
and l :: "'b::topological_space" |
| 53282 | 1998 |
shows "(f ---> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) ---> l) (at 0)" |
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
1999 |
using LIM_offset_zero LIM_offset_zero_cancel .. |
| 33175 | 2000 |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44076
diff
changeset
|
2001 |
text{* It's also sometimes useful to extract the limit point from the filter. *}
|
| 33175 | 2002 |
|
| 52624 | 2003 |
abbreviation netlimit :: "'a::t2_space filter \<Rightarrow> 'a" |
2004 |
where "netlimit F \<equiv> Lim F (\<lambda>x. x)" |
|
| 33175 | 2005 |
|
| 53282 | 2006 |
lemma netlimit_within: "\<not> trivial_limit (at a within S) \<Longrightarrow> netlimit (at a within S) = a" |
| 51365 | 2007 |
by (rule tendsto_Lim) (auto intro: tendsto_intros) |
| 33175 | 2008 |
|
2009 |
lemma netlimit_at: |
|
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2010 |
fixes a :: "'a::{perfect_space,t2_space}"
|
| 33175 | 2011 |
shows "netlimit (at a) = a" |
| 45031 | 2012 |
using netlimit_within [of a UNIV] by simp |
| 33175 | 2013 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2014 |
lemma lim_within_interior: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2015 |
"x \<in> interior S \<Longrightarrow> (f ---> l) (at x within S) \<longleftrightarrow> (f ---> l) (at x)" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
2016 |
by (metis at_within_interior) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2017 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2018 |
lemma netlimit_within_interior: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2019 |
fixes x :: "'a::{t2_space,perfect_space}"
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2020 |
assumes "x \<in> interior S" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2021 |
shows "netlimit (at x within S) = x" |
| 52624 | 2022 |
using assms by (metis at_within_interior netlimit_at) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2023 |
|
| 33175 | 2024 |
text{* Transformation of limit. *}
|
2025 |
||
2026 |
lemma Lim_transform: |
|
2027 |
fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector" |
|
2028 |
assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net" |
|
2029 |
shows "(g ---> l) net" |
|
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
2030 |
using tendsto_diff [OF assms(2) assms(1)] by simp |
| 33175 | 2031 |
|
2032 |
lemma Lim_transform_eventually: |
|
| 36667 | 2033 |
"eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net \<Longrightarrow> (g ---> l) net" |
| 33175 | 2034 |
apply (rule topological_tendstoI) |
2035 |
apply (drule (2) topological_tendstoD) |
|
2036 |
apply (erule (1) eventually_elim2, simp) |
|
2037 |
done |
|
2038 |
||
2039 |
lemma Lim_transform_within: |
|
| 53282 | 2040 |
assumes "0 < d" |
2041 |
and "\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'" |
|
2042 |
and "(f ---> l) (at x within S)" |
|
| 36667 | 2043 |
shows "(g ---> l) (at x within S)" |
2044 |
proof (rule Lim_transform_eventually) |
|
2045 |
show "eventually (\<lambda>x. f x = g x) (at x within S)" |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
2046 |
using assms(1,2) by (auto simp: dist_nz eventually_at) |
| 36667 | 2047 |
show "(f ---> l) (at x within S)" by fact |
2048 |
qed |
|
| 33175 | 2049 |
|
2050 |
lemma Lim_transform_at: |
|
| 53282 | 2051 |
assumes "0 < d" |
2052 |
and "\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'" |
|
2053 |
and "(f ---> l) (at x)" |
|
| 36667 | 2054 |
shows "(g ---> l) (at x)" |
| 53282 | 2055 |
using _ assms(3) |
| 36667 | 2056 |
proof (rule Lim_transform_eventually) |
2057 |
show "eventually (\<lambda>x. f x = g x) (at x)" |
|
|
51530
609914f0934a
rename eventually_at / _within, to distinguish them from the lemmas in the HOL image
hoelzl
parents:
51518
diff
changeset
|
2058 |
unfolding eventually_at2 |
| 36667 | 2059 |
using assms(1,2) by auto |
2060 |
qed |
|
| 33175 | 2061 |
|
2062 |
text{* Common case assuming being away from some crucial point like 0. *}
|
|
2063 |
||
2064 |
lemma Lim_transform_away_within: |
|
| 36669 | 2065 |
fixes a b :: "'a::t1_space" |
| 53282 | 2066 |
assumes "a \<noteq> b" |
2067 |
and "\<forall>x\<in>S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" |
|
2068 |
and "(f ---> l) (at a within S)" |
|
| 33175 | 2069 |
shows "(g ---> l) (at a within S)" |
| 36669 | 2070 |
proof (rule Lim_transform_eventually) |
2071 |
show "(f ---> l) (at a within S)" by fact |
|
2072 |
show "eventually (\<lambda>x. f x = g x) (at a within S)" |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
2073 |
unfolding eventually_at_topological |
| 36669 | 2074 |
by (rule exI [where x="- {b}"], simp add: open_Compl assms)
|
| 33175 | 2075 |
qed |
2076 |
||
2077 |
lemma Lim_transform_away_at: |
|
| 36669 | 2078 |
fixes a b :: "'a::t1_space" |
| 52624 | 2079 |
assumes ab: "a\<noteq>b" |
2080 |
and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x" |
|
2081 |
and fl: "(f ---> l) (at a)" |
|
| 33175 | 2082 |
shows "(g ---> l) (at a)" |
| 52624 | 2083 |
using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp |
| 33175 | 2084 |
|
2085 |
text{* Alternatively, within an open set. *}
|
|
2086 |
||
2087 |
lemma Lim_transform_within_open: |
|
| 53282 | 2088 |
assumes "open S" and "a \<in> S" |
2089 |
and "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x" |
|
2090 |
and "(f ---> l) (at a)" |
|
| 33175 | 2091 |
shows "(g ---> l) (at a)" |
| 36667 | 2092 |
proof (rule Lim_transform_eventually) |
2093 |
show "eventually (\<lambda>x. f x = g x) (at a)" |
|
2094 |
unfolding eventually_at_topological |
|
2095 |
using assms(1,2,3) by auto |
|
2096 |
show "(f ---> l) (at a)" by fact |
|
| 33175 | 2097 |
qed |
2098 |
||
2099 |
text{* A congruence rule allowing us to transform limits assuming not at point. *}
|
|
2100 |
||
2101 |
(* FIXME: Only one congruence rule for tendsto can be used at a time! *) |
|
2102 |
||
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
2103 |
lemma Lim_cong_within(*[cong add]*): |
| 53282 | 2104 |
assumes "a = b" |
2105 |
and "x = y" |
|
2106 |
and "S = T" |
|
2107 |
and "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x" |
|
| 43338 | 2108 |
shows "(f ---> x) (at a within S) \<longleftrightarrow> (g ---> y) (at b within T)" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
2109 |
unfolding tendsto_def eventually_at_topological |
| 36667 | 2110 |
using assms by simp |
2111 |
||
2112 |
lemma Lim_cong_at(*[cong add]*): |
|
| 43338 | 2113 |
assumes "a = b" "x = y" |
| 53282 | 2114 |
and "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x" |
| 43338 | 2115 |
shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))" |
| 36667 | 2116 |
unfolding tendsto_def eventually_at_topological |
2117 |
using assms by simp |
|
| 33175 | 2118 |
|
2119 |
text{* Useful lemmas on closure and set of possible sequential limits.*}
|
|
2120 |
||
2121 |
lemma closure_sequential: |
|
| 50883 | 2122 |
fixes l :: "'a::first_countable_topology" |
| 53291 | 2123 |
shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)" |
2124 |
(is "?lhs = ?rhs") |
|
| 33175 | 2125 |
proof |
| 53282 | 2126 |
assume "?lhs" |
2127 |
moreover |
|
2128 |
{
|
|
2129 |
assume "l \<in> S" |
|
2130 |
then have "?rhs" using tendsto_const[of l sequentially] by auto |
|
| 52624 | 2131 |
} |
2132 |
moreover |
|
| 53282 | 2133 |
{
|
2134 |
assume "l islimpt S" |
|
2135 |
then have "?rhs" unfolding islimpt_sequential by auto |
|
| 52624 | 2136 |
} |
2137 |
ultimately show "?rhs" |
|
2138 |
unfolding closure_def by auto |
|
| 33175 | 2139 |
next |
2140 |
assume "?rhs" |
|
| 53282 | 2141 |
then show "?lhs" unfolding closure_def islimpt_sequential by auto |
| 33175 | 2142 |
qed |
2143 |
||
2144 |
lemma closed_sequential_limits: |
|
| 50883 | 2145 |
fixes S :: "'a::first_countable_topology set" |
| 33175 | 2146 |
shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)" |
| 55775 | 2147 |
by (metis closure_sequential closure_subset_eq subset_iff) |
| 33175 | 2148 |
|
2149 |
lemma closure_approachable: |
|
2150 |
fixes S :: "'a::metric_space set" |
|
2151 |
shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)" |
|
2152 |
apply (auto simp add: closure_def islimpt_approachable) |
|
| 52624 | 2153 |
apply (metis dist_self) |
2154 |
done |
|
| 33175 | 2155 |
|
2156 |
lemma closed_approachable: |
|
2157 |
fixes S :: "'a::metric_space set" |
|
| 53291 | 2158 |
shows "closed S \<Longrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S" |
| 33175 | 2159 |
by (metis closure_closed closure_approachable) |
2160 |
||
| 51351 | 2161 |
lemma closure_contains_Inf: |
2162 |
fixes S :: "real set" |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2163 |
assumes "S \<noteq> {}" "bdd_below S"
|
| 51351 | 2164 |
shows "Inf S \<in> closure S" |
| 52624 | 2165 |
proof - |
| 51351 | 2166 |
have *: "\<forall>x\<in>S. Inf S \<le> x" |
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2167 |
using cInf_lower[of _ S] assms by metis |
| 52624 | 2168 |
{
|
| 53282 | 2169 |
fix e :: real |
2170 |
assume "e > 0" |
|
| 52624 | 2171 |
then have "Inf S < Inf S + e" by simp |
2172 |
with assms obtain x where "x \<in> S" "x < Inf S + e" |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2173 |
by (subst (asm) cInf_less_iff) auto |
| 52624 | 2174 |
with * have "\<exists>x\<in>S. dist x (Inf S) < e" |
2175 |
by (intro bexI[of _ x]) (auto simp add: dist_real_def) |
|
2176 |
} |
|
2177 |
then show ?thesis unfolding closure_approachable by auto |
|
| 51351 | 2178 |
qed |
2179 |
||
2180 |
lemma closed_contains_Inf: |
|
2181 |
fixes S :: "real set" |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2182 |
shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
|
| 51351 | 2183 |
by (metis closure_contains_Inf closure_closed assms) |
2184 |
||
2185 |
lemma not_trivial_limit_within_ball: |
|
| 53640 | 2186 |
"\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
|
| 51351 | 2187 |
(is "?lhs = ?rhs") |
2188 |
proof - |
|
| 53282 | 2189 |
{
|
2190 |
assume "?lhs" |
|
2191 |
{
|
|
2192 |
fix e :: real |
|
2193 |
assume "e > 0" |
|
| 53640 | 2194 |
then obtain y where "y \<in> S - {x}" and "dist y x < e"
|
| 51351 | 2195 |
using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
|
2196 |
by auto |
|
| 53640 | 2197 |
then have "y \<in> S \<inter> ball x e - {x}"
|
| 51351 | 2198 |
unfolding ball_def by (simp add: dist_commute) |
| 53640 | 2199 |
then have "S \<inter> ball x e - {x} \<noteq> {}" by blast
|
| 52624 | 2200 |
} |
2201 |
then have "?rhs" by auto |
|
| 51351 | 2202 |
} |
2203 |
moreover |
|
| 53282 | 2204 |
{
|
2205 |
assume "?rhs" |
|
2206 |
{
|
|
2207 |
fix e :: real |
|
2208 |
assume "e > 0" |
|
| 53640 | 2209 |
then obtain y where "y \<in> S \<inter> ball x e - {x}"
|
| 53282 | 2210 |
using `?rhs` by blast |
| 53640 | 2211 |
then have "y \<in> S - {x}" and "dist y x < e"
|
2212 |
unfolding ball_def by (simp_all add: dist_commute) |
|
2213 |
then have "\<exists>y \<in> S - {x}. dist y x < e"
|
|
| 53282 | 2214 |
by auto |
| 51351 | 2215 |
} |
2216 |
then have "?lhs" |
|
| 53282 | 2217 |
using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
|
2218 |
by auto |
|
| 51351 | 2219 |
} |
2220 |
ultimately show ?thesis by auto |
|
2221 |
qed |
|
2222 |
||
| 52624 | 2223 |
|
| 50087 | 2224 |
subsection {* Infimum Distance *}
|
2225 |
||
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2226 |
definition "infdist x A = (if A = {} then 0 else INF a:A. dist x a)"
|
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2227 |
|
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2228 |
lemma bdd_below_infdist[intro, simp]: "bdd_below (dist x`A)" |
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2229 |
by (auto intro!: zero_le_dist) |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2230 |
|
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2231 |
lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = (INF a:A. dist x a)"
|
| 50087 | 2232 |
by (simp add: infdist_def) |
2233 |
||
| 52624 | 2234 |
lemma infdist_nonneg: "0 \<le> infdist x A" |
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2235 |
by (auto simp add: infdist_def intro: cINF_greatest) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2236 |
|
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2237 |
lemma infdist_le: "a \<in> A \<Longrightarrow> infdist x A \<le> dist x a" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2238 |
by (auto intro: cINF_lower simp add: infdist_def) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2239 |
|
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2240 |
lemma infdist_le2: "a \<in> A \<Longrightarrow> dist x a \<le> d \<Longrightarrow> infdist x A \<le> d" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2241 |
by (auto intro!: cINF_lower2 simp add: infdist_def) |
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2242 |
|
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2243 |
lemma infdist_zero[simp]: "a \<in> A \<Longrightarrow> infdist a A = 0" |
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2244 |
by (auto intro!: antisym infdist_nonneg infdist_le2) |
| 50087 | 2245 |
|
| 52624 | 2246 |
lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y" |
| 53640 | 2247 |
proof (cases "A = {}")
|
2248 |
case True |
|
| 53282 | 2249 |
then show ?thesis by (simp add: infdist_def) |
| 50087 | 2250 |
next |
| 53640 | 2251 |
case False |
| 52624 | 2252 |
then obtain a where "a \<in> A" by auto |
| 50087 | 2253 |
have "infdist x A \<le> Inf {dist x y + dist y a |a. a \<in> A}"
|
|
51475
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents:
51473
diff
changeset
|
2254 |
proof (rule cInf_greatest) |
| 53282 | 2255 |
from `A \<noteq> {}` show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}"
|
2256 |
by simp |
|
2257 |
fix d |
|
2258 |
assume "d \<in> {dist x y + dist y a |a. a \<in> A}"
|
|
2259 |
then obtain a where d: "d = dist x y + dist y a" "a \<in> A" |
|
2260 |
by auto |
|
| 50087 | 2261 |
show "infdist x A \<le> d" |
2262 |
unfolding infdist_notempty[OF `A \<noteq> {}`]
|
|
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2263 |
proof (rule cINF_lower2) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2264 |
show "a \<in> A" by fact |
| 53282 | 2265 |
show "dist x a \<le> d" |
2266 |
unfolding d by (rule dist_triangle) |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2267 |
qed simp |
| 50087 | 2268 |
qed |
2269 |
also have "\<dots> = dist x y + infdist y A" |
|
|
51475
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents:
51473
diff
changeset
|
2270 |
proof (rule cInf_eq, safe) |
| 53282 | 2271 |
fix a |
2272 |
assume "a \<in> A" |
|
2273 |
then show "dist x y + infdist y A \<le> dist x y + dist y a" |
|
2274 |
by (auto intro: infdist_le) |
|
| 50087 | 2275 |
next |
| 53282 | 2276 |
fix i |
2277 |
assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
|
|
2278 |
then have "i - dist x y \<le> infdist y A" |
|
2279 |
unfolding infdist_notempty[OF `A \<noteq> {}`] using `a \<in> A`
|
|
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2280 |
by (intro cINF_greatest) (auto simp: field_simps) |
| 53282 | 2281 |
then show "i \<le> dist x y + infdist y A" |
2282 |
by simp |
|
| 50087 | 2283 |
qed |
2284 |
finally show ?thesis by simp |
|
2285 |
qed |
|
2286 |
||
|
51475
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents:
51473
diff
changeset
|
2287 |
lemma in_closure_iff_infdist_zero: |
| 50087 | 2288 |
assumes "A \<noteq> {}"
|
2289 |
shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0" |
|
2290 |
proof |
|
2291 |
assume "x \<in> closure A" |
|
2292 |
show "infdist x A = 0" |
|
2293 |
proof (rule ccontr) |
|
2294 |
assume "infdist x A \<noteq> 0" |
|
| 53282 | 2295 |
with infdist_nonneg[of x A] have "infdist x A > 0" |
2296 |
by auto |
|
2297 |
then have "ball x (infdist x A) \<inter> closure A = {}"
|
|
| 52624 | 2298 |
apply auto |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
2299 |
apply (metis `x \<in> closure A` closure_approachable dist_commute infdist_le not_less) |
| 52624 | 2300 |
done |
| 53282 | 2301 |
then have "x \<notin> closure A" |
| 52624 | 2302 |
by (metis `0 < infdist x A` centre_in_ball disjoint_iff_not_equal) |
| 53282 | 2303 |
then show False using `x \<in> closure A` by simp |
| 50087 | 2304 |
qed |
2305 |
next |
|
2306 |
assume x: "infdist x A = 0" |
|
| 53282 | 2307 |
then obtain a where "a \<in> A" |
2308 |
by atomize_elim (metis all_not_in_conv assms) |
|
2309 |
show "x \<in> closure A" |
|
2310 |
unfolding closure_approachable |
|
2311 |
apply safe |
|
2312 |
proof (rule ccontr) |
|
2313 |
fix e :: real |
|
2314 |
assume "e > 0" |
|
| 50087 | 2315 |
assume "\<not> (\<exists>y\<in>A. dist y x < e)" |
| 53282 | 2316 |
then have "infdist x A \<ge> e" using `a \<in> A` |
| 50087 | 2317 |
unfolding infdist_def |
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
2318 |
by (force simp: dist_commute intro: cINF_greatest) |
| 53282 | 2319 |
with x `e > 0` show False by auto |
| 50087 | 2320 |
qed |
2321 |
qed |
|
2322 |
||
|
51475
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
hoelzl
parents:
51473
diff
changeset
|
2323 |
lemma in_closed_iff_infdist_zero: |
| 50087 | 2324 |
assumes "closed A" "A \<noteq> {}"
|
2325 |
shows "x \<in> A \<longleftrightarrow> infdist x A = 0" |
|
2326 |
proof - |
|
2327 |
have "x \<in> closure A \<longleftrightarrow> infdist x A = 0" |
|
2328 |
by (rule in_closure_iff_infdist_zero) fact |
|
2329 |
with assms show ?thesis by simp |
|
2330 |
qed |
|
2331 |
||
2332 |
lemma tendsto_infdist [tendsto_intros]: |
|
2333 |
assumes f: "(f ---> l) F" |
|
2334 |
shows "((\<lambda>x. infdist (f x) A) ---> infdist l A) F" |
|
2335 |
proof (rule tendstoI) |
|
| 53282 | 2336 |
fix e ::real |
2337 |
assume "e > 0" |
|
| 50087 | 2338 |
from tendstoD[OF f this] |
2339 |
show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F" |
|
2340 |
proof (eventually_elim) |
|
2341 |
fix x |
|
2342 |
from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l] |
|
2343 |
have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l" |
|
2344 |
by (simp add: dist_commute dist_real_def) |
|
2345 |
also assume "dist (f x) l < e" |
|
2346 |
finally show "dist (infdist (f x) A) (infdist l A) < e" . |
|
2347 |
qed |
|
2348 |
qed |
|
2349 |
||
| 33175 | 2350 |
text{* Some other lemmas about sequences. *}
|
2351 |
||
| 53597 | 2352 |
lemma sequentially_offset: (* TODO: move to Topological_Spaces.thy *) |
| 36441 | 2353 |
assumes "eventually (\<lambda>i. P i) sequentially" |
2354 |
shows "eventually (\<lambda>i. P (i + k)) sequentially" |
|
| 53597 | 2355 |
using assms by (rule eventually_sequentially_seg [THEN iffD2]) |
2356 |
||
2357 |
lemma seq_offset_neg: (* TODO: move to Topological_Spaces.thy *) |
|
| 53291 | 2358 |
"(f ---> l) sequentially \<Longrightarrow> ((\<lambda>i. f(i - k)) ---> l) sequentially" |
| 53597 | 2359 |
apply (erule filterlim_compose) |
2360 |
apply (simp add: filterlim_def le_sequentially eventually_filtermap eventually_sequentially) |
|
| 52624 | 2361 |
apply arith |
2362 |
done |
|
| 33175 | 2363 |
|
2364 |
lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially" |
|
| 53597 | 2365 |
using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc) (* TODO: move to Limits.thy *) |
| 33175 | 2366 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2367 |
subsection {* More properties of closed balls *}
|
| 33175 | 2368 |
|
| 54070 | 2369 |
lemma closed_vimage: (* TODO: move to Topological_Spaces.thy *) |
2370 |
assumes "closed s" and "continuous_on UNIV f" |
|
2371 |
shows "closed (vimage f s)" |
|
2372 |
using assms unfolding continuous_on_closed_vimage [OF closed_UNIV] |
|
2373 |
by simp |
|
2374 |
||
| 33175 | 2375 |
lemma closed_cball: "closed (cball x e)" |
| 54070 | 2376 |
proof - |
2377 |
have "closed (dist x -` {..e})"
|
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
2378 |
by (intro closed_vimage closed_atMost continuous_intros) |
| 54070 | 2379 |
also have "dist x -` {..e} = cball x e"
|
2380 |
by auto |
|
2381 |
finally show ?thesis . |
|
2382 |
qed |
|
| 33175 | 2383 |
|
2384 |
lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. cball x e \<subseteq> S)" |
|
| 52624 | 2385 |
proof - |
2386 |
{
|
|
2387 |
fix x and e::real |
|
2388 |
assume "x\<in>S" "e>0" "ball x e \<subseteq> S" |
|
| 53282 | 2389 |
then have "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) |
| 52624 | 2390 |
} |
2391 |
moreover |
|
2392 |
{
|
|
2393 |
fix x and e::real |
|
2394 |
assume "x\<in>S" "e>0" "cball x e \<subseteq> S" |
|
| 53282 | 2395 |
then have "\<exists>d>0. ball x d \<subseteq> S" |
| 52624 | 2396 |
unfolding subset_eq |
2397 |
apply(rule_tac x="e/2" in exI) |
|
2398 |
apply auto |
|
2399 |
done |
|
2400 |
} |
|
2401 |
ultimately show ?thesis |
|
2402 |
unfolding open_contains_ball by auto |
|
| 33175 | 2403 |
qed |
2404 |
||
| 53291 | 2405 |
lemma open_contains_cball_eq: "open S \<Longrightarrow> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))" |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
2406 |
by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball) |
| 33175 | 2407 |
|
2408 |
lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)" |
|
2409 |
apply (simp add: interior_def, safe) |
|
2410 |
apply (force simp add: open_contains_cball) |
|
2411 |
apply (rule_tac x="ball x e" in exI) |
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
2412 |
apply (simp add: subset_trans [OF ball_subset_cball]) |
| 33175 | 2413 |
done |
2414 |
||
2415 |
lemma islimpt_ball: |
|
2416 |
fixes x y :: "'a::{real_normed_vector,perfect_space}"
|
|
| 53291 | 2417 |
shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e" |
2418 |
(is "?lhs = ?rhs") |
|
| 33175 | 2419 |
proof |
2420 |
assume "?lhs" |
|
| 53282 | 2421 |
{
|
2422 |
assume "e \<le> 0" |
|
2423 |
then have *:"ball x e = {}"
|
|
2424 |
using ball_eq_empty[of x e] by auto |
|
2425 |
have False using `?lhs` |
|
2426 |
unfolding * using islimpt_EMPTY[of y] by auto |
|
| 33175 | 2427 |
} |
| 53282 | 2428 |
then have "e > 0" by (metis not_less) |
| 33175 | 2429 |
moreover |
| 52624 | 2430 |
have "y \<in> cball x e" |
2431 |
using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"] |
|
2432 |
ball_subset_cball[of x e] `?lhs` |
|
2433 |
unfolding closed_limpt by auto |
|
| 33175 | 2434 |
ultimately show "?rhs" by auto |
2435 |
next |
|
| 53282 | 2436 |
assume "?rhs" |
| 53640 | 2437 |
then have "e > 0" by auto |
| 53282 | 2438 |
{
|
2439 |
fix d :: real |
|
2440 |
assume "d > 0" |
|
| 33175 | 2441 |
have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" |
| 53282 | 2442 |
proof (cases "d \<le> dist x y") |
2443 |
case True |
|
2444 |
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" |
|
2445 |
proof (cases "x = y") |
|
2446 |
case True |
|
2447 |
then have False |
|
2448 |
using `d \<le> dist x y` `d>0` by auto |
|
2449 |
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" |
|
2450 |
by auto |
|
| 33175 | 2451 |
next |
2452 |
case False |
|
| 53282 | 2453 |
have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) = |
2454 |
norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))" |
|
| 53291 | 2455 |
unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[symmetric] |
| 53282 | 2456 |
by auto |
| 33175 | 2457 |
also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)" |
| 53291 | 2458 |
using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", symmetric, of "y - x"] |
| 33175 | 2459 |
unfolding scaleR_minus_left scaleR_one |
2460 |
by (auto simp add: norm_minus_commute) |
|
2461 |
also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>" |
|
2462 |
unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]] |
|
| 53282 | 2463 |
unfolding distrib_right using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm] |
2464 |
by auto |
|
2465 |
also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs` |
|
2466 |
by (auto simp add: dist_norm) |
|
2467 |
finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0` |
|
2468 |
by auto |
|
| 33175 | 2469 |
moreover |
2470 |
have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0" |
|
| 53282 | 2471 |
using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff |
2472 |
by (auto simp add: dist_commute) |
|
| 33175 | 2473 |
moreover |
| 53282 | 2474 |
have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d" |
2475 |
unfolding dist_norm |
|
2476 |
apply simp |
|
2477 |
unfolding norm_minus_cancel |
|
2478 |
using `d > 0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y] |
|
2479 |
unfolding dist_norm |
|
2480 |
apply auto |
|
2481 |
done |
|
2482 |
ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" |
|
2483 |
apply (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI) |
|
2484 |
apply auto |
|
2485 |
done |
|
| 33175 | 2486 |
qed |
2487 |
next |
|
| 53282 | 2488 |
case False |
2489 |
then have "d > dist x y" by auto |
|
2490 |
show "\<exists>x' \<in> ball x e. x' \<noteq> y \<and> dist x' y < d" |
|
2491 |
proof (cases "x = y") |
|
| 33175 | 2492 |
case True |
2493 |
obtain z where **: "z \<noteq> y" "dist z y < min e d" |
|
2494 |
using perfect_choose_dist[of "min e d" y] |
|
2495 |
using `d > 0` `e>0` by auto |
|
2496 |
show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" |
|
2497 |
unfolding `x = y` |
|
2498 |
using `z \<noteq> y` ** |
|
| 53282 | 2499 |
apply (rule_tac x=z in bexI) |
2500 |
apply (auto simp add: dist_commute) |
|
2501 |
done |
|
| 33175 | 2502 |
next |
| 53282 | 2503 |
case False |
2504 |
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d" |
|
2505 |
using `d>0` `d > dist x y` `?rhs` |
|
2506 |
apply (rule_tac x=x in bexI) |
|
2507 |
apply auto |
|
2508 |
done |
|
| 33175 | 2509 |
qed |
| 53282 | 2510 |
qed |
2511 |
} |
|
2512 |
then show "?lhs" |
|
2513 |
unfolding mem_cball islimpt_approachable mem_ball by auto |
|
| 33175 | 2514 |
qed |
2515 |
||
2516 |
lemma closure_ball_lemma: |
|
2517 |
fixes x y :: "'a::real_normed_vector" |
|
| 53282 | 2518 |
assumes "x \<noteq> y" |
2519 |
shows "y islimpt ball x (dist x y)" |
|
| 33175 | 2520 |
proof (rule islimptI) |
| 53282 | 2521 |
fix T |
2522 |
assume "y \<in> T" "open T" |
|
| 33175 | 2523 |
then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T" |
2524 |
unfolding open_dist by fast |
|
2525 |
(* choose point between x and y, within distance r of y. *) |
|
2526 |
def k \<equiv> "min 1 (r / (2 * dist x y))" |
|
2527 |
def z \<equiv> "y + scaleR k (x - y)" |
|
2528 |
have z_def2: "z = x + scaleR (1 - k) (y - x)" |
|
2529 |
unfolding z_def by (simp add: algebra_simps) |
|
2530 |
have "dist z y < r" |
|
2531 |
unfolding z_def k_def using `0 < r` |
|
2532 |
by (simp add: dist_norm min_def) |
|
| 53282 | 2533 |
then have "z \<in> T" |
2534 |
using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp |
|
| 33175 | 2535 |
have "dist x z < dist x y" |
2536 |
unfolding z_def2 dist_norm |
|
2537 |
apply (simp add: norm_minus_commute) |
|
2538 |
apply (simp only: dist_norm [symmetric]) |
|
2539 |
apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp) |
|
2540 |
apply (rule mult_strict_right_mono) |
|
| 56541 | 2541 |
apply (simp add: k_def zero_less_dist_iff `0 < r` `x \<noteq> y`) |
| 33175 | 2542 |
apply (simp add: zero_less_dist_iff `x \<noteq> y`) |
2543 |
done |
|
| 53282 | 2544 |
then have "z \<in> ball x (dist x y)" |
2545 |
by simp |
|
| 33175 | 2546 |
have "z \<noteq> y" |
2547 |
unfolding z_def k_def using `x \<noteq> y` `0 < r` |
|
2548 |
by (simp add: min_def) |
|
2549 |
show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y" |
|
2550 |
using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y` |
|
2551 |
by fast |
|
2552 |
qed |
|
2553 |
||
2554 |
lemma closure_ball: |
|
2555 |
fixes x :: "'a::real_normed_vector" |
|
2556 |
shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e" |
|
| 52624 | 2557 |
apply (rule equalityI) |
2558 |
apply (rule closure_minimal) |
|
2559 |
apply (rule ball_subset_cball) |
|
2560 |
apply (rule closed_cball) |
|
2561 |
apply (rule subsetI, rename_tac y) |
|
2562 |
apply (simp add: le_less [where 'a=real]) |
|
2563 |
apply (erule disjE) |
|
2564 |
apply (rule subsetD [OF closure_subset], simp) |
|
2565 |
apply (simp add: closure_def) |
|
2566 |
apply clarify |
|
2567 |
apply (rule closure_ball_lemma) |
|
2568 |
apply (simp add: zero_less_dist_iff) |
|
2569 |
done |
|
| 33175 | 2570 |
|
2571 |
(* In a trivial vector space, this fails for e = 0. *) |
|
2572 |
lemma interior_cball: |
|
2573 |
fixes x :: "'a::{real_normed_vector, perfect_space}"
|
|
2574 |
shows "interior (cball x e) = ball x e" |
|
| 53640 | 2575 |
proof (cases "e \<ge> 0") |
| 33175 | 2576 |
case False note cs = this |
| 53282 | 2577 |
from cs have "ball x e = {}"
|
2578 |
using ball_empty[of e x] by auto |
|
2579 |
moreover |
|
2580 |
{
|
|
2581 |
fix y |
|
2582 |
assume "y \<in> cball x e" |
|
2583 |
then have False |
|
2584 |
unfolding mem_cball using dist_nz[of x y] cs by auto |
|
2585 |
} |
|
2586 |
then have "cball x e = {}" by auto
|
|
2587 |
then have "interior (cball x e) = {}"
|
|
2588 |
using interior_empty by auto |
|
| 33175 | 2589 |
ultimately show ?thesis by blast |
2590 |
next |
|
2591 |
case True note cs = this |
|
| 53282 | 2592 |
have "ball x e \<subseteq> cball x e" |
2593 |
using ball_subset_cball by auto |
|
2594 |
moreover |
|
2595 |
{
|
|
2596 |
fix S y |
|
2597 |
assume as: "S \<subseteq> cball x e" "open S" "y\<in>S" |
|
2598 |
then obtain d where "d>0" and d: "\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S" |
|
2599 |
unfolding open_dist by blast |
|
| 33175 | 2600 |
then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d" |
2601 |
using perfect_choose_dist [of d] by auto |
|
| 53282 | 2602 |
have "xa \<in> S" |
2603 |
using d[THEN spec[where x = xa]] |
|
2604 |
using xa by (auto simp add: dist_commute) |
|
2605 |
then have xa_cball: "xa \<in> cball x e" |
|
2606 |
using as(1) by auto |
|
2607 |
then have "y \<in> ball x e" |
|
2608 |
proof (cases "x = y") |
|
| 33175 | 2609 |
case True |
| 53282 | 2610 |
then have "e > 0" |
2611 |
using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball] |
|
| 52624 | 2612 |
by (auto simp add: dist_commute) |
| 53282 | 2613 |
then show "y \<in> ball x e" |
2614 |
using `x = y ` by simp |
|
| 33175 | 2615 |
next |
2616 |
case False |
|
| 53282 | 2617 |
have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d" |
2618 |
unfolding dist_norm |
|
| 33175 | 2619 |
using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto |
| 53282 | 2620 |
then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e" |
| 52624 | 2621 |
using d as(1)[unfolded subset_eq] by blast |
| 33175 | 2622 |
have "y - x \<noteq> 0" using `x \<noteq> y` by auto |
| 56541 | 2623 |
hence **:"d / (2 * norm (y - x)) > 0" |
2624 |
unfolding zero_less_norm_iff[symmetric] using `d>0` by auto |
|
| 53282 | 2625 |
have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x = |
2626 |
norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)" |
|
| 33175 | 2627 |
by (auto simp add: dist_norm algebra_simps) |
2628 |
also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))" |
|
2629 |
by (auto simp add: algebra_simps) |
|
2630 |
also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)" |
|
2631 |
using ** by auto |
|
| 53282 | 2632 |
also have "\<dots> = (dist y x) + d/2" |
2633 |
using ** by (auto simp add: distrib_right dist_norm) |
|
2634 |
finally have "e \<ge> dist x y +d/2" |
|
2635 |
using *[unfolded mem_cball] by (auto simp add: dist_commute) |
|
2636 |
then show "y \<in> ball x e" |
|
2637 |
unfolding mem_ball using `d>0` by auto |
|
| 52624 | 2638 |
qed |
2639 |
} |
|
| 53282 | 2640 |
then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e" |
2641 |
by auto |
|
| 52624 | 2642 |
ultimately show ?thesis |
| 53640 | 2643 |
using interior_unique[of "ball x e" "cball x e"] |
2644 |
using open_ball[of x e] |
|
2645 |
by auto |
|
| 33175 | 2646 |
qed |
2647 |
||
2648 |
lemma frontier_ball: |
|
2649 |
fixes a :: "'a::real_normed_vector" |
|
| 53291 | 2650 |
shows "0 < e \<Longrightarrow> frontier(ball a e) = {x. dist a x = e}"
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
2651 |
apply (simp add: frontier_def closure_ball interior_open order_less_imp_le) |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
2652 |
apply (simp add: set_eq_iff) |
| 52624 | 2653 |
apply arith |
2654 |
done |
|
| 33175 | 2655 |
|
2656 |
lemma frontier_cball: |
|
2657 |
fixes a :: "'a::{real_normed_vector, perfect_space}"
|
|
| 53640 | 2658 |
shows "frontier (cball a e) = {x. dist a x = e}"
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
2659 |
apply (simp add: frontier_def interior_cball closed_cball order_less_imp_le) |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
2660 |
apply (simp add: set_eq_iff) |
| 52624 | 2661 |
apply arith |
2662 |
done |
|
| 33175 | 2663 |
|
| 53640 | 2664 |
lemma cball_eq_empty: "cball x e = {} \<longleftrightarrow> e < 0"
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
2665 |
apply (simp add: set_eq_iff not_le) |
| 52624 | 2666 |
apply (metis zero_le_dist dist_self order_less_le_trans) |
2667 |
done |
|
2668 |
||
| 53282 | 2669 |
lemma cball_empty: "e < 0 \<Longrightarrow> cball x e = {}"
|
| 52624 | 2670 |
by (simp add: cball_eq_empty) |
| 33175 | 2671 |
|
2672 |
lemma cball_eq_sing: |
|
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2673 |
fixes x :: "'a::{metric_space,perfect_space}"
|
| 53640 | 2674 |
shows "cball x e = {x} \<longleftrightarrow> e = 0"
|
| 33175 | 2675 |
proof (rule linorder_cases) |
2676 |
assume e: "0 < e" |
|
2677 |
obtain a where "a \<noteq> x" "dist a x < e" |
|
2678 |
using perfect_choose_dist [OF e] by auto |
|
| 53282 | 2679 |
then have "a \<noteq> x" "dist x a \<le> e" |
2680 |
by (auto simp add: dist_commute) |
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
2681 |
with e show ?thesis by (auto simp add: set_eq_iff) |
| 33175 | 2682 |
qed auto |
2683 |
||
2684 |
lemma cball_sing: |
|
2685 |
fixes x :: "'a::metric_space" |
|
| 53291 | 2686 |
shows "e = 0 \<Longrightarrow> cball x e = {x}"
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
2687 |
by (auto simp add: set_eq_iff) |
| 33175 | 2688 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2689 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2690 |
subsection {* Boundedness *}
|
| 33175 | 2691 |
|
2692 |
(* FIXME: This has to be unified with BSEQ!! *) |
|
| 52624 | 2693 |
definition (in metric_space) bounded :: "'a set \<Rightarrow> bool" |
2694 |
where "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)" |
|
| 33175 | 2695 |
|
| 50998 | 2696 |
lemma bounded_subset_cball: "bounded S \<longleftrightarrow> (\<exists>e x. S \<subseteq> cball x e)" |
2697 |
unfolding bounded_def subset_eq by auto |
|
2698 |
||
| 33175 | 2699 |
lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)" |
| 52624 | 2700 |
unfolding bounded_def |
| 55775 | 2701 |
by auto (metis add_commute add_le_cancel_right dist_commute dist_triangle_le) |
| 33175 | 2702 |
|
2703 |
lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)" |
|
| 52624 | 2704 |
unfolding bounded_any_center [where a=0] |
2705 |
by (simp add: dist_norm) |
|
| 33175 | 2706 |
|
| 53282 | 2707 |
lemma bounded_realI: |
2708 |
assumes "\<forall>x\<in>s. abs (x::real) \<le> B" |
|
2709 |
shows "bounded s" |
|
2710 |
unfolding bounded_def dist_real_def |
|
| 55775 | 2711 |
by (metis abs_minus_commute assms diff_0_right) |
| 50104 | 2712 |
|
| 50948 | 2713 |
lemma bounded_empty [simp]: "bounded {}"
|
2714 |
by (simp add: bounded_def) |
|
2715 |
||
| 53291 | 2716 |
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> bounded S" |
| 33175 | 2717 |
by (metis bounded_def subset_eq) |
2718 |
||
| 53291 | 2719 |
lemma bounded_interior[intro]: "bounded S \<Longrightarrow> bounded(interior S)" |
| 33175 | 2720 |
by (metis bounded_subset interior_subset) |
2721 |
||
| 52624 | 2722 |
lemma bounded_closure[intro]: |
2723 |
assumes "bounded S" |
|
2724 |
shows "bounded (closure S)" |
|
2725 |
proof - |
|
2726 |
from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a" |
|
2727 |
unfolding bounded_def by auto |
|
2728 |
{
|
|
2729 |
fix y |
|
2730 |
assume "y \<in> closure S" |
|
| 33175 | 2731 |
then obtain f where f: "\<forall>n. f n \<in> S" "(f ---> y) sequentially" |
2732 |
unfolding closure_sequential by auto |
|
2733 |
have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp |
|
| 53282 | 2734 |
then have "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially" |
| 33175 | 2735 |
by (rule eventually_mono, simp add: f(1)) |
2736 |
have "dist x y \<le> a" |
|
2737 |
apply (rule Lim_dist_ubound [of sequentially f]) |
|
2738 |
apply (rule trivial_limit_sequentially) |
|
2739 |
apply (rule f(2)) |
|
2740 |
apply fact |
|
2741 |
done |
|
2742 |
} |
|
| 53282 | 2743 |
then show ?thesis |
2744 |
unfolding bounded_def by auto |
|
| 33175 | 2745 |
qed |
2746 |
||
2747 |
lemma bounded_cball[simp,intro]: "bounded (cball x e)" |
|
2748 |
apply (simp add: bounded_def) |
|
2749 |
apply (rule_tac x=x in exI) |
|
2750 |
apply (rule_tac x=e in exI) |
|
2751 |
apply auto |
|
2752 |
done |
|
2753 |
||
| 53640 | 2754 |
lemma bounded_ball[simp,intro]: "bounded (ball x e)" |
| 33175 | 2755 |
by (metis ball_subset_cball bounded_cball bounded_subset) |
2756 |
||
2757 |
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T" |
|
2758 |
apply (auto simp add: bounded_def) |
|
| 55775 | 2759 |
by (metis Un_iff add_le_cancel_left dist_triangle le_max_iff_disj max.order_iff) |
| 33175 | 2760 |
|
| 53640 | 2761 |
lemma bounded_Union[intro]: "finite F \<Longrightarrow> \<forall>S\<in>F. bounded S \<Longrightarrow> bounded (\<Union>F)" |
| 52624 | 2762 |
by (induct rule: finite_induct[of F]) auto |
| 33175 | 2763 |
|
| 50955 | 2764 |
lemma bounded_UN [intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. bounded (B x) \<Longrightarrow> bounded (\<Union>x\<in>A. B x)" |
| 52624 | 2765 |
by (induct set: finite) auto |
| 50955 | 2766 |
|
| 50948 | 2767 |
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S" |
2768 |
proof - |
|
| 53640 | 2769 |
have "\<forall>y\<in>{x}. dist x y \<le> 0"
|
2770 |
by simp |
|
2771 |
then have "bounded {x}"
|
|
2772 |
unfolding bounded_def by fast |
|
2773 |
then show ?thesis |
|
2774 |
by (metis insert_is_Un bounded_Un) |
|
| 50948 | 2775 |
qed |
2776 |
||
2777 |
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S" |
|
| 52624 | 2778 |
by (induct set: finite) simp_all |
| 50948 | 2779 |
|
| 53640 | 2780 |
lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x \<le> b)" |
| 33175 | 2781 |
apply (simp add: bounded_iff) |
| 53640 | 2782 |
apply (subgoal_tac "\<And>x (y::real). 0 < 1 + abs y \<and> (x \<le> y \<longrightarrow> x \<le> 1 + abs y)") |
| 52624 | 2783 |
apply metis |
2784 |
apply arith |
|
2785 |
done |
|
| 33175 | 2786 |
|
| 53640 | 2787 |
lemma Bseq_eq_bounded: |
2788 |
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
2789 |
shows "Bseq f \<longleftrightarrow> bounded (range f)" |
|
| 50972 | 2790 |
unfolding Bseq_def bounded_pos by auto |
2791 |
||
| 33175 | 2792 |
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)" |
2793 |
by (metis Int_lower1 Int_lower2 bounded_subset) |
|
2794 |
||
| 53291 | 2795 |
lemma bounded_diff[intro]: "bounded S \<Longrightarrow> bounded (S - T)" |
| 52624 | 2796 |
by (metis Diff_subset bounded_subset) |
| 33175 | 2797 |
|
2798 |
lemma not_bounded_UNIV[simp, intro]: |
|
2799 |
"\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
|
|
| 53640 | 2800 |
proof (auto simp add: bounded_pos not_le) |
| 33175 | 2801 |
obtain x :: 'a where "x \<noteq> 0" |
2802 |
using perfect_choose_dist [OF zero_less_one] by fast |
|
| 53640 | 2803 |
fix b :: real |
2804 |
assume b: "b >0" |
|
2805 |
have b1: "b +1 \<ge> 0" |
|
2806 |
using b by simp |
|
| 33175 | 2807 |
with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))" |
2808 |
by (simp add: norm_sgn) |
|
2809 |
then show "\<exists>x::'a. b < norm x" .. |
|
2810 |
qed |
|
2811 |
||
2812 |
lemma bounded_linear_image: |
|
| 53291 | 2813 |
assumes "bounded S" |
2814 |
and "bounded_linear f" |
|
2815 |
shows "bounded (f ` S)" |
|
| 52624 | 2816 |
proof - |
| 53640 | 2817 |
from assms(1) obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b" |
| 52624 | 2818 |
unfolding bounded_pos by auto |
| 53640 | 2819 |
from assms(2) obtain B where B: "B > 0" "\<forall>x. norm (f x) \<le> B * norm x" |
| 52624 | 2820 |
using bounded_linear.pos_bounded by (auto simp add: mult_ac) |
2821 |
{
|
|
| 53282 | 2822 |
fix x |
| 53640 | 2823 |
assume "x \<in> S" |
2824 |
then have "norm x \<le> b" |
|
2825 |
using b by auto |
|
2826 |
then have "norm (f x) \<le> B * b" |
|
2827 |
using B(2) |
|
| 52624 | 2828 |
apply (erule_tac x=x in allE) |
2829 |
apply (metis B(1) B(2) order_trans mult_le_cancel_left_pos) |
|
2830 |
done |
|
| 33175 | 2831 |
} |
| 53282 | 2832 |
then show ?thesis |
2833 |
unfolding bounded_pos |
|
| 52624 | 2834 |
apply (rule_tac x="b*B" in exI) |
| 56544 | 2835 |
using b B by (auto simp add: mult_commute) |
| 33175 | 2836 |
qed |
2837 |
||
2838 |
lemma bounded_scaling: |
|
2839 |
fixes S :: "'a::real_normed_vector set" |
|
2840 |
shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)" |
|
| 53291 | 2841 |
apply (rule bounded_linear_image) |
2842 |
apply assumption |
|
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44252
diff
changeset
|
2843 |
apply (rule bounded_linear_scaleR_right) |
| 33175 | 2844 |
done |
2845 |
||
2846 |
lemma bounded_translation: |
|
2847 |
fixes S :: "'a::real_normed_vector set" |
|
| 52624 | 2848 |
assumes "bounded S" |
2849 |
shows "bounded ((\<lambda>x. a + x) ` S)" |
|
| 53282 | 2850 |
proof - |
| 53640 | 2851 |
from assms obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b" |
| 52624 | 2852 |
unfolding bounded_pos by auto |
2853 |
{
|
|
2854 |
fix x |
|
| 53640 | 2855 |
assume "x \<in> S" |
| 53282 | 2856 |
then have "norm (a + x) \<le> b + norm a" |
| 52624 | 2857 |
using norm_triangle_ineq[of a x] b by auto |
| 33175 | 2858 |
} |
| 53282 | 2859 |
then show ?thesis |
| 52624 | 2860 |
unfolding bounded_pos |
2861 |
using norm_ge_zero[of a] b(1) and add_strict_increasing[of b 0 "norm a"] |
|
|
48048
87b94fb75198
remove stray reference to no-longer-existing theorem 'add'
huffman
parents:
47108
diff
changeset
|
2862 |
by (auto intro!: exI[of _ "b + norm a"]) |
| 33175 | 2863 |
qed |
2864 |
||
2865 |
||
2866 |
text{* Some theorems on sups and infs using the notion "bounded". *}
|
|
2867 |
||
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2868 |
lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)" |
| 33175 | 2869 |
by (simp add: bounded_iff) |
2870 |
||
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2871 |
lemma bounded_imp_bdd_above: "bounded S \<Longrightarrow> bdd_above (S :: real set)" |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2872 |
by (auto simp: bounded_def bdd_above_def dist_real_def) |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2873 |
(metis abs_le_D1 abs_minus_commute diff_le_eq) |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2874 |
|
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2875 |
lemma bounded_imp_bdd_below: "bounded S \<Longrightarrow> bdd_below (S :: real set)" |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2876 |
by (auto simp: bounded_def bdd_below_def dist_real_def) |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2877 |
(metis abs_le_D1 add_commute diff_le_eq) |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2878 |
|
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2879 |
(* TODO: remove the following lemmas about Inf and Sup, is now in conditionally complete lattice *) |
|
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2880 |
|
| 33270 | 2881 |
lemma bounded_has_Sup: |
2882 |
fixes S :: "real set" |
|
| 53640 | 2883 |
assumes "bounded S" |
2884 |
and "S \<noteq> {}"
|
|
| 53282 | 2885 |
shows "\<forall>x\<in>S. x \<le> Sup S" |
2886 |
and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" |
|
| 33270 | 2887 |
proof |
| 53282 | 2888 |
show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b" |
2889 |
using assms by (metis cSup_least) |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2890 |
qed (metis cSup_upper assms(1) bounded_imp_bdd_above) |
| 33270 | 2891 |
|
2892 |
lemma Sup_insert: |
|
2893 |
fixes S :: "real set" |
|
| 53291 | 2894 |
shows "bounded S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2895 |
by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If) |
| 33270 | 2896 |
|
2897 |
lemma Sup_insert_finite: |
|
2898 |
fixes S :: "real set" |
|
| 53291 | 2899 |
shows "finite S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
|
| 33270 | 2900 |
apply (rule Sup_insert) |
2901 |
apply (rule finite_imp_bounded) |
|
| 52624 | 2902 |
apply simp |
2903 |
done |
|
| 33270 | 2904 |
|
2905 |
lemma bounded_has_Inf: |
|
2906 |
fixes S :: "real set" |
|
| 53640 | 2907 |
assumes "bounded S" |
2908 |
and "S \<noteq> {}"
|
|
| 53282 | 2909 |
shows "\<forall>x\<in>S. x \<ge> Inf S" |
2910 |
and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b" |
|
| 33175 | 2911 |
proof |
| 53640 | 2912 |
show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b" |
| 53282 | 2913 |
using assms by (metis cInf_greatest) |
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
2914 |
qed (metis cInf_lower assms(1) bounded_imp_bdd_below) |
| 33270 | 2915 |
|
2916 |
lemma Inf_insert: |
|
2917 |
fixes S :: "real set" |
|
| 53291 | 2918 |
shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
|
|
54259
71c701dc5bf9
add SUP and INF for conditionally complete lattices
hoelzl
parents:
54258
diff
changeset
|
2919 |
by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If) |
| 50944 | 2920 |
|
| 33270 | 2921 |
lemma Inf_insert_finite: |
2922 |
fixes S :: "real set" |
|
| 53291 | 2923 |
shows "finite S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
|
| 53282 | 2924 |
apply (rule Inf_insert) |
2925 |
apply (rule finite_imp_bounded) |
|
2926 |
apply simp |
|
2927 |
done |
|
| 33270 | 2928 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2929 |
subsection {* Compactness *}
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2930 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2931 |
subsubsection {* Bolzano-Weierstrass property *}
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2932 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2933 |
lemma heine_borel_imp_bolzano_weierstrass: |
| 53640 | 2934 |
assumes "compact s" |
2935 |
and "infinite t" |
|
2936 |
and "t \<subseteq> s" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2937 |
shows "\<exists>x \<in> s. x islimpt t" |
| 53291 | 2938 |
proof (rule ccontr) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2939 |
assume "\<not> (\<exists>x \<in> s. x islimpt t)" |
| 53640 | 2940 |
then obtain f where f: "\<forall>x\<in>s. x \<in> f x \<and> open (f x) \<and> (\<forall>y\<in>t. y \<in> f x \<longrightarrow> y = x)" |
| 52624 | 2941 |
unfolding islimpt_def |
| 53282 | 2942 |
using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"] |
2943 |
by auto |
|
| 53640 | 2944 |
obtain g where g: "g \<subseteq> {t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
|
| 52624 | 2945 |
using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]]
|
2946 |
using f by auto |
|
| 53640 | 2947 |
from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa" |
2948 |
by auto |
|
| 52624 | 2949 |
{
|
2950 |
fix x y |
|
| 53640 | 2951 |
assume "x \<in> t" "y \<in> t" "f x = f y" |
| 53282 | 2952 |
then have "x \<in> f x" "y \<in> f x \<longrightarrow> y = x" |
| 53640 | 2953 |
using f[THEN bspec[where x=x]] and `t \<subseteq> s` by auto |
| 53282 | 2954 |
then have "x = y" |
| 53640 | 2955 |
using `f x = f y` and f[THEN bspec[where x=y]] and `y \<in> t` and `t \<subseteq> s` |
2956 |
by auto |
|
| 52624 | 2957 |
} |
| 53282 | 2958 |
then have "inj_on f t" |
| 52624 | 2959 |
unfolding inj_on_def by simp |
| 53282 | 2960 |
then have "infinite (f ` t)" |
| 52624 | 2961 |
using assms(2) using finite_imageD by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2962 |
moreover |
| 52624 | 2963 |
{
|
2964 |
fix x |
|
| 53640 | 2965 |
assume "x \<in> t" "f x \<notin> g" |
2966 |
from g(3) assms(3) `x \<in> t` obtain h where "h \<in> g" and "x \<in> h" |
|
2967 |
by auto |
|
2968 |
then obtain y where "y \<in> s" "h = f y" |
|
| 52624 | 2969 |
using g'[THEN bspec[where x=h]] by auto |
| 53282 | 2970 |
then have "y = x" |
| 53640 | 2971 |
using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`] |
2972 |
by auto |
|
| 53282 | 2973 |
then have False |
| 53640 | 2974 |
using `f x \<notin> g` `h \<in> g` unfolding `h = f y` |
2975 |
by auto |
|
| 52624 | 2976 |
} |
| 53282 | 2977 |
then have "f ` t \<subseteq> g" by auto |
| 52624 | 2978 |
ultimately show False |
2979 |
using g(2) using finite_subset by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2980 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2981 |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2982 |
lemma acc_point_range_imp_convergent_subsequence: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2983 |
fixes l :: "'a :: first_countable_topology" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2984 |
assumes l: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> range f)" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2985 |
shows "\<exists>r. subseq r \<and> (f \<circ> r) ----> l" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2986 |
proof - |
| 55522 | 2987 |
from countable_basis_at_decseq[of l] |
2988 |
obtain A where A: |
|
2989 |
"\<And>i. open (A i)" |
|
2990 |
"\<And>i. l \<in> A i" |
|
2991 |
"\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
|
2992 |
by blast |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
2993 |
def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> f j \<in> A (Suc n)" |
| 52624 | 2994 |
{
|
2995 |
fix n i |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2996 |
have "infinite (A (Suc n) \<inter> range f - f`{.. i})"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2997 |
using l A by auto |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2998 |
then have "\<exists>x. x \<in> A (Suc n) \<inter> range f - f`{.. i}"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
2999 |
unfolding ex_in_conv by (intro notI) simp |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3000 |
then have "\<exists>j. f j \<in> A (Suc n) \<and> j \<notin> {.. i}"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3001 |
by auto |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3002 |
then have "\<exists>a. i < a \<and> f a \<in> A (Suc n)" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3003 |
by (auto simp: not_le) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3004 |
then have "i < s n i" "f (s n i) \<in> A (Suc n)" |
| 52624 | 3005 |
unfolding s_def by (auto intro: someI2_ex) |
3006 |
} |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3007 |
note s = this |
| 55415 | 3008 |
def r \<equiv> "rec_nat (s 0 0) s" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3009 |
have "subseq r" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3010 |
by (auto simp: r_def s subseq_Suc_iff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3011 |
moreover |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3012 |
have "(\<lambda>n. f (r n)) ----> l" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3013 |
proof (rule topological_tendstoI) |
| 53282 | 3014 |
fix S |
3015 |
assume "open S" "l \<in> S" |
|
| 53640 | 3016 |
with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
3017 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3018 |
moreover |
| 52624 | 3019 |
{
|
3020 |
fix i |
|
| 53282 | 3021 |
assume "Suc 0 \<le> i" |
3022 |
then have "f (r i) \<in> A i" |
|
| 52624 | 3023 |
by (cases i) (simp_all add: r_def s) |
3024 |
} |
|
3025 |
then have "eventually (\<lambda>i. f (r i) \<in> A i) sequentially" |
|
3026 |
by (auto simp: eventually_sequentially) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3027 |
ultimately show "eventually (\<lambda>i. f (r i) \<in> S) sequentially" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3028 |
by eventually_elim auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3029 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3030 |
ultimately show "\<exists>r. subseq r \<and> (f \<circ> r) ----> l" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3031 |
by (auto simp: convergent_def comp_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3032 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3033 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3034 |
lemma sequence_infinite_lemma: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3035 |
fixes f :: "nat \<Rightarrow> 'a::t1_space" |
| 53282 | 3036 |
assumes "\<forall>n. f n \<noteq> l" |
3037 |
and "(f ---> l) sequentially" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3038 |
shows "infinite (range f)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3039 |
proof |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3040 |
assume "finite (range f)" |
| 53640 | 3041 |
then have "closed (range f)" |
3042 |
by (rule finite_imp_closed) |
|
3043 |
then have "open (- range f)" |
|
3044 |
by (rule open_Compl) |
|
3045 |
from assms(1) have "l \<in> - range f" |
|
3046 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3047 |
from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially" |
| 53640 | 3048 |
using `open (- range f)` `l \<in> - range f` |
3049 |
by (rule topological_tendstoD) |
|
3050 |
then show False |
|
3051 |
unfolding eventually_sequentially |
|
3052 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3053 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3054 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3055 |
lemma closure_insert: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3056 |
fixes x :: "'a::t1_space" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3057 |
shows "closure (insert x s) = insert x (closure s)" |
| 52624 | 3058 |
apply (rule closure_unique) |
3059 |
apply (rule insert_mono [OF closure_subset]) |
|
3060 |
apply (rule closed_insert [OF closed_closure]) |
|
3061 |
apply (simp add: closure_minimal) |
|
3062 |
done |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3063 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3064 |
lemma islimpt_insert: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3065 |
fixes x :: "'a::t1_space" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3066 |
shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3067 |
proof |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3068 |
assume *: "x islimpt (insert a s)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3069 |
show "x islimpt s" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3070 |
proof (rule islimptI) |
| 53282 | 3071 |
fix t |
3072 |
assume t: "x \<in> t" "open t" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3073 |
show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3074 |
proof (cases "x = a") |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3075 |
case True |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3076 |
obtain y where "y \<in> insert a s" "y \<in> t" "y \<noteq> x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3077 |
using * t by (rule islimptE) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3078 |
with `x = a` show ?thesis by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3079 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3080 |
case False |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3081 |
with t have t': "x \<in> t - {a}" "open (t - {a})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3082 |
by (simp_all add: open_Diff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3083 |
obtain y where "y \<in> insert a s" "y \<in> t - {a}" "y \<noteq> x"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3084 |
using * t' by (rule islimptE) |
| 53282 | 3085 |
then show ?thesis by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3086 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3087 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3088 |
next |
| 53282 | 3089 |
assume "x islimpt s" |
3090 |
then show "x islimpt (insert a s)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3091 |
by (rule islimpt_subset) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3092 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3093 |
|
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3094 |
lemma islimpt_finite: |
|
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3095 |
fixes x :: "'a::t1_space" |
|
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3096 |
shows "finite s \<Longrightarrow> \<not> x islimpt s" |
| 52624 | 3097 |
by (induct set: finite) (simp_all add: islimpt_insert) |
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3098 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3099 |
lemma islimpt_union_finite: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3100 |
fixes x :: "'a::t1_space" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3101 |
shows "finite s \<Longrightarrow> x islimpt (s \<union> t) \<longleftrightarrow> x islimpt t" |
| 52624 | 3102 |
by (simp add: islimpt_Un islimpt_finite) |
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3103 |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3104 |
lemma islimpt_eq_acc_point: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3105 |
fixes l :: "'a :: t1_space" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3106 |
shows "l islimpt S \<longleftrightarrow> (\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S))" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3107 |
proof (safe intro!: islimptI) |
| 53282 | 3108 |
fix U |
3109 |
assume "l islimpt S" "l \<in> U" "open U" "finite (U \<inter> S)" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3110 |
then have "l islimpt S" "l \<in> (U - (U \<inter> S - {l}))" "open (U - (U \<inter> S - {l}))"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3111 |
by (auto intro: finite_imp_closed) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3112 |
then show False |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3113 |
by (rule islimptE) auto |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3114 |
next |
| 53282 | 3115 |
fix T |
3116 |
assume *: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S)" "l \<in> T" "open T" |
|
3117 |
then have "infinite (T \<inter> S - {l})"
|
|
3118 |
by auto |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3119 |
then have "\<exists>x. x \<in> (T \<inter> S - {l})"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3120 |
unfolding ex_in_conv by (intro notI) simp |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3121 |
then show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> l" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3122 |
by auto |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3123 |
qed |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3124 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3125 |
lemma islimpt_range_imp_convergent_subsequence: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3126 |
fixes l :: "'a :: {t1_space, first_countable_topology}"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3127 |
assumes l: "l islimpt (range f)" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3128 |
shows "\<exists>r. subseq r \<and> (f \<circ> r) ----> l" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3129 |
using l unfolding islimpt_eq_acc_point |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3130 |
by (rule acc_point_range_imp_convergent_subsequence) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3131 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3132 |
lemma sequence_unique_limpt: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3133 |
fixes f :: "nat \<Rightarrow> 'a::t2_space" |
| 53282 | 3134 |
assumes "(f ---> l) sequentially" |
3135 |
and "l' islimpt (range f)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3136 |
shows "l' = l" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3137 |
proof (rule ccontr) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3138 |
assume "l' \<noteq> l" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3139 |
obtain s t where "open s" "open t" "l' \<in> s" "l \<in> t" "s \<inter> t = {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3140 |
using hausdorff [OF `l' \<noteq> l`] by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3141 |
have "eventually (\<lambda>n. f n \<in> t) sequentially" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3142 |
using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3143 |
then obtain N where "\<forall>n\<ge>N. f n \<in> t" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3144 |
unfolding eventually_sequentially by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3145 |
|
| 53282 | 3146 |
have "UNIV = {..<N} \<union> {N..}"
|
3147 |
by auto |
|
3148 |
then have "l' islimpt (f ` ({..<N} \<union> {N..}))"
|
|
3149 |
using assms(2) by simp |
|
3150 |
then have "l' islimpt (f ` {..<N} \<union> f ` {N..})"
|
|
3151 |
by (simp add: image_Un) |
|
3152 |
then have "l' islimpt (f ` {N..})"
|
|
3153 |
by (simp add: islimpt_union_finite) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3154 |
then obtain y where "y \<in> f ` {N..}" "y \<in> s" "y \<noteq> l'"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3155 |
using `l' \<in> s` `open s` by (rule islimptE) |
| 53282 | 3156 |
then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'" |
3157 |
by auto |
|
3158 |
with `\<forall>n\<ge>N. f n \<in> t` have "f n \<in> s \<inter> t" |
|
3159 |
by simp |
|
3160 |
with `s \<inter> t = {}` show False
|
|
3161 |
by simp |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3162 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3163 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3164 |
lemma bolzano_weierstrass_imp_closed: |
| 53640 | 3165 |
fixes s :: "'a::{first_countable_topology,t2_space} set"
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3166 |
assumes "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3167 |
shows "closed s" |
| 52624 | 3168 |
proof - |
3169 |
{
|
|
3170 |
fix x l |
|
3171 |
assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially" |
|
| 53282 | 3172 |
then have "l \<in> s" |
| 52624 | 3173 |
proof (cases "\<forall>n. x n \<noteq> l") |
3174 |
case False |
|
| 53282 | 3175 |
then show "l\<in>s" using as(1) by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3176 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3177 |
case True note cas = this |
| 52624 | 3178 |
with as(2) have "infinite (range x)" |
3179 |
using sequence_infinite_lemma[of x l] by auto |
|
3180 |
then obtain l' where "l'\<in>s" "l' islimpt (range x)" |
|
3181 |
using assms[THEN spec[where x="range x"]] as(1) by auto |
|
| 53282 | 3182 |
then show "l\<in>s" using sequence_unique_limpt[of x l l'] |
| 52624 | 3183 |
using as cas by auto |
3184 |
qed |
|
3185 |
} |
|
| 53282 | 3186 |
then show ?thesis |
3187 |
unfolding closed_sequential_limits by fast |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3188 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3189 |
|
| 50944 | 3190 |
lemma compact_imp_bounded: |
| 52624 | 3191 |
assumes "compact U" |
3192 |
shows "bounded U" |
|
| 50944 | 3193 |
proof - |
| 52624 | 3194 |
have "compact U" "\<forall>x\<in>U. open (ball x 1)" "U \<subseteq> (\<Union>x\<in>U. ball x 1)" |
3195 |
using assms by auto |
|
| 50944 | 3196 |
then obtain D where D: "D \<subseteq> U" "finite D" "U \<subseteq> (\<Union>x\<in>D. ball x 1)" |
| 52624 | 3197 |
by (rule compactE_image) |
| 50955 | 3198 |
from `finite D` have "bounded (\<Union>x\<in>D. ball x 1)" |
3199 |
by (simp add: bounded_UN) |
|
| 53282 | 3200 |
then show "bounded U" using `U \<subseteq> (\<Union>x\<in>D. ball x 1)` |
| 50955 | 3201 |
by (rule bounded_subset) |
| 50944 | 3202 |
qed |
3203 |
||
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3204 |
text{* In particular, some common special cases. *}
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3205 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3206 |
lemma compact_union [intro]: |
| 53291 | 3207 |
assumes "compact s" |
3208 |
and "compact t" |
|
| 53282 | 3209 |
shows " compact (s \<union> t)" |
| 50898 | 3210 |
proof (rule compactI) |
| 52624 | 3211 |
fix f |
3212 |
assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3213 |
from * `compact s` obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'" |
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
55927
diff
changeset
|
3214 |
unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) |
| 52624 | 3215 |
moreover |
3216 |
from * `compact t` obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'" |
|
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
55927
diff
changeset
|
3217 |
unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3218 |
ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<union> t \<subseteq> \<Union>f'" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3219 |
by (auto intro!: exI[of _ "s' \<union> t'"]) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3220 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3221 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3222 |
lemma compact_Union [intro]: "finite S \<Longrightarrow> (\<And>T. T \<in> S \<Longrightarrow> compact T) \<Longrightarrow> compact (\<Union>S)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3223 |
by (induct set: finite) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3224 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3225 |
lemma compact_UN [intro]: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3226 |
"finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> compact (B x)) \<Longrightarrow> compact (\<Union>x\<in>A. B x)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3227 |
unfolding SUP_def by (rule compact_Union) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3228 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3229 |
lemma closed_inter_compact [intro]: |
| 53282 | 3230 |
assumes "closed s" |
3231 |
and "compact t" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3232 |
shows "compact (s \<inter> t)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3233 |
using compact_inter_closed [of t s] assms |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3234 |
by (simp add: Int_commute) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3235 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3236 |
lemma compact_inter [intro]: |
| 50898 | 3237 |
fixes s t :: "'a :: t2_space set" |
| 53282 | 3238 |
assumes "compact s" |
3239 |
and "compact t" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3240 |
shows "compact (s \<inter> t)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3241 |
using assms by (intro compact_inter_closed compact_imp_closed) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3242 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3243 |
lemma compact_sing [simp]: "compact {a}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3244 |
unfolding compact_eq_heine_borel by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3245 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3246 |
lemma compact_insert [simp]: |
| 53282 | 3247 |
assumes "compact s" |
3248 |
shows "compact (insert x s)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3249 |
proof - |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3250 |
have "compact ({x} \<union> s)"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3251 |
using compact_sing assms by (rule compact_union) |
| 53282 | 3252 |
then show ?thesis by simp |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3253 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3254 |
|
| 52624 | 3255 |
lemma finite_imp_compact: "finite s \<Longrightarrow> compact s" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3256 |
by (induct set: finite) simp_all |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3257 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3258 |
lemma open_delete: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3259 |
fixes s :: "'a::t1_space set" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3260 |
shows "open s \<Longrightarrow> open (s - {x})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3261 |
by (simp add: open_Diff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3262 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3263 |
text{*Compactness expressed with filters*}
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3264 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3265 |
definition "filter_from_subbase B = Abs_filter (\<lambda>P. \<exists>X \<subseteq> B. finite X \<and> Inf X \<le> P)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3266 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3267 |
lemma eventually_filter_from_subbase: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3268 |
"eventually P (filter_from_subbase B) \<longleftrightarrow> (\<exists>X \<subseteq> B. finite X \<and> Inf X \<le> P)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3269 |
(is "_ \<longleftrightarrow> ?R P") |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3270 |
unfolding filter_from_subbase_def |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3271 |
proof (rule eventually_Abs_filter is_filter.intro)+ |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3272 |
show "?R (\<lambda>x. True)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3273 |
by (rule exI[of _ "{}"]) (simp add: le_fun_def)
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3274 |
next |
| 55522 | 3275 |
fix P Q |
3276 |
assume "?R P" then guess X .. |
|
3277 |
moreover |
|
3278 |
assume "?R Q" then guess Y .. |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3279 |
ultimately show "?R (\<lambda>x. P x \<and> Q x)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3280 |
by (intro exI[of _ "X \<union> Y"]) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3281 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3282 |
fix P Q |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3283 |
assume "?R P" then guess X .. |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3284 |
moreover assume "\<forall>x. P x \<longrightarrow> Q x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3285 |
ultimately show "?R Q" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3286 |
by (intro exI[of _ X]) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3287 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3288 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3289 |
lemma eventually_filter_from_subbaseI: "P \<in> B \<Longrightarrow> eventually P (filter_from_subbase B)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3290 |
by (subst eventually_filter_from_subbase) (auto intro!: exI[of _ "{P}"])
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3291 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3292 |
lemma filter_from_subbase_not_bot: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3293 |
"\<forall>X \<subseteq> B. finite X \<longrightarrow> Inf X \<noteq> bot \<Longrightarrow> filter_from_subbase B \<noteq> bot" |
|
56742
678a52e676b6
more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents:
56544
diff
changeset
|
3294 |
unfolding trivial_limit_def eventually_filter_from_subbase |
|
678a52e676b6
more complete classical rules for Inf and Sup, modelled after theiry counterparts on Inter and Union (and INF and SUP)
haftmann
parents:
56544
diff
changeset
|
3295 |
bot_bool_def [symmetric] bot_fun_def [symmetric] bot_unique by simp |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3296 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3297 |
lemma closure_iff_nhds_not_empty: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3298 |
"x \<in> closure X \<longleftrightarrow> (\<forall>A. \<forall>S\<subseteq>A. open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3299 |
proof safe |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3300 |
assume x: "x \<in> closure X" |
| 53282 | 3301 |
fix S A |
3302 |
assume "open S" "x \<in> S" "X \<inter> A = {}" "S \<subseteq> A"
|
|
3303 |
then have "x \<notin> closure (-S)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3304 |
by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3305 |
with x have "x \<in> closure X - closure (-S)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3306 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3307 |
also have "\<dots> \<subseteq> closure (X \<inter> S)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3308 |
using `open S` open_inter_closure_subset[of S X] by (simp add: closed_Compl ac_simps) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3309 |
finally have "X \<inter> S \<noteq> {}" by auto
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3310 |
then show False using `X \<inter> A = {}` `S \<subseteq> A` by auto
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3311 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3312 |
assume "\<forall>A S. S \<subseteq> A \<longrightarrow> open S \<longrightarrow> x \<in> S \<longrightarrow> X \<inter> A \<noteq> {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3313 |
from this[THEN spec, of "- X", THEN spec, of "- closure X"] |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3314 |
show "x \<in> closure X" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3315 |
by (simp add: closure_subset open_Compl) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3316 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3317 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3318 |
lemma compact_filter: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3319 |
"compact U \<longleftrightarrow> (\<forall>F. F \<noteq> bot \<longrightarrow> eventually (\<lambda>x. x \<in> U) F \<longrightarrow> (\<exists>x\<in>U. inf (nhds x) F \<noteq> bot))" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3320 |
proof (intro allI iffI impI compact_fip[THEN iffD2] notI) |
| 53282 | 3321 |
fix F |
3322 |
assume "compact U" |
|
3323 |
assume F: "F \<noteq> bot" "eventually (\<lambda>x. x \<in> U) F" |
|
3324 |
then have "U \<noteq> {}"
|
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3325 |
by (auto simp: eventually_False) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3326 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3327 |
def Z \<equiv> "closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3328 |
then have "\<forall>z\<in>Z. closed z" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3329 |
by auto |
| 53282 | 3330 |
moreover |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3331 |
have ev_Z: "\<And>z. z \<in> Z \<Longrightarrow> eventually (\<lambda>x. x \<in> z) F" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3332 |
unfolding Z_def by (auto elim: eventually_elim1 intro: set_mp[OF closure_subset]) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3333 |
have "(\<forall>B \<subseteq> Z. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3334 |
proof (intro allI impI) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3335 |
fix B assume "finite B" "B \<subseteq> Z" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3336 |
with `finite B` ev_Z have "eventually (\<lambda>x. \<forall>b\<in>B. x \<in> b) F" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3337 |
by (auto intro!: eventually_Ball_finite) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3338 |
with F(2) have "eventually (\<lambda>x. x \<in> U \<inter> (\<Inter>B)) F" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3339 |
by eventually_elim auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3340 |
with F show "U \<inter> \<Inter>B \<noteq> {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3341 |
by (intro notI) (simp add: eventually_False) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3342 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3343 |
ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3344 |
using `compact U` unfolding compact_fip by blast |
| 53282 | 3345 |
then obtain x where "x \<in> U" and x: "\<And>z. z \<in> Z \<Longrightarrow> x \<in> z" |
3346 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3347 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3348 |
have "\<And>P. eventually P (inf (nhds x) F) \<Longrightarrow> P \<noteq> bot" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3349 |
unfolding eventually_inf eventually_nhds |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3350 |
proof safe |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3351 |
fix P Q R S |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3352 |
assume "eventually R F" "open S" "x \<in> S" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3353 |
with open_inter_closure_eq_empty[of S "{x. R x}"] x[of "closure {x. R x}"]
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3354 |
have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3355 |
moreover assume "Ball S Q" "\<forall>x. Q x \<and> R x \<longrightarrow> bot x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3356 |
ultimately show False by (auto simp: set_eq_iff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3357 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3358 |
with `x \<in> U` show "\<exists>x\<in>U. inf (nhds x) F \<noteq> bot" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3359 |
by (metis eventually_bot) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3360 |
next |
| 53282 | 3361 |
fix A |
3362 |
assume A: "\<forall>a\<in>A. closed a" "\<forall>B\<subseteq>A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}" "U \<inter> \<Inter>A = {}"
|
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3363 |
def P' \<equiv> "(\<lambda>a (x::'a). x \<in> a)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3364 |
then have inj_P': "\<And>A. inj_on P' A" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3365 |
by (auto intro!: inj_onI simp: fun_eq_iff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3366 |
def F \<equiv> "filter_from_subbase (P' ` insert U A)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3367 |
have "F \<noteq> bot" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3368 |
unfolding F_def |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3369 |
proof (safe intro!: filter_from_subbase_not_bot) |
| 53282 | 3370 |
fix X |
3371 |
assume "X \<subseteq> P' ` insert U A" "finite X" "Inf X = bot" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3372 |
then obtain B where "B \<subseteq> insert U A" "finite B" and B: "Inf (P' ` B) = bot" |
| 56166 | 3373 |
unfolding subset_image_iff by (auto intro: inj_P' finite_imageD simp del: Inf_image_eq) |
| 53282 | 3374 |
with A(2)[THEN spec, of "B - {U}"] have "U \<inter> \<Inter>(B - {U}) \<noteq> {}"
|
3375 |
by auto |
|
3376 |
with B show False |
|
3377 |
by (auto simp: P'_def fun_eq_iff) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3378 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3379 |
moreover have "eventually (\<lambda>x. x \<in> U) F" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3380 |
unfolding F_def by (rule eventually_filter_from_subbaseI) (auto simp: P'_def) |
| 53282 | 3381 |
moreover |
3382 |
assume "\<forall>F. F \<noteq> bot \<longrightarrow> eventually (\<lambda>x. x \<in> U) F \<longrightarrow> (\<exists>x\<in>U. inf (nhds x) F \<noteq> bot)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3383 |
ultimately obtain x where "x \<in> U" and x: "inf (nhds x) F \<noteq> bot" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3384 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3385 |
|
| 53282 | 3386 |
{
|
3387 |
fix V |
|
3388 |
assume "V \<in> A" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3389 |
then have V: "eventually (\<lambda>x. x \<in> V) F" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3390 |
by (auto simp add: F_def image_iff P'_def intro!: eventually_filter_from_subbaseI) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3391 |
have "x \<in> closure V" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3392 |
unfolding closure_iff_nhds_not_empty |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3393 |
proof (intro impI allI) |
| 53282 | 3394 |
fix S A |
3395 |
assume "open S" "x \<in> S" "S \<subseteq> A" |
|
3396 |
then have "eventually (\<lambda>x. x \<in> A) (nhds x)" |
|
3397 |
by (auto simp: eventually_nhds) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3398 |
with V have "eventually (\<lambda>x. x \<in> V \<inter> A) (inf (nhds x) F)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3399 |
by (auto simp: eventually_inf) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3400 |
with x show "V \<inter> A \<noteq> {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3401 |
by (auto simp del: Int_iff simp add: trivial_limit_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3402 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3403 |
then have "x \<in> V" |
| 53282 | 3404 |
using `V \<in> A` A(1) by simp |
3405 |
} |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3406 |
with `x\<in>U` have "x \<in> U \<inter> \<Inter>A" by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3407 |
with `U \<inter> \<Inter>A = {}` show False by auto
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3408 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3409 |
|
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3410 |
definition "countably_compact U \<longleftrightarrow> |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3411 |
(\<forall>A. countable A \<longrightarrow> (\<forall>a\<in>A. open a) \<longrightarrow> U \<subseteq> \<Union>A \<longrightarrow> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T))" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3412 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3413 |
lemma countably_compactE: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3414 |
assumes "countably_compact s" and "\<forall>t\<in>C. open t" and "s \<subseteq> \<Union>C" "countable C" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3415 |
obtains C' where "C' \<subseteq> C" and "finite C'" and "s \<subseteq> \<Union>C'" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3416 |
using assms unfolding countably_compact_def by metis |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3417 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3418 |
lemma countably_compactI: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3419 |
assumes "\<And>C. \<forall>t\<in>C. open t \<Longrightarrow> s \<subseteq> \<Union>C \<Longrightarrow> countable C \<Longrightarrow> (\<exists>C'\<subseteq>C. finite C' \<and> s \<subseteq> \<Union>C')" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3420 |
shows "countably_compact s" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3421 |
using assms unfolding countably_compact_def by metis |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3422 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3423 |
lemma compact_imp_countably_compact: "compact U \<Longrightarrow> countably_compact U" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3424 |
by (auto simp: compact_eq_heine_borel countably_compact_def) |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3425 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3426 |
lemma countably_compact_imp_compact: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3427 |
assumes "countably_compact U" |
| 53282 | 3428 |
and ccover: "countable B" "\<forall>b\<in>B. open b" |
3429 |
and basis: "\<And>T x. open T \<Longrightarrow> x \<in> T \<Longrightarrow> x \<in> U \<Longrightarrow> \<exists>b\<in>B. x \<in> b \<and> b \<inter> U \<subseteq> T" |
|
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3430 |
shows "compact U" |
| 53282 | 3431 |
using `countably_compact U` |
3432 |
unfolding compact_eq_heine_borel countably_compact_def |
|
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3433 |
proof safe |
| 53282 | 3434 |
fix A |
3435 |
assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3436 |
assume *: "\<forall>A. countable A \<longrightarrow> (\<forall>a\<in>A. open a) \<longrightarrow> U \<subseteq> \<Union>A \<longrightarrow> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3437 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3438 |
moreover def C \<equiv> "{b\<in>B. \<exists>a\<in>A. b \<inter> U \<subseteq> a}"
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3439 |
ultimately have "countable C" "\<forall>a\<in>C. open a" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3440 |
unfolding C_def using ccover by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3441 |
moreover |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3442 |
have "\<Union>A \<inter> U \<subseteq> \<Union>C" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3443 |
proof safe |
| 53282 | 3444 |
fix x a |
3445 |
assume "x \<in> U" "x \<in> a" "a \<in> A" |
|
3446 |
with basis[of a x] A obtain b where "b \<in> B" "x \<in> b" "b \<inter> U \<subseteq> a" |
|
3447 |
by blast |
|
3448 |
with `a \<in> A` show "x \<in> \<Union>C" |
|
3449 |
unfolding C_def by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3450 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3451 |
then have "U \<subseteq> \<Union>C" using `U \<subseteq> \<Union>A` by auto |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
3452 |
ultimately obtain T where T: "T\<subseteq>C" "finite T" "U \<subseteq> \<Union>T" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3453 |
using * by metis |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
3454 |
then have "\<forall>t\<in>T. \<exists>a\<in>A. t \<inter> U \<subseteq> a" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3455 |
by (auto simp: C_def) |
| 55522 | 3456 |
then obtain f where "\<forall>t\<in>T. f t \<in> A \<and> t \<inter> U \<subseteq> f t" |
3457 |
unfolding bchoice_iff Bex_def .. |
|
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
3458 |
with T show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3459 |
unfolding C_def by (intro exI[of _ "f`T"]) fastforce |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3460 |
qed |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3461 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3462 |
lemma countably_compact_imp_compact_second_countable: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3463 |
"countably_compact U \<Longrightarrow> compact (U :: 'a :: second_countable_topology set)" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3464 |
proof (rule countably_compact_imp_compact) |
| 53282 | 3465 |
fix T and x :: 'a |
3466 |
assume "open T" "x \<in> T" |
|
| 55522 | 3467 |
from topological_basisE[OF is_basis this] obtain b where |
3468 |
"b \<in> (SOME B. countable B \<and> topological_basis B)" "x \<in> b" "b \<subseteq> T" . |
|
| 53282 | 3469 |
then show "\<exists>b\<in>SOME B. countable B \<and> topological_basis B. x \<in> b \<and> b \<inter> U \<subseteq> T" |
| 55522 | 3470 |
by blast |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3471 |
qed (insert countable_basis topological_basis_open[OF is_basis], auto) |
| 36437 | 3472 |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3473 |
lemma countably_compact_eq_compact: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3474 |
"countably_compact U \<longleftrightarrow> compact (U :: 'a :: second_countable_topology set)" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3475 |
using countably_compact_imp_compact_second_countable compact_imp_countably_compact by blast |
| 53282 | 3476 |
|
| 36437 | 3477 |
subsubsection{* Sequential compactness *}
|
| 33175 | 3478 |
|
| 53282 | 3479 |
definition seq_compact :: "'a::topological_space set \<Rightarrow> bool" |
3480 |
where "seq_compact S \<longleftrightarrow> |
|
| 53640 | 3481 |
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially))" |
| 33175 | 3482 |
|
| 54070 | 3483 |
lemma seq_compactI: |
3484 |
assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
|
3485 |
shows "seq_compact S" |
|
3486 |
unfolding seq_compact_def using assms by fast |
|
3487 |
||
3488 |
lemma seq_compactE: |
|
3489 |
assumes "seq_compact S" "\<forall>n. f n \<in> S" |
|
3490 |
obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially" |
|
3491 |
using assms unfolding seq_compact_def by fast |
|
3492 |
||
3493 |
lemma closed_sequentially: (* TODO: move upwards *) |
|
3494 |
assumes "closed s" and "\<forall>n. f n \<in> s" and "f ----> l" |
|
3495 |
shows "l \<in> s" |
|
3496 |
proof (rule ccontr) |
|
3497 |
assume "l \<notin> s" |
|
3498 |
with `closed s` and `f ----> l` have "eventually (\<lambda>n. f n \<in> - s) sequentially" |
|
3499 |
by (fast intro: topological_tendstoD) |
|
3500 |
with `\<forall>n. f n \<in> s` show "False" |
|
3501 |
by simp |
|
3502 |
qed |
|
3503 |
||
3504 |
lemma seq_compact_inter_closed: |
|
3505 |
assumes "seq_compact s" and "closed t" |
|
3506 |
shows "seq_compact (s \<inter> t)" |
|
3507 |
proof (rule seq_compactI) |
|
3508 |
fix f assume "\<forall>n::nat. f n \<in> s \<inter> t" |
|
3509 |
hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t" |
|
3510 |
by simp_all |
|
3511 |
from `seq_compact s` and `\<forall>n. f n \<in> s` |
|
3512 |
obtain l r where "l \<in> s" and r: "subseq r" and l: "(f \<circ> r) ----> l" |
|
3513 |
by (rule seq_compactE) |
|
3514 |
from `\<forall>n. f n \<in> t` have "\<forall>n. (f \<circ> r) n \<in> t" |
|
3515 |
by simp |
|
3516 |
from `closed t` and this and l have "l \<in> t" |
|
3517 |
by (rule closed_sequentially) |
|
3518 |
with `l \<in> s` and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> (f \<circ> r) ----> l" |
|
3519 |
by fast |
|
3520 |
qed |
|
3521 |
||
3522 |
lemma seq_compact_closed_subset: |
|
3523 |
assumes "closed s" and "s \<subseteq> t" and "seq_compact t" |
|
3524 |
shows "seq_compact s" |
|
3525 |
using assms seq_compact_inter_closed [of t s] by (simp add: Int_absorb1) |
|
3526 |
||
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3527 |
lemma seq_compact_imp_countably_compact: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3528 |
fixes U :: "'a :: first_countable_topology set" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3529 |
assumes "seq_compact U" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3530 |
shows "countably_compact U" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3531 |
proof (safe intro!: countably_compactI) |
| 52624 | 3532 |
fix A |
3533 |
assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A" "countable A" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3534 |
have subseq: "\<And>X. range X \<subseteq> U \<Longrightarrow> \<exists>r x. x \<in> U \<and> subseq r \<and> (X \<circ> r) ----> x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3535 |
using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3536 |
show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3537 |
proof cases |
| 52624 | 3538 |
assume "finite A" |
3539 |
with A show ?thesis by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3540 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3541 |
assume "infinite A" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3542 |
then have "A \<noteq> {}" by auto
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3543 |
show ?thesis |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3544 |
proof (rule ccontr) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3545 |
assume "\<not> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)" |
| 53282 | 3546 |
then have "\<forall>T. \<exists>x. T \<subseteq> A \<and> finite T \<longrightarrow> (x \<in> U - \<Union>T)" |
3547 |
by auto |
|
3548 |
then obtain X' where T: "\<And>T. T \<subseteq> A \<Longrightarrow> finite T \<Longrightarrow> X' T \<in> U - \<Union>T" |
|
3549 |
by metis |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3550 |
def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3551 |
have X: "\<And>n. X n \<in> U - (\<Union>i\<le>n. from_nat_into A i)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3552 |
using `A \<noteq> {}` unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into)
|
| 53282 | 3553 |
then have "range X \<subseteq> U" |
3554 |
by auto |
|
3555 |
with subseq[of X] obtain r x where "x \<in> U" and r: "subseq r" "(X \<circ> r) ----> x" |
|
3556 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3557 |
from `x\<in>U` `U \<subseteq> \<Union>A` from_nat_into_surj[OF `countable A`] |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3558 |
obtain n where "x \<in> from_nat_into A n" by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3559 |
with r(2) A(1) from_nat_into[OF `A \<noteq> {}`, of n]
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3560 |
have "eventually (\<lambda>i. X (r i) \<in> from_nat_into A n) sequentially" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3561 |
unfolding tendsto_def by (auto simp: comp_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3562 |
then obtain N where "\<And>i. N \<le> i \<Longrightarrow> X (r i) \<in> from_nat_into A n" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3563 |
by (auto simp: eventually_sequentially) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3564 |
moreover from X have "\<And>i. n \<le> r i \<Longrightarrow> X (r i) \<notin> from_nat_into A n" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3565 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3566 |
moreover from `subseq r`[THEN seq_suble, of "max n N"] have "\<exists>i. n \<le> r i \<and> N \<le> i" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3567 |
by (auto intro!: exI[of _ "max n N"]) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3568 |
ultimately show False |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3569 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3570 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3571 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3572 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3573 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3574 |
lemma compact_imp_seq_compact: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3575 |
fixes U :: "'a :: first_countable_topology set" |
| 53282 | 3576 |
assumes "compact U" |
3577 |
shows "seq_compact U" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3578 |
unfolding seq_compact_def |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3579 |
proof safe |
| 52624 | 3580 |
fix X :: "nat \<Rightarrow> 'a" |
3581 |
assume "\<forall>n. X n \<in> U" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3582 |
then have "eventually (\<lambda>x. x \<in> U) (filtermap X sequentially)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3583 |
by (auto simp: eventually_filtermap) |
| 52624 | 3584 |
moreover |
3585 |
have "filtermap X sequentially \<noteq> bot" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3586 |
by (simp add: trivial_limit_def eventually_filtermap) |
| 52624 | 3587 |
ultimately |
3588 |
obtain x where "x \<in> U" and x: "inf (nhds x) (filtermap X sequentially) \<noteq> bot" (is "?F \<noteq> _") |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3589 |
using `compact U` by (auto simp: compact_filter) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3590 |
|
| 55522 | 3591 |
from countable_basis_at_decseq[of x] |
3592 |
obtain A where A: |
|
3593 |
"\<And>i. open (A i)" |
|
3594 |
"\<And>i. x \<in> A i" |
|
3595 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
|
3596 |
by blast |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3597 |
def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> X j \<in> A (Suc n)" |
| 52624 | 3598 |
{
|
3599 |
fix n i |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3600 |
have "\<exists>a. i < a \<and> X a \<in> A (Suc n)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3601 |
proof (rule ccontr) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3602 |
assume "\<not> (\<exists>a>i. X a \<in> A (Suc n))" |
| 53282 | 3603 |
then have "\<And>a. Suc i \<le> a \<Longrightarrow> X a \<notin> A (Suc n)" |
3604 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3605 |
then have "eventually (\<lambda>x. x \<notin> A (Suc n)) (filtermap X sequentially)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3606 |
by (auto simp: eventually_filtermap eventually_sequentially) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3607 |
moreover have "eventually (\<lambda>x. x \<in> A (Suc n)) (nhds x)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3608 |
using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3609 |
ultimately have "eventually (\<lambda>x. False) ?F" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3610 |
by (auto simp add: eventually_inf) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3611 |
with x show False |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3612 |
by (simp add: eventually_False) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3613 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3614 |
then have "i < s n i" "X (s n i) \<in> A (Suc n)" |
| 52624 | 3615 |
unfolding s_def by (auto intro: someI2_ex) |
3616 |
} |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3617 |
note s = this |
| 55415 | 3618 |
def r \<equiv> "rec_nat (s 0 0) s" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3619 |
have "subseq r" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3620 |
by (auto simp: r_def s subseq_Suc_iff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3621 |
moreover |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3622 |
have "(\<lambda>n. X (r n)) ----> x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3623 |
proof (rule topological_tendstoI) |
| 52624 | 3624 |
fix S |
3625 |
assume "open S" "x \<in> S" |
|
| 53282 | 3626 |
with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
3627 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3628 |
moreover |
| 52624 | 3629 |
{
|
3630 |
fix i |
|
3631 |
assume "Suc 0 \<le> i" |
|
3632 |
then have "X (r i) \<in> A i" |
|
3633 |
by (cases i) (simp_all add: r_def s) |
|
3634 |
} |
|
3635 |
then have "eventually (\<lambda>i. X (r i) \<in> A i) sequentially" |
|
3636 |
by (auto simp: eventually_sequentially) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3637 |
ultimately show "eventually (\<lambda>i. X (r i) \<in> S) sequentially" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3638 |
by eventually_elim auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3639 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3640 |
ultimately show "\<exists>x \<in> U. \<exists>r. subseq r \<and> (X \<circ> r) ----> x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3641 |
using `x \<in> U` by (auto simp: convergent_def comp_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3642 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3643 |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3644 |
lemma countably_compact_imp_acc_point: |
| 53291 | 3645 |
assumes "countably_compact s" |
3646 |
and "countable t" |
|
3647 |
and "infinite t" |
|
3648 |
and "t \<subseteq> s" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3649 |
shows "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3650 |
proof (rule ccontr) |
| 53282 | 3651 |
def C \<equiv> "(\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3652 |
note `countably_compact s` |
| 53282 | 3653 |
moreover have "\<forall>t\<in>C. open t" |
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3654 |
by (auto simp: C_def) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3655 |
moreover |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3656 |
assume "\<not> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3657 |
then have s: "\<And>x. x \<in> s \<Longrightarrow> \<exists>U. x\<in>U \<and> open U \<and> finite (U \<inter> t)" by metis |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3658 |
have "s \<subseteq> \<Union>C" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3659 |
using `t \<subseteq> s` |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3660 |
unfolding C_def Union_image_eq |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3661 |
apply (safe dest!: s) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3662 |
apply (rule_tac a="U \<inter> t" in UN_I) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3663 |
apply (auto intro!: interiorI simp add: finite_subset) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3664 |
done |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3665 |
moreover |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3666 |
from `countable t` have "countable C" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3667 |
unfolding C_def by (auto intro: countable_Collect_finite_subset) |
| 55522 | 3668 |
ultimately |
3669 |
obtain D where "D \<subseteq> C" "finite D" "s \<subseteq> \<Union>D" |
|
3670 |
by (rule countably_compactE) |
|
| 53282 | 3671 |
then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E"
|
3672 |
and s: "s \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- t)))" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3673 |
by (metis (lifting) Union_image_eq finite_subset_image C_def) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3674 |
from s `t \<subseteq> s` have "t \<subseteq> \<Union>E" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3675 |
using interior_subset by blast |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3676 |
moreover have "finite (\<Union>E)" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3677 |
using E by auto |
| 53282 | 3678 |
ultimately show False using `infinite t` |
3679 |
by (auto simp: finite_subset) |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3680 |
qed |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3681 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3682 |
lemma countable_acc_point_imp_seq_compact: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3683 |
fixes s :: "'a::first_countable_topology set" |
| 53291 | 3684 |
assumes "\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s \<longrightarrow> |
3685 |
(\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t))" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3686 |
shows "seq_compact s" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3687 |
proof - |
| 52624 | 3688 |
{
|
3689 |
fix f :: "nat \<Rightarrow> 'a" |
|
3690 |
assume f: "\<forall>n. f n \<in> s" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3691 |
have "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3692 |
proof (cases "finite (range f)") |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3693 |
case True |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3694 |
obtain l where "infinite {n. f n = f l}"
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3695 |
using pigeonhole_infinite[OF _ True] by auto |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3696 |
then obtain r where "subseq r" and fr: "\<forall>n. f (r n) = f l" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3697 |
using infinite_enumerate by blast |
| 53282 | 3698 |
then have "subseq r \<and> (f \<circ> r) ----> f l" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3699 |
by (simp add: fr tendsto_const o_def) |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3700 |
with f show "\<exists>l\<in>s. \<exists>r. subseq r \<and> (f \<circ> r) ----> l" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
3701 |
by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3702 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3703 |
case False |
| 53282 | 3704 |
with f assms have "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" |
3705 |
by auto |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3706 |
then obtain l where "l \<in> s" "\<forall>U. l\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" .. |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3707 |
from this(2) have "\<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3708 |
using acc_point_range_imp_convergent_subsequence[of l f] by auto |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3709 |
with `l \<in> s` show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" .. |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3710 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3711 |
} |
| 53282 | 3712 |
then show ?thesis |
3713 |
unfolding seq_compact_def by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3714 |
qed |
| 44075 | 3715 |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3716 |
lemma seq_compact_eq_countably_compact: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3717 |
fixes U :: "'a :: first_countable_topology set" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3718 |
shows "seq_compact U \<longleftrightarrow> countably_compact U" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3719 |
using |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3720 |
countable_acc_point_imp_seq_compact |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3721 |
countably_compact_imp_acc_point |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3722 |
seq_compact_imp_countably_compact |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3723 |
by metis |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3724 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3725 |
lemma seq_compact_eq_acc_point: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3726 |
fixes s :: "'a :: first_countable_topology set" |
| 53291 | 3727 |
shows "seq_compact s \<longleftrightarrow> |
3728 |
(\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s --> (\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> t)))" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3729 |
using |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3730 |
countable_acc_point_imp_seq_compact[of s] |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3731 |
countably_compact_imp_acc_point[of s] |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3732 |
seq_compact_imp_countably_compact[of s] |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3733 |
by metis |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3734 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3735 |
lemma seq_compact_eq_compact: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3736 |
fixes U :: "'a :: second_countable_topology set" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3737 |
shows "seq_compact U \<longleftrightarrow> compact U" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3738 |
using seq_compact_eq_countably_compact countably_compact_eq_compact by blast |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3739 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3740 |
lemma bolzano_weierstrass_imp_seq_compact: |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3741 |
fixes s :: "'a::{t1_space, first_countable_topology} set"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3742 |
shows "\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> seq_compact s" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3743 |
by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3744 |
|
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3745 |
subsubsection{* Total boundedness *}
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3746 |
|
| 53282 | 3747 |
lemma cauchy_def: "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N --> dist(s m)(s n) < e)" |
| 52624 | 3748 |
unfolding Cauchy_def by metis |
3749 |
||
| 53282 | 3750 |
fun helper_1 :: "('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a"
|
3751 |
where |
|
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3752 |
"helper_1 s e n = (SOME y::'a. y \<in> s \<and> (\<forall>m<n. \<not> (dist (helper_1 s e m) y < e)))" |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3753 |
declare helper_1.simps[simp del] |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3754 |
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3755 |
lemma seq_compact_imp_totally_bounded: |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3756 |
assumes "seq_compact s" |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3757 |
shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))" |
| 52624 | 3758 |
proof (rule, rule, rule ccontr) |
3759 |
fix e::real |
|
| 53282 | 3760 |
assume "e > 0" |
3761 |
assume assm: "\<not> (\<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> \<Union>((\<lambda>x. ball x e) ` k))" |
|
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3762 |
def x \<equiv> "helper_1 s e" |
| 52624 | 3763 |
{
|
3764 |
fix n |
|
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3765 |
have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" |
| 52624 | 3766 |
proof (induct n rule: nat_less_induct) |
3767 |
fix n |
|
3768 |
def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))" |
|
3769 |
assume as: "\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)" |
|
3770 |
have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)"
|
|
3771 |
using assm |
|
3772 |
apply simp |
|
3773 |
apply (erule_tac x="x ` {0 ..< n}" in allE)
|
|
3774 |
using as |
|
3775 |
apply auto |
|
3776 |
done |
|
3777 |
then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)"
|
|
3778 |
unfolding subset_eq by auto |
|
3779 |
have "Q (x n)" |
|
3780 |
unfolding x_def and helper_1.simps[of s e n] |
|
3781 |
apply (rule someI2[where a=z]) |
|
3782 |
unfolding x_def[symmetric] and Q_def |
|
3783 |
using z |
|
3784 |
apply auto |
|
3785 |
done |
|
| 53282 | 3786 |
then show "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)" |
| 52624 | 3787 |
unfolding Q_def by auto |
3788 |
qed |
|
3789 |
} |
|
| 53282 | 3790 |
then have "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)" |
| 52624 | 3791 |
by blast+ |
3792 |
then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially" |
|
3793 |
using assms(1)[unfolded seq_compact_def, THEN spec[where x=x]] by auto |
|
3794 |
from this(3) have "Cauchy (x \<circ> r)" |
|
3795 |
using LIMSEQ_imp_Cauchy by auto |
|
3796 |
then obtain N::nat where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e" |
|
3797 |
unfolding cauchy_def using `e>0` by auto |
|
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3798 |
show False |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3799 |
using N[THEN spec[where x=N], THEN spec[where x="N+1"]] |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3800 |
using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]] |
| 52624 | 3801 |
using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]] |
3802 |
by auto |
|
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3803 |
qed |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3804 |
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3805 |
subsubsection{* Heine-Borel theorem *}
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3806 |
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3807 |
lemma seq_compact_imp_heine_borel: |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3808 |
fixes s :: "'a :: metric_space set" |
| 53282 | 3809 |
assumes "seq_compact s" |
3810 |
shows "compact s" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3811 |
proof - |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3812 |
from seq_compact_imp_totally_bounded[OF `seq_compact s`] |
| 55522 | 3813 |
obtain f where f: "\<forall>e>0. finite (f e) \<and> f e \<subseteq> s \<and> s \<subseteq> \<Union>((\<lambda>x. ball x e) ` f e)" |
3814 |
unfolding choice_iff' .. |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3815 |
def K \<equiv> "(\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3816 |
have "countably_compact s" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3817 |
using `seq_compact s` by (rule seq_compact_imp_countably_compact) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3818 |
then show "compact s" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3819 |
proof (rule countably_compact_imp_compact) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3820 |
show "countable K" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3821 |
unfolding K_def using f |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3822 |
by (auto intro: countable_finite countable_subset countable_rat |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3823 |
intro!: countable_image countable_SIGMA countable_UN) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3824 |
show "\<forall>b\<in>K. open b" by (auto simp: K_def) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3825 |
next |
| 53282 | 3826 |
fix T x |
3827 |
assume T: "open T" "x \<in> T" and x: "x \<in> s" |
|
3828 |
from openE[OF T] obtain e where "0 < e" "ball x e \<subseteq> T" |
|
3829 |
by auto |
|
3830 |
then have "0 < e / 2" "ball x (e / 2) \<subseteq> T" |
|
3831 |
by auto |
|
3832 |
from Rats_dense_in_real[OF `0 < e / 2`] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2" |
|
3833 |
by auto |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3834 |
from f[rule_format, of r] `0 < r` `x \<in> s` obtain k where "k \<in> f r" "x \<in> ball k r" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3835 |
unfolding Union_image_eq by auto |
| 53282 | 3836 |
from `r \<in> \<rat>` `0 < r` `k \<in> f r` have "ball k r \<in> K" |
3837 |
by (auto simp: K_def) |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3838 |
then show "\<exists>b\<in>K. x \<in> b \<and> b \<inter> s \<subseteq> T" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3839 |
proof (rule bexI[rotated], safe) |
| 53282 | 3840 |
fix y |
3841 |
assume "y \<in> ball k r" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3842 |
with `r < e / 2` `x \<in> ball k r` have "dist x y < e" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3843 |
by (intro dist_double[where x = k and d=e]) (auto simp: dist_commute) |
| 53282 | 3844 |
with `ball x e \<subseteq> T` show "y \<in> T" |
3845 |
by auto |
|
3846 |
next |
|
3847 |
show "x \<in> ball k r" by fact |
|
3848 |
qed |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
3849 |
qed |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3850 |
qed |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3851 |
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3852 |
lemma compact_eq_seq_compact_metric: |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3853 |
"compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s" |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3854 |
using compact_imp_seq_compact seq_compact_imp_heine_borel by blast |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3855 |
|
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3856 |
lemma compact_def: |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3857 |
"compact (S :: 'a::metric_space set) \<longleftrightarrow> |
| 53640 | 3858 |
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> (f \<circ> r) ----> l))" |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3859 |
unfolding compact_eq_seq_compact_metric seq_compact_def by auto |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
3860 |
|
| 50944 | 3861 |
subsubsection {* Complete the chain of compactness variants *}
|
3862 |
||
3863 |
lemma compact_eq_bolzano_weierstrass: |
|
3864 |
fixes s :: "'a::metric_space set" |
|
| 53282 | 3865 |
shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" |
3866 |
(is "?lhs = ?rhs") |
|
| 50944 | 3867 |
proof |
| 52624 | 3868 |
assume ?lhs |
| 53282 | 3869 |
then show ?rhs |
3870 |
using heine_borel_imp_bolzano_weierstrass[of s] by auto |
|
| 50944 | 3871 |
next |
| 52624 | 3872 |
assume ?rhs |
| 53282 | 3873 |
then show ?lhs |
| 50944 | 3874 |
unfolding compact_eq_seq_compact_metric by (rule bolzano_weierstrass_imp_seq_compact) |
3875 |
qed |
|
3876 |
||
3877 |
lemma bolzano_weierstrass_imp_bounded: |
|
| 53282 | 3878 |
"\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s" |
| 50944 | 3879 |
using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass . |
3880 |
||
| 54070 | 3881 |
subsection {* Metric spaces with the Heine-Borel property *}
|
3882 |
||
| 33175 | 3883 |
text {*
|
3884 |
A metric space (or topological vector space) is said to have the |
|
3885 |
Heine-Borel property if every closed and bounded subset is compact. |
|
3886 |
*} |
|
3887 |
||
|
44207
ea99698c2070
locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
huffman
parents:
44170
diff
changeset
|
3888 |
class heine_borel = metric_space + |
| 33175 | 3889 |
assumes bounded_imp_convergent_subsequence: |
| 50998 | 3890 |
"bounded (range f) \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
| 33175 | 3891 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3892 |
lemma bounded_closed_imp_seq_compact: |
| 33175 | 3893 |
fixes s::"'a::heine_borel set" |
| 53282 | 3894 |
assumes "bounded s" |
3895 |
and "closed s" |
|
3896 |
shows "seq_compact s" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3897 |
proof (unfold seq_compact_def, clarify) |
| 53282 | 3898 |
fix f :: "nat \<Rightarrow> 'a" |
3899 |
assume f: "\<forall>n. f n \<in> s" |
|
3900 |
with `bounded s` have "bounded (range f)" |
|
3901 |
by (auto intro: bounded_subset) |
|
| 33175 | 3902 |
obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially" |
| 50998 | 3903 |
using bounded_imp_convergent_subsequence [OF `bounded (range f)`] by auto |
| 53282 | 3904 |
from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" |
3905 |
by simp |
|
| 33175 | 3906 |
have "l \<in> s" using `closed s` fr l |
| 54070 | 3907 |
by (rule closed_sequentially) |
| 33175 | 3908 |
show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
3909 |
using `l \<in> s` r l by blast |
|
3910 |
qed |
|
3911 |
||
| 50944 | 3912 |
lemma compact_eq_bounded_closed: |
3913 |
fixes s :: "'a::heine_borel set" |
|
| 53291 | 3914 |
shows "compact s \<longleftrightarrow> bounded s \<and> closed s" |
3915 |
(is "?lhs = ?rhs") |
|
| 50944 | 3916 |
proof |
| 52624 | 3917 |
assume ?lhs |
| 53282 | 3918 |
then show ?rhs |
| 52624 | 3919 |
using compact_imp_closed compact_imp_bounded |
3920 |
by blast |
|
| 50944 | 3921 |
next |
| 52624 | 3922 |
assume ?rhs |
| 53282 | 3923 |
then show ?lhs |
| 52624 | 3924 |
using bounded_closed_imp_seq_compact[of s] |
3925 |
unfolding compact_eq_seq_compact_metric |
|
3926 |
by auto |
|
| 50944 | 3927 |
qed |
3928 |
||
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
3929 |
(* TODO: is this lemma necessary? *) |
| 50972 | 3930 |
lemma bounded_increasing_convergent: |
3931 |
fixes s :: "nat \<Rightarrow> real" |
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
3932 |
shows "bounded {s n| n. True} \<Longrightarrow> \<forall>n. s n \<le> s (Suc n) \<Longrightarrow> \<exists>l. s ----> l"
|
| 50972 | 3933 |
using Bseq_mono_convergent[of s] incseq_Suc_iff[of s] |
3934 |
by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) |
|
| 33175 | 3935 |
|
3936 |
instance real :: heine_borel |
|
3937 |
proof |
|
| 50998 | 3938 |
fix f :: "nat \<Rightarrow> real" |
3939 |
assume f: "bounded (range f)" |
|
| 50972 | 3940 |
obtain r where r: "subseq r" "monoseq (f \<circ> r)" |
3941 |
unfolding comp_def by (metis seq_monosub) |
|
3942 |
then have "Bseq (f \<circ> r)" |
|
| 50998 | 3943 |
unfolding Bseq_eq_bounded using f by (auto intro: bounded_subset) |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
3944 |
with r show "\<exists>l r. subseq r \<and> (f \<circ> r) ----> l" |
| 50972 | 3945 |
using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def) |
| 33175 | 3946 |
qed |
3947 |
||
3948 |
lemma compact_lemma: |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
3949 |
fixes f :: "nat \<Rightarrow> 'a::euclidean_space" |
| 50998 | 3950 |
assumes "bounded (range f)" |
| 53291 | 3951 |
shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. |
3952 |
subseq r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)" |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3953 |
proof safe |
| 52624 | 3954 |
fix d :: "'a set" |
| 53282 | 3955 |
assume d: "d \<subseteq> Basis" |
3956 |
with finite_Basis have "finite d" |
|
3957 |
by (blast intro: finite_subset) |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3958 |
from this d show "\<exists>l::'a. \<exists>r. subseq r \<and> |
| 52624 | 3959 |
(\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)" |
3960 |
proof (induct d) |
|
3961 |
case empty |
|
| 53282 | 3962 |
then show ?case |
3963 |
unfolding subseq_def by auto |
|
| 52624 | 3964 |
next |
3965 |
case (insert k d) |
|
| 53282 | 3966 |
have k[intro]: "k \<in> Basis" |
3967 |
using insert by auto |
|
3968 |
have s': "bounded ((\<lambda>x. x \<bullet> k) ` range f)" |
|
3969 |
using `bounded (range f)` |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3970 |
by (auto intro!: bounded_linear_image bounded_linear_inner_left) |
| 53282 | 3971 |
obtain l1::"'a" and r1 where r1: "subseq r1" |
3972 |
and lr1: "\<forall>e > 0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially" |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
3973 |
using insert(3) using insert(4) by auto |
| 53282 | 3974 |
have f': "\<forall>n. f (r1 n) \<bullet> k \<in> (\<lambda>x. x \<bullet> k) ` range f" |
3975 |
by simp |
|
| 50998 | 3976 |
have "bounded (range (\<lambda>i. f (r1 i) \<bullet> k))" |
3977 |
by (metis (lifting) bounded_subset f' image_subsetI s') |
|
3978 |
then obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) \<bullet> k) ---> l2) sequentially" |
|
| 53282 | 3979 |
using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) \<bullet> k"] |
3980 |
by (auto simp: o_def) |
|
3981 |
def r \<equiv> "r1 \<circ> r2" |
|
3982 |
have r:"subseq r" |
|
| 33175 | 3983 |
using r1 and r2 unfolding r_def o_def subseq_def by auto |
3984 |
moreover |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3985 |
def l \<equiv> "(\<Sum>i\<in>Basis. (if i = k then l2 else l1\<bullet>i) *\<^sub>R i)::'a" |
| 52624 | 3986 |
{
|
3987 |
fix e::real |
|
| 53282 | 3988 |
assume "e > 0" |
3989 |
from lr1 `e > 0` have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) \<bullet> i) (l1 \<bullet> i) < e) sequentially" |
|
| 52624 | 3990 |
by blast |
| 53282 | 3991 |
from lr2 `e > 0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) \<bullet> k) l2 < e) sequentially" |
| 52624 | 3992 |
by (rule tendstoD) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3993 |
from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) \<bullet> i) (l1 \<bullet> i) < e) sequentially" |
| 33175 | 3994 |
by (rule eventually_subseq) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3995 |
have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially" |
| 53282 | 3996 |
using N1' N2 |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
3997 |
by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def) |
| 33175 | 3998 |
} |
3999 |
ultimately show ?case by auto |
|
4000 |
qed |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
4001 |
qed |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
4002 |
|
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
4003 |
instance euclidean_space \<subseteq> heine_borel |
| 33175 | 4004 |
proof |
| 50998 | 4005 |
fix f :: "nat \<Rightarrow> 'a" |
4006 |
assume f: "bounded (range f)" |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
4007 |
then obtain l::'a and r where r: "subseq r" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
4008 |
and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially" |
| 50998 | 4009 |
using compact_lemma [OF f] by blast |
| 52624 | 4010 |
{
|
4011 |
fix e::real |
|
| 53282 | 4012 |
assume "e > 0" |
| 56541 | 4013 |
hence "e / real_of_nat DIM('a) > 0" by (simp add: DIM_positive)
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
4014 |
with l have "eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))) sequentially"
|
| 33175 | 4015 |
by simp |
4016 |
moreover |
|
| 52624 | 4017 |
{
|
4018 |
fix n |
|
4019 |
assume n: "\<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e / (real_of_nat DIM('a))"
|
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
4020 |
have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))" |
| 52624 | 4021 |
apply (subst euclidean_dist_l2) |
4022 |
using zero_le_dist |
|
| 53282 | 4023 |
apply (rule setL2_le_setsum) |
4024 |
done |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
4025 |
also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
|
| 52624 | 4026 |
apply (rule setsum_strict_mono) |
4027 |
using n |
|
| 53282 | 4028 |
apply auto |
4029 |
done |
|
4030 |
finally have "dist (f (r n)) l < e" |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
4031 |
by auto |
| 33175 | 4032 |
} |
4033 |
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" |
|
4034 |
by (rule eventually_elim1) |
|
4035 |
} |
|
| 53282 | 4036 |
then have *: "((f \<circ> r) ---> l) sequentially" |
| 52624 | 4037 |
unfolding o_def tendsto_iff by simp |
4038 |
with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
|
4039 |
by auto |
|
| 33175 | 4040 |
qed |
4041 |
||
4042 |
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)" |
|
| 52624 | 4043 |
unfolding bounded_def |
| 55775 | 4044 |
by (metis (erased, hide_lams) dist_fst_le image_iff order_trans) |
| 33175 | 4045 |
|
4046 |
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)" |
|
| 52624 | 4047 |
unfolding bounded_def |
| 55775 | 4048 |
by (metis (no_types, hide_lams) dist_snd_le image_iff order.trans) |
| 33175 | 4049 |
|
|
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37649
diff
changeset
|
4050 |
instance prod :: (heine_borel, heine_borel) heine_borel |
| 33175 | 4051 |
proof |
| 50998 | 4052 |
fix f :: "nat \<Rightarrow> 'a \<times> 'b" |
4053 |
assume f: "bounded (range f)" |
|
|
56154
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4054 |
then have "bounded (fst ` range f)" |
|
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4055 |
by (rule bounded_fst) |
|
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4056 |
then have s1: "bounded (range (fst \<circ> f))" |
|
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4057 |
by (simp add: image_comp) |
| 50998 | 4058 |
obtain l1 r1 where r1: "subseq r1" and l1: "(\<lambda>n. fst (f (r1 n))) ----> l1" |
4059 |
using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast |
|
4060 |
from f have s2: "bounded (range (snd \<circ> f \<circ> r1))" |
|
4061 |
by (auto simp add: image_comp intro: bounded_snd bounded_subset) |
|
| 53282 | 4062 |
obtain l2 r2 where r2: "subseq r2" and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially" |
| 50998 | 4063 |
using bounded_imp_convergent_subsequence [OF s2] |
| 33175 | 4064 |
unfolding o_def by fast |
4065 |
have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially" |
|
| 50972 | 4066 |
using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def . |
| 33175 | 4067 |
have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially" |
4068 |
using tendsto_Pair [OF l1' l2] unfolding o_def by simp |
|
4069 |
have r: "subseq (r1 \<circ> r2)" |
|
4070 |
using r1 r2 unfolding subseq_def by simp |
|
4071 |
show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially" |
|
4072 |
using l r by fast |
|
4073 |
qed |
|
4074 |
||
| 54070 | 4075 |
subsubsection {* Completeness *}
|
| 33175 | 4076 |
|
| 52624 | 4077 |
definition complete :: "'a::metric_space set \<Rightarrow> bool" |
4078 |
where "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>s. f ----> l))" |
|
4079 |
||
| 54070 | 4080 |
lemma completeI: |
4081 |
assumes "\<And>f. \<forall>n. f n \<in> s \<Longrightarrow> Cauchy f \<Longrightarrow> \<exists>l\<in>s. f ----> l" |
|
4082 |
shows "complete s" |
|
4083 |
using assms unfolding complete_def by fast |
|
4084 |
||
4085 |
lemma completeE: |
|
4086 |
assumes "complete s" and "\<forall>n. f n \<in> s" and "Cauchy f" |
|
4087 |
obtains l where "l \<in> s" and "f ----> l" |
|
4088 |
using assms unfolding complete_def by fast |
|
4089 |
||
| 52624 | 4090 |
lemma compact_imp_complete: |
4091 |
assumes "compact s" |
|
4092 |
shows "complete s" |
|
4093 |
proof - |
|
4094 |
{
|
|
4095 |
fix f |
|
4096 |
assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f" |
|
| 50971 | 4097 |
from as(1) obtain l r where lr: "l\<in>s" "subseq r" "(f \<circ> r) ----> l" |
4098 |
using assms unfolding compact_def by blast |
|
4099 |
||
4100 |
note lr' = seq_suble [OF lr(2)] |
|
| 52624 | 4101 |
{
|
| 53282 | 4102 |
fix e :: real |
4103 |
assume "e > 0" |
|
| 52624 | 4104 |
from as(2) obtain N where N:"\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (f m) (f n) < e/2" |
4105 |
unfolding cauchy_def |
|
| 53282 | 4106 |
using `e > 0` |
4107 |
apply (erule_tac x="e/2" in allE) |
|
| 52624 | 4108 |
apply auto |
4109 |
done |
|
4110 |
from lr(3)[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] |
|
| 53282 | 4111 |
obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" |
4112 |
using `e > 0` by auto |
|
| 52624 | 4113 |
{
|
| 53282 | 4114 |
fix n :: nat |
4115 |
assume n: "n \<ge> max N M" |
|
4116 |
have "dist ((f \<circ> r) n) l < e/2" |
|
4117 |
using n M by auto |
|
4118 |
moreover have "r n \<ge> N" |
|
4119 |
using lr'[of n] n by auto |
|
4120 |
then have "dist (f n) ((f \<circ> r) n) < e / 2" |
|
4121 |
using N and n by auto |
|
| 52624 | 4122 |
ultimately have "dist (f n) l < e" |
| 53282 | 4123 |
using dist_triangle_half_r[of "f (r n)" "f n" e l] |
4124 |
by (auto simp add: dist_commute) |
|
| 52624 | 4125 |
} |
| 53282 | 4126 |
then have "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast |
| 52624 | 4127 |
} |
| 53282 | 4128 |
then have "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s` |
| 52624 | 4129 |
unfolding LIMSEQ_def by auto |
4130 |
} |
|
| 53282 | 4131 |
then show ?thesis unfolding complete_def by auto |
| 50971 | 4132 |
qed |
4133 |
||
4134 |
lemma nat_approx_posE: |
|
4135 |
fixes e::real |
|
4136 |
assumes "0 < e" |
|
| 53282 | 4137 |
obtains n :: nat where "1 / (Suc n) < e" |
| 50971 | 4138 |
proof atomize_elim |
4139 |
have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))" |
|
| 56544 | 4140 |
by (rule divide_strict_left_mono) (auto simp: `0 < e`) |
| 50971 | 4141 |
also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)" |
| 56541 | 4142 |
by (rule divide_left_mono) (auto simp: `0 < e`) |
| 50971 | 4143 |
also have "\<dots> = e" by simp |
4144 |
finally show "\<exists>n. 1 / real (Suc n) < e" .. |
|
4145 |
qed |
|
4146 |
||
4147 |
lemma compact_eq_totally_bounded: |
|
4148 |
"compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k)))" |
|
4149 |
(is "_ \<longleftrightarrow> ?rhs") |
|
4150 |
proof |
|
4151 |
assume assms: "?rhs" |
|
4152 |
then obtain k where k: "\<And>e. 0 < e \<Longrightarrow> finite (k e)" "\<And>e. 0 < e \<Longrightarrow> s \<subseteq> (\<Union>x\<in>k e. ball x e)" |
|
4153 |
by (auto simp: choice_iff') |
|
4154 |
||
4155 |
show "compact s" |
|
4156 |
proof cases |
|
| 53282 | 4157 |
assume "s = {}"
|
4158 |
then show "compact s" by (simp add: compact_def) |
|
| 50971 | 4159 |
next |
4160 |
assume "s \<noteq> {}"
|
|
4161 |
show ?thesis |
|
4162 |
unfolding compact_def |
|
4163 |
proof safe |
|
| 53282 | 4164 |
fix f :: "nat \<Rightarrow> 'a" |
4165 |
assume f: "\<forall>n. f n \<in> s" |
|
4166 |
||
| 50971 | 4167 |
def e \<equiv> "\<lambda>n. 1 / (2 * Suc n)" |
4168 |
then have [simp]: "\<And>n. 0 < e n" by auto |
|
4169 |
def B \<equiv> "\<lambda>n U. SOME b. infinite {n. f n \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
|
|
| 53282 | 4170 |
{
|
4171 |
fix n U |
|
4172 |
assume "infinite {n. f n \<in> U}"
|
|
| 50971 | 4173 |
then have "\<exists>b\<in>k (e n). infinite {i\<in>{n. f n \<in> U}. f i \<in> ball b (e n)}"
|
4174 |
using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq) |
|
| 55522 | 4175 |
then obtain a where |
4176 |
"a \<in> k (e n)" |
|
4177 |
"infinite {i \<in> {n. f n \<in> U}. f i \<in> ball a (e n)}" ..
|
|
| 50971 | 4178 |
then have "\<exists>b. infinite {i. f i \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
|
4179 |
by (intro exI[of _ "ball a (e n) \<inter> U"] exI[of _ a]) (auto simp: ac_simps) |
|
4180 |
from someI_ex[OF this] |
|
4181 |
have "infinite {i. f i \<in> B n U}" "\<exists>x. B n U \<subseteq> ball x (e n) \<inter> U"
|
|
| 53282 | 4182 |
unfolding B_def by auto |
4183 |
} |
|
| 50971 | 4184 |
note B = this |
4185 |
||
| 55415 | 4186 |
def F \<equiv> "rec_nat (B 0 UNIV) B" |
| 53282 | 4187 |
{
|
4188 |
fix n |
|
4189 |
have "infinite {i. f i \<in> F n}"
|
|
4190 |
by (induct n) (auto simp: F_def B) |
|
4191 |
} |
|
| 50971 | 4192 |
then have F: "\<And>n. \<exists>x. F (Suc n) \<subseteq> ball x (e n) \<inter> F n" |
4193 |
using B by (simp add: F_def) |
|
4194 |
then have F_dec: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m" |
|
4195 |
using decseq_SucI[of F] by (auto simp: decseq_def) |
|
4196 |
||
4197 |
obtain sel where sel: "\<And>k i. i < sel k i" "\<And>k i. f (sel k i) \<in> F k" |
|
4198 |
proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI) |
|
4199 |
fix k i |
|
4200 |
have "infinite ({n. f n \<in> F k} - {.. i})"
|
|
4201 |
using `infinite {n. f n \<in> F k}` by auto
|
|
4202 |
from infinite_imp_nonempty[OF this] |
|
4203 |
show "\<exists>x>i. f x \<in> F k" |
|
4204 |
by (simp add: set_eq_iff not_le conj_commute) |
|
4205 |
qed |
|
4206 |
||
| 55415 | 4207 |
def t \<equiv> "rec_nat (sel 0 0) (\<lambda>n i. sel (Suc n) i)" |
| 50971 | 4208 |
have "subseq t" |
4209 |
unfolding subseq_Suc_iff by (simp add: t_def sel) |
|
4210 |
moreover have "\<forall>i. (f \<circ> t) i \<in> s" |
|
4211 |
using f by auto |
|
4212 |
moreover |
|
| 53282 | 4213 |
{
|
4214 |
fix n |
|
4215 |
have "(f \<circ> t) n \<in> F n" |
|
4216 |
by (cases n) (simp_all add: t_def sel) |
|
4217 |
} |
|
| 50971 | 4218 |
note t = this |
4219 |
||
4220 |
have "Cauchy (f \<circ> t)" |
|
4221 |
proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE) |
|
| 53282 | 4222 |
fix r :: real and N n m |
4223 |
assume "1 / Suc N < r" "Suc N \<le> n" "Suc N \<le> m" |
|
| 50971 | 4224 |
then have "(f \<circ> t) n \<in> F (Suc N)" "(f \<circ> t) m \<in> F (Suc N)" "2 * e N < r" |
4225 |
using F_dec t by (auto simp: e_def field_simps real_of_nat_Suc) |
|
4226 |
with F[of N] obtain x where "dist x ((f \<circ> t) n) < e N" "dist x ((f \<circ> t) m) < e N" |
|
4227 |
by (auto simp: subset_eq) |
|
4228 |
with dist_triangle[of "(f \<circ> t) m" "(f \<circ> t) n" x] `2 * e N < r` |
|
4229 |
show "dist ((f \<circ> t) m) ((f \<circ> t) n) < r" |
|
4230 |
by (simp add: dist_commute) |
|
4231 |
qed |
|
4232 |
||
4233 |
ultimately show "\<exists>l\<in>s. \<exists>r. subseq r \<and> (f \<circ> r) ----> l" |
|
4234 |
using assms unfolding complete_def by blast |
|
4235 |
qed |
|
4236 |
qed |
|
4237 |
qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded) |
|
| 33175 | 4238 |
|
4239 |
lemma cauchy: "Cauchy s \<longleftrightarrow> (\<forall>e>0.\<exists> N::nat. \<forall>n\<ge>N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") |
|
| 53282 | 4240 |
proof - |
4241 |
{
|
|
4242 |
assume ?rhs |
|
4243 |
{
|
|
4244 |
fix e::real |
|
| 33175 | 4245 |
assume "e>0" |
4246 |
with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2" |
|
4247 |
by (erule_tac x="e/2" in allE) auto |
|
| 53282 | 4248 |
{
|
4249 |
fix n m |
|
| 33175 | 4250 |
assume nm:"N \<le> m \<and> N \<le> n" |
| 53282 | 4251 |
then have "dist (s m) (s n) < e" using N |
| 33175 | 4252 |
using dist_triangle_half_l[of "s m" "s N" "e" "s n"] |
4253 |
by blast |
|
4254 |
} |
|
| 53282 | 4255 |
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e" |
| 33175 | 4256 |
by blast |
4257 |
} |
|
| 53282 | 4258 |
then have ?lhs |
| 33175 | 4259 |
unfolding cauchy_def |
4260 |
by blast |
|
4261 |
} |
|
| 53282 | 4262 |
then show ?thesis |
| 33175 | 4263 |
unfolding cauchy_def |
4264 |
using dist_triangle_half_l |
|
4265 |
by blast |
|
4266 |
qed |
|
4267 |
||
| 53282 | 4268 |
lemma cauchy_imp_bounded: |
4269 |
assumes "Cauchy s" |
|
4270 |
shows "bounded (range s)" |
|
4271 |
proof - |
|
4272 |
from assms obtain N :: nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1" |
|
| 52624 | 4273 |
unfolding cauchy_def |
4274 |
apply (erule_tac x= 1 in allE) |
|
4275 |
apply auto |
|
4276 |
done |
|
| 53282 | 4277 |
then have N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto |
| 33175 | 4278 |
moreover |
| 52624 | 4279 |
have "bounded (s ` {0..N})"
|
4280 |
using finite_imp_bounded[of "s ` {1..N}"] by auto
|
|
| 33175 | 4281 |
then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
|
4282 |
unfolding bounded_any_center [where a="s N"] by auto |
|
4283 |
ultimately show "?thesis" |
|
4284 |
unfolding bounded_any_center [where a="s N"] |
|
| 52624 | 4285 |
apply (rule_tac x="max a 1" in exI) |
4286 |
apply auto |
|
4287 |
apply (erule_tac x=y in allE) |
|
4288 |
apply (erule_tac x=y in ballE) |
|
4289 |
apply auto |
|
4290 |
done |
|
| 33175 | 4291 |
qed |
4292 |
||
4293 |
instance heine_borel < complete_space |
|
4294 |
proof |
|
4295 |
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f" |
|
| 53282 | 4296 |
then have "bounded (range f)" |
| 34104 | 4297 |
by (rule cauchy_imp_bounded) |
| 53282 | 4298 |
then have "compact (closure (range f))" |
| 50971 | 4299 |
unfolding compact_eq_bounded_closed by auto |
| 53282 | 4300 |
then have "complete (closure (range f))" |
| 50971 | 4301 |
by (rule compact_imp_complete) |
| 33175 | 4302 |
moreover have "\<forall>n. f n \<in> closure (range f)" |
4303 |
using closure_subset [of "range f"] by auto |
|
4304 |
ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially" |
|
4305 |
using `Cauchy f` unfolding complete_def by auto |
|
4306 |
then show "convergent f" |
|
|
36660
1cc4ab4b7ff7
make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents:
36659
diff
changeset
|
4307 |
unfolding convergent_def by auto |
| 33175 | 4308 |
qed |
4309 |
||
| 44632 | 4310 |
instance euclidean_space \<subseteq> banach .. |
4311 |
||
| 54070 | 4312 |
lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
|
4313 |
proof (rule completeI) |
|
4314 |
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f" |
|
| 53282 | 4315 |
then have "convergent f" by (rule Cauchy_convergent) |
| 54070 | 4316 |
then show "\<exists>l\<in>UNIV. f ----> l" unfolding convergent_def by simp |
| 53282 | 4317 |
qed |
4318 |
||
4319 |
lemma complete_imp_closed: |
|
4320 |
assumes "complete s" |
|
4321 |
shows "closed s" |
|
| 54070 | 4322 |
proof (unfold closed_sequential_limits, clarify) |
4323 |
fix f x assume "\<forall>n. f n \<in> s" and "f ----> x" |
|
4324 |
from `f ----> x` have "Cauchy f" |
|
4325 |
by (rule LIMSEQ_imp_Cauchy) |
|
4326 |
with `complete s` and `\<forall>n. f n \<in> s` obtain l where "l \<in> s" and "f ----> l" |
|
4327 |
by (rule completeE) |
|
4328 |
from `f ----> x` and `f ----> l` have "x = l" |
|
4329 |
by (rule LIMSEQ_unique) |
|
4330 |
with `l \<in> s` show "x \<in> s" |
|
4331 |
by simp |
|
4332 |
qed |
|
4333 |
||
4334 |
lemma complete_inter_closed: |
|
4335 |
assumes "complete s" and "closed t" |
|
4336 |
shows "complete (s \<inter> t)" |
|
4337 |
proof (rule completeI) |
|
4338 |
fix f assume "\<forall>n. f n \<in> s \<inter> t" and "Cauchy f" |
|
4339 |
then have "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t" |
|
4340 |
by simp_all |
|
4341 |
from `complete s` obtain l where "l \<in> s" and "f ----> l" |
|
4342 |
using `\<forall>n. f n \<in> s` and `Cauchy f` by (rule completeE) |
|
4343 |
from `closed t` and `\<forall>n. f n \<in> t` and `f ----> l` have "l \<in> t" |
|
4344 |
by (rule closed_sequentially) |
|
4345 |
with `l \<in> s` and `f ----> l` show "\<exists>l\<in>s \<inter> t. f ----> l" |
|
4346 |
by fast |
|
4347 |
qed |
|
4348 |
||
4349 |
lemma complete_closed_subset: |
|
4350 |
assumes "closed s" and "s \<subseteq> t" and "complete t" |
|
4351 |
shows "complete s" |
|
4352 |
using assms complete_inter_closed [of t s] by (simp add: Int_absorb1) |
|
| 33175 | 4353 |
|
4354 |
lemma complete_eq_closed: |
|
| 54070 | 4355 |
fixes s :: "('a::complete_space) set"
|
4356 |
shows "complete s \<longleftrightarrow> closed s" |
|
| 33175 | 4357 |
proof |
| 54070 | 4358 |
assume "closed s" then show "complete s" |
4359 |
using subset_UNIV complete_UNIV by (rule complete_closed_subset) |
|
| 33175 | 4360 |
next |
| 54070 | 4361 |
assume "complete s" then show "closed s" |
4362 |
by (rule complete_imp_closed) |
|
| 33175 | 4363 |
qed |
4364 |
||
4365 |
lemma convergent_eq_cauchy: |
|
4366 |
fixes s :: "nat \<Rightarrow> 'a::complete_space" |
|
| 44632 | 4367 |
shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s" |
4368 |
unfolding Cauchy_convergent_iff convergent_def .. |
|
| 33175 | 4369 |
|
4370 |
lemma convergent_imp_bounded: |
|
4371 |
fixes s :: "nat \<Rightarrow> 'a::metric_space" |
|
| 44632 | 4372 |
shows "(s ---> l) sequentially \<Longrightarrow> bounded (range s)" |
|
50939
ae7cd20ed118
replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
hoelzl
parents:
50938
diff
changeset
|
4373 |
by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy) |
| 33175 | 4374 |
|
4375 |
lemma compact_cball[simp]: |
|
4376 |
fixes x :: "'a::heine_borel" |
|
| 54070 | 4377 |
shows "compact (cball x e)" |
| 33175 | 4378 |
using compact_eq_bounded_closed bounded_cball closed_cball |
4379 |
by blast |
|
4380 |
||
4381 |
lemma compact_frontier_bounded[intro]: |
|
4382 |
fixes s :: "'a::heine_borel set" |
|
| 54070 | 4383 |
shows "bounded s \<Longrightarrow> compact (frontier s)" |
| 33175 | 4384 |
unfolding frontier_def |
4385 |
using compact_eq_bounded_closed |
|
4386 |
by blast |
|
4387 |
||
4388 |
lemma compact_frontier[intro]: |
|
4389 |
fixes s :: "'a::heine_borel set" |
|
| 53291 | 4390 |
shows "compact s \<Longrightarrow> compact (frontier s)" |
| 33175 | 4391 |
using compact_eq_bounded_closed compact_frontier_bounded |
4392 |
by blast |
|
4393 |
||
4394 |
lemma frontier_subset_compact: |
|
4395 |
fixes s :: "'a::heine_borel set" |
|
| 53291 | 4396 |
shows "compact s \<Longrightarrow> frontier s \<subseteq> s" |
| 33175 | 4397 |
using frontier_subset_closed compact_eq_bounded_closed |
4398 |
by blast |
|
4399 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
4400 |
subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
|
| 33175 | 4401 |
|
4402 |
lemma bounded_closed_nest: |
|
| 54070 | 4403 |
fixes s :: "nat \<Rightarrow> ('a::heine_borel) set"
|
4404 |
assumes "\<forall>n. closed (s n)" |
|
4405 |
and "\<forall>n. s n \<noteq> {}"
|
|
4406 |
and "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m" |
|
4407 |
and "bounded (s 0)" |
|
4408 |
shows "\<exists>a. \<forall>n. a \<in> s n" |
|
| 52624 | 4409 |
proof - |
| 54070 | 4410 |
from assms(2) obtain x where x: "\<forall>n. x n \<in> s n" |
4411 |
using choice[of "\<lambda>n x. x \<in> s n"] by auto |
|
4412 |
from assms(4,1) have "seq_compact (s 0)" |
|
4413 |
by (simp add: bounded_closed_imp_seq_compact) |
|
4414 |
then obtain l r where lr: "l \<in> s 0" "subseq r" "(x \<circ> r) ----> l" |
|
4415 |
using x and assms(3) unfolding seq_compact_def by blast |
|
4416 |
have "\<forall>n. l \<in> s n" |
|
4417 |
proof |
|
| 53282 | 4418 |
fix n :: nat |
| 54070 | 4419 |
have "closed (s n)" |
4420 |
using assms(1) by simp |
|
4421 |
moreover have "\<forall>i. (x \<circ> r) i \<in> s i" |
|
4422 |
using x and assms(3) and lr(2) [THEN seq_suble] by auto |
|
4423 |
then have "\<forall>i. (x \<circ> r) (i + n) \<in> s n" |
|
4424 |
using assms(3) by (fast intro!: le_add2) |
|
4425 |
moreover have "(\<lambda>i. (x \<circ> r) (i + n)) ----> l" |
|
4426 |
using lr(3) by (rule LIMSEQ_ignore_initial_segment) |
|
4427 |
ultimately show "l \<in> s n" |
|
4428 |
by (rule closed_sequentially) |
|
4429 |
qed |
|
4430 |
then show ?thesis .. |
|
| 33175 | 4431 |
qed |
4432 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
4433 |
text {* Decreasing case does not even need compactness, just completeness. *}
|
| 33175 | 4434 |
|
4435 |
lemma decreasing_closed_nest: |
|
| 54070 | 4436 |
fixes s :: "nat \<Rightarrow> ('a::complete_space) set"
|
| 53282 | 4437 |
assumes |
| 54070 | 4438 |
"\<forall>n. closed (s n)" |
4439 |
"\<forall>n. s n \<noteq> {}"
|
|
4440 |
"\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m" |
|
4441 |
"\<forall>e>0. \<exists>n. \<forall>x\<in>s n. \<forall>y\<in>s n. dist x y < e" |
|
4442 |
shows "\<exists>a. \<forall>n. a \<in> s n" |
|
4443 |
proof - |
|
4444 |
have "\<forall>n. \<exists>x. x \<in> s n" |
|
| 53282 | 4445 |
using assms(2) by auto |
4446 |
then have "\<exists>t. \<forall>n. t n \<in> s n" |
|
| 54070 | 4447 |
using choice[of "\<lambda>n x. x \<in> s n"] by auto |
| 33175 | 4448 |
then obtain t where t: "\<forall>n. t n \<in> s n" by auto |
| 53282 | 4449 |
{
|
4450 |
fix e :: real |
|
4451 |
assume "e > 0" |
|
4452 |
then obtain N where N:"\<forall>x\<in>s N. \<forall>y\<in>s N. dist x y < e" |
|
4453 |
using assms(4) by auto |
|
4454 |
{
|
|
4455 |
fix m n :: nat |
|
4456 |
assume "N \<le> m \<and> N \<le> n" |
|
4457 |
then have "t m \<in> s N" "t n \<in> s N" |
|
4458 |
using assms(3) t unfolding subset_eq t by blast+ |
|
4459 |
then have "dist (t m) (t n) < e" |
|
4460 |
using N by auto |
|
| 33175 | 4461 |
} |
| 53282 | 4462 |
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e" |
4463 |
by auto |
|
| 33175 | 4464 |
} |
| 53282 | 4465 |
then have "Cauchy t" |
4466 |
unfolding cauchy_def by auto |
|
4467 |
then obtain l where l:"(t ---> l) sequentially" |
|
| 54070 | 4468 |
using complete_UNIV unfolding complete_def by auto |
| 53282 | 4469 |
{
|
4470 |
fix n :: nat |
|
4471 |
{
|
|
4472 |
fix e :: real |
|
4473 |
assume "e > 0" |
|
4474 |
then obtain N :: nat where N: "\<forall>n\<ge>N. dist (t n) l < e" |
|
4475 |
using l[unfolded LIMSEQ_def] by auto |
|
4476 |
have "t (max n N) \<in> s n" |
|
4477 |
using assms(3) |
|
4478 |
unfolding subset_eq |
|
4479 |
apply (erule_tac x=n in allE) |
|
4480 |
apply (erule_tac x="max n N" in allE) |
|
4481 |
using t |
|
4482 |
apply auto |
|
4483 |
done |
|
4484 |
then have "\<exists>y\<in>s n. dist y l < e" |
|
4485 |
apply (rule_tac x="t (max n N)" in bexI) |
|
4486 |
using N |
|
4487 |
apply auto |
|
4488 |
done |
|
| 33175 | 4489 |
} |
| 53282 | 4490 |
then have "l \<in> s n" |
4491 |
using closed_approachable[of "s n" l] assms(1) by auto |
|
| 33175 | 4492 |
} |
4493 |
then show ?thesis by auto |
|
4494 |
qed |
|
4495 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
4496 |
text {* Strengthen it to the intersection actually being a singleton. *}
|
| 33175 | 4497 |
|
4498 |
lemma decreasing_closed_nest_sing: |
|
| 44632 | 4499 |
fixes s :: "nat \<Rightarrow> 'a::complete_space set" |
| 53282 | 4500 |
assumes |
4501 |
"\<forall>n. closed(s n)" |
|
4502 |
"\<forall>n. s n \<noteq> {}"
|
|
| 54070 | 4503 |
"\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m" |
| 53282 | 4504 |
"\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e" |
| 34104 | 4505 |
shows "\<exists>a. \<Inter>(range s) = {a}"
|
| 53282 | 4506 |
proof - |
4507 |
obtain a where a: "\<forall>n. a \<in> s n" |
|
4508 |
using decreasing_closed_nest[of s] using assms by auto |
|
4509 |
{
|
|
4510 |
fix b |
|
4511 |
assume b: "b \<in> \<Inter>(range s)" |
|
4512 |
{
|
|
4513 |
fix e :: real |
|
4514 |
assume "e > 0" |
|
4515 |
then have "dist a b < e" |
|
4516 |
using assms(4) and b and a by blast |
|
| 33175 | 4517 |
} |
| 53282 | 4518 |
then have "dist a b = 0" |
4519 |
by (metis dist_eq_0_iff dist_nz less_le) |
|
| 33175 | 4520 |
} |
| 53282 | 4521 |
with a have "\<Inter>(range s) = {a}"
|
4522 |
unfolding image_def by auto |
|
4523 |
then show ?thesis .. |
|
| 33175 | 4524 |
qed |
4525 |
||
4526 |
text{* Cauchy-type criteria for uniform convergence. *}
|
|
4527 |
||
| 53282 | 4528 |
lemma uniformly_convergent_eq_cauchy: |
4529 |
fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::complete_space" |
|
4530 |
shows |
|
| 53291 | 4531 |
"(\<exists>l. \<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e) \<longleftrightarrow> |
4532 |
(\<forall>e>0. \<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e)" |
|
| 53282 | 4533 |
(is "?lhs = ?rhs") |
4534 |
proof |
|
| 33175 | 4535 |
assume ?lhs |
| 53282 | 4536 |
then obtain l where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e" |
4537 |
by auto |
|
4538 |
{
|
|
4539 |
fix e :: real |
|
4540 |
assume "e > 0" |
|
4541 |
then obtain N :: nat where N: "\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2" |
|
4542 |
using l[THEN spec[where x="e/2"]] by auto |
|
4543 |
{
|
|
4544 |
fix n m :: nat and x :: "'b" |
|
4545 |
assume "N \<le> m \<and> N \<le> n \<and> P x" |
|
4546 |
then have "dist (s m x) (s n x) < e" |
|
| 33175 | 4547 |
using N[THEN spec[where x=m], THEN spec[where x=x]] |
4548 |
using N[THEN spec[where x=n], THEN spec[where x=x]] |
|
| 53282 | 4549 |
using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto |
4550 |
} |
|
4551 |
then have "\<exists>N. \<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x --> dist (s m x) (s n x) < e" by auto |
|
4552 |
} |
|
4553 |
then show ?rhs by auto |
|
| 33175 | 4554 |
next |
4555 |
assume ?rhs |
|
| 53282 | 4556 |
then have "\<forall>x. P x \<longrightarrow> Cauchy (\<lambda>n. s n x)" |
4557 |
unfolding cauchy_def |
|
4558 |
apply auto |
|
4559 |
apply (erule_tac x=e in allE) |
|
4560 |
apply auto |
|
4561 |
done |
|
4562 |
then obtain l where l: "\<forall>x. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l x) sequentially" |
|
| 53291 | 4563 |
unfolding convergent_eq_cauchy[symmetric] |
| 53282 | 4564 |
using choice[of "\<lambda>x l. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l) sequentially"] |
4565 |
by auto |
|
4566 |
{
|
|
4567 |
fix e :: real |
|
4568 |
assume "e > 0" |
|
| 33175 | 4569 |
then obtain N where N:"\<forall>m n x. N \<le> m \<and> N \<le> n \<and> P x \<longrightarrow> dist (s m x) (s n x) < e/2" |
4570 |
using `?rhs`[THEN spec[where x="e/2"]] by auto |
|
| 53282 | 4571 |
{
|
4572 |
fix x |
|
4573 |
assume "P x" |
|
| 33175 | 4574 |
then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2" |
| 53282 | 4575 |
using l[THEN spec[where x=x], unfolded LIMSEQ_def] and `e > 0` |
4576 |
by (auto elim!: allE[where x="e/2"]) |
|
4577 |
fix n :: nat |
|
4578 |
assume "n \<ge> N" |
|
4579 |
then have "dist(s n x)(l x) < e" |
|
4580 |
using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]] |
|
4581 |
using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"] |
|
4582 |
by (auto simp add: dist_commute) |
|
4583 |
} |
|
4584 |
then have "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" |
|
4585 |
by auto |
|
4586 |
} |
|
4587 |
then show ?lhs by auto |
|
| 33175 | 4588 |
qed |
4589 |
||
4590 |
lemma uniformly_cauchy_imp_uniformly_convergent: |
|
| 51102 | 4591 |
fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::complete_space" |
| 33175 | 4592 |
assumes "\<forall>e>0.\<exists>N. \<forall>m (n::nat) x. N \<le> m \<and> N \<le> n \<and> P x --> dist(s m x)(s n x) < e" |
| 53291 | 4593 |
and "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n \<longrightarrow> dist(s n x)(l x) < e)" |
4594 |
shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e" |
|
| 53282 | 4595 |
proof - |
| 33175 | 4596 |
obtain l' where l:"\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l' x) < e" |
| 53291 | 4597 |
using assms(1) unfolding uniformly_convergent_eq_cauchy[symmetric] by auto |
| 33175 | 4598 |
moreover |
| 53282 | 4599 |
{
|
4600 |
fix x |
|
4601 |
assume "P x" |
|
4602 |
then have "l x = l' x" |
|
4603 |
using tendsto_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"] |
|
4604 |
using l and assms(2) unfolding LIMSEQ_def by blast |
|
4605 |
} |
|
| 33175 | 4606 |
ultimately show ?thesis by auto |
4607 |
qed |
|
4608 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
4609 |
|
| 36437 | 4610 |
subsection {* Continuity *}
|
4611 |
||
| 33175 | 4612 |
text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
|
4613 |
||
4614 |
lemma continuous_within_eps_delta: |
|
4615 |
"continuous (at x within s) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'\<in> s. dist x' x < d --> dist (f x') (f x) < e)" |
|
4616 |
unfolding continuous_within and Lim_within |
|
| 53282 | 4617 |
apply auto |
| 55775 | 4618 |
apply (metis dist_nz dist_self) |
4619 |
apply blast |
|
| 53282 | 4620 |
done |
4621 |
||
4622 |
lemma continuous_at_eps_delta: |
|
4623 |
"continuous (at x) f \<longleftrightarrow> (\<forall>e > 0. \<exists>d > 0. \<forall>x'. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)" |
|
| 45031 | 4624 |
using continuous_within_eps_delta [of x UNIV f] by simp |
| 33175 | 4625 |
|
4626 |
text{* Versions in terms of open balls. *}
|
|
4627 |
||
4628 |
lemma continuous_within_ball: |
|
| 53282 | 4629 |
"continuous (at x within s) f \<longleftrightarrow> |
4630 |
(\<forall>e > 0. \<exists>d > 0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" |
|
4631 |
(is "?lhs = ?rhs") |
|
| 33175 | 4632 |
proof |
4633 |
assume ?lhs |
|
| 53282 | 4634 |
{
|
4635 |
fix e :: real |
|
4636 |
assume "e > 0" |
|
| 33175 | 4637 |
then obtain d where d: "d>0" "\<forall>xa\<in>s. 0 < dist xa x \<and> dist xa x < d \<longrightarrow> dist (f xa) (f x) < e" |
4638 |
using `?lhs`[unfolded continuous_within Lim_within] by auto |
|
| 53282 | 4639 |
{
|
4640 |
fix y |
|
4641 |
assume "y \<in> f ` (ball x d \<inter> s)" |
|
4642 |
then have "y \<in> ball (f x) e" |
|
4643 |
using d(2) |
|
| 53291 | 4644 |
unfolding dist_nz[symmetric] |
| 53282 | 4645 |
apply (auto simp add: dist_commute) |
4646 |
apply (erule_tac x=xa in ballE) |
|
4647 |
apply auto |
|
4648 |
using `e > 0` |
|
4649 |
apply auto |
|
4650 |
done |
|
| 33175 | 4651 |
} |
| 53282 | 4652 |
then have "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" |
4653 |
using `d > 0` |
|
4654 |
unfolding subset_eq ball_def by (auto simp add: dist_commute) |
|
4655 |
} |
|
4656 |
then show ?rhs by auto |
|
| 33175 | 4657 |
next |
| 53282 | 4658 |
assume ?rhs |
4659 |
then show ?lhs |
|
4660 |
unfolding continuous_within Lim_within ball_def subset_eq |
|
4661 |
apply (auto simp add: dist_commute) |
|
4662 |
apply (erule_tac x=e in allE) |
|
4663 |
apply auto |
|
4664 |
done |
|
| 33175 | 4665 |
qed |
4666 |
||
4667 |
lemma continuous_at_ball: |
|
4668 |
"continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs") |
|
4669 |
proof |
|
| 53282 | 4670 |
assume ?lhs |
4671 |
then show ?rhs |
|
4672 |
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball |
|
4673 |
apply auto |
|
4674 |
apply (erule_tac x=e in allE) |
|
4675 |
apply auto |
|
4676 |
apply (rule_tac x=d in exI) |
|
4677 |
apply auto |
|
4678 |
apply (erule_tac x=xa in allE) |
|
4679 |
apply (auto simp add: dist_commute dist_nz) |
|
| 53291 | 4680 |
unfolding dist_nz[symmetric] |
| 53282 | 4681 |
apply auto |
4682 |
done |
|
| 33175 | 4683 |
next |
| 53282 | 4684 |
assume ?rhs |
4685 |
then show ?lhs |
|
4686 |
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball |
|
4687 |
apply auto |
|
4688 |
apply (erule_tac x=e in allE) |
|
4689 |
apply auto |
|
4690 |
apply (rule_tac x=d in exI) |
|
4691 |
apply auto |
|
4692 |
apply (erule_tac x="f xa" in allE) |
|
4693 |
apply (auto simp add: dist_commute dist_nz) |
|
4694 |
done |
|
| 33175 | 4695 |
qed |
4696 |
||
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4697 |
text{* Define setwise continuity in terms of limits within the set. *}
|
| 33175 | 4698 |
|
| 36359 | 4699 |
lemma continuous_on_iff: |
4700 |
"continuous_on s f \<longleftrightarrow> |
|
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4701 |
(\<forall>x\<in>s. \<forall>e>0. \<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)" |
| 53282 | 4702 |
unfolding continuous_on_def Lim_within |
| 55775 | 4703 |
by (metis dist_pos_lt dist_self) |
| 53282 | 4704 |
|
4705 |
definition uniformly_continuous_on :: "'a set \<Rightarrow> ('a::metric_space \<Rightarrow> 'b::metric_space) \<Rightarrow> bool"
|
|
4706 |
where "uniformly_continuous_on s f \<longleftrightarrow> |
|
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4707 |
(\<forall>e>0. \<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e)" |
|
35172
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents:
35028
diff
changeset
|
4708 |
|
| 33175 | 4709 |
text{* Some simple consequential lemmas. *}
|
4710 |
||
4711 |
lemma uniformly_continuous_imp_continuous: |
|
| 53282 | 4712 |
"uniformly_continuous_on s f \<Longrightarrow> continuous_on s f" |
| 36359 | 4713 |
unfolding uniformly_continuous_on_def continuous_on_iff by blast |
| 33175 | 4714 |
|
4715 |
lemma continuous_at_imp_continuous_within: |
|
| 53282 | 4716 |
"continuous (at x) f \<Longrightarrow> continuous (at x within s) f" |
| 33175 | 4717 |
unfolding continuous_within continuous_at using Lim_at_within by auto |
4718 |
||
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4719 |
lemma Lim_trivial_limit: "trivial_limit net \<Longrightarrow> (f ---> l) net" |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
4720 |
by simp |
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4721 |
|
|
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4722 |
lemmas continuous_on = continuous_on_def -- "legacy theorem name" |
| 33175 | 4723 |
|
4724 |
lemma continuous_within_subset: |
|
| 53282 | 4725 |
"continuous (at x within s) f \<Longrightarrow> t \<subseteq> s \<Longrightarrow> continuous (at x within t) f" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
4726 |
unfolding continuous_within by(metis tendsto_within_subset) |
| 33175 | 4727 |
|
4728 |
lemma continuous_on_interior: |
|
| 53282 | 4729 |
"continuous_on s f \<Longrightarrow> x \<in> interior s \<Longrightarrow> continuous (at x) f" |
| 55775 | 4730 |
by (metis continuous_on_eq_continuous_at continuous_on_subset interiorE) |
| 33175 | 4731 |
|
4732 |
lemma continuous_on_eq: |
|
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4733 |
"(\<forall>x \<in> s. f x = g x) \<Longrightarrow> continuous_on s f \<Longrightarrow> continuous_on s g" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
4734 |
unfolding continuous_on_def tendsto_def eventually_at_topological |
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
4735 |
by simp |
| 33175 | 4736 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
4737 |
text {* Characterization of various kinds of continuity in terms of sequences. *}
|
| 33175 | 4738 |
|
4739 |
lemma continuous_within_sequentially: |
|
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4740 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 33175 | 4741 |
shows "continuous (at a within s) f \<longleftrightarrow> |
| 53282 | 4742 |
(\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially |
| 53640 | 4743 |
\<longrightarrow> ((f \<circ> x) ---> f a) sequentially)" |
| 53282 | 4744 |
(is "?lhs = ?rhs") |
| 33175 | 4745 |
proof |
4746 |
assume ?lhs |
|
| 53282 | 4747 |
{
|
4748 |
fix x :: "nat \<Rightarrow> 'a" |
|
4749 |
assume x: "\<forall>n. x n \<in> s" "\<forall>e>0. eventually (\<lambda>n. dist (x n) a < e) sequentially" |
|
4750 |
fix T :: "'b set" |
|
4751 |
assume "open T" and "f a \<in> T" |
|
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4752 |
with `?lhs` obtain d where "d>0" and d:"\<forall>x\<in>s. 0 < dist x a \<and> dist x a < d \<longrightarrow> f x \<in> T" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
4753 |
unfolding continuous_within tendsto_def eventually_at by (auto simp: dist_nz) |
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4754 |
have "eventually (\<lambda>n. dist (x n) a < d) sequentially" |
|
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4755 |
using x(2) `d>0` by simp |
| 53282 | 4756 |
then have "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially" |
| 46887 | 4757 |
proof eventually_elim |
| 53282 | 4758 |
case (elim n) |
4759 |
then show ?case |
|
| 53291 | 4760 |
using d x(1) `f a \<in> T` unfolding dist_nz[symmetric] by auto |
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4761 |
qed |
| 33175 | 4762 |
} |
| 53282 | 4763 |
then show ?rhs |
4764 |
unfolding tendsto_iff tendsto_def by simp |
|
| 33175 | 4765 |
next |
| 53282 | 4766 |
assume ?rhs |
4767 |
then show ?lhs |
|
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4768 |
unfolding continuous_within tendsto_def [where l="f a"] |
|
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4769 |
by (simp add: sequentially_imp_eventually_within) |
| 33175 | 4770 |
qed |
4771 |
||
4772 |
lemma continuous_at_sequentially: |
|
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4773 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 53291 | 4774 |
shows "continuous (at a) f \<longleftrightarrow> |
| 53640 | 4775 |
(\<forall>x. (x ---> a) sequentially --> ((f \<circ> x) ---> f a) sequentially)" |
| 45031 | 4776 |
using continuous_within_sequentially[of a UNIV f] by simp |
| 33175 | 4777 |
|
4778 |
lemma continuous_on_sequentially: |
|
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
4779 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 36359 | 4780 |
shows "continuous_on s f \<longleftrightarrow> |
4781 |
(\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially |
|
| 53640 | 4782 |
--> ((f \<circ> x) ---> f a) sequentially)" |
| 53291 | 4783 |
(is "?lhs = ?rhs") |
| 33175 | 4784 |
proof |
| 53282 | 4785 |
assume ?rhs |
4786 |
then show ?lhs |
|
4787 |
using continuous_within_sequentially[of _ s f] |
|
4788 |
unfolding continuous_on_eq_continuous_within |
|
4789 |
by auto |
|
| 33175 | 4790 |
next |
| 53282 | 4791 |
assume ?lhs |
4792 |
then show ?rhs |
|
4793 |
unfolding continuous_on_eq_continuous_within |
|
4794 |
using continuous_within_sequentially[of _ s f] |
|
4795 |
by auto |
|
| 33175 | 4796 |
qed |
4797 |
||
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4798 |
lemma uniformly_continuous_on_sequentially: |
| 36441 | 4799 |
"uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and> |
4800 |
((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially |
|
4801 |
\<longrightarrow> ((\<lambda>n. dist (f(x n)) (f(y n))) ---> 0) sequentially)" (is "?lhs = ?rhs") |
|
| 33175 | 4802 |
proof |
4803 |
assume ?lhs |
|
| 53282 | 4804 |
{
|
4805 |
fix x y |
|
4806 |
assume x: "\<forall>n. x n \<in> s" |
|
4807 |
and y: "\<forall>n. y n \<in> s" |
|
4808 |
and xy: "((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially" |
|
4809 |
{
|
|
4810 |
fix e :: real |
|
4811 |
assume "e > 0" |
|
4812 |
then obtain d where "d > 0" and d: "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" |
|
| 33175 | 4813 |
using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto |
| 53282 | 4814 |
obtain N where N: "\<forall>n\<ge>N. dist (x n) (y n) < d" |
4815 |
using xy[unfolded LIMSEQ_def dist_norm] and `d>0` by auto |
|
4816 |
{
|
|
4817 |
fix n |
|
4818 |
assume "n\<ge>N" |
|
4819 |
then have "dist (f (x n)) (f (y n)) < e" |
|
4820 |
using N[THEN spec[where x=n]] |
|
4821 |
using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] |
|
4822 |
using x and y |
|
4823 |
unfolding dist_commute |
|
4824 |
by simp |
|
4825 |
} |
|
4826 |
then have "\<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e" |
|
4827 |
by auto |
|
4828 |
} |
|
4829 |
then have "((\<lambda>n. dist (f(x n)) (f(y n))) ---> 0) sequentially" |
|
4830 |
unfolding LIMSEQ_def and dist_real_def by auto |
|
4831 |
} |
|
4832 |
then show ?rhs by auto |
|
| 33175 | 4833 |
next |
4834 |
assume ?rhs |
|
| 53282 | 4835 |
{
|
4836 |
assume "\<not> ?lhs" |
|
4837 |
then obtain e where "e > 0" "\<forall>d>0. \<exists>x\<in>s. \<exists>x'\<in>s. dist x' x < d \<and> \<not> dist (f x') (f x) < e" |
|
4838 |
unfolding uniformly_continuous_on_def by auto |
|
4839 |
then obtain fa where fa: |
|
4840 |
"\<forall>x. 0 < x \<longrightarrow> fst (fa x) \<in> s \<and> snd (fa x) \<in> s \<and> dist (fst (fa x)) (snd (fa x)) < x \<and> \<not> dist (f (fst (fa x))) (f (snd (fa x))) < e" |
|
4841 |
using choice[of "\<lambda>d x. d>0 \<longrightarrow> fst x \<in> s \<and> snd x \<in> s \<and> dist (snd x) (fst x) < d \<and> \<not> dist (f (snd x)) (f (fst x)) < e"] |
|
4842 |
unfolding Bex_def |
|
| 33175 | 4843 |
by (auto simp add: dist_commute) |
4844 |
def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))" |
|
4845 |
def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))" |
|
| 53282 | 4846 |
have xyn: "\<forall>n. x n \<in> s \<and> y n \<in> s" |
4847 |
and xy0: "\<forall>n. dist (x n) (y n) < inverse (real n + 1)" |
|
4848 |
and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e" |
|
4849 |
unfolding x_def and y_def using fa |
|
4850 |
by auto |
|
4851 |
{
|
|
4852 |
fix e :: real |
|
4853 |
assume "e > 0" |
|
4854 |
then obtain N :: nat where "N \<noteq> 0" and N: "0 < inverse (real N) \<and> inverse (real N) < e" |
|
4855 |
unfolding real_arch_inv[of e] by auto |
|
4856 |
{
|
|
4857 |
fix n :: nat |
|
4858 |
assume "n \<ge> N" |
|
4859 |
then have "inverse (real n + 1) < inverse (real N)" |
|
4860 |
using real_of_nat_ge_zero and `N\<noteq>0` by auto |
|
| 33175 | 4861 |
also have "\<dots> < e" using N by auto |
4862 |
finally have "inverse (real n + 1) < e" by auto |
|
| 53282 | 4863 |
then have "dist (x n) (y n) < e" |
4864 |
using xy0[THEN spec[where x=n]] by auto |
|
4865 |
} |
|
4866 |
then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto |
|
4867 |
} |
|
4868 |
then have "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e" |
|
4869 |
using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn |
|
4870 |
unfolding LIMSEQ_def dist_real_def by auto |
|
4871 |
then have False using fxy and `e>0` by auto |
|
4872 |
} |
|
4873 |
then show ?lhs |
|
4874 |
unfolding uniformly_continuous_on_def by blast |
|
| 33175 | 4875 |
qed |
4876 |
||
4877 |
text{* The usual transformation theorems. *}
|
|
4878 |
||
4879 |
lemma continuous_transform_within: |
|
| 36667 | 4880 |
fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 53282 | 4881 |
assumes "0 < d" |
4882 |
and "x \<in> s" |
|
4883 |
and "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'" |
|
4884 |
and "continuous (at x within s) f" |
|
| 33175 | 4885 |
shows "continuous (at x within s) g" |
| 53282 | 4886 |
unfolding continuous_within |
| 36667 | 4887 |
proof (rule Lim_transform_within) |
4888 |
show "0 < d" by fact |
|
4889 |
show "\<forall>x'\<in>s. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'" |
|
4890 |
using assms(3) by auto |
|
4891 |
have "f x = g x" |
|
4892 |
using assms(1,2,3) by auto |
|
| 53282 | 4893 |
then show "(f ---> g x) (at x within s)" |
| 36667 | 4894 |
using assms(4) unfolding continuous_within by simp |
| 33175 | 4895 |
qed |
4896 |
||
4897 |
lemma continuous_transform_at: |
|
| 36667 | 4898 |
fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 53282 | 4899 |
assumes "0 < d" |
4900 |
and "\<forall>x'. dist x' x < d --> f x' = g x'" |
|
4901 |
and "continuous (at x) f" |
|
| 33175 | 4902 |
shows "continuous (at x) g" |
| 45031 | 4903 |
using continuous_transform_within [of d x UNIV f g] assms by simp |
| 33175 | 4904 |
|
| 53282 | 4905 |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4906 |
subsubsection {* Structural rules for pointwise continuity *}
|
| 33175 | 4907 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
4908 |
lemmas continuous_within_id = continuous_ident |
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
4909 |
|
|
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
4910 |
lemmas continuous_at_id = isCont_ident |
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4911 |
|
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
4912 |
lemma continuous_infdist[continuous_intros]: |
| 50087 | 4913 |
assumes "continuous F f" |
4914 |
shows "continuous F (\<lambda>x. infdist (f x) A)" |
|
4915 |
using assms unfolding continuous_def by (rule tendsto_infdist) |
|
4916 |
||
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
4917 |
lemma continuous_infnorm[continuous_intros]: |
| 53282 | 4918 |
"continuous F f \<Longrightarrow> continuous F (\<lambda>x. infnorm (f x))" |
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4919 |
unfolding continuous_def by (rule tendsto_infnorm) |
| 33175 | 4920 |
|
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
4921 |
lemma continuous_inner[continuous_intros]: |
| 53282 | 4922 |
assumes "continuous F f" |
4923 |
and "continuous F g" |
|
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4924 |
shows "continuous F (\<lambda>x. inner (f x) (g x))" |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4925 |
using assms unfolding continuous_def by (rule tendsto_inner) |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4926 |
|
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
4927 |
lemmas continuous_at_inverse = isCont_inverse |
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4928 |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4929 |
subsubsection {* Structural rules for setwise continuity *}
|
| 33175 | 4930 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4931 |
lemma continuous_on_infnorm[continuous_intros]: |
| 53282 | 4932 |
"continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. infnorm (f x))" |
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4933 |
unfolding continuous_on by (fast intro: tendsto_infnorm) |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4934 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4935 |
lemma continuous_on_inner[continuous_intros]: |
|
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
4936 |
fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner" |
| 53282 | 4937 |
assumes "continuous_on s f" |
4938 |
and "continuous_on s g" |
|
|
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
4939 |
shows "continuous_on s (\<lambda>x. inner (f x) (g x))" |
|
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
4940 |
using bounded_bilinear_inner assms |
|
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
4941 |
by (rule bounded_bilinear.continuous_on) |
|
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
4942 |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4943 |
subsubsection {* Structural rules for uniform continuity *}
|
| 33175 | 4944 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4945 |
lemma uniformly_continuous_on_id[continuous_intros]: |
| 53282 | 4946 |
"uniformly_continuous_on s (\<lambda>x. x)" |
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4947 |
unfolding uniformly_continuous_on_def by auto |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
4948 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4949 |
lemma uniformly_continuous_on_const[continuous_intros]: |
| 53282 | 4950 |
"uniformly_continuous_on s (\<lambda>x. c)" |
| 33175 | 4951 |
unfolding uniformly_continuous_on_def by simp |
4952 |
||
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4953 |
lemma uniformly_continuous_on_dist[continuous_intros]: |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4954 |
fixes f g :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4955 |
assumes "uniformly_continuous_on s f" |
| 53282 | 4956 |
and "uniformly_continuous_on s g" |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4957 |
shows "uniformly_continuous_on s (\<lambda>x. dist (f x) (g x))" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4958 |
proof - |
| 53282 | 4959 |
{
|
4960 |
fix a b c d :: 'b |
|
4961 |
have "\<bar>dist a b - dist c d\<bar> \<le> dist a c + dist b d" |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4962 |
using dist_triangle2 [of a b c] dist_triangle2 [of b c d] |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4963 |
using dist_triangle3 [of c d a] dist_triangle [of a d b] |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4964 |
by arith |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4965 |
} note le = this |
| 53282 | 4966 |
{
|
4967 |
fix x y |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4968 |
assume f: "(\<lambda>n. dist (f (x n)) (f (y n))) ----> 0" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4969 |
assume g: "(\<lambda>n. dist (g (x n)) (g (y n))) ----> 0" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4970 |
have "(\<lambda>n. \<bar>dist (f (x n)) (g (x n)) - dist (f (y n)) (g (y n))\<bar>) ----> 0" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4971 |
by (rule Lim_transform_bound [OF _ tendsto_add_zero [OF f g]], |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4972 |
simp add: le) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4973 |
} |
| 53282 | 4974 |
then show ?thesis |
4975 |
using assms unfolding uniformly_continuous_on_sequentially |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4976 |
unfolding dist_real_def by simp |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4977 |
qed |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4978 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4979 |
lemma uniformly_continuous_on_norm[continuous_intros]: |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4980 |
assumes "uniformly_continuous_on s f" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4981 |
shows "uniformly_continuous_on s (\<lambda>x. norm (f x))" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4982 |
unfolding norm_conv_dist using assms |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4983 |
by (intro uniformly_continuous_on_dist uniformly_continuous_on_const) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4984 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4985 |
lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4986 |
assumes "uniformly_continuous_on s g" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4987 |
shows "uniformly_continuous_on s (\<lambda>x. f (g x))" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4988 |
using assms unfolding uniformly_continuous_on_sequentially |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4989 |
unfolding dist_norm tendsto_norm_zero_iff diff[symmetric] |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4990 |
by (auto intro: tendsto_zero) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4991 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
4992 |
lemma uniformly_continuous_on_cmul[continuous_intros]: |
| 36441 | 4993 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
| 33175 | 4994 |
assumes "uniformly_continuous_on s f" |
4995 |
shows "uniformly_continuous_on s (\<lambda>x. c *\<^sub>R f(x))" |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4996 |
using bounded_linear_scaleR_right assms |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
4997 |
by (rule bounded_linear.uniformly_continuous_on) |
| 33175 | 4998 |
|
4999 |
lemma dist_minus: |
|
5000 |
fixes x y :: "'a::real_normed_vector" |
|
5001 |
shows "dist (- x) (- y) = dist x y" |
|
5002 |
unfolding dist_norm minus_diff_minus norm_minus_cancel .. |
|
5003 |
||
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5004 |
lemma uniformly_continuous_on_minus[continuous_intros]: |
| 33175 | 5005 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5006 |
shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. - f x)" |
| 33175 | 5007 |
unfolding uniformly_continuous_on_def dist_minus . |
5008 |
||
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5009 |
lemma uniformly_continuous_on_add[continuous_intros]: |
| 36441 | 5010 |
fixes f g :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5011 |
assumes "uniformly_continuous_on s f" |
| 53282 | 5012 |
and "uniformly_continuous_on s g" |
| 33175 | 5013 |
shows "uniformly_continuous_on s (\<lambda>x. f x + g x)" |
| 53282 | 5014 |
using assms |
5015 |
unfolding uniformly_continuous_on_sequentially |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5016 |
unfolding dist_norm tendsto_norm_zero_iff add_diff_add |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5017 |
by (auto intro: tendsto_add_zero) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5018 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5019 |
lemma uniformly_continuous_on_diff[continuous_intros]: |
| 36441 | 5020 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
| 53282 | 5021 |
assumes "uniformly_continuous_on s f" |
5022 |
and "uniformly_continuous_on s g" |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5023 |
shows "uniformly_continuous_on s (\<lambda>x. f x - g x)" |
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
54070
diff
changeset
|
5024 |
using assms uniformly_continuous_on_add [of s f "- g"] |
|
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
54070
diff
changeset
|
5025 |
by (simp add: fun_Compl_def uniformly_continuous_on_minus) |
| 33175 | 5026 |
|
5027 |
text{* Continuity of all kinds is preserved under composition. *}
|
|
5028 |
||
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
5029 |
lemmas continuous_at_compose = isCont_o |
| 33175 | 5030 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5031 |
lemma uniformly_continuous_on_compose[continuous_intros]: |
| 33175 | 5032 |
assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g" |
| 53640 | 5033 |
shows "uniformly_continuous_on s (g \<circ> f)" |
5034 |
proof - |
|
| 53282 | 5035 |
{
|
5036 |
fix e :: real |
|
5037 |
assume "e > 0" |
|
5038 |
then obtain d where "d > 0" |
|
5039 |
and d: "\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" |
|
5040 |
using assms(2) unfolding uniformly_continuous_on_def by auto |
|
5041 |
obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d" |
|
5042 |
using `d > 0` using assms(1) unfolding uniformly_continuous_on_def by auto |
|
5043 |
then have "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist ((g \<circ> f) x') ((g \<circ> f) x) < e" |
|
5044 |
using `d>0` using d by auto |
|
5045 |
} |
|
5046 |
then show ?thesis |
|
5047 |
using assms unfolding uniformly_continuous_on_def by auto |
|
| 33175 | 5048 |
qed |
5049 |
||
5050 |
text{* Continuity in terms of open preimages. *}
|
|
5051 |
||
5052 |
lemma continuous_at_open: |
|
| 53282 | 5053 |
"continuous (at x) f \<longleftrightarrow> (\<forall>t. open t \<and> f x \<in> t --> (\<exists>s. open s \<and> x \<in> s \<and> (\<forall>x' \<in> s. (f x') \<in> t)))" |
5054 |
unfolding continuous_within_topological [of x UNIV f] |
|
5055 |
unfolding imp_conjL |
|
5056 |
by (intro all_cong imp_cong ex_cong conj_cong refl) auto |
|
| 33175 | 5057 |
|
| 51351 | 5058 |
lemma continuous_imp_tendsto: |
| 53282 | 5059 |
assumes "continuous (at x0) f" |
5060 |
and "x ----> x0" |
|
| 51351 | 5061 |
shows "(f \<circ> x) ----> (f x0)" |
5062 |
proof (rule topological_tendstoI) |
|
5063 |
fix S |
|
5064 |
assume "open S" "f x0 \<in> S" |
|
5065 |
then obtain T where T_def: "open T" "x0 \<in> T" "\<forall>x\<in>T. f x \<in> S" |
|
5066 |
using assms continuous_at_open by metis |
|
5067 |
then have "eventually (\<lambda>n. x n \<in> T) sequentially" |
|
5068 |
using assms T_def by (auto simp: tendsto_def) |
|
5069 |
then show "eventually (\<lambda>n. (f \<circ> x) n \<in> S) sequentially" |
|
5070 |
using T_def by (auto elim!: eventually_elim1) |
|
5071 |
qed |
|
5072 |
||
| 33175 | 5073 |
lemma continuous_on_open: |
|
51481
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5074 |
"continuous_on s f \<longleftrightarrow> |
| 53282 | 5075 |
(\<forall>t. openin (subtopology euclidean (f ` s)) t \<longrightarrow> |
5076 |
openin (subtopology euclidean s) {x \<in> s. f x \<in> t})"
|
|
|
51481
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5077 |
unfolding continuous_on_open_invariant openin_open Int_def vimage_def Int_commute |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5078 |
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) |
| 36441 | 5079 |
|
5080 |
text {* Similarly in terms of closed sets. *}
|
|
| 33175 | 5081 |
|
5082 |
lemma continuous_on_closed: |
|
| 53282 | 5083 |
"continuous_on s f \<longleftrightarrow> |
5084 |
(\<forall>t. closedin (subtopology euclidean (f ` s)) t \<longrightarrow> |
|
5085 |
closedin (subtopology euclidean s) {x \<in> s. f x \<in> t})"
|
|
|
51481
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5086 |
unfolding continuous_on_closed_invariant closedin_closed Int_def vimage_def Int_commute |
|
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5087 |
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) |
| 33175 | 5088 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5089 |
text {* Half-global and completely global cases. *}
|
| 33175 | 5090 |
|
5091 |
lemma continuous_open_in_preimage: |
|
5092 |
assumes "continuous_on s f" "open t" |
|
5093 |
shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
|
|
| 53282 | 5094 |
proof - |
5095 |
have *: "\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" |
|
5096 |
by auto |
|
| 33175 | 5097 |
have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)" |
5098 |
using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto |
|
| 53282 | 5099 |
then show ?thesis |
5100 |
using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]] |
|
5101 |
using * by auto |
|
| 33175 | 5102 |
qed |
5103 |
||
5104 |
lemma continuous_closed_in_preimage: |
|
| 53291 | 5105 |
assumes "continuous_on s f" and "closed t" |
| 33175 | 5106 |
shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
|
| 53282 | 5107 |
proof - |
5108 |
have *: "\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)" |
|
5109 |
by auto |
|
| 33175 | 5110 |
have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)" |
| 53282 | 5111 |
using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute |
5112 |
by auto |
|
5113 |
then show ?thesis |
|
5114 |
using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]] |
|
5115 |
using * by auto |
|
| 33175 | 5116 |
qed |
5117 |
||
5118 |
lemma continuous_open_preimage: |
|
| 53291 | 5119 |
assumes "continuous_on s f" |
5120 |
and "open s" |
|
5121 |
and "open t" |
|
| 33175 | 5122 |
shows "open {x \<in> s. f x \<in> t}"
|
5123 |
proof- |
|
5124 |
obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
|
|
5125 |
using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto |
|
| 53282 | 5126 |
then show ?thesis |
5127 |
using open_Int[of s T, OF assms(2)] by auto |
|
| 33175 | 5128 |
qed |
5129 |
||
5130 |
lemma continuous_closed_preimage: |
|
| 53291 | 5131 |
assumes "continuous_on s f" |
5132 |
and "closed s" |
|
5133 |
and "closed t" |
|
| 33175 | 5134 |
shows "closed {x \<in> s. f x \<in> t}"
|
5135 |
proof- |
|
| 53282 | 5136 |
obtain T where "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
|
5137 |
using continuous_closed_in_preimage[OF assms(1,3)] |
|
5138 |
unfolding closedin_closed by auto |
|
5139 |
then show ?thesis using closed_Int[of s T, OF assms(2)] by auto |
|
| 33175 | 5140 |
qed |
5141 |
||
5142 |
lemma continuous_open_preimage_univ: |
|
| 53282 | 5143 |
"\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
|
| 33175 | 5144 |
using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto |
5145 |
||
5146 |
lemma continuous_closed_preimage_univ: |
|
| 53291 | 5147 |
"(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s \<Longrightarrow> closed {x. f x \<in> s}"
|
| 33175 | 5148 |
using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto |
5149 |
||
| 53282 | 5150 |
lemma continuous_open_vimage: "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)" |
| 33175 | 5151 |
unfolding vimage_def by (rule continuous_open_preimage_univ) |
5152 |
||
| 53282 | 5153 |
lemma continuous_closed_vimage: "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)" |
| 33175 | 5154 |
unfolding vimage_def by (rule continuous_closed_preimage_univ) |
5155 |
||
| 36441 | 5156 |
lemma interior_image_subset: |
| 53291 | 5157 |
assumes "\<forall>x. continuous (at x) f" |
5158 |
and "inj f" |
|
|
35172
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents:
35028
diff
changeset
|
5159 |
shows "interior (f ` s) \<subseteq> f ` (interior s)" |
| 44519 | 5160 |
proof |
5161 |
fix x assume "x \<in> interior (f ` s)" |
|
5162 |
then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` s" .. |
|
| 53282 | 5163 |
then have "x \<in> f ` s" by auto |
| 44519 | 5164 |
then obtain y where y: "y \<in> s" "x = f y" by auto |
5165 |
have "open (vimage f T)" |
|
5166 |
using assms(1) `open T` by (rule continuous_open_vimage) |
|
5167 |
moreover have "y \<in> vimage f T" |
|
5168 |
using `x = f y` `x \<in> T` by simp |
|
5169 |
moreover have "vimage f T \<subseteq> s" |
|
5170 |
using `T \<subseteq> image f s` `inj f` unfolding inj_on_def subset_eq by auto |
|
5171 |
ultimately have "y \<in> interior s" .. |
|
5172 |
with `x = f y` show "x \<in> f ` interior s" .. |
|
5173 |
qed |
|
|
35172
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
himmelma
parents:
35028
diff
changeset
|
5174 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5175 |
text {* Equality of continuous functions on closure and related results. *}
|
| 33175 | 5176 |
|
5177 |
lemma continuous_closed_in_preimage_constant: |
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5178 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
| 53291 | 5179 |
shows "continuous_on s f \<Longrightarrow> closedin (subtopology euclidean s) {x \<in> s. f x = a}"
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5180 |
using continuous_closed_in_preimage[of s f "{a}"] by auto
|
| 33175 | 5181 |
|
5182 |
lemma continuous_closed_preimage_constant: |
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5183 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
| 53291 | 5184 |
shows "continuous_on s f \<Longrightarrow> closed s \<Longrightarrow> closed {x \<in> s. f x = a}"
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5185 |
using continuous_closed_preimage[of s f "{a}"] by auto
|
| 33175 | 5186 |
|
5187 |
lemma continuous_constant_on_closure: |
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5188 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
| 33175 | 5189 |
assumes "continuous_on (closure s) f" |
| 53282 | 5190 |
and "\<forall>x \<in> s. f x = a" |
| 33175 | 5191 |
shows "\<forall>x \<in> (closure s). f x = a" |
5192 |
using continuous_closed_preimage_constant[of "closure s" f a] |
|
| 53282 | 5193 |
assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset
|
5194 |
unfolding subset_eq |
|
5195 |
by auto |
|
| 33175 | 5196 |
|
5197 |
lemma image_closure_subset: |
|
| 53291 | 5198 |
assumes "continuous_on (closure s) f" |
5199 |
and "closed t" |
|
5200 |
and "(f ` s) \<subseteq> t" |
|
| 33175 | 5201 |
shows "f ` (closure s) \<subseteq> t" |
| 53282 | 5202 |
proof - |
5203 |
have "s \<subseteq> {x \<in> closure s. f x \<in> t}"
|
|
5204 |
using assms(3) closure_subset by auto |
|
| 33175 | 5205 |
moreover have "closed {x \<in> closure s. f x \<in> t}"
|
5206 |
using continuous_closed_preimage[OF assms(1)] and assms(2) by auto |
|
5207 |
ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
|
|
5208 |
using closure_minimal[of s "{x \<in> closure s. f x \<in> t}"] by auto
|
|
| 53282 | 5209 |
then show ?thesis by auto |
| 33175 | 5210 |
qed |
5211 |
||
5212 |
lemma continuous_on_closure_norm_le: |
|
5213 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
|
| 53282 | 5214 |
assumes "continuous_on (closure s) f" |
5215 |
and "\<forall>y \<in> s. norm(f y) \<le> b" |
|
5216 |
and "x \<in> (closure s)" |
|
| 53291 | 5217 |
shows "norm (f x) \<le> b" |
| 53282 | 5218 |
proof - |
5219 |
have *: "f ` s \<subseteq> cball 0 b" |
|
| 53291 | 5220 |
using assms(2)[unfolded mem_cball_0[symmetric]] by auto |
| 33175 | 5221 |
show ?thesis |
5222 |
using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3) |
|
| 53282 | 5223 |
unfolding subset_eq |
5224 |
apply (erule_tac x="f x" in ballE) |
|
5225 |
apply (auto simp add: dist_norm) |
|
5226 |
done |
|
| 33175 | 5227 |
qed |
5228 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5229 |
text {* Making a continuous function avoid some value in a neighbourhood. *}
|
| 33175 | 5230 |
|
5231 |
lemma continuous_within_avoid: |
|
| 50898 | 5232 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" |
| 53282 | 5233 |
assumes "continuous (at x within s) f" |
5234 |
and "f x \<noteq> a" |
|
| 33175 | 5235 |
shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a" |
| 53291 | 5236 |
proof - |
| 50898 | 5237 |
obtain U where "open U" and "f x \<in> U" and "a \<notin> U" |
5238 |
using t1_space [OF `f x \<noteq> a`] by fast |
|
5239 |
have "(f ---> f x) (at x within s)" |
|
5240 |
using assms(1) by (simp add: continuous_within) |
|
| 53282 | 5241 |
then have "eventually (\<lambda>y. f y \<in> U) (at x within s)" |
| 50898 | 5242 |
using `open U` and `f x \<in> U` |
5243 |
unfolding tendsto_def by fast |
|
| 53282 | 5244 |
then have "eventually (\<lambda>y. f y \<noteq> a) (at x within s)" |
| 50898 | 5245 |
using `a \<notin> U` by (fast elim: eventually_mono [rotated]) |
| 53282 | 5246 |
then show ?thesis |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
5247 |
using `f x \<noteq> a` by (auto simp: dist_commute zero_less_dist_iff eventually_at) |
| 33175 | 5248 |
qed |
5249 |
||
5250 |
lemma continuous_at_avoid: |
|
| 50898 | 5251 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" |
| 53282 | 5252 |
assumes "continuous (at x) f" |
5253 |
and "f x \<noteq> a" |
|
| 33175 | 5254 |
shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a" |
| 45031 | 5255 |
using assms continuous_within_avoid[of x UNIV f a] by simp |
| 33175 | 5256 |
|
5257 |
lemma continuous_on_avoid: |
|
| 50898 | 5258 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" |
| 53282 | 5259 |
assumes "continuous_on s f" |
5260 |
and "x \<in> s" |
|
5261 |
and "f x \<noteq> a" |
|
| 33175 | 5262 |
shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a" |
| 53282 | 5263 |
using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x], |
5264 |
OF assms(2)] continuous_within_avoid[of x s f a] |
|
5265 |
using assms(3) |
|
5266 |
by auto |
|
| 33175 | 5267 |
|
5268 |
lemma continuous_on_open_avoid: |
|
| 50898 | 5269 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space" |
| 53291 | 5270 |
assumes "continuous_on s f" |
5271 |
and "open s" |
|
5272 |
and "x \<in> s" |
|
5273 |
and "f x \<noteq> a" |
|
| 33175 | 5274 |
shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a" |
| 53282 | 5275 |
using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)] |
5276 |
using continuous_at_avoid[of x f a] assms(4) |
|
5277 |
by auto |
|
| 33175 | 5278 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5279 |
text {* Proving a function is constant by proving open-ness of level set. *}
|
| 33175 | 5280 |
|
5281 |
lemma continuous_levelset_open_in_cases: |
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5282 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
| 36359 | 5283 |
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow> |
| 33175 | 5284 |
openin (subtopology euclidean s) {x \<in> s. f x = a}
|
| 53282 | 5285 |
\<Longrightarrow> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)" |
5286 |
unfolding connected_clopen |
|
5287 |
using continuous_closed_in_preimage_constant by auto |
|
| 33175 | 5288 |
|
5289 |
lemma continuous_levelset_open_in: |
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5290 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
| 36359 | 5291 |
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow> |
| 33175 | 5292 |
openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
|
| 53291 | 5293 |
(\<exists>x \<in> s. f x = a) \<Longrightarrow> (\<forall>x \<in> s. f x = a)" |
| 53282 | 5294 |
using continuous_levelset_open_in_cases[of s f ] |
5295 |
by meson |
|
| 33175 | 5296 |
|
5297 |
lemma continuous_levelset_open: |
|
|
36668
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
huffman
parents:
36667
diff
changeset
|
5298 |
fixes f :: "_ \<Rightarrow> 'b::t1_space" |
| 53282 | 5299 |
assumes "connected s" |
5300 |
and "continuous_on s f" |
|
5301 |
and "open {x \<in> s. f x = a}"
|
|
5302 |
and "\<exists>x \<in> s. f x = a" |
|
| 33175 | 5303 |
shows "\<forall>x \<in> s. f x = a" |
| 53282 | 5304 |
using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open] |
5305 |
using assms (3,4) |
|
5306 |
by fast |
|
| 33175 | 5307 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5308 |
text {* Some arithmetical combinations (more to prove). *}
|
| 33175 | 5309 |
|
5310 |
lemma open_scaling[intro]: |
|
5311 |
fixes s :: "'a::real_normed_vector set" |
|
| 53291 | 5312 |
assumes "c \<noteq> 0" |
5313 |
and "open s" |
|
| 33175 | 5314 |
shows "open((\<lambda>x. c *\<^sub>R x) ` s)" |
| 53282 | 5315 |
proof - |
5316 |
{
|
|
5317 |
fix x |
|
5318 |
assume "x \<in> s" |
|
5319 |
then obtain e where "e>0" |
|
5320 |
and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]] |
|
5321 |
by auto |
|
5322 |
have "e * abs c > 0" |
|
| 56544 | 5323 |
using assms(1)[unfolded zero_less_abs_iff[symmetric]] `e>0` by auto |
| 33175 | 5324 |
moreover |
| 53282 | 5325 |
{
|
5326 |
fix y |
|
5327 |
assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>" |
|
5328 |
then have "norm ((1 / c) *\<^sub>R y - x) < e" |
|
5329 |
unfolding dist_norm |
|
| 33175 | 5330 |
using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1) |
| 53291 | 5331 |
assms(1)[unfolded zero_less_abs_iff[symmetric]] by (simp del:zero_less_abs_iff) |
| 53282 | 5332 |
then have "y \<in> op *\<^sub>R c ` s" |
5333 |
using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"] |
|
5334 |
using e[THEN spec[where x="(1 / c) *\<^sub>R y"]] |
|
5335 |
using assms(1) |
|
5336 |
unfolding dist_norm scaleR_scaleR |
|
5337 |
by auto |
|
5338 |
} |
|
5339 |
ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s" |
|
5340 |
apply (rule_tac x="e * abs c" in exI) |
|
5341 |
apply auto |
|
5342 |
done |
|
5343 |
} |
|
5344 |
then show ?thesis unfolding open_dist by auto |
|
| 33175 | 5345 |
qed |
5346 |
||
5347 |
lemma minus_image_eq_vimage: |
|
5348 |
fixes A :: "'a::ab_group_add set" |
|
5349 |
shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A" |
|
5350 |
by (auto intro!: image_eqI [where f="\<lambda>x. - x"]) |
|
5351 |
||
5352 |
lemma open_negations: |
|
5353 |
fixes s :: "'a::real_normed_vector set" |
|
|
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54263
diff
changeset
|
5354 |
shows "open s \<Longrightarrow> open ((\<lambda>x. - x) ` s)" |
|
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
54263
diff
changeset
|
5355 |
using open_scaling [of "- 1" s] by simp |
| 33175 | 5356 |
|
5357 |
lemma open_translation: |
|
5358 |
fixes s :: "'a::real_normed_vector set" |
|
| 53291 | 5359 |
assumes "open s" |
5360 |
shows "open((\<lambda>x. a + x) ` s)" |
|
| 53282 | 5361 |
proof - |
5362 |
{
|
|
5363 |
fix x |
|
5364 |
have "continuous (at x) (\<lambda>x. x - a)" |
|
5365 |
by (intro continuous_diff continuous_at_id continuous_const) |
|
5366 |
} |
|
5367 |
moreover have "{x. x - a \<in> s} = op + a ` s"
|
|
5368 |
by force |
|
5369 |
ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s] |
|
5370 |
using assms by auto |
|
| 33175 | 5371 |
qed |
5372 |
||
5373 |
lemma open_affinity: |
|
5374 |
fixes s :: "'a::real_normed_vector set" |
|
5375 |
assumes "open s" "c \<noteq> 0" |
|
5376 |
shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)" |
|
| 53282 | 5377 |
proof - |
5378 |
have *: "(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)" |
|
5379 |
unfolding o_def .. |
|
5380 |
have "op + a ` op *\<^sub>R c ` s = (op + a \<circ> op *\<^sub>R c) ` s" |
|
5381 |
by auto |
|
5382 |
then show ?thesis |
|
5383 |
using assms open_translation[of "op *\<^sub>R c ` s" a] |
|
5384 |
unfolding * |
|
5385 |
by auto |
|
| 33175 | 5386 |
qed |
5387 |
||
5388 |
lemma interior_translation: |
|
5389 |
fixes s :: "'a::real_normed_vector set" |
|
5390 |
shows "interior ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (interior s)" |
|
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
5391 |
proof (rule set_eqI, rule) |
| 53282 | 5392 |
fix x |
5393 |
assume "x \<in> interior (op + a ` s)" |
|
5394 |
then obtain e where "e > 0" and e: "ball x e \<subseteq> op + a ` s" |
|
5395 |
unfolding mem_interior by auto |
|
5396 |
then have "ball (x - a) e \<subseteq> s" |
|
5397 |
unfolding subset_eq Ball_def mem_ball dist_norm |
|
5398 |
apply auto |
|
5399 |
apply (erule_tac x="a + xa" in allE) |
|
| 53291 | 5400 |
unfolding ab_group_add_class.diff_diff_eq[symmetric] |
| 53282 | 5401 |
apply auto |
5402 |
done |
|
5403 |
then show "x \<in> op + a ` interior s" |
|
5404 |
unfolding image_iff |
|
5405 |
apply (rule_tac x="x - a" in bexI) |
|
5406 |
unfolding mem_interior |
|
5407 |
using `e > 0` |
|
5408 |
apply auto |
|
5409 |
done |
|
| 33175 | 5410 |
next |
| 53282 | 5411 |
fix x |
5412 |
assume "x \<in> op + a ` interior s" |
|
5413 |
then obtain y e where "e > 0" and e: "ball y e \<subseteq> s" and y: "x = a + y" |
|
5414 |
unfolding image_iff Bex_def mem_interior by auto |
|
5415 |
{
|
|
5416 |
fix z |
|
5417 |
have *: "a + y - z = y + a - z" by auto |
|
5418 |
assume "z \<in> ball x e" |
|
5419 |
then have "z - a \<in> s" |
|
5420 |
using e[unfolded subset_eq, THEN bspec[where x="z - a"]] |
|
5421 |
unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 * |
|
5422 |
by auto |
|
5423 |
then have "z \<in> op + a ` s" |
|
5424 |
unfolding image_iff by (auto intro!: bexI[where x="z - a"]) |
|
5425 |
} |
|
5426 |
then have "ball x e \<subseteq> op + a ` s" |
|
5427 |
unfolding subset_eq by auto |
|
5428 |
then show "x \<in> interior (op + a ` s)" |
|
5429 |
unfolding mem_interior using `e > 0` by auto |
|
| 33175 | 5430 |
qed |
5431 |
||
| 36437 | 5432 |
text {* Topological properties of linear functions. *}
|
5433 |
||
5434 |
lemma linear_lim_0: |
|
| 53282 | 5435 |
assumes "bounded_linear f" |
5436 |
shows "(f ---> 0) (at (0))" |
|
5437 |
proof - |
|
| 36437 | 5438 |
interpret f: bounded_linear f by fact |
5439 |
have "(f ---> f 0) (at 0)" |
|
5440 |
using tendsto_ident_at by (rule f.tendsto) |
|
| 53282 | 5441 |
then show ?thesis unfolding f.zero . |
| 36437 | 5442 |
qed |
5443 |
||
5444 |
lemma linear_continuous_at: |
|
| 53282 | 5445 |
assumes "bounded_linear f" |
5446 |
shows "continuous (at a) f" |
|
| 36437 | 5447 |
unfolding continuous_at using assms |
5448 |
apply (rule bounded_linear.tendsto) |
|
5449 |
apply (rule tendsto_ident_at) |
|
5450 |
done |
|
5451 |
||
5452 |
lemma linear_continuous_within: |
|
| 53291 | 5453 |
"bounded_linear f \<Longrightarrow> continuous (at x within s) f" |
| 36437 | 5454 |
using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto |
5455 |
||
5456 |
lemma linear_continuous_on: |
|
| 53291 | 5457 |
"bounded_linear f \<Longrightarrow> continuous_on s f" |
| 36437 | 5458 |
using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto |
5459 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5460 |
text {* Also bilinear functions, in composition form. *}
|
| 36437 | 5461 |
|
5462 |
lemma bilinear_continuous_at_compose: |
|
| 53282 | 5463 |
"continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h \<Longrightarrow> |
5464 |
continuous (at x) (\<lambda>x. h (f x) (g x))" |
|
5465 |
unfolding continuous_at |
|
5466 |
using Lim_bilinear[of f "f x" "(at x)" g "g x" h] |
|
5467 |
by auto |
|
| 36437 | 5468 |
|
5469 |
lemma bilinear_continuous_within_compose: |
|
| 53282 | 5470 |
"continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bounded_bilinear h \<Longrightarrow> |
5471 |
continuous (at x within s) (\<lambda>x. h (f x) (g x))" |
|
5472 |
unfolding continuous_within |
|
5473 |
using Lim_bilinear[of f "f x"] |
|
5474 |
by auto |
|
| 36437 | 5475 |
|
5476 |
lemma bilinear_continuous_on_compose: |
|
| 53282 | 5477 |
"continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h \<Longrightarrow> |
5478 |
continuous_on s (\<lambda>x. h (f x) (g x))" |
|
| 36441 | 5479 |
unfolding continuous_on_def |
5480 |
by (fast elim: bounded_bilinear.tendsto) |
|
| 36437 | 5481 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5482 |
text {* Preservation of compactness and connectedness under continuous function. *}
|
| 33175 | 5483 |
|
| 50898 | 5484 |
lemma compact_eq_openin_cover: |
5485 |
"compact S \<longleftrightarrow> |
|
5486 |
(\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow> |
|
5487 |
(\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))" |
|
5488 |
proof safe |
|
5489 |
fix C |
|
5490 |
assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C" |
|
| 53282 | 5491 |
then have "\<forall>c\<in>{T. open T \<and> S \<inter> T \<in> C}. open c" and "S \<subseteq> \<Union>{T. open T \<and> S \<inter> T \<in> C}"
|
| 50898 | 5492 |
unfolding openin_open by force+ |
5493 |
with `compact S` obtain D where "D \<subseteq> {T. open T \<and> S \<inter> T \<in> C}" and "finite D" and "S \<subseteq> \<Union>D"
|
|
5494 |
by (rule compactE) |
|
| 53282 | 5495 |
then have "image (\<lambda>T. S \<inter> T) D \<subseteq> C \<and> finite (image (\<lambda>T. S \<inter> T) D) \<and> S \<subseteq> \<Union>(image (\<lambda>T. S \<inter> T) D)" |
| 50898 | 5496 |
by auto |
| 53282 | 5497 |
then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" .. |
| 50898 | 5498 |
next |
5499 |
assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow> |
|
5500 |
(\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)" |
|
5501 |
show "compact S" |
|
5502 |
proof (rule compactI) |
|
5503 |
fix C |
|
5504 |
let ?C = "image (\<lambda>T. S \<inter> T) C" |
|
5505 |
assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C" |
|
| 53282 | 5506 |
then have "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C" |
| 50898 | 5507 |
unfolding openin_open by auto |
5508 |
with 1 obtain D where "D \<subseteq> ?C" and "finite D" and "S \<subseteq> \<Union>D" |
|
5509 |
by metis |
|
5510 |
let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D" |
|
5511 |
have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D" |
|
5512 |
proof (intro conjI) |
|
5513 |
from `D \<subseteq> ?C` show "?D \<subseteq> C" |
|
5514 |
by (fast intro: inv_into_into) |
|
5515 |
from `finite D` show "finite ?D" |
|
5516 |
by (rule finite_imageI) |
|
5517 |
from `S \<subseteq> \<Union>D` show "S \<subseteq> \<Union>?D" |
|
5518 |
apply (rule subset_trans) |
|
5519 |
apply clarsimp |
|
5520 |
apply (frule subsetD [OF `D \<subseteq> ?C`, THEN f_inv_into_f]) |
|
5521 |
apply (erule rev_bexI, fast) |
|
5522 |
done |
|
5523 |
qed |
|
| 53282 | 5524 |
then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" .. |
| 50898 | 5525 |
qed |
5526 |
qed |
|
5527 |
||
| 33175 | 5528 |
lemma connected_continuous_image: |
| 53291 | 5529 |
assumes "continuous_on s f" |
5530 |
and "connected s" |
|
| 33175 | 5531 |
shows "connected(f ` s)" |
| 53282 | 5532 |
proof - |
5533 |
{
|
|
5534 |
fix T |
|
| 53291 | 5535 |
assume as: |
5536 |
"T \<noteq> {}"
|
|
5537 |
"T \<noteq> f ` s" |
|
5538 |
"openin (subtopology euclidean (f ` s)) T" |
|
5539 |
"closedin (subtopology euclidean (f ` s)) T" |
|
| 33175 | 5540 |
have "{x \<in> s. f x \<in> T} = {} \<or> {x \<in> s. f x \<in> T} = s"
|
5541 |
using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]] |
|
5542 |
using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]] |
|
5543 |
using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
|
|
| 53282 | 5544 |
then have False using as(1,2) |
5545 |
using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto |
|
5546 |
} |
|
5547 |
then show ?thesis |
|
5548 |
unfolding connected_clopen by auto |
|
| 33175 | 5549 |
qed |
5550 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5551 |
text {* Continuity implies uniform continuity on a compact domain. *}
|
| 53282 | 5552 |
|
| 33175 | 5553 |
lemma compact_uniformly_continuous: |
| 53291 | 5554 |
assumes f: "continuous_on s f" |
5555 |
and s: "compact s" |
|
| 33175 | 5556 |
shows "uniformly_continuous_on s f" |
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5557 |
unfolding uniformly_continuous_on_def |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5558 |
proof (cases, safe) |
| 53282 | 5559 |
fix e :: real |
5560 |
assume "0 < e" "s \<noteq> {}"
|
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5561 |
def [simp]: R \<equiv> "{(y, d). y \<in> s \<and> 0 < d \<and> ball y d \<inter> s \<subseteq> {x \<in> s. f x \<in> ball (f y) (e/2) } }"
|
| 50944 | 5562 |
let ?b = "(\<lambda>(y, d). ball y (d/2))" |
5563 |
have "(\<forall>r\<in>R. open (?b r))" "s \<subseteq> (\<Union>r\<in>R. ?b r)" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5564 |
proof safe |
| 53282 | 5565 |
fix y |
5566 |
assume "y \<in> s" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5567 |
from continuous_open_in_preimage[OF f open_ball] |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5568 |
obtain T where "open T" and T: "{x \<in> s. f x \<in> ball (f y) (e/2)} = T \<inter> s"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5569 |
unfolding openin_subtopology open_openin by metis |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5570 |
then obtain d where "ball y d \<subseteq> T" "0 < d" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5571 |
using `0 < e` `y \<in> s` by (auto elim!: openE) |
| 50944 | 5572 |
with T `y \<in> s` show "y \<in> (\<Union>r\<in>R. ?b r)" |
5573 |
by (intro UN_I[of "(y, d)"]) auto |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5574 |
qed auto |
| 50944 | 5575 |
with s obtain D where D: "finite D" "D \<subseteq> R" "s \<subseteq> (\<Union>(y, d)\<in>D. ball y (d/2))" |
5576 |
by (rule compactE_image) |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5577 |
with `s \<noteq> {}` have [simp]: "\<And>x. x < Min (snd ` D) \<longleftrightarrow> (\<forall>(y, d)\<in>D. x < d)"
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5578 |
by (subst Min_gr_iff) auto |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5579 |
show "\<exists>d>0. \<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5580 |
proof (rule, safe) |
| 53282 | 5581 |
fix x x' |
5582 |
assume in_s: "x' \<in> s" "x \<in> s" |
|
|
50943
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5583 |
with D obtain y d where x: "x \<in> ball y (d/2)" "(y, d) \<in> D" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5584 |
by blast |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5585 |
moreover assume "dist x x' < Min (snd`D) / 2" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5586 |
ultimately have "dist y x' < d" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5587 |
by (intro dist_double[where x=x and d=d]) (auto simp: dist_commute) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5588 |
with D x in_s show "dist (f x) (f x') < e" |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5589 |
by (intro dist_double[where x="f y" and d=e]) (auto simp: dist_commute subset_eq) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5590 |
qed (insert D, auto) |
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
5591 |
qed auto |
| 33175 | 5592 |
|
| 36437 | 5593 |
text {* A uniformly convergent limit of continuous functions is continuous. *}
|
| 33175 | 5594 |
|
5595 |
lemma continuous_uniform_limit: |
|
|
44212
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5596 |
fixes f :: "'a \<Rightarrow> 'b::metric_space \<Rightarrow> 'c::metric_space" |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5597 |
assumes "\<not> trivial_limit F" |
| 53282 | 5598 |
and "eventually (\<lambda>n. continuous_on s (f n)) F" |
5599 |
and "\<forall>e>0. eventually (\<lambda>n. \<forall>x\<in>s. dist (f n x) (g x) < e) F" |
|
| 33175 | 5600 |
shows "continuous_on s g" |
| 53282 | 5601 |
proof - |
5602 |
{
|
|
5603 |
fix x and e :: real |
|
5604 |
assume "x\<in>s" "e>0" |
|
|
44212
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5605 |
have "eventually (\<lambda>n. \<forall>x\<in>s. dist (f n x) (g x) < e / 3) F" |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5606 |
using `e>0` assms(3)[THEN spec[where x="e/3"]] by auto |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5607 |
from eventually_happens [OF eventually_conj [OF this assms(2)]] |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5608 |
obtain n where n:"\<forall>x\<in>s. dist (f n x) (g x) < e / 3" "continuous_on s (f n)" |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5609 |
using assms(1) by blast |
| 33175 | 5610 |
have "e / 3 > 0" using `e>0` by auto |
5611 |
then obtain d where "d>0" and d:"\<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f n x') (f n x) < e / 3" |
|
| 36359 | 5612 |
using n(2)[unfolded continuous_on_iff, THEN bspec[where x=x], OF `x\<in>s`, THEN spec[where x="e/3"]] by blast |
| 53282 | 5613 |
{
|
5614 |
fix y |
|
5615 |
assume "y \<in> s" and "dist y x < d" |
|
5616 |
then have "dist (f n y) (f n x) < e / 3" |
|
|
44212
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5617 |
by (rule d [rule_format]) |
| 53282 | 5618 |
then have "dist (f n y) (g x) < 2 * e / 3" |
|
44212
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5619 |
using dist_triangle [of "f n y" "g x" "f n x"] |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5620 |
using n(1)[THEN bspec[where x=x], OF `x\<in>s`] |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5621 |
by auto |
| 53282 | 5622 |
then have "dist (g y) (g x) < e" |
|
44212
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5623 |
using n(1)[THEN bspec[where x=y], OF `y\<in>s`] |
|
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
huffman
parents:
44211
diff
changeset
|
5624 |
using dist_triangle3 [of "g y" "g x" "f n y"] |
| 53282 | 5625 |
by auto |
5626 |
} |
|
5627 |
then have "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e" |
|
5628 |
using `d>0` by auto |
|
5629 |
} |
|
5630 |
then show ?thesis |
|
5631 |
unfolding continuous_on_iff by auto |
|
| 33175 | 5632 |
qed |
5633 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5634 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5635 |
subsection {* Topological stuff lifted from and dropped to R *}
|
| 33175 | 5636 |
|
5637 |
lemma open_real: |
|
| 53282 | 5638 |
fixes s :: "real set" |
5639 |
shows "open s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)" |
|
| 33175 | 5640 |
unfolding open_dist dist_norm by simp |
5641 |
||
5642 |
lemma islimpt_approachable_real: |
|
5643 |
fixes s :: "real set" |
|
5644 |
shows "x islimpt s \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)" |
|
5645 |
unfolding islimpt_approachable dist_norm by simp |
|
5646 |
||
5647 |
lemma closed_real: |
|
5648 |
fixes s :: "real set" |
|
| 53282 | 5649 |
shows "closed s \<longleftrightarrow> (\<forall>x. (\<forall>e>0. \<exists>x' \<in> s. x' \<noteq> x \<and> abs(x' - x) < e) \<longrightarrow> x \<in> s)" |
| 33175 | 5650 |
unfolding closed_limpt islimpt_approachable dist_norm by simp |
5651 |
||
5652 |
lemma continuous_at_real_range: |
|
5653 |
fixes f :: "'a::real_normed_vector \<Rightarrow> real" |
|
| 53282 | 5654 |
shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)" |
5655 |
unfolding continuous_at |
|
5656 |
unfolding Lim_at |
|
| 53291 | 5657 |
unfolding dist_nz[symmetric] |
| 53282 | 5658 |
unfolding dist_norm |
5659 |
apply auto |
|
5660 |
apply (erule_tac x=e in allE) |
|
5661 |
apply auto |
|
5662 |
apply (rule_tac x=d in exI) |
|
5663 |
apply auto |
|
5664 |
apply (erule_tac x=x' in allE) |
|
5665 |
apply auto |
|
5666 |
apply (erule_tac x=e in allE) |
|
5667 |
apply auto |
|
5668 |
done |
|
| 33175 | 5669 |
|
5670 |
lemma continuous_on_real_range: |
|
5671 |
fixes f :: "'a::real_normed_vector \<Rightarrow> real" |
|
| 53282 | 5672 |
shows "continuous_on s f \<longleftrightarrow> |
5673 |
(\<forall>x \<in> s. \<forall>e>0. \<exists>d>0. (\<forall>x' \<in> s. norm(x' - x) < d \<longrightarrow> abs(f x' - f x) < e))" |
|
| 36359 | 5674 |
unfolding continuous_on_iff dist_norm by simp |
| 33175 | 5675 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5676 |
text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
|
| 33175 | 5677 |
|
5678 |
lemma distance_attains_sup: |
|
5679 |
assumes "compact s" "s \<noteq> {}"
|
|
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5680 |
shows "\<exists>x\<in>s. \<forall>y\<in>s. dist a y \<le> dist a x" |
| 33175 | 5681 |
proof (rule continuous_attains_sup [OF assms]) |
| 53282 | 5682 |
{
|
5683 |
fix x |
|
5684 |
assume "x\<in>s" |
|
| 33175 | 5685 |
have "(dist a ---> dist a x) (at x within s)" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
5686 |
by (intro tendsto_dist tendsto_const tendsto_ident_at) |
| 33175 | 5687 |
} |
| 53282 | 5688 |
then show "continuous_on s (dist a)" |
| 33175 | 5689 |
unfolding continuous_on .. |
5690 |
qed |
|
5691 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5692 |
text {* For \emph{minimal} distance, we only need closure, not compactness. *}
|
| 33175 | 5693 |
|
5694 |
lemma distance_attains_inf: |
|
5695 |
fixes a :: "'a::heine_borel" |
|
| 53291 | 5696 |
assumes "closed s" |
5697 |
and "s \<noteq> {}"
|
|
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5698 |
shows "\<exists>x\<in>s. \<forall>y\<in>s. dist a x \<le> dist a y" |
| 53282 | 5699 |
proof - |
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5700 |
from assms(2) obtain b where "b \<in> s" by auto |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5701 |
let ?B = "s \<inter> cball a (dist b a)" |
| 53282 | 5702 |
have "?B \<noteq> {}" using `b \<in> s`
|
5703 |
by (auto simp add: dist_commute) |
|
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5704 |
moreover have "continuous_on ?B (dist a)" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5705 |
by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_at_id continuous_const) |
| 33175 | 5706 |
moreover have "compact ?B" |
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5707 |
by (intro closed_inter_compact `closed s` compact_cball) |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5708 |
ultimately obtain x where "x \<in> ?B" "\<forall>y\<in>?B. dist a x \<le> dist a y" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
5709 |
by (metis continuous_attains_inf) |
| 53282 | 5710 |
then show ?thesis by fastforce |
| 33175 | 5711 |
qed |
5712 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5713 |
|
| 36437 | 5714 |
subsection {* Pasted sets *}
|
| 33175 | 5715 |
|
5716 |
lemma bounded_Times: |
|
| 53282 | 5717 |
assumes "bounded s" "bounded t" |
5718 |
shows "bounded (s \<times> t)" |
|
5719 |
proof - |
|
| 33175 | 5720 |
obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b" |
5721 |
using assms [unfolded bounded_def] by auto |
|
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52625
diff
changeset
|
5722 |
then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<^sup>2 + b\<^sup>2)" |
| 33175 | 5723 |
by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono) |
| 53282 | 5724 |
then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto |
| 33175 | 5725 |
qed |
5726 |
||
5727 |
lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B" |
|
| 53282 | 5728 |
by (induct x) simp |
| 33175 | 5729 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
5730 |
lemma seq_compact_Times: "seq_compact s \<Longrightarrow> seq_compact t \<Longrightarrow> seq_compact (s \<times> t)" |
| 53282 | 5731 |
unfolding seq_compact_def |
5732 |
apply clarify |
|
5733 |
apply (drule_tac x="fst \<circ> f" in spec) |
|
5734 |
apply (drule mp, simp add: mem_Times_iff) |
|
5735 |
apply (clarify, rename_tac l1 r1) |
|
5736 |
apply (drule_tac x="snd \<circ> f \<circ> r1" in spec) |
|
5737 |
apply (drule mp, simp add: mem_Times_iff) |
|
5738 |
apply (clarify, rename_tac l2 r2) |
|
5739 |
apply (rule_tac x="(l1, l2)" in rev_bexI, simp) |
|
5740 |
apply (rule_tac x="r1 \<circ> r2" in exI) |
|
5741 |
apply (rule conjI, simp add: subseq_def) |
|
5742 |
apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption) |
|
5743 |
apply (drule (1) tendsto_Pair) back |
|
5744 |
apply (simp add: o_def) |
|
5745 |
done |
|
5746 |
||
5747 |
lemma compact_Times: |
|
| 51349 | 5748 |
assumes "compact s" "compact t" |
5749 |
shows "compact (s \<times> t)" |
|
5750 |
proof (rule compactI) |
|
| 53282 | 5751 |
fix C |
5752 |
assume C: "\<forall>t\<in>C. open t" "s \<times> t \<subseteq> \<Union>C" |
|
| 51349 | 5753 |
have "\<forall>x\<in>s. \<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)" |
5754 |
proof |
|
| 53282 | 5755 |
fix x |
5756 |
assume "x \<in> s" |
|
| 51349 | 5757 |
have "\<forall>y\<in>t. \<exists>a b c. c \<in> C \<and> open a \<and> open b \<and> x \<in> a \<and> y \<in> b \<and> a \<times> b \<subseteq> c" (is "\<forall>y\<in>t. ?P y") |
| 53282 | 5758 |
proof |
5759 |
fix y |
|
5760 |
assume "y \<in> t" |
|
| 51349 | 5761 |
with `x \<in> s` C obtain c where "c \<in> C" "(x, y) \<in> c" "open c" by auto |
5762 |
then show "?P y" by (auto elim!: open_prod_elim) |
|
5763 |
qed |
|
5764 |
then obtain a b c where b: "\<And>y. y \<in> t \<Longrightarrow> open (b y)" |
|
5765 |
and c: "\<And>y. y \<in> t \<Longrightarrow> c y \<in> C \<and> open (a y) \<and> open (b y) \<and> x \<in> a y \<and> y \<in> b y \<and> a y \<times> b y \<subseteq> c y" |
|
5766 |
by metis |
|
5767 |
then have "\<forall>y\<in>t. open (b y)" "t \<subseteq> (\<Union>y\<in>t. b y)" by auto |
|
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
5768 |
from compactE_image[OF `compact t` this] obtain D where D: "D \<subseteq> t" "finite D" "t \<subseteq> (\<Union>y\<in>D. b y)" |
| 51349 | 5769 |
by auto |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
5770 |
moreover from D c have "(\<Inter>y\<in>D. a y) \<times> t \<subseteq> (\<Union>y\<in>D. c y)" |
| 51349 | 5771 |
by (fastforce simp: subset_eq) |
5772 |
ultimately show "\<exists>a. open a \<and> x \<in> a \<and> (\<exists>d\<subseteq>C. finite d \<and> a \<times> t \<subseteq> \<Union>d)" |
|
|
52141
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
haftmann
parents:
51773
diff
changeset
|
5773 |
using c by (intro exI[of _ "c`D"] exI[of _ "\<Inter>(a`D)"] conjI) (auto intro!: open_INT) |
| 51349 | 5774 |
qed |
5775 |
then obtain a d where a: "\<forall>x\<in>s. open (a x)" "s \<subseteq> (\<Union>x\<in>s. a x)" |
|
5776 |
and d: "\<And>x. x \<in> s \<Longrightarrow> d x \<subseteq> C \<and> finite (d x) \<and> a x \<times> t \<subseteq> \<Union>d x" |
|
5777 |
unfolding subset_eq UN_iff by metis |
|
| 53282 | 5778 |
moreover |
5779 |
from compactE_image[OF `compact s` a] |
|
5780 |
obtain e where e: "e \<subseteq> s" "finite e" and s: "s \<subseteq> (\<Union>x\<in>e. a x)" |
|
5781 |
by auto |
|
| 51349 | 5782 |
moreover |
| 53282 | 5783 |
{
|
5784 |
from s have "s \<times> t \<subseteq> (\<Union>x\<in>e. a x \<times> t)" |
|
5785 |
by auto |
|
5786 |
also have "\<dots> \<subseteq> (\<Union>x\<in>e. \<Union>d x)" |
|
5787 |
using d `e \<subseteq> s` by (intro UN_mono) auto |
|
5788 |
finally have "s \<times> t \<subseteq> (\<Union>x\<in>e. \<Union>d x)" . |
|
5789 |
} |
|
| 51349 | 5790 |
ultimately show "\<exists>C'\<subseteq>C. finite C' \<and> s \<times> t \<subseteq> \<Union>C'" |
5791 |
by (intro exI[of _ "(\<Union>x\<in>e. d x)"]) (auto simp add: subset_eq) |
|
5792 |
qed |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
5793 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5794 |
text{* Hence some useful properties follow quite easily. *}
|
| 33175 | 5795 |
|
5796 |
lemma compact_scaling: |
|
5797 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5798 |
assumes "compact s" |
5799 |
shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)" |
|
5800 |
proof - |
|
| 33175 | 5801 |
let ?f = "\<lambda>x. scaleR c x" |
| 53282 | 5802 |
have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right) |
5803 |
show ?thesis |
|
5804 |
using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f] |
|
5805 |
using linear_continuous_at[OF *] assms |
|
5806 |
by auto |
|
| 33175 | 5807 |
qed |
5808 |
||
5809 |
lemma compact_negations: |
|
5810 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5811 |
assumes "compact s" |
| 53291 | 5812 |
shows "compact ((\<lambda>x. - x) ` s)" |
| 33175 | 5813 |
using compact_scaling [OF assms, of "- 1"] by auto |
5814 |
||
5815 |
lemma compact_sums: |
|
5816 |
fixes s t :: "'a::real_normed_vector set" |
|
| 53291 | 5817 |
assumes "compact s" |
5818 |
and "compact t" |
|
| 53282 | 5819 |
shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
|
5820 |
proof - |
|
5821 |
have *: "{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
|
|
5822 |
apply auto |
|
5823 |
unfolding image_iff |
|
5824 |
apply (rule_tac x="(xa, y)" in bexI) |
|
5825 |
apply auto |
|
5826 |
done |
|
| 33175 | 5827 |
have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)" |
5828 |
unfolding continuous_on by (rule ballI) (intro tendsto_intros) |
|
| 53282 | 5829 |
then show ?thesis |
5830 |
unfolding * using compact_continuous_image compact_Times [OF assms] by auto |
|
| 33175 | 5831 |
qed |
5832 |
||
5833 |
lemma compact_differences: |
|
5834 |
fixes s t :: "'a::real_normed_vector set" |
|
| 53291 | 5835 |
assumes "compact s" |
5836 |
and "compact t" |
|
5837 |
shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
|
|
| 33175 | 5838 |
proof- |
5839 |
have "{x - y | x y. x\<in>s \<and> y \<in> t} = {x + y | x y. x \<in> s \<and> y \<in> (uminus ` t)}"
|
|
| 53282 | 5840 |
apply auto |
5841 |
apply (rule_tac x= xa in exI) |
|
5842 |
apply auto |
|
5843 |
done |
|
5844 |
then show ?thesis |
|
5845 |
using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto |
|
| 33175 | 5846 |
qed |
5847 |
||
5848 |
lemma compact_translation: |
|
5849 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5850 |
assumes "compact s" |
5851 |
shows "compact ((\<lambda>x. a + x) ` s)" |
|
5852 |
proof - |
|
5853 |
have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s"
|
|
5854 |
by auto |
|
5855 |
then show ?thesis |
|
5856 |
using compact_sums[OF assms compact_sing[of a]] by auto |
|
| 33175 | 5857 |
qed |
5858 |
||
5859 |
lemma compact_affinity: |
|
5860 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5861 |
assumes "compact s" |
5862 |
shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)" |
|
5863 |
proof - |
|
5864 |
have "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" |
|
5865 |
by auto |
|
5866 |
then show ?thesis |
|
5867 |
using compact_translation[OF compact_scaling[OF assms], of a c] by auto |
|
| 33175 | 5868 |
qed |
5869 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5870 |
text {* Hence we get the following. *}
|
| 33175 | 5871 |
|
5872 |
lemma compact_sup_maxdistance: |
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5873 |
fixes s :: "'a::metric_space set" |
| 53291 | 5874 |
assumes "compact s" |
5875 |
and "s \<noteq> {}"
|
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5876 |
shows "\<exists>x\<in>s. \<exists>y\<in>s. \<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y" |
| 53282 | 5877 |
proof - |
5878 |
have "compact (s \<times> s)" |
|
5879 |
using `compact s` by (intro compact_Times) |
|
5880 |
moreover have "s \<times> s \<noteq> {}"
|
|
5881 |
using `s \<noteq> {}` by auto
|
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5882 |
moreover have "continuous_on (s \<times> s) (\<lambda>x. dist (fst x) (snd x))" |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
5883 |
by (intro continuous_at_imp_continuous_on ballI continuous_intros) |
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5884 |
ultimately show ?thesis |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5885 |
using continuous_attains_sup[of "s \<times> s" "\<lambda>x. dist (fst x) (snd x)"] by auto |
| 33175 | 5886 |
qed |
5887 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5888 |
text {* We can state this in terms of diameter of a set. *}
|
| 33175 | 5889 |
|
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5890 |
definition diameter :: "'a::metric_space set \<Rightarrow> real" where |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5891 |
"diameter S = (if S = {} then 0 else SUP (x,y):S\<times>S. dist x y)"
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5892 |
|
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5893 |
lemma diameter_bounded_bound: |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5894 |
fixes s :: "'a :: metric_space set" |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5895 |
assumes s: "bounded s" "x \<in> s" "y \<in> s" |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5896 |
shows "dist x y \<le> diameter s" |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5897 |
proof - |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5898 |
from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d" |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5899 |
unfolding bounded_def by auto |
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5900 |
have "bdd_above (split dist ` (s\<times>s))" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5901 |
proof (intro bdd_aboveI, safe) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5902 |
fix a b |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5903 |
assume "a \<in> s" "b \<in> s" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5904 |
with z[of a] z[of b] dist_triangle[of a b z] |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5905 |
show "dist a b \<le> 2 * d" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5906 |
by (simp add: dist_commute) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5907 |
qed |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5908 |
moreover have "(x,y) \<in> s\<times>s" using s by auto |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5909 |
ultimately have "dist x y \<le> (SUP (x,y):s\<times>s. dist x y)" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5910 |
by (rule cSUP_upper2) simp |
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5911 |
with `x \<in> s` show ?thesis |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5912 |
by (auto simp add: diameter_def) |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5913 |
qed |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5914 |
|
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5915 |
lemma diameter_lower_bounded: |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5916 |
fixes s :: "'a :: metric_space set" |
| 53282 | 5917 |
assumes s: "bounded s" |
5918 |
and d: "0 < d" "d < diameter s" |
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5919 |
shows "\<exists>x\<in>s. \<exists>y\<in>s. d < dist x y" |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5920 |
proof (rule ccontr) |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5921 |
assume contr: "\<not> ?thesis" |
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5922 |
moreover have "s \<noteq> {}"
|
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5923 |
using d by (auto simp add: diameter_def) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5924 |
ultimately have "diameter s \<le> d" |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5925 |
by (auto simp: not_less diameter_def intro!: cSUP_least) |
|
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5926 |
with `d < diameter s` show False by auto |
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5927 |
qed |
| 33175 | 5928 |
|
5929 |
lemma diameter_bounded: |
|
5930 |
assumes "bounded s" |
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5931 |
shows "\<forall>x\<in>s. \<forall>y\<in>s. dist x y \<le> diameter s" |
| 53291 | 5932 |
and "\<forall>d>0. d < diameter s \<longrightarrow> (\<exists>x\<in>s. \<exists>y\<in>s. dist x y > d)" |
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5933 |
using diameter_bounded_bound[of s] diameter_lower_bounded[of s] assms |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5934 |
by auto |
| 33175 | 5935 |
|
5936 |
lemma diameter_compact_attained: |
|
| 53291 | 5937 |
assumes "compact s" |
5938 |
and "s \<noteq> {}"
|
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5939 |
shows "\<exists>x\<in>s. \<exists>y\<in>s. dist x y = diameter s" |
|
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5940 |
proof - |
| 53282 | 5941 |
have b: "bounded s" using assms(1) |
5942 |
by (rule compact_imp_bounded) |
|
| 53291 | 5943 |
then obtain x y where xys: "x\<in>s" "y\<in>s" |
5944 |
and xy: "\<forall>u\<in>s. \<forall>v\<in>s. dist u v \<le> dist x y" |
|
|
50973
4a2c82644889
generalized diameter from real_normed_vector to metric_space
hoelzl
parents:
50972
diff
changeset
|
5945 |
using compact_sup_maxdistance[OF assms] by auto |
| 53282 | 5946 |
then have "diameter s \<le> dist x y" |
5947 |
unfolding diameter_def |
|
5948 |
apply clarsimp |
|
|
54260
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
hoelzl
parents:
54259
diff
changeset
|
5949 |
apply (rule cSUP_least) |
| 53282 | 5950 |
apply fast+ |
5951 |
done |
|
5952 |
then show ?thesis |
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
5953 |
by (metis b diameter_bounded_bound order_antisym xys) |
| 33175 | 5954 |
qed |
5955 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
5956 |
text {* Related results with closure as the conclusion. *}
|
| 33175 | 5957 |
|
5958 |
lemma closed_scaling: |
|
5959 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5960 |
assumes "closed s" |
5961 |
shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)" |
|
| 53813 | 5962 |
proof (cases "c = 0") |
5963 |
case True then show ?thesis |
|
5964 |
by (auto simp add: image_constant_conv) |
|
| 33175 | 5965 |
next |
5966 |
case False |
|
| 53813 | 5967 |
from assms have "closed ((\<lambda>x. inverse c *\<^sub>R x) -` s)" |
5968 |
by (simp add: continuous_closed_vimage) |
|
5969 |
also have "(\<lambda>x. inverse c *\<^sub>R x) -` s = (\<lambda>x. c *\<^sub>R x) ` s" |
|
5970 |
using `c \<noteq> 0` by (auto elim: image_eqI [rotated]) |
|
5971 |
finally show ?thesis . |
|
| 33175 | 5972 |
qed |
5973 |
||
5974 |
lemma closed_negations: |
|
5975 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5976 |
assumes "closed s" |
5977 |
shows "closed ((\<lambda>x. -x) ` s)" |
|
| 33175 | 5978 |
using closed_scaling[OF assms, of "- 1"] by simp |
5979 |
||
5980 |
lemma compact_closed_sums: |
|
5981 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 5982 |
assumes "compact s" and "closed t" |
5983 |
shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
|
|
5984 |
proof - |
|
| 33175 | 5985 |
let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
|
| 53282 | 5986 |
{
|
5987 |
fix x l |
|
5988 |
assume as: "\<forall>n. x n \<in> ?S" "(x ---> l) sequentially" |
|
5989 |
from as(1) obtain f where f: "\<forall>n. x n = fst (f n) + snd (f n)" "\<forall>n. fst (f n) \<in> s" "\<forall>n. snd (f n) \<in> t" |
|
| 33175 | 5990 |
using choice[of "\<lambda>n y. x n = (fst y) + (snd y) \<and> fst y \<in> s \<and> snd y \<in> t"] by auto |
| 53282 | 5991 |
obtain l' r where "l'\<in>s" and r: "subseq r" and lr: "(((\<lambda>n. fst (f n)) \<circ> r) ---> l') sequentially" |
| 33175 | 5992 |
using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto |
5993 |
have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially" |
|
| 53282 | 5994 |
using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1) |
5995 |
unfolding o_def |
|
5996 |
by auto |
|
5997 |
then have "l - l' \<in> t" |
|
| 53291 | 5998 |
using assms(2)[unfolded closed_sequential_limits, |
5999 |
THEN spec[where x="\<lambda> n. snd (f (r n))"], |
|
6000 |
THEN spec[where x="l - l'"]] |
|
| 53282 | 6001 |
using f(3) |
6002 |
by auto |
|
6003 |
then have "l \<in> ?S" |
|
6004 |
using `l' \<in> s` |
|
6005 |
apply auto |
|
6006 |
apply (rule_tac x=l' in exI) |
|
6007 |
apply (rule_tac x="l - l'" in exI) |
|
6008 |
apply auto |
|
6009 |
done |
|
| 33175 | 6010 |
} |
| 53282 | 6011 |
then show ?thesis |
6012 |
unfolding closed_sequential_limits by fast |
|
| 33175 | 6013 |
qed |
6014 |
||
6015 |
lemma closed_compact_sums: |
|
6016 |
fixes s t :: "'a::real_normed_vector set" |
|
| 53291 | 6017 |
assumes "closed s" |
6018 |
and "compact t" |
|
| 33175 | 6019 |
shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
|
| 53282 | 6020 |
proof - |
6021 |
have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}"
|
|
6022 |
apply auto |
|
6023 |
apply (rule_tac x=y in exI) |
|
6024 |
apply auto |
|
6025 |
apply (rule_tac x=y in exI) |
|
6026 |
apply auto |
|
6027 |
done |
|
6028 |
then show ?thesis |
|
6029 |
using compact_closed_sums[OF assms(2,1)] by simp |
|
| 33175 | 6030 |
qed |
6031 |
||
6032 |
lemma compact_closed_differences: |
|
6033 |
fixes s t :: "'a::real_normed_vector set" |
|
| 53291 | 6034 |
assumes "compact s" |
6035 |
and "closed t" |
|
| 33175 | 6036 |
shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
|
| 53282 | 6037 |
proof - |
| 33175 | 6038 |
have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
|
| 53282 | 6039 |
apply auto |
6040 |
apply (rule_tac x=xa in exI) |
|
6041 |
apply auto |
|
6042 |
apply (rule_tac x=xa in exI) |
|
6043 |
apply auto |
|
6044 |
done |
|
6045 |
then show ?thesis |
|
6046 |
using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto |
|
| 33175 | 6047 |
qed |
6048 |
||
6049 |
lemma closed_compact_differences: |
|
6050 |
fixes s t :: "'a::real_normed_vector set" |
|
| 53291 | 6051 |
assumes "closed s" |
6052 |
and "compact t" |
|
| 33175 | 6053 |
shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
|
| 53282 | 6054 |
proof - |
| 33175 | 6055 |
have "{x + y |x y. x \<in> s \<and> y \<in> uminus ` t} = {x - y |x y. x \<in> s \<and> y \<in> t}"
|
| 53282 | 6056 |
apply auto |
6057 |
apply (rule_tac x=xa in exI) |
|
6058 |
apply auto |
|
6059 |
apply (rule_tac x=xa in exI) |
|
6060 |
apply auto |
|
6061 |
done |
|
6062 |
then show ?thesis |
|
6063 |
using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp |
|
| 33175 | 6064 |
qed |
6065 |
||
6066 |
lemma closed_translation: |
|
6067 |
fixes a :: "'a::real_normed_vector" |
|
| 53282 | 6068 |
assumes "closed s" |
6069 |
shows "closed ((\<lambda>x. a + x) ` s)" |
|
6070 |
proof - |
|
| 33175 | 6071 |
have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
|
| 53282 | 6072 |
then show ?thesis |
6073 |
using compact_closed_sums[OF compact_sing[of a] assms] by auto |
|
| 33175 | 6074 |
qed |
6075 |
||
| 34105 | 6076 |
lemma translation_Compl: |
6077 |
fixes a :: "'a::ab_group_add" |
|
6078 |
shows "(\<lambda>x. a + x) ` (- t) = - ((\<lambda>x. a + x) ` t)" |
|
| 53282 | 6079 |
apply (auto simp add: image_iff) |
6080 |
apply (rule_tac x="x - a" in bexI) |
|
6081 |
apply auto |
|
6082 |
done |
|
| 34105 | 6083 |
|
| 33175 | 6084 |
lemma translation_UNIV: |
| 53282 | 6085 |
fixes a :: "'a::ab_group_add" |
6086 |
shows "range (\<lambda>x. a + x) = UNIV" |
|
6087 |
apply (auto simp add: image_iff) |
|
6088 |
apply (rule_tac x="x - a" in exI) |
|
6089 |
apply auto |
|
6090 |
done |
|
| 33175 | 6091 |
|
6092 |
lemma translation_diff: |
|
6093 |
fixes a :: "'a::ab_group_add" |
|
6094 |
shows "(\<lambda>x. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)" |
|
6095 |
by auto |
|
6096 |
||
6097 |
lemma closure_translation: |
|
6098 |
fixes a :: "'a::real_normed_vector" |
|
6099 |
shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)" |
|
| 53282 | 6100 |
proof - |
6101 |
have *: "op + a ` (- s) = - op + a ` s" |
|
6102 |
apply auto |
|
6103 |
unfolding image_iff |
|
6104 |
apply (rule_tac x="x - a" in bexI) |
|
6105 |
apply auto |
|
6106 |
done |
|
6107 |
show ?thesis |
|
6108 |
unfolding closure_interior translation_Compl |
|
6109 |
using interior_translation[of a "- s"] |
|
6110 |
unfolding * |
|
6111 |
by auto |
|
| 33175 | 6112 |
qed |
6113 |
||
6114 |
lemma frontier_translation: |
|
6115 |
fixes a :: "'a::real_normed_vector" |
|
6116 |
shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)" |
|
| 53282 | 6117 |
unfolding frontier_def translation_diff interior_translation closure_translation |
6118 |
by auto |
|
| 33175 | 6119 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6120 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6121 |
subsection {* Separation between points and sets *}
|
| 33175 | 6122 |
|
6123 |
lemma separate_point_closed: |
|
6124 |
fixes s :: "'a::heine_borel set" |
|
| 53291 | 6125 |
assumes "closed s" |
6126 |
and "a \<notin> s" |
|
| 53282 | 6127 |
shows "\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x" |
6128 |
proof (cases "s = {}")
|
|
| 33175 | 6129 |
case True |
| 53282 | 6130 |
then show ?thesis by(auto intro!: exI[where x=1]) |
| 33175 | 6131 |
next |
6132 |
case False |
|
| 53282 | 6133 |
from assms obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y" |
6134 |
using `s \<noteq> {}` distance_attains_inf [of s a] by blast
|
|
6135 |
with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s` |
|
6136 |
by blast |
|
| 33175 | 6137 |
qed |
6138 |
||
6139 |
lemma separate_compact_closed: |
|
| 50949 | 6140 |
fixes s t :: "'a::heine_borel set" |
| 53282 | 6141 |
assumes "compact s" |
6142 |
and t: "closed t" "s \<inter> t = {}"
|
|
| 33175 | 6143 |
shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y" |
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6144 |
proof cases |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6145 |
assume "s \<noteq> {} \<and> t \<noteq> {}"
|
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6146 |
then have "s \<noteq> {}" "t \<noteq> {}" by auto
|
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6147 |
let ?inf = "\<lambda>x. infdist x t" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6148 |
have "continuous_on s ?inf" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6149 |
by (auto intro!: continuous_at_imp_continuous_on continuous_infdist continuous_at_id) |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6150 |
then obtain x where x: "x \<in> s" "\<forall>y\<in>s. ?inf x \<le> ?inf y" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6151 |
using continuous_attains_inf[OF `compact s` `s \<noteq> {}`] by auto
|
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6152 |
then have "0 < ?inf x" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6153 |
using t `t \<noteq> {}` in_closed_iff_infdist_zero by (auto simp: less_le infdist_nonneg)
|
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6154 |
moreover have "\<forall>x'\<in>s. \<forall>y\<in>t. ?inf x \<le> dist x' y" |
|
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6155 |
using x by (auto intro: order_trans infdist_le) |
| 53282 | 6156 |
ultimately show ?thesis by auto |
|
51346
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
hoelzl
parents:
51345
diff
changeset
|
6157 |
qed (auto intro!: exI[of _ 1]) |
| 33175 | 6158 |
|
6159 |
lemma separate_closed_compact: |
|
| 50949 | 6160 |
fixes s t :: "'a::heine_borel set" |
| 53282 | 6161 |
assumes "closed s" |
6162 |
and "compact t" |
|
6163 |
and "s \<inter> t = {}"
|
|
| 33175 | 6164 |
shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y" |
| 53282 | 6165 |
proof - |
6166 |
have *: "t \<inter> s = {}"
|
|
6167 |
using assms(3) by auto |
|
6168 |
show ?thesis |
|
6169 |
using separate_compact_closed[OF assms(2,1) *] |
|
6170 |
apply auto |
|
6171 |
apply (rule_tac x=d in exI) |
|
6172 |
apply auto |
|
6173 |
apply (erule_tac x=y in ballE) |
|
6174 |
apply (auto simp add: dist_commute) |
|
6175 |
done |
|
| 33175 | 6176 |
qed |
6177 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6178 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6179 |
subsection {* Closure of halfspaces and hyperplanes *}
|
| 33175 | 6180 |
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6181 |
lemma isCont_open_vimage: |
| 53282 | 6182 |
assumes "\<And>x. isCont f x" |
6183 |
and "open s" |
|
6184 |
shows "open (f -` s)" |
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6185 |
proof - |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6186 |
from assms(1) have "continuous_on UNIV f" |
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
6187 |
unfolding isCont_def continuous_on_def by simp |
| 53282 | 6188 |
then have "open {x \<in> UNIV. f x \<in> s}"
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6189 |
using open_UNIV `open s` by (rule continuous_open_preimage) |
| 53282 | 6190 |
then show "open (f -` s)" |
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6191 |
by (simp add: vimage_def) |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6192 |
qed |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6193 |
|
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6194 |
lemma isCont_closed_vimage: |
| 53282 | 6195 |
assumes "\<And>x. isCont f x" |
6196 |
and "closed s" |
|
6197 |
shows "closed (f -` s)" |
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6198 |
using assms unfolding closed_def vimage_Compl [symmetric] |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6199 |
by (rule isCont_open_vimage) |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6200 |
|
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6201 |
lemma open_Collect_less: |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
6202 |
fixes f g :: "'a::t2_space \<Rightarrow> real" |
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6203 |
assumes f: "\<And>x. isCont f x" |
| 53282 | 6204 |
and g: "\<And>x. isCont g x" |
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6205 |
shows "open {x. f x < g x}"
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6206 |
proof - |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6207 |
have "open ((\<lambda>x. g x - f x) -` {0<..})"
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6208 |
using isCont_diff [OF g f] open_real_greaterThan |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6209 |
by (rule isCont_open_vimage) |
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6210 |
also have "((\<lambda>x. g x - f x) -` {0<..}) = {x. f x < g x}"
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6211 |
by auto |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6212 |
finally show ?thesis . |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6213 |
qed |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6214 |
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6215 |
lemma closed_Collect_le: |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
6216 |
fixes f g :: "'a::t2_space \<Rightarrow> real" |
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6217 |
assumes f: "\<And>x. isCont f x" |
| 53282 | 6218 |
and g: "\<And>x. isCont g x" |
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6219 |
shows "closed {x. f x \<le> g x}"
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6220 |
proof - |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6221 |
have "closed ((\<lambda>x. g x - f x) -` {0..})"
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6222 |
using isCont_diff [OF g f] closed_real_atLeast |
|
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6223 |
by (rule isCont_closed_vimage) |
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6224 |
also have "((\<lambda>x. g x - f x) -` {0..}) = {x. f x \<le> g x}"
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6225 |
by auto |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6226 |
finally show ?thesis . |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6227 |
qed |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6228 |
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6229 |
lemma closed_Collect_eq: |
|
51478
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents:
51475
diff
changeset
|
6230 |
fixes f g :: "'a::t2_space \<Rightarrow> 'b::t2_space" |
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6231 |
assumes f: "\<And>x. isCont f x" |
| 53282 | 6232 |
and g: "\<And>x. isCont g x" |
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6233 |
shows "closed {x. f x = g x}"
|
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6234 |
proof - |
| 44216 | 6235 |
have "open {(x::'b, y::'b). x \<noteq> y}"
|
6236 |
unfolding open_prod_def by (auto dest!: hausdorff) |
|
| 53282 | 6237 |
then have "closed {(x::'b, y::'b). x = y}"
|
| 44216 | 6238 |
unfolding closed_def split_def Collect_neg_eq . |
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6239 |
with isCont_Pair [OF f g] |
| 44216 | 6240 |
have "closed ((\<lambda>x. (f x, g x)) -` {(x, y). x = y})"
|
|
44219
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
huffman
parents:
44216
diff
changeset
|
6241 |
by (rule isCont_closed_vimage) |
| 44216 | 6242 |
also have "\<dots> = {x. f x = g x}" by auto
|
|
44213
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6243 |
finally show ?thesis . |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6244 |
qed |
|
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
huffman
parents:
44212
diff
changeset
|
6245 |
|
| 33175 | 6246 |
lemma continuous_at_inner: "continuous (at x) (inner a)" |
6247 |
unfolding continuous_at by (intro tendsto_intros) |
|
6248 |
||
6249 |
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
|
|
| 44233 | 6250 |
by (simp add: closed_Collect_le) |
| 33175 | 6251 |
|
6252 |
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
|
|
| 44233 | 6253 |
by (simp add: closed_Collect_le) |
| 33175 | 6254 |
|
6255 |
lemma closed_hyperplane: "closed {x. inner a x = b}"
|
|
| 44233 | 6256 |
by (simp add: closed_Collect_eq) |
| 33175 | 6257 |
|
| 53282 | 6258 |
lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
|
| 44233 | 6259 |
by (simp add: closed_Collect_le) |
| 33175 | 6260 |
|
| 53282 | 6261 |
lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
|
| 44233 | 6262 |
by (simp add: closed_Collect_le) |
| 33175 | 6263 |
|
| 53813 | 6264 |
lemma closed_interval_left: |
6265 |
fixes b :: "'a::euclidean_space" |
|
6266 |
shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}"
|
|
6267 |
by (simp add: Collect_ball_eq closed_INT closed_Collect_le) |
|
6268 |
||
6269 |
lemma closed_interval_right: |
|
6270 |
fixes a :: "'a::euclidean_space" |
|
6271 |
shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
|
|
6272 |
by (simp add: Collect_ball_eq closed_INT closed_Collect_le) |
|
6273 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6274 |
text {* Openness of halfspaces. *}
|
| 33175 | 6275 |
|
6276 |
lemma open_halfspace_lt: "open {x. inner a x < b}"
|
|
| 44233 | 6277 |
by (simp add: open_Collect_less) |
| 33175 | 6278 |
|
6279 |
lemma open_halfspace_gt: "open {x. inner a x > b}"
|
|
| 44233 | 6280 |
by (simp add: open_Collect_less) |
| 33175 | 6281 |
|
| 53282 | 6282 |
lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\<bullet>i < a}"
|
| 44233 | 6283 |
by (simp add: open_Collect_less) |
| 33175 | 6284 |
|
| 53282 | 6285 |
lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\<bullet>i > a}"
|
| 44233 | 6286 |
by (simp add: open_Collect_less) |
| 33175 | 6287 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6288 |
text {* This gives a simple derivation of limit component bounds. *}
|
| 33175 | 6289 |
|
| 53282 | 6290 |
lemma Lim_component_le: |
6291 |
fixes f :: "'a \<Rightarrow> 'b::euclidean_space" |
|
6292 |
assumes "(f ---> l) net" |
|
6293 |
and "\<not> (trivial_limit net)" |
|
6294 |
and "eventually (\<lambda>x. f(x)\<bullet>i \<le> b) net" |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
6295 |
shows "l\<bullet>i \<le> b" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
6296 |
by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
6297 |
|
| 53282 | 6298 |
lemma Lim_component_ge: |
6299 |
fixes f :: "'a \<Rightarrow> 'b::euclidean_space" |
|
6300 |
assumes "(f ---> l) net" |
|
6301 |
and "\<not> (trivial_limit net)" |
|
6302 |
and "eventually (\<lambda>x. b \<le> (f x)\<bullet>i) net" |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
6303 |
shows "b \<le> l\<bullet>i" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
6304 |
by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) |
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
6305 |
|
| 53282 | 6306 |
lemma Lim_component_eq: |
6307 |
fixes f :: "'a \<Rightarrow> 'b::euclidean_space" |
|
| 53640 | 6308 |
assumes net: "(f ---> l) net" "\<not> trivial_limit net" |
| 53282 | 6309 |
and ev:"eventually (\<lambda>x. f(x)\<bullet>i = b) net" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
6310 |
shows "l\<bullet>i = b" |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
6311 |
using ev[unfolded order_eq_iff eventually_conj_iff] |
| 53282 | 6312 |
using Lim_component_ge[OF net, of b i] |
6313 |
using Lim_component_le[OF net, of i b] |
|
6314 |
by auto |
|
6315 |
||
6316 |
text {* Limits relative to a union. *}
|
|
| 33175 | 6317 |
|
6318 |
lemma eventually_within_Un: |
|
| 53282 | 6319 |
"eventually P (at x within (s \<union> t)) \<longleftrightarrow> |
6320 |
eventually P (at x within s) \<and> eventually P (at x within t)" |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
6321 |
unfolding eventually_at_filter |
| 33175 | 6322 |
by (auto elim!: eventually_rev_mp) |
6323 |
||
6324 |
lemma Lim_within_union: |
|
|
51641
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
6325 |
"(f ---> l) (at x within (s \<union> t)) \<longleftrightarrow> |
|
cd05e9fcc63d
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents:
51530
diff
changeset
|
6326 |
(f ---> l) (at x within s) \<and> (f ---> l) (at x within t)" |
| 33175 | 6327 |
unfolding tendsto_def |
6328 |
by (auto simp add: eventually_within_Un) |
|
6329 |
||
| 36442 | 6330 |
lemma Lim_topological: |
| 53282 | 6331 |
"(f ---> l) net \<longleftrightarrow> |
6332 |
trivial_limit net \<or> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)" |
|
| 36442 | 6333 |
unfolding tendsto_def trivial_limit_eq by auto |
6334 |
||
| 53282 | 6335 |
text{* Some more convenient intermediate-value theorem formulations. *}
|
| 33175 | 6336 |
|
6337 |
lemma connected_ivt_hyperplane: |
|
| 53291 | 6338 |
assumes "connected s" |
6339 |
and "x \<in> s" |
|
6340 |
and "y \<in> s" |
|
6341 |
and "inner a x \<le> b" |
|
6342 |
and "b \<le> inner a y" |
|
| 33175 | 6343 |
shows "\<exists>z \<in> s. inner a z = b" |
| 53282 | 6344 |
proof (rule ccontr) |
| 33175 | 6345 |
assume as:"\<not> (\<exists>z\<in>s. inner a z = b)" |
6346 |
let ?A = "{x. inner a x < b}"
|
|
6347 |
let ?B = "{x. inner a x > b}"
|
|
| 53282 | 6348 |
have "open ?A" "open ?B" |
6349 |
using open_halfspace_lt and open_halfspace_gt by auto |
|
| 53291 | 6350 |
moreover |
6351 |
have "?A \<inter> ?B = {}" by auto
|
|
6352 |
moreover |
|
6353 |
have "s \<subseteq> ?A \<union> ?B" using as by auto |
|
6354 |
ultimately |
|
6355 |
show False |
|
| 53282 | 6356 |
using assms(1)[unfolded connected_def not_ex, |
6357 |
THEN spec[where x="?A"], THEN spec[where x="?B"]] |
|
6358 |
using assms(2-5) |
|
| 52625 | 6359 |
by auto |
6360 |
qed |
|
6361 |
||
6362 |
lemma connected_ivt_component: |
|
6363 |
fixes x::"'a::euclidean_space" |
|
6364 |
shows "connected s \<Longrightarrow> |
|
6365 |
x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> |
|
6366 |
x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>s. z\<bullet>k = a)" |
|
6367 |
using connected_ivt_hyperplane[of s x y "k::'a" a] |
|
6368 |
by (auto simp: inner_commute) |
|
| 33175 | 6369 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
6370 |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6371 |
subsection {* Intervals *}
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6372 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6373 |
lemma open_box[intro]: "open (box a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6374 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6375 |
have "open (\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6376 |
by (auto intro!: continuous_open_vimage continuous_inner continuous_at_id continuous_const) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6377 |
also have "(\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i}) = box a b"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6378 |
by (auto simp add: box_def inner_commute) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6379 |
finally show ?thesis . |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6380 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6381 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6382 |
instance euclidean_space \<subseteq> second_countable_topology |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6383 |
proof |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6384 |
def a \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. fst (f i) *\<^sub>R i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6385 |
then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6386 |
by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6387 |
def b \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. snd (f i) *\<^sub>R i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6388 |
then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6389 |
by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6390 |
def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6391 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6392 |
have "Ball B open" by (simp add: B_def open_box) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6393 |
moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6394 |
proof safe |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6395 |
fix A::"'a set" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6396 |
assume "open A" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6397 |
show "\<exists>B'\<subseteq>B. \<Union>B' = A" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6398 |
apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6399 |
apply (subst (3) open_UNION_box[OF `open A`]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6400 |
apply (auto simp add: a b B_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6401 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6402 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6403 |
ultimately |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6404 |
have "topological_basis B" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6405 |
unfolding topological_basis_def by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6406 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6407 |
have "countable B" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6408 |
unfolding B_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6409 |
by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6410 |
ultimately show "\<exists>B::'a set set. countable B \<and> open = generate_topology B" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6411 |
by (blast intro: topological_basis_imp_subbasis) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6412 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6413 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6414 |
instance euclidean_space \<subseteq> polish_space .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6415 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6416 |
lemma closed_cbox[intro]: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6417 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6418 |
shows "closed (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6419 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6420 |
have "closed (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6421 |
by (intro closed_INT ballI continuous_closed_vimage allI |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6422 |
linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6423 |
also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i}) = cbox a b"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6424 |
by (auto simp add: cbox_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6425 |
finally show "closed (cbox a b)" . |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6426 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6427 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6428 |
lemma interior_cbox [intro]: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6429 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6430 |
shows "interior (cbox a b) = box a b" (is "?L = ?R") |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6431 |
proof(rule subset_antisym) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6432 |
show "?R \<subseteq> ?L" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6433 |
using box_subset_cbox open_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6434 |
by (rule interior_maximal) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6435 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6436 |
fix x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6437 |
assume "x \<in> interior (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6438 |
then obtain s where s: "open s" "x \<in> s" "s \<subseteq> cbox a b" .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6439 |
then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6440 |
unfolding open_dist and subset_eq by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6441 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6442 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6443 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6444 |
have "dist (x - (e / 2) *\<^sub>R i) x < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6445 |
and "dist (x + (e / 2) *\<^sub>R i) x < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6446 |
unfolding dist_norm |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6447 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6448 |
unfolding norm_minus_cancel |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6449 |
using norm_Basis[OF i] `e>0` |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6450 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6451 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6452 |
then have "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i" and "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6453 |
using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6454 |
and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6455 |
unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6456 |
using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6457 |
by blast+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6458 |
then have "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6459 |
using `e>0` i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6460 |
by (auto simp: inner_diff_left inner_Basis inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6461 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6462 |
then have "x \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6463 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6464 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6465 |
then show "?L \<subseteq> ?R" .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6466 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6467 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6468 |
lemma bounded_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6469 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6470 |
shows "bounded (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6471 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6472 |
let ?b = "\<Sum>i\<in>Basis. \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6473 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6474 |
fix x :: "'a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6475 |
assume x: "\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6476 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6477 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6478 |
assume "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6479 |
then have "\<bar>x\<bullet>i\<bar> \<le> \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6480 |
using x[THEN bspec[where x=i]] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6481 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6482 |
then have "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6483 |
apply - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6484 |
apply (rule setsum_mono) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6485 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6486 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6487 |
then have "norm x \<le> ?b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6488 |
using norm_le_l1[of x] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6489 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6490 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6491 |
unfolding cbox_def bounded_iff by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6492 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6493 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6494 |
lemma bounded_box: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6495 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6496 |
shows "bounded (box a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6497 |
using bounded_cbox[of a b] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6498 |
using box_subset_cbox[of a b] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6499 |
using bounded_subset[of "cbox a b" "box a b"] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6500 |
by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6501 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6502 |
lemma not_interval_univ: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6503 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6504 |
shows "cbox a b \<noteq> UNIV" "box a b \<noteq> UNIV" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6505 |
using bounded_box[of a b] bounded_cbox[of a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6506 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6507 |
lemma compact_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6508 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6509 |
shows "compact (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6510 |
using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6511 |
by (auto simp: compact_eq_seq_compact_metric) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6512 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6513 |
lemma box_midpoint: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6514 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6515 |
assumes "box a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6516 |
shows "((1/2) *\<^sub>R (a + b)) \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6517 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6518 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6519 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6520 |
assume "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6521 |
then have "a \<bullet> i < ((1 / 2) *\<^sub>R (a + b)) \<bullet> i \<and> ((1 / 2) *\<^sub>R (a + b)) \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6522 |
using assms[unfolded box_ne_empty, THEN bspec[where x=i]] by (auto simp: inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6523 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6524 |
then show ?thesis unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6525 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6526 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6527 |
lemma open_cbox_convex: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6528 |
fixes x :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6529 |
assumes x: "x \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6530 |
and y: "y \<in> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6531 |
and e: "0 < e" "e \<le> 1" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6532 |
shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6533 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6534 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6535 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6536 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6537 |
have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6538 |
unfolding left_diff_distrib by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6539 |
also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6540 |
apply (rule add_less_le_mono) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6541 |
using e unfolding mult_less_cancel_left and mult_le_cancel_left |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6542 |
apply simp_all |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6543 |
using x unfolding mem_box using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6544 |
apply simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6545 |
using y unfolding mem_box using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6546 |
apply simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6547 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6548 |
finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6549 |
unfolding inner_simps by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6550 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6551 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6552 |
have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6553 |
unfolding left_diff_distrib by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6554 |
also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6555 |
apply (rule add_less_le_mono) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6556 |
using e unfolding mult_less_cancel_left and mult_le_cancel_left |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6557 |
apply simp_all |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6558 |
using x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6559 |
unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6560 |
using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6561 |
apply simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6562 |
using y |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6563 |
unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6564 |
using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6565 |
apply simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6566 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6567 |
finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6568 |
unfolding inner_simps by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6569 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6570 |
ultimately have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6571 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6572 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6573 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6574 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6575 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6576 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6577 |
lemma closure_box: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6578 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6579 |
assumes "box a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6580 |
shows "closure (box a b) = cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6581 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6582 |
have ab: "a <e b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6583 |
using assms by (simp add: eucl_less_def box_ne_empty) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6584 |
let ?c = "(1 / 2) *\<^sub>R (a + b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6585 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6586 |
fix x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6587 |
assume as:"x \<in> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6588 |
def f \<equiv> "\<lambda>n::nat. x + (inverse (real n + 1)) *\<^sub>R (?c - x)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6589 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6590 |
fix n |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6591 |
assume fn: "f n <e b \<longrightarrow> a <e f n \<longrightarrow> f n = x" and xc: "x \<noteq> ?c" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6592 |
have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6593 |
unfolding inverse_le_1_iff by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6594 |
have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6595 |
x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6596 |
by (auto simp add: algebra_simps) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6597 |
then have "f n <e b" and "a <e f n" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6598 |
using open_cbox_convex[OF box_midpoint[OF assms] as *] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6599 |
unfolding f_def by (auto simp: box_def eucl_less_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6600 |
then have False |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6601 |
using fn unfolding f_def using xc by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6602 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6603 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6604 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6605 |
assume "\<not> (f ---> x) sequentially" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6606 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6607 |
fix e :: real |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6608 |
assume "e > 0" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6609 |
then have "\<exists>N::nat. inverse (real (N + 1)) < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6610 |
using real_arch_inv[of e] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6611 |
apply (auto simp add: Suc_pred') |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6612 |
apply (rule_tac x="n - 1" in exI) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6613 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6614 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6615 |
then obtain N :: nat where "inverse (real (N + 1)) < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6616 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6617 |
then have "\<forall>n\<ge>N. inverse (real n + 1) < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6618 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6619 |
apply (metis Suc_le_mono le_SucE less_imp_inverse_less nat_le_real_less order_less_trans |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6620 |
real_of_nat_Suc real_of_nat_Suc_gt_zero) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6621 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6622 |
then have "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6623 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6624 |
then have "((\<lambda>n. inverse (real n + 1)) ---> 0) sequentially" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6625 |
unfolding LIMSEQ_def by(auto simp add: dist_norm) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6626 |
then have "(f ---> x) sequentially" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6627 |
unfolding f_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6628 |
using tendsto_add[OF tendsto_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6629 |
using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6630 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6631 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6632 |
ultimately have "x \<in> closure (box a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6633 |
using as and box_midpoint[OF assms] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6634 |
unfolding closure_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6635 |
unfolding islimpt_sequential |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6636 |
by (cases "x=?c") (auto simp: in_box_eucl_less) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6637 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6638 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6639 |
using closure_minimal[OF box_subset_cbox, of a b] by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6640 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6641 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6642 |
lemma bounded_subset_box_symmetric: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6643 |
fixes s::"('a::euclidean_space) set"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6644 |
assumes "bounded s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6645 |
shows "\<exists>a. s \<subseteq> box (-a) a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6646 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6647 |
obtain b where "b>0" and b: "\<forall>x\<in>s. norm x \<le> b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6648 |
using assms[unfolded bounded_pos] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6649 |
def a \<equiv> "(\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)::'a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6650 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6651 |
fix x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6652 |
assume "x \<in> s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6653 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6654 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6655 |
then have "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6656 |
using b[THEN bspec[where x=x], OF `x\<in>s`] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6657 |
using Basis_le_norm[OF i, of x] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6658 |
unfolding inner_simps and a_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6659 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6660 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6661 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6662 |
by (auto intro: exI[where x=a] simp add: box_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6663 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6664 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6665 |
lemma bounded_subset_open_interval: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6666 |
fixes s :: "('a::euclidean_space) set"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6667 |
shows "bounded s \<Longrightarrow> (\<exists>a b. s \<subseteq> box a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6668 |
by (auto dest!: bounded_subset_box_symmetric) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6669 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6670 |
lemma bounded_subset_cbox_symmetric: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6671 |
fixes s :: "('a::euclidean_space) set"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6672 |
assumes "bounded s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6673 |
shows "\<exists>a. s \<subseteq> cbox (-a) a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6674 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6675 |
obtain a where "s \<subseteq> box (-a) a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6676 |
using bounded_subset_box_symmetric[OF assms] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6677 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6678 |
using box_subset_cbox[of "-a" a] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6679 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6680 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6681 |
lemma bounded_subset_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6682 |
fixes s :: "('a::euclidean_space) set"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6683 |
shows "bounded s \<Longrightarrow> \<exists>a b. s \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6684 |
using bounded_subset_cbox_symmetric[of s] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6685 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6686 |
lemma frontier_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6687 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6688 |
shows "frontier (cbox a b) = cbox a b - box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6689 |
unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6690 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6691 |
lemma frontier_box: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6692 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6693 |
shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6694 |
proof (cases "box a b = {}")
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6695 |
case True |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6696 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6697 |
using frontier_empty by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6698 |
next |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6699 |
case False |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6700 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6701 |
unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6702 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6703 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6704 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6705 |
lemma inter_interval_mixed_eq_empty: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6706 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6707 |
assumes "box c d \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6708 |
shows "box a b \<inter> cbox c d = {} \<longleftrightarrow> box a b \<inter> box c d = {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6709 |
unfolding closure_box[OF assms, symmetric] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6710 |
unfolding open_inter_closure_eq_empty[OF open_box] .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6711 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6712 |
lemma diameter_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6713 |
fixes a b::"'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6714 |
shows "(\<forall>i \<in> Basis. a \<bullet> i \<le> b \<bullet> i) \<Longrightarrow> diameter (cbox a b) = dist a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6715 |
by (force simp add: diameter_def SUP_def simp del: Sup_image_eq intro!: cSup_eq_maximum setL2_mono |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6716 |
simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6717 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6718 |
lemma eucl_less_eq_halfspaces: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6719 |
fixes a :: "'a\<Colon>euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6720 |
shows "{x. x <e a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6721 |
"{x. a <e x} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6722 |
by (auto simp: eucl_less_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6723 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6724 |
lemma eucl_le_eq_halfspaces: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6725 |
fixes a :: "'a\<Colon>euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6726 |
shows "{x. \<forall>i\<in>Basis. x \<bullet> i \<le> a \<bullet> i} = (\<Inter>i\<in>Basis. {x. x \<bullet> i \<le> a \<bullet> i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6727 |
"{x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i} = (\<Inter>i\<in>Basis. {x. a \<bullet> i \<le> x \<bullet> i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6728 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6729 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6730 |
lemma open_Collect_eucl_less[simp, intro]: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6731 |
fixes a :: "'a\<Colon>euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6732 |
shows "open {x. x <e a}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6733 |
"open {x. a <e x}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6734 |
by (auto simp: eucl_less_eq_halfspaces open_halfspace_component_lt open_halfspace_component_gt) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6735 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6736 |
lemma closed_Collect_eucl_le[simp, intro]: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6737 |
fixes a :: "'a\<Colon>euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6738 |
shows "closed {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6739 |
"closed {x. \<forall>i\<in>Basis. x \<bullet> i \<le> a \<bullet> i}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6740 |
unfolding eucl_le_eq_halfspaces |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6741 |
by (simp_all add: closed_INT closed_Collect_le) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6742 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6743 |
lemma image_affinity_cbox: fixes m::real |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6744 |
fixes a b c :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6745 |
shows "(\<lambda>x. m *\<^sub>R x + c) ` cbox a b = |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6746 |
(if cbox a b = {} then {}
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6747 |
else (if 0 \<le> m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6748 |
else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6749 |
proof (cases "m = 0") |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6750 |
case True |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6751 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6752 |
fix x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6753 |
assume "\<forall>i\<in>Basis. x \<bullet> i \<le> c \<bullet> i" "\<forall>i\<in>Basis. c \<bullet> i \<le> x \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6754 |
then have "x = c" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6755 |
apply - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6756 |
apply (subst euclidean_eq_iff) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6757 |
apply (auto intro: order_antisym) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6758 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6759 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6760 |
moreover have "c \<in> cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6761 |
unfolding True by (auto simp add: cbox_sing) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6762 |
ultimately show ?thesis using True by (auto simp: cbox_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6763 |
next |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6764 |
case False |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6765 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6766 |
fix y |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6767 |
assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m > 0" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6768 |
then have "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6769 |
by (auto simp: inner_distrib) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6770 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6771 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6772 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6773 |
fix y |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6774 |
assume "\<forall>i\<in>Basis. a \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> b \<bullet> i" "m < 0" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6775 |
then have "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> (m *\<^sub>R y + c) \<bullet> i" and "\<forall>i\<in>Basis. (m *\<^sub>R y + c) \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6776 |
by (auto simp add: mult_left_mono_neg inner_distrib) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6777 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6778 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6779 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6780 |
fix y |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6781 |
assume "m > 0" and "\<forall>i\<in>Basis. (m *\<^sub>R a + c) \<bullet> i \<le> y \<bullet> i" and "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R b + c) \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6782 |
then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6783 |
unfolding image_iff Bex_def mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6784 |
apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6785 |
apply (auto simp add: pos_le_divide_eq pos_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6786 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6787 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6788 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6789 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6790 |
fix y |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6791 |
assume "\<forall>i\<in>Basis. (m *\<^sub>R b + c) \<bullet> i \<le> y \<bullet> i" "\<forall>i\<in>Basis. y \<bullet> i \<le> (m *\<^sub>R a + c) \<bullet> i" "m < 0" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6792 |
then have "y \<in> (\<lambda>x. m *\<^sub>R x + c) ` cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6793 |
unfolding image_iff Bex_def mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6794 |
apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6795 |
apply (auto simp add: neg_le_divide_eq neg_divide_le_eq mult_commute diff_le_iff inner_distrib inner_diff_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6796 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6797 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6798 |
ultimately show ?thesis using False by (auto simp: cbox_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6799 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6800 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6801 |
lemma image_smult_cbox:"(\<lambda>x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6802 |
(if cbox a b = {} then {} else if 0 \<le> m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6803 |
using image_affinity_cbox[of m 0 a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6804 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6805 |
|
| 36437 | 6806 |
subsection {* Homeomorphisms *}
|
| 33175 | 6807 |
|
| 52625 | 6808 |
definition "homeomorphism s t f g \<longleftrightarrow> |
6809 |
(\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and> |
|
6810 |
(\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g" |
|
| 33175 | 6811 |
|
| 53640 | 6812 |
definition homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool" |
| 53282 | 6813 |
(infixr "homeomorphic" 60) |
6814 |
where "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)" |
|
| 33175 | 6815 |
|
6816 |
lemma homeomorphic_refl: "s homeomorphic s" |
|
6817 |
unfolding homeomorphic_def |
|
6818 |
unfolding homeomorphism_def |
|
6819 |
using continuous_on_id |
|
| 53282 | 6820 |
apply (rule_tac x = "(\<lambda>x. x)" in exI) |
6821 |
apply (rule_tac x = "(\<lambda>x. x)" in exI) |
|
| 52625 | 6822 |
apply blast |
6823 |
done |
|
6824 |
||
6825 |
lemma homeomorphic_sym: "s homeomorphic t \<longleftrightarrow> t homeomorphic s" |
|
6826 |
unfolding homeomorphic_def |
|
6827 |
unfolding homeomorphism_def |
|
| 53282 | 6828 |
by blast |
| 33175 | 6829 |
|
6830 |
lemma homeomorphic_trans: |
|
| 53282 | 6831 |
assumes "s homeomorphic t" |
6832 |
and "t homeomorphic u" |
|
| 52625 | 6833 |
shows "s homeomorphic u" |
| 53282 | 6834 |
proof - |
6835 |
obtain f1 g1 where fg1: "\<forall>x\<in>s. g1 (f1 x) = x" "f1 ` s = t" |
|
6836 |
"continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1" |
|
| 33175 | 6837 |
using assms(1) unfolding homeomorphic_def homeomorphism_def by auto |
| 53282 | 6838 |
obtain f2 g2 where fg2: "\<forall>x\<in>t. g2 (f2 x) = x" "f2 ` t = u" "continuous_on t f2" |
6839 |
"\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2" |
|
| 33175 | 6840 |
using assms(2) unfolding homeomorphic_def homeomorphism_def by auto |
| 52625 | 6841 |
{
|
6842 |
fix x |
|
6843 |
assume "x\<in>s" |
|
| 53282 | 6844 |
then have "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x" |
| 52625 | 6845 |
using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2) |
6846 |
by auto |
|
6847 |
} |
|
6848 |
moreover have "(f2 \<circ> f1) ` s = u" |
|
6849 |
using fg1(2) fg2(2) by auto |
|
6850 |
moreover have "continuous_on s (f2 \<circ> f1)" |
|
6851 |
using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto |
|
| 53282 | 6852 |
moreover |
6853 |
{
|
|
| 52625 | 6854 |
fix y |
6855 |
assume "y\<in>u" |
|
| 53282 | 6856 |
then have "(f2 \<circ> f1) ((g1 \<circ> g2) y) = y" |
| 52625 | 6857 |
using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5) |
6858 |
by auto |
|
6859 |
} |
|
| 33175 | 6860 |
moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto |
| 52625 | 6861 |
moreover have "continuous_on u (g1 \<circ> g2)" |
6862 |
using continuous_on_compose[OF fg2(6)] and fg1(6) |
|
6863 |
unfolding fg2(5) |
|
6864 |
by auto |
|
6865 |
ultimately show ?thesis |
|
6866 |
unfolding homeomorphic_def homeomorphism_def |
|
6867 |
apply (rule_tac x="f2 \<circ> f1" in exI) |
|
6868 |
apply (rule_tac x="g1 \<circ> g2" in exI) |
|
6869 |
apply auto |
|
6870 |
done |
|
| 33175 | 6871 |
qed |
6872 |
||
6873 |
lemma homeomorphic_minimal: |
|
| 52625 | 6874 |
"s homeomorphic t \<longleftrightarrow> |
| 33175 | 6875 |
(\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and> |
6876 |
(\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and> |
|
6877 |
continuous_on s f \<and> continuous_on t g)" |
|
| 52625 | 6878 |
unfolding homeomorphic_def homeomorphism_def |
6879 |
apply auto |
|
6880 |
apply (rule_tac x=f in exI) |
|
6881 |
apply (rule_tac x=g in exI) |
|
6882 |
apply auto |
|
6883 |
apply (rule_tac x=f in exI) |
|
6884 |
apply (rule_tac x=g in exI) |
|
6885 |
apply auto |
|
6886 |
unfolding image_iff |
|
6887 |
apply (erule_tac x="g x" in ballE) |
|
6888 |
apply (erule_tac x="x" in ballE) |
|
6889 |
apply auto |
|
6890 |
apply (rule_tac x="g x" in bexI) |
|
6891 |
apply auto |
|
6892 |
apply (erule_tac x="f x" in ballE) |
|
6893 |
apply (erule_tac x="x" in ballE) |
|
6894 |
apply auto |
|
6895 |
apply (rule_tac x="f x" in bexI) |
|
6896 |
apply auto |
|
6897 |
done |
|
| 33175 | 6898 |
|
| 36437 | 6899 |
text {* Relatively weak hypotheses if a set is compact. *}
|
| 33175 | 6900 |
|
6901 |
lemma homeomorphism_compact: |
|
| 50898 | 6902 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space" |
| 33175 | 6903 |
assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s" |
6904 |
shows "\<exists>g. homeomorphism s t f g" |
|
| 53282 | 6905 |
proof - |
| 33175 | 6906 |
def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x" |
| 52625 | 6907 |
have g: "\<forall>x\<in>s. g (f x) = x" |
6908 |
using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto |
|
6909 |
{
|
|
| 53282 | 6910 |
fix y |
6911 |
assume "y \<in> t" |
|
6912 |
then obtain x where x:"f x = y" "x\<in>s" |
|
6913 |
using assms(3) by auto |
|
6914 |
then have "g (f x) = x" using g by auto |
|
| 53291 | 6915 |
then have "f (g y) = y" unfolding x(1)[symmetric] by auto |
| 52625 | 6916 |
} |
| 53282 | 6917 |
then have g':"\<forall>x\<in>t. f (g x) = x" by auto |
| 33175 | 6918 |
moreover |
| 52625 | 6919 |
{
|
6920 |
fix x |
|
6921 |
have "x\<in>s \<Longrightarrow> x \<in> g ` t" |
|
6922 |
using g[THEN bspec[where x=x]] |
|
6923 |
unfolding image_iff |
|
6924 |
using assms(3) |
|
6925 |
by (auto intro!: bexI[where x="f x"]) |
|
| 33175 | 6926 |
moreover |
| 52625 | 6927 |
{
|
6928 |
assume "x\<in>g ` t" |
|
| 33175 | 6929 |
then obtain y where y:"y\<in>t" "g y = x" by auto |
| 52625 | 6930 |
then obtain x' where x':"x'\<in>s" "f x' = y" |
6931 |
using assms(3) by auto |
|
| 53282 | 6932 |
then have "x \<in> s" |
| 52625 | 6933 |
unfolding g_def |
6934 |
using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"] |
|
| 53291 | 6935 |
unfolding y(2)[symmetric] and g_def |
| 52625 | 6936 |
by auto |
6937 |
} |
|
6938 |
ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" .. |
|
6939 |
} |
|
| 53282 | 6940 |
then have "g ` t = s" by auto |
| 52625 | 6941 |
ultimately show ?thesis |
6942 |
unfolding homeomorphism_def homeomorphic_def |
|
6943 |
apply (rule_tac x=g in exI) |
|
6944 |
using g and assms(3) and continuous_on_inv[OF assms(2,1), of g, unfolded assms(3)] and assms(2) |
|
6945 |
apply auto |
|
6946 |
done |
|
| 33175 | 6947 |
qed |
6948 |
||
6949 |
lemma homeomorphic_compact: |
|
| 50898 | 6950 |
fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space" |
| 53282 | 6951 |
shows "compact s \<Longrightarrow> continuous_on s f \<Longrightarrow> (f ` s = t) \<Longrightarrow> inj_on f s \<Longrightarrow> s homeomorphic t" |
|
37486
b993fac7985b
beta-eta was too much, because it transformed SOME x. P x into Eps P, which caused problems later;
blanchet
parents:
37452
diff
changeset
|
6952 |
unfolding homeomorphic_def by (metis homeomorphism_compact) |
| 33175 | 6953 |
|
| 53282 | 6954 |
text{* Preservation of topological properties. *}
|
| 33175 | 6955 |
|
| 52625 | 6956 |
lemma homeomorphic_compactness: "s homeomorphic t \<Longrightarrow> (compact s \<longleftrightarrow> compact t)" |
6957 |
unfolding homeomorphic_def homeomorphism_def |
|
6958 |
by (metis compact_continuous_image) |
|
| 33175 | 6959 |
|
| 53282 | 6960 |
text{* Results on translation, scaling etc. *}
|
| 33175 | 6961 |
|
6962 |
lemma homeomorphic_scaling: |
|
6963 |
fixes s :: "'a::real_normed_vector set" |
|
| 53282 | 6964 |
assumes "c \<noteq> 0" |
6965 |
shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)" |
|
| 33175 | 6966 |
unfolding homeomorphic_minimal |
| 52625 | 6967 |
apply (rule_tac x="\<lambda>x. c *\<^sub>R x" in exI) |
6968 |
apply (rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI) |
|
6969 |
using assms |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
6970 |
apply (auto simp add: continuous_intros) |
| 52625 | 6971 |
done |
| 33175 | 6972 |
|
6973 |
lemma homeomorphic_translation: |
|
6974 |
fixes s :: "'a::real_normed_vector set" |
|
6975 |
shows "s homeomorphic ((\<lambda>x. a + x) ` s)" |
|
6976 |
unfolding homeomorphic_minimal |
|
| 52625 | 6977 |
apply (rule_tac x="\<lambda>x. a + x" in exI) |
6978 |
apply (rule_tac x="\<lambda>x. -a + x" in exI) |
|
|
54230
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
54070
diff
changeset
|
6979 |
using continuous_on_add [OF continuous_on_const continuous_on_id, of s a] |
|
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
54070
diff
changeset
|
6980 |
continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"] |
| 52625 | 6981 |
apply auto |
6982 |
done |
|
| 33175 | 6983 |
|
6984 |
lemma homeomorphic_affinity: |
|
6985 |
fixes s :: "'a::real_normed_vector set" |
|
| 52625 | 6986 |
assumes "c \<noteq> 0" |
6987 |
shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)" |
|
| 53282 | 6988 |
proof - |
| 52625 | 6989 |
have *: "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto |
| 33175 | 6990 |
show ?thesis |
6991 |
using homeomorphic_trans |
|
6992 |
using homeomorphic_scaling[OF assms, of s] |
|
| 52625 | 6993 |
using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a] |
6994 |
unfolding * |
|
6995 |
by auto |
|
| 33175 | 6996 |
qed |
6997 |
||
6998 |
lemma homeomorphic_balls: |
|
| 50898 | 6999 |
fixes a b ::"'a::real_normed_vector" |
| 33175 | 7000 |
assumes "0 < d" "0 < e" |
7001 |
shows "(ball a d) homeomorphic (ball b e)" (is ?th) |
|
| 53282 | 7002 |
and "(cball a d) homeomorphic (cball b e)" (is ?cth) |
7003 |
proof - |
|
| 33175 | 7004 |
show ?th unfolding homeomorphic_minimal |
7005 |
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) |
|
7006 |
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) |
|
| 51364 | 7007 |
using assms |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
7008 |
apply (auto intro!: continuous_intros |
| 52625 | 7009 |
simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono) |
| 51364 | 7010 |
done |
| 33175 | 7011 |
show ?cth unfolding homeomorphic_minimal |
7012 |
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI) |
|
7013 |
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI) |
|
| 51364 | 7014 |
using assms |
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
7015 |
apply (auto intro!: continuous_intros |
| 52625 | 7016 |
simp: dist_commute dist_norm pos_divide_le_eq mult_strict_left_mono) |
| 51364 | 7017 |
done |
| 33175 | 7018 |
qed |
7019 |
||
7020 |
text{* "Isometry" (up to constant bounds) of injective linear map etc. *}
|
|
7021 |
||
7022 |
lemma cauchy_isometric: |
|
| 53640 | 7023 |
assumes e: "e > 0" |
| 52625 | 7024 |
and s: "subspace s" |
7025 |
and f: "bounded_linear f" |
|
| 53640 | 7026 |
and normf: "\<forall>x\<in>s. norm (f x) \<ge> e * norm x" |
7027 |
and xs: "\<forall>n. x n \<in> s" |
|
7028 |
and cf: "Cauchy (f \<circ> x)" |
|
| 33175 | 7029 |
shows "Cauchy x" |
| 52625 | 7030 |
proof - |
| 33175 | 7031 |
interpret f: bounded_linear f by fact |
| 52625 | 7032 |
{
|
| 53291 | 7033 |
fix d :: real |
7034 |
assume "d > 0" |
|
| 33175 | 7035 |
then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d" |
| 56544 | 7036 |
using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]] e |
| 52625 | 7037 |
by auto |
7038 |
{
|
|
7039 |
fix n |
|
7040 |
assume "n\<ge>N" |
|
|
45270
d5b5c9259afd
fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents:
45051
diff
changeset
|
7041 |
have "e * norm (x n - x N) \<le> norm (f (x n - x N))" |
| 52625 | 7042 |
using subspace_sub[OF s, of "x n" "x N"] |
7043 |
using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]] |
|
7044 |
using normf[THEN bspec[where x="x n - x N"]] |
|
7045 |
by auto |
|
|
45270
d5b5c9259afd
fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
huffman
parents:
45051
diff
changeset
|
7046 |
also have "norm (f (x n - x N)) < e * d" |
| 53291 | 7047 |
using `N \<le> n` N unfolding f.diff[symmetric] by auto |
| 52625 | 7048 |
finally have "norm (x n - x N) < d" using `e>0` by simp |
7049 |
} |
|
| 53282 | 7050 |
then have "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto |
| 52625 | 7051 |
} |
| 53282 | 7052 |
then show ?thesis unfolding cauchy and dist_norm by auto |
| 33175 | 7053 |
qed |
7054 |
||
7055 |
lemma complete_isometric_image: |
|
| 52625 | 7056 |
assumes "0 < e" |
7057 |
and s: "subspace s" |
|
7058 |
and f: "bounded_linear f" |
|
7059 |
and normf: "\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)" |
|
7060 |
and cs: "complete s" |
|
| 53291 | 7061 |
shows "complete (f ` s)" |
| 52625 | 7062 |
proof - |
7063 |
{
|
|
7064 |
fix g |
|
7065 |
assume as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g" |
|
| 53282 | 7066 |
then obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)" |
| 53640 | 7067 |
using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"] |
7068 |
by auto |
|
7069 |
then have x:"\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)" |
|
7070 |
by auto |
|
7071 |
then have "f \<circ> x = g" |
|
7072 |
unfolding fun_eq_iff |
|
7073 |
by auto |
|
| 33175 | 7074 |
then obtain l where "l\<in>s" and l:"(x ---> l) sequentially" |
7075 |
using cs[unfolded complete_def, THEN spec[where x="x"]] |
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
7076 |
using cauchy_isometric[OF `0 < e` s f normf] and cfg and x(1) |
| 53640 | 7077 |
by auto |
| 53282 | 7078 |
then have "\<exists>l\<in>f ` s. (g ---> l) sequentially" |
| 33175 | 7079 |
using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l] |
| 53640 | 7080 |
unfolding `f \<circ> x = g` |
7081 |
by auto |
|
| 52625 | 7082 |
} |
| 53640 | 7083 |
then show ?thesis |
7084 |
unfolding complete_def by auto |
|
| 33175 | 7085 |
qed |
7086 |
||
| 52625 | 7087 |
lemma injective_imp_isometric: |
7088 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
7089 |
assumes s: "closed s" "subspace s" |
|
| 53640 | 7090 |
and f: "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" |
7091 |
shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm x" |
|
| 52625 | 7092 |
proof (cases "s \<subseteq> {0::'a}")
|
| 33175 | 7093 |
case True |
| 52625 | 7094 |
{
|
7095 |
fix x |
|
7096 |
assume "x \<in> s" |
|
| 53282 | 7097 |
then have "x = 0" using True by auto |
7098 |
then have "norm x \<le> norm (f x)" by auto |
|
| 52625 | 7099 |
} |
| 53282 | 7100 |
then show ?thesis by (auto intro!: exI[where x=1]) |
| 33175 | 7101 |
next |
7102 |
interpret f: bounded_linear f by fact |
|
7103 |
case False |
|
| 53640 | 7104 |
then obtain a where a: "a \<noteq> 0" "a \<in> s" |
7105 |
by auto |
|
7106 |
from False have "s \<noteq> {}"
|
|
7107 |
by auto |
|
| 33175 | 7108 |
let ?S = "{f x| x. (x \<in> s \<and> norm x = norm a)}"
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
7109 |
let ?S' = "{x::'a. x\<in>s \<and> norm x = norm a}"
|
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
7110 |
let ?S'' = "{x::'a. norm x = norm a}"
|
| 33175 | 7111 |
|
| 52625 | 7112 |
have "?S'' = frontier(cball 0 (norm a))" |
7113 |
unfolding frontier_cball and dist_norm by auto |
|
| 53282 | 7114 |
then have "compact ?S''" |
| 52625 | 7115 |
using compact_frontier[OF compact_cball, of 0 "norm a"] by auto |
| 33175 | 7116 |
moreover have "?S' = s \<inter> ?S''" by auto |
| 52625 | 7117 |
ultimately have "compact ?S'" |
7118 |
using closed_inter_compact[of s ?S''] using s(1) by auto |
|
| 33175 | 7119 |
moreover have *:"f ` ?S' = ?S" by auto |
| 52625 | 7120 |
ultimately have "compact ?S" |
7121 |
using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto |
|
| 53282 | 7122 |
then have "closed ?S" using compact_imp_closed by auto |
| 33175 | 7123 |
moreover have "?S \<noteq> {}" using a by auto
|
| 52625 | 7124 |
ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y" |
7125 |
using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto |
|
| 53282 | 7126 |
then obtain b where "b\<in>s" |
7127 |
and ba: "norm b = norm a" |
|
7128 |
and b: "\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)"
|
|
| 53291 | 7129 |
unfolding *[symmetric] unfolding image_iff by auto |
| 33175 | 7130 |
|
7131 |
let ?e = "norm (f b) / norm b" |
|
7132 |
have "norm b > 0" using ba and a and norm_ge_zero by auto |
|
| 52625 | 7133 |
moreover have "norm (f b) > 0" |
7134 |
using f(2)[THEN bspec[where x=b], OF `b\<in>s`] |
|
7135 |
using `norm b >0` |
|
7136 |
unfolding zero_less_norm_iff |
|
7137 |
by auto |
|
| 56541 | 7138 |
ultimately have "0 < norm (f b) / norm b" by simp |
| 33175 | 7139 |
moreover |
| 52625 | 7140 |
{
|
7141 |
fix x |
|
7142 |
assume "x\<in>s" |
|
| 53282 | 7143 |
then have "norm (f b) / norm b * norm x \<le> norm (f x)" |
| 52625 | 7144 |
proof (cases "x=0") |
7145 |
case True |
|
| 53282 | 7146 |
then show "norm (f b) / norm b * norm x \<le> norm (f x)" by auto |
| 33175 | 7147 |
next |
7148 |
case False |
|
| 53282 | 7149 |
then have *: "0 < norm a / norm x" |
| 52625 | 7150 |
using `a\<noteq>0` |
| 56541 | 7151 |
unfolding zero_less_norm_iff[symmetric] by simp |
| 52625 | 7152 |
have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s" |
7153 |
using s[unfolded subspace_def] by auto |
|
| 53282 | 7154 |
then have "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}"
|
| 52625 | 7155 |
using `x\<in>s` and `x\<noteq>0` by auto |
| 53282 | 7156 |
then show "norm (f b) / norm b * norm x \<le> norm (f x)" |
| 52625 | 7157 |
using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] |
| 33175 | 7158 |
unfolding f.scaleR and ba using `x\<noteq>0` `a\<noteq>0` |
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36669
diff
changeset
|
7159 |
by (auto simp add: mult_commute pos_le_divide_eq pos_divide_le_eq) |
| 52625 | 7160 |
qed |
7161 |
} |
|
7162 |
ultimately show ?thesis by auto |
|
| 33175 | 7163 |
qed |
7164 |
||
7165 |
lemma closed_injective_image_subspace: |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
7166 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
| 53282 | 7167 |
assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" "closed s" |
| 33175 | 7168 |
shows "closed(f ` s)" |
| 53282 | 7169 |
proof - |
7170 |
obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)" |
|
| 52625 | 7171 |
using injective_imp_isometric[OF assms(4,1,2,3)] by auto |
7172 |
show ?thesis |
|
7173 |
using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4) |
|
| 53291 | 7174 |
unfolding complete_eq_closed[symmetric] by auto |
| 33175 | 7175 |
qed |
7176 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
7177 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
7178 |
subsection {* Some properties of a canonical subspace *}
|
| 33175 | 7179 |
|
7180 |
lemma subspace_substandard: |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7181 |
"subspace {x::'a::euclidean_space. (\<forall>i\<in>Basis. P i \<longrightarrow> x\<bullet>i = 0)}"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7182 |
unfolding subspace_def by (auto simp: inner_add_left) |
| 33175 | 7183 |
|
7184 |
lemma closed_substandard: |
|
| 52625 | 7185 |
"closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i --> x\<bullet>i = 0}" (is "closed ?A")
|
7186 |
proof - |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7187 |
let ?D = "{i\<in>Basis. P i}"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7188 |
have "closed (\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0})"
|
|
44457
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
huffman
parents:
44365
diff
changeset
|
7189 |
by (simp add: closed_INT closed_Collect_eq) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7190 |
also have "(\<Inter>i\<in>?D. {x::'a. x\<bullet>i = 0}) = ?A"
|
|
44457
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
huffman
parents:
44365
diff
changeset
|
7191 |
by auto |
|
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
huffman
parents:
44365
diff
changeset
|
7192 |
finally show "closed ?A" . |
| 33175 | 7193 |
qed |
7194 |
||
| 52625 | 7195 |
lemma dim_substandard: |
7196 |
assumes d: "d \<subseteq> Basis" |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7197 |
shows "dim {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0} = card d" (is "dim ?A = _")
|
| 53813 | 7198 |
proof (rule dim_unique) |
7199 |
show "d \<subseteq> ?A" |
|
7200 |
using d by (auto simp: inner_Basis) |
|
7201 |
show "independent d" |
|
7202 |
using independent_mono [OF independent_Basis d] . |
|
7203 |
show "?A \<subseteq> span d" |
|
7204 |
proof (clarify) |
|
7205 |
fix x assume x: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0" |
|
7206 |
have "finite d" |
|
7207 |
using finite_subset [OF d finite_Basis] . |
|
7208 |
then have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) \<in> span d" |
|
7209 |
by (simp add: span_setsum span_clauses) |
|
7210 |
also have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)" |
|
7211 |
by (rule setsum_mono_zero_cong_left [OF finite_Basis d]) (auto simp add: x) |
|
7212 |
finally show "x \<in> span d" |
|
7213 |
unfolding euclidean_representation . |
|
7214 |
qed |
|
7215 |
qed simp |
|
| 33175 | 7216 |
|
| 53282 | 7217 |
text{* Hence closure and completeness of all subspaces. *}
|
7218 |
||
7219 |
lemma ex_card: |
|
7220 |
assumes "n \<le> card A" |
|
7221 |
shows "\<exists>S\<subseteq>A. card S = n" |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7222 |
proof cases |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7223 |
assume "finite A" |
| 55522 | 7224 |
from ex_bij_betw_nat_finite[OF this] obtain f where f: "bij_betw f {0..<card A} A" ..
|
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
7225 |
moreover from f `n \<le> card A` have "{..< n} \<subseteq> {..< card A}" "inj_on f {..< n}"
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7226 |
by (auto simp: bij_betw_def intro: subset_inj_on) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7227 |
ultimately have "f ` {..< n} \<subseteq> A" "card (f ` {..< n}) = n"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7228 |
by (auto simp: bij_betw_def card_image) |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7229 |
then show ?thesis by blast |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7230 |
next |
| 52625 | 7231 |
assume "\<not> finite A" |
7232 |
with `n \<le> card A` show ?thesis by force |
|
7233 |
qed |
|
7234 |
||
7235 |
lemma closed_subspace: |
|
| 53291 | 7236 |
fixes s :: "'a::euclidean_space set" |
| 52625 | 7237 |
assumes "subspace s" |
7238 |
shows "closed s" |
|
7239 |
proof - |
|
7240 |
have "dim s \<le> card (Basis :: 'a set)" |
|
7241 |
using dim_subset_UNIV by auto |
|
7242 |
with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis" |
|
7243 |
by auto |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7244 |
let ?t = "{x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0}"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7245 |
have "\<exists>f. linear f \<and> f ` {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s \<and>
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7246 |
inj_on f {x::'a. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
|
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7247 |
using dim_substandard[of d] t d assms |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
7248 |
by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_Basis) |
| 55522 | 7249 |
then obtain f where f: |
7250 |
"linear f" |
|
7251 |
"f ` {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0} = s"
|
|
7252 |
"inj_on f {x. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0}"
|
|
7253 |
by blast |
|
| 52625 | 7254 |
interpret f: bounded_linear f |
7255 |
using f unfolding linear_conv_bounded_linear by auto |
|
7256 |
{
|
|
7257 |
fix x |
|
7258 |
have "x\<in>?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0" |
|
7259 |
using f.zero d f(3)[THEN inj_onD, of x 0] by auto |
|
7260 |
} |
|
| 33175 | 7261 |
moreover have "closed ?t" using closed_substandard . |
7262 |
moreover have "subspace ?t" using subspace_substandard . |
|
| 52625 | 7263 |
ultimately show ?thesis |
7264 |
using closed_injective_image_subspace[of ?t f] |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
7265 |
unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto |
| 33175 | 7266 |
qed |
7267 |
||
7268 |
lemma complete_subspace: |
|
| 52625 | 7269 |
fixes s :: "('a::euclidean_space) set"
|
7270 |
shows "subspace s \<Longrightarrow> complete s" |
|
7271 |
using complete_eq_closed closed_subspace by auto |
|
| 33175 | 7272 |
|
7273 |
lemma dim_closure: |
|
|
37489
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
7274 |
fixes s :: "('a::euclidean_space) set"
|
| 33175 | 7275 |
shows "dim(closure s) = dim s" (is "?dc = ?d") |
| 52625 | 7276 |
proof - |
| 33175 | 7277 |
have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s] |
7278 |
using closed_subspace[OF subspace_span, of s] |
|
| 52625 | 7279 |
using dim_subset[of "closure s" "span s"] |
7280 |
unfolding dim_span |
|
7281 |
by auto |
|
| 53282 | 7282 |
then show ?thesis using dim_subset[OF closure_subset, of s] |
| 52625 | 7283 |
by auto |
| 33175 | 7284 |
qed |
7285 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
7286 |
|
| 36437 | 7287 |
subsection {* Affine transformations of intervals *}
|
| 33175 | 7288 |
|
7289 |
lemma real_affinity_le: |
|
| 53291 | 7290 |
"0 < (m::'a::linordered_field) \<Longrightarrow> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))" |
| 33175 | 7291 |
by (simp add: field_simps inverse_eq_divide) |
7292 |
||
7293 |
lemma real_le_affinity: |
|
| 53291 | 7294 |
"0 < (m::'a::linordered_field) \<Longrightarrow> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)" |
| 33175 | 7295 |
by (simp add: field_simps inverse_eq_divide) |
7296 |
||
7297 |
lemma real_affinity_lt: |
|
| 53291 | 7298 |
"0 < (m::'a::linordered_field) \<Longrightarrow> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))" |
| 33175 | 7299 |
by (simp add: field_simps inverse_eq_divide) |
7300 |
||
7301 |
lemma real_lt_affinity: |
|
| 53291 | 7302 |
"0 < (m::'a::linordered_field) \<Longrightarrow> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)" |
| 33175 | 7303 |
by (simp add: field_simps inverse_eq_divide) |
7304 |
||
7305 |
lemma real_affinity_eq: |
|
| 53291 | 7306 |
"(m::'a::linordered_field) \<noteq> 0 \<Longrightarrow> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))" |
| 33175 | 7307 |
by (simp add: field_simps inverse_eq_divide) |
7308 |
||
7309 |
lemma real_eq_affinity: |
|
| 53291 | 7310 |
"(m::'a::linordered_field) \<noteq> 0 \<Longrightarrow> (y = m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) = x)" |
| 33175 | 7311 |
by (simp add: field_simps inverse_eq_divide) |
7312 |
||
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
7313 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
7314 |
subsection {* Banach fixed point theorem (not really topological...) *}
|
| 33175 | 7315 |
|
7316 |
lemma banach_fix: |
|
| 53282 | 7317 |
assumes s: "complete s" "s \<noteq> {}"
|
7318 |
and c: "0 \<le> c" "c < 1" |
|
7319 |
and f: "(f ` s) \<subseteq> s" |
|
| 53291 | 7320 |
and lipschitz: "\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y" |
7321 |
shows "\<exists>!x\<in>s. f x = x" |
|
| 53282 | 7322 |
proof - |
| 33175 | 7323 |
have "1 - c > 0" using c by auto |
7324 |
||
7325 |
from s(2) obtain z0 where "z0 \<in> s" by auto |
|
7326 |
def z \<equiv> "\<lambda>n. (f ^^ n) z0" |
|
| 53282 | 7327 |
{
|
7328 |
fix n :: nat |
|
| 33175 | 7329 |
have "z n \<in> s" unfolding z_def |
| 52625 | 7330 |
proof (induct n) |
7331 |
case 0 |
|
| 53282 | 7332 |
then show ?case using `z0 \<in> s` by auto |
| 52625 | 7333 |
next |
7334 |
case Suc |
|
| 53282 | 7335 |
then show ?case using f by auto qed |
| 52625 | 7336 |
} note z_in_s = this |
| 33175 | 7337 |
|
7338 |
def d \<equiv> "dist (z 0) (z 1)" |
|
7339 |
||
7340 |
have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto |
|
| 52625 | 7341 |
{
|
| 53282 | 7342 |
fix n :: nat |
| 33175 | 7343 |
have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d" |
| 52625 | 7344 |
proof (induct n) |
| 53282 | 7345 |
case 0 |
7346 |
then show ?case |
|
| 52625 | 7347 |
unfolding d_def by auto |
| 33175 | 7348 |
next |
7349 |
case (Suc m) |
|
| 53282 | 7350 |
then have "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d" |
| 52625 | 7351 |
using `0 \<le> c` |
7352 |
using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c] |
|
7353 |
by auto |
|
| 53282 | 7354 |
then show ?case |
| 52625 | 7355 |
using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s] |
7356 |
unfolding fzn and mult_le_cancel_left |
|
7357 |
by auto |
|
| 33175 | 7358 |
qed |
7359 |
} note cf_z = this |
|
7360 |
||
| 52625 | 7361 |
{
|
| 53282 | 7362 |
fix n m :: nat |
| 33175 | 7363 |
have "(1 - c) * dist (z m) (z (m+n)) \<le> (c ^ m) * d * (1 - c ^ n)" |
| 52625 | 7364 |
proof (induct n) |
| 53282 | 7365 |
case 0 |
7366 |
show ?case by auto |
|
| 33175 | 7367 |
next |
7368 |
case (Suc k) |
|
| 52625 | 7369 |
have "(1 - c) * dist (z m) (z (m + Suc k)) \<le> |
7370 |
(1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))" |
|
7371 |
using dist_triangle and c by (auto simp add: dist_triangle) |
|
| 33175 | 7372 |
also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)" |
7373 |
using cf_z[of "m + k"] and c by auto |
|
7374 |
also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d" |
|
| 36350 | 7375 |
using Suc by (auto simp add: field_simps) |
| 33175 | 7376 |
also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)" |
| 36350 | 7377 |
unfolding power_add by (auto simp add: field_simps) |
| 33175 | 7378 |
also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)" |
| 36350 | 7379 |
using c by (auto simp add: field_simps) |
| 33175 | 7380 |
finally show ?case by auto |
7381 |
qed |
|
7382 |
} note cf_z2 = this |
|
| 52625 | 7383 |
{
|
| 53282 | 7384 |
fix e :: real |
7385 |
assume "e > 0" |
|
7386 |
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e" |
|
| 52625 | 7387 |
proof (cases "d = 0") |
| 33175 | 7388 |
case True |
| 41863 | 7389 |
have *: "\<And>x. ((1 - c) * x \<le> 0) = (x \<le> 0)" using `1 - c > 0` |
|
45051
c478d1876371
discontinued legacy theorem names from RealDef.thy
huffman
parents:
45031
diff
changeset
|
7390 |
by (metis mult_zero_left mult_commute real_mult_le_cancel_iff1) |
| 41863 | 7391 |
from True have "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def |
7392 |
by (simp add: *) |
|
| 53282 | 7393 |
then show ?thesis using `e>0` by auto |
| 33175 | 7394 |
next |
| 52625 | 7395 |
case False |
| 53282 | 7396 |
then have "d>0" unfolding d_def using zero_le_dist[of "z 0" "z 1"] |
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36669
diff
changeset
|
7397 |
by (metis False d_def less_le) |
| 56541 | 7398 |
hence "0 < e * (1 - c) / d" |
| 56544 | 7399 |
using `e>0` and `1-c>0` by auto |
| 52625 | 7400 |
then obtain N where N:"c ^ N < e * (1 - c) / d" |
7401 |
using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto |
|
7402 |
{
|
|
7403 |
fix m n::nat |
|
7404 |
assume "m>n" and as:"m\<ge>N" "n\<ge>N" |
|
7405 |
have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c |
|
7406 |
using power_decreasing[OF `n\<ge>N`, of c] by auto |
|
7407 |
have "1 - c ^ (m - n) > 0" |
|
7408 |
using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto |
|
| 56541 | 7409 |
hence **: "d * (1 - c ^ (m - n)) / (1 - c) > 0" |
| 56544 | 7410 |
using `d>0` `0 < 1 - c` by auto |
| 33175 | 7411 |
|
7412 |
have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)" |
|
| 52625 | 7413 |
using cf_z2[of n "m - n"] and `m>n` |
7414 |
unfolding pos_le_divide_eq[OF `1-c>0`] |
|
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36669
diff
changeset
|
7415 |
by (auto simp add: mult_commute dist_commute) |
| 33175 | 7416 |
also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)" |
7417 |
using mult_right_mono[OF * order_less_imp_le[OF **]] |
|
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36669
diff
changeset
|
7418 |
unfolding mult_assoc by auto |
| 33175 | 7419 |
also have "\<dots> < (e * (1 - c) / d) * d * (1 - c ^ (m - n)) / (1 - c)" |
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36669
diff
changeset
|
7420 |
using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto |
| 53282 | 7421 |
also have "\<dots> = e * (1 - c ^ (m - n))" |
7422 |
using c and `d>0` and `1 - c > 0` by auto |
|
7423 |
also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0` |
|
7424 |
using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto |
|
| 33175 | 7425 |
finally have "dist (z m) (z n) < e" by auto |
7426 |
} note * = this |
|
| 52625 | 7427 |
{
|
| 53282 | 7428 |
fix m n :: nat |
7429 |
assume as: "N \<le> m" "N \<le> n" |
|
7430 |
then have "dist (z n) (z m) < e" |
|
| 52625 | 7431 |
proof (cases "n = m") |
7432 |
case True |
|
| 53282 | 7433 |
then show ?thesis using `e>0` by auto |
| 33175 | 7434 |
next |
| 52625 | 7435 |
case False |
| 53282 | 7436 |
then show ?thesis using as and *[of n m] *[of m n] |
| 52625 | 7437 |
unfolding nat_neq_iff by (auto simp add: dist_commute) |
7438 |
qed |
|
7439 |
} |
|
| 53282 | 7440 |
then show ?thesis by auto |
| 33175 | 7441 |
qed |
7442 |
} |
|
| 53282 | 7443 |
then have "Cauchy z" |
7444 |
unfolding cauchy_def by auto |
|
| 52625 | 7445 |
then obtain x where "x\<in>s" and x:"(z ---> x) sequentially" |
7446 |
using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto |
|
| 33175 | 7447 |
|
7448 |
def e \<equiv> "dist (f x) x" |
|
| 52625 | 7449 |
have "e = 0" |
7450 |
proof (rule ccontr) |
|
7451 |
assume "e \<noteq> 0" |
|
| 53282 | 7452 |
then have "e > 0" |
7453 |
unfolding e_def using zero_le_dist[of "f x" x] |
|
| 33175 | 7454 |
by (metis dist_eq_0_iff dist_nz e_def) |
7455 |
then obtain N where N:"\<forall>n\<ge>N. dist (z n) x < e / 2" |
|
|
44907
93943da0a010
remove redundant lemma Lim_sequentially in favor of lemma LIMSEQ_def
huffman
parents:
44905
diff
changeset
|
7456 |
using x[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto |
| 53282 | 7457 |
then have N':"dist (z N) x < e / 2" by auto |
7458 |
||
7459 |
have *: "c * dist (z N) x \<le> dist (z N) x" |
|
| 52625 | 7460 |
unfolding mult_le_cancel_right2 |
| 33175 | 7461 |
using zero_le_dist[of "z N" x] and c |
|
36778
739a9379e29b
avoid using real-specific versions of generic lemmas
huffman
parents:
36669
diff
changeset
|
7462 |
by (metis dist_eq_0_iff dist_nz order_less_asym less_le) |
| 52625 | 7463 |
have "dist (f (z N)) (f x) \<le> c * dist (z N) x" |
7464 |
using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]] |
|
7465 |
using z_in_s[of N] `x\<in>s` |
|
7466 |
using c |
|
7467 |
by auto |
|
7468 |
also have "\<dots> < e / 2" |
|
7469 |
using N' and c using * by auto |
|
7470 |
finally show False |
|
7471 |
unfolding fzn |
|
| 33175 | 7472 |
using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x] |
| 52625 | 7473 |
unfolding e_def |
7474 |
by auto |
|
| 33175 | 7475 |
qed |
| 53282 | 7476 |
then have "f x = x" unfolding e_def by auto |
| 33175 | 7477 |
moreover |
| 52625 | 7478 |
{
|
7479 |
fix y |
|
7480 |
assume "f y = y" "y\<in>s" |
|
| 53282 | 7481 |
then have "dist x y \<le> c * dist x y" |
| 52625 | 7482 |
using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]] |
7483 |
using `x\<in>s` and `f x = x` |
|
7484 |
by auto |
|
| 53282 | 7485 |
then have "dist x y = 0" |
| 52625 | 7486 |
unfolding mult_le_cancel_right1 |
7487 |
using c and zero_le_dist[of x y] |
|
7488 |
by auto |
|
| 53282 | 7489 |
then have "y = x" by auto |
| 33175 | 7490 |
} |
|
34999
5312d2ffee3b
Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
hoelzl
parents:
34964
diff
changeset
|
7491 |
ultimately show ?thesis using `x\<in>s` by blast+ |
| 33175 | 7492 |
qed |
7493 |
||
| 53282 | 7494 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
7495 |
subsection {* Edelstein fixed point theorem *}
|
| 33175 | 7496 |
|
7497 |
lemma edelstein_fix: |
|
|
50970
3e5b67f85bf9
generalized theorem edelstein_fix to class metric_space
huffman
parents:
50955
diff
changeset
|
7498 |
fixes s :: "'a::metric_space set" |
| 52625 | 7499 |
assumes s: "compact s" "s \<noteq> {}"
|
7500 |
and gs: "(g ` s) \<subseteq> s" |
|
7501 |
and dist: "\<forall>x\<in>s. \<forall>y\<in>s. x \<noteq> y \<longrightarrow> dist (g x) (g y) < dist x y" |
|
|
51347
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7502 |
shows "\<exists>!x\<in>s. g x = x" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7503 |
proof - |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7504 |
let ?D = "(\<lambda>x. (x, x)) ` s" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7505 |
have D: "compact ?D" "?D \<noteq> {}"
|
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7506 |
by (rule compact_continuous_image) |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7507 |
(auto intro!: s continuous_Pair continuous_within_id simp: continuous_on_eq_continuous_within) |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7508 |
|
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7509 |
have "\<And>x y e. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> 0 < e \<Longrightarrow> dist y x < e \<Longrightarrow> dist (g y) (g x) < e" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7510 |
using dist by fastforce |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7511 |
then have "continuous_on s g" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7512 |
unfolding continuous_on_iff by auto |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7513 |
then have cont: "continuous_on ?D (\<lambda>x. dist ((g \<circ> fst) x) (snd x))" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7514 |
unfolding continuous_on_eq_continuous_within |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7515 |
by (intro continuous_dist ballI continuous_within_compose) |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7516 |
(auto intro!: continuous_fst continuous_snd continuous_within_id simp: image_image) |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7517 |
|
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7518 |
obtain a where "a \<in> s" and le: "\<And>x. x \<in> s \<Longrightarrow> dist (g a) a \<le> dist (g x) x" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7519 |
using continuous_attains_inf[OF D cont] by auto |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7520 |
|
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7521 |
have "g a = a" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7522 |
proof (rule ccontr) |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7523 |
assume "g a \<noteq> a" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7524 |
with `a \<in> s` gs have "dist (g (g a)) (g a) < dist (g a) a" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7525 |
by (intro dist[rule_format]) auto |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7526 |
moreover have "dist (g a) a \<le> dist (g (g a)) (g a)" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7527 |
using `a \<in> s` gs by (intro le) auto |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7528 |
ultimately show False by auto |
| 33175 | 7529 |
qed |
|
51347
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7530 |
moreover have "\<And>x. x \<in> s \<Longrightarrow> g x = x \<Longrightarrow> x = a" |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7531 |
using dist[THEN bspec[where x=a]] `g a = a` and `a\<in>s` by auto |
|
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
hoelzl
parents:
51346
diff
changeset
|
7532 |
ultimately show "\<exists>!x\<in>s. g x = x" using `a \<in> s` by blast |
| 33175 | 7533 |
qed |
7534 |
||
|
44131
5fc334b94e00
declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents:
44129
diff
changeset
|
7535 |
declare tendsto_const [intro] (* FIXME: move *) |
|
5fc334b94e00
declare tendsto_const [intro] (accidentally removed in 230a8665c919)
huffman
parents:
44129
diff
changeset
|
7536 |
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
7537 |
no_notation |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
7538 |
eucl_less (infix "<e" 50) |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
7539 |
|
| 33175 | 7540 |
end |