| author | blanchet | 
| Fri, 10 Jan 2014 14:39:37 +0100 | |
| changeset 54971 | b4b828025880 | 
| parent 54863 | 82acc20ded73 | 
| child 55415 | 05f5fdb8d093 | 
| 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'"
 | 
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
142  | 
  shows "(\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>B' \<in> B. f B' \<in> X))"
 | 
| 
 
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> {}`]]
 | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
147  | 
guess B' . note B' = this  | 
| 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)  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
169  | 
moreover from topological_basisE[OF A a] guess A0 .  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
170  | 
moreover from topological_basisE[OF B b] guess B0 .  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
171  | 
    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
 | 
172  | 
by (intro UN_I[of "(A0, B0)"]) auto  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
173  | 
qed auto  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
174  | 
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
 | 
175  | 
|
| 53255 | 176  | 
|
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
177  | 
subsection {* Countable Basis *}
 | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
178  | 
|
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
179  | 
locale countable_basis =  | 
| 53640 | 180  | 
fixes B :: "'a::topological_space set set"  | 
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
181  | 
assumes is_basis: "topological_basis B"  | 
| 53282 | 182  | 
and countable_basis: "countable B"  | 
| 33175 | 183  | 
begin  | 
184  | 
||
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
185  | 
lemma open_countable_basis_ex:  | 
| 50087 | 186  | 
assumes "open X"  | 
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
187  | 
shows "\<exists>B' \<subseteq> B. X = Union B'"  | 
| 53255 | 188  | 
using assms countable_basis is_basis  | 
189  | 
unfolding topological_basis_def by blast  | 
|
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
190  | 
|
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
191  | 
lemma open_countable_basisE:  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
192  | 
assumes "open X"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
193  | 
obtains B' where "B' \<subseteq> B" "X = Union B'"  | 
| 53255 | 194  | 
using assms open_countable_basis_ex  | 
195  | 
by (atomize_elim) simp  | 
|
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
196  | 
|
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
197  | 
lemma countable_dense_exists:  | 
| 53291 | 198  | 
  "\<exists>D::'a set. countable D \<and> (\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>d \<in> D. d \<in> X))"
 | 
| 50087 | 199  | 
proof -  | 
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
200  | 
let ?f = "(\<lambda>B'. SOME x. x \<in> B')"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
201  | 
have "countable (?f ` B)" using countable_basis by simp  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
202  | 
with basis_dense[OF is_basis, of ?f] show ?thesis  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
203  | 
by (intro exI[where x="?f ` B"]) (metis (mono_tags) all_not_in_conv imageI someI)  | 
| 50087 | 204  | 
qed  | 
205  | 
||
206  | 
lemma countable_dense_setE:  | 
|
| 
50245
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
207  | 
obtains D :: "'a set"  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
208  | 
  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
 | 
209  | 
using countable_dense_exists by blast  | 
| 
 
dea9363887a6
based countable topological basis on Countable_Set
 
immler 
parents: 
50105 
diff
changeset
 | 
210  | 
|
| 50087 | 211  | 
end  | 
212  | 
||
| 50883 | 213  | 
lemma (in first_countable_topology) first_countable_basisE:  | 
214  | 
obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"  | 
|
215  | 
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)"  | 
|
216  | 
using first_countable_basis[of x]  | 
|
| 51473 | 217  | 
apply atomize_elim  | 
218  | 
apply (elim exE)  | 
|
219  | 
apply (rule_tac x="range A" in exI)  | 
|
220  | 
apply auto  | 
|
221  | 
done  | 
|
| 50883 | 222  | 
|
| 51105 | 223  | 
lemma (in first_countable_topology) first_countable_basis_Int_stableE:  | 
224  | 
obtains A where "countable A" "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"  | 
|
225  | 
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>a\<in>A. a \<subseteq> S)"  | 
|
226  | 
"\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A"  | 
|
227  | 
proof atomize_elim  | 
|
228  | 
from first_countable_basisE[of x] guess A' . note A' = this  | 
|
229  | 
def A \<equiv> "(\<lambda>N. \<Inter>((\<lambda>n. from_nat_into A' n) ` N)) ` (Collect finite::nat set set)"  | 
|
| 53255 | 230  | 
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 | 231  | 
(\<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)"  | 
232  | 
proof (safe intro!: exI[where x=A])  | 
|
| 53255 | 233  | 
show "countable A"  | 
234  | 
unfolding A_def by (intro countable_image countable_Collect_finite)  | 
|
235  | 
fix a  | 
|
236  | 
assume "a \<in> A"  | 
|
237  | 
then show "x \<in> a" "open a"  | 
|
238  | 
using A'(4)[OF open_UNIV] by (auto simp: A_def intro: A' from_nat_into)  | 
|
| 51105 | 239  | 
next  | 
| 
52141
 
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
 
haftmann 
parents: 
51773 
diff
changeset
 | 
240  | 
let ?int = "\<lambda>N. \<Inter>(from_nat_into A' ` N)"  | 
| 53255 | 241  | 
fix a b  | 
242  | 
assume "a \<in> A" "b \<in> A"  | 
|
243  | 
then obtain N M where "a = ?int N" "b = ?int M" "finite (N \<union> M)"  | 
|
244  | 
by (auto simp: A_def)  | 
|
245  | 
then show "a \<inter> b \<in> A"  | 
|
246  | 
by (auto simp: A_def intro!: image_eqI[where x="N \<union> M"])  | 
|
| 51105 | 247  | 
next  | 
| 53255 | 248  | 
fix S  | 
249  | 
assume "open S" "x \<in> S"  | 
|
250  | 
then obtain a where a: "a\<in>A'" "a \<subseteq> S" using A' by blast  | 
|
251  | 
then show "\<exists>a\<in>A. a \<subseteq> S" using a A'  | 
|
| 51105 | 252  | 
      by (intro bexI[where x=a]) (auto simp: A_def intro: image_eqI[where x="{to_nat_on A' a}"])
 | 
253  | 
qed  | 
|
254  | 
qed  | 
|
255  | 
||
| 51473 | 256  | 
lemma (in topological_space) first_countableI:  | 
| 53255 | 257  | 
assumes "countable A"  | 
258  | 
and 1: "\<And>a. a \<in> A \<Longrightarrow> x \<in> a" "\<And>a. a \<in> A \<Longrightarrow> open a"  | 
|
259  | 
and 2: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>a\<in>A. a \<subseteq> S"  | 
|
| 51473 | 260  | 
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))"  | 
261  | 
proof (safe intro!: exI[of _ "from_nat_into A"])  | 
|
| 53255 | 262  | 
fix i  | 
| 51473 | 263  | 
  have "A \<noteq> {}" using 2[of UNIV] by auto
 | 
| 53255 | 264  | 
show "x \<in> from_nat_into A i" "open (from_nat_into A i)"  | 
265  | 
    using range_from_nat_into_subset[OF `A \<noteq> {}`] 1 by auto
 | 
|
266  | 
next  | 
|
267  | 
fix S  | 
|
268  | 
assume "open S" "x\<in>S" from 2[OF this]  | 
|
269  | 
show "\<exists>i. from_nat_into A i \<subseteq> S"  | 
|
270  | 
using subset_range_from_nat_into[OF `countable A`] by auto  | 
|
| 51473 | 271  | 
qed  | 
| 51350 | 272  | 
|
| 50883 | 273  | 
instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology  | 
274  | 
proof  | 
|
275  | 
fix x :: "'a \<times> 'b"  | 
|
276  | 
from first_countable_basisE[of "fst x"] guess A :: "'a set set" . note A = this  | 
|
277  | 
from first_countable_basisE[of "snd x"] guess B :: "'b set set" . note B = this  | 
|
| 53282 | 278  | 
  show "\<exists>A::nat \<Rightarrow> ('a \<times> 'b) set.
 | 
279  | 
(\<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 | 280  | 
proof (rule first_countableI[of "(\<lambda>(a, b). a \<times> b) ` (A \<times> B)"], safe)  | 
| 53255 | 281  | 
fix a b  | 
282  | 
assume x: "a \<in> A" "b \<in> B"  | 
|
| 53640 | 283  | 
with A(2, 3)[of a] B(2, 3)[of b] show "x \<in> a \<times> b" and "open (a \<times> b)"  | 
284  | 
unfolding mem_Times_iff  | 
|
285  | 
by (auto intro: open_Times)  | 
|
| 50883 | 286  | 
next  | 
| 53255 | 287  | 
fix S  | 
288  | 
assume "open S" "x \<in> S"  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
289  | 
from open_prod_elim[OF this] guess a' b' . note a'b' = this  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
290  | 
moreover from a'b' A(4)[of a'] B(4)[of b']  | 
| 50883 | 291  | 
obtain a b where "a \<in> A" "a \<subseteq> a'" "b \<in> B" "b \<subseteq> b'" by auto  | 
292  | 
ultimately show "\<exists>a\<in>(\<lambda>(a, b). a \<times> b) ` (A \<times> B). a \<subseteq> S"  | 
|
293  | 
by (auto intro!: bexI[of _ "a \<times> b"] bexI[of _ a] bexI[of _ b])  | 
|
294  | 
qed (simp add: A B)  | 
|
295  | 
qed  | 
|
296  | 
||
| 
50881
 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
297  | 
class second_countable_topology = topological_space +  | 
| 53282 | 298  | 
assumes ex_countable_subbasis:  | 
299  | 
"\<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
 | 
300  | 
begin  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
301  | 
|
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
302  | 
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
 | 
303  | 
proof -  | 
| 53255 | 304  | 
from ex_countable_subbasis obtain B where B: "countable B" "open = generate_topology B"  | 
305  | 
by blast  | 
|
| 
51343
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
306  | 
  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
 | 
307  | 
|
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
308  | 
show ?thesis  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
309  | 
proof (intro exI conjI)  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
310  | 
show "countable ?B"  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
311  | 
by (intro countable_image countable_Collect_finite_subset B)  | 
| 53255 | 312  | 
    {
 | 
313  | 
fix S  | 
|
314  | 
assume "open S"  | 
|
| 
51343
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
315  | 
      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
 | 
316  | 
unfolding B  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
317  | 
proof induct  | 
| 53255 | 318  | 
case UNIV  | 
319  | 
        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
 | 
320  | 
next  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
321  | 
case (Int a b)  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
322  | 
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
 | 
323  | 
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
 | 
324  | 
by blast  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
325  | 
show ?case  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
326  | 
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
 | 
327  | 
          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
 | 
328  | 
next  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
329  | 
case (UN K)  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
330  | 
        then have "\<forall>k\<in>K. \<exists>B'\<subseteq>{b. finite b \<and> b \<subseteq> B}. UNION B' Inter = k" by auto
 | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
331  | 
then guess k unfolding bchoice_iff ..  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
332  | 
        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
 | 
333  | 
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
 | 
334  | 
next  | 
| 53255 | 335  | 
case (Basis S)  | 
336  | 
then show ?case  | 
|
| 
51343
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
337  | 
          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
 | 
338  | 
qed  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
339  | 
      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
 | 
340  | 
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
 | 
341  | 
then show "topological_basis ?B"  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
342  | 
unfolding topological_space_class.topological_basis_def  | 
| 53282 | 343  | 
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
 | 
344  | 
(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
 | 
345  | 
qed  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
346  | 
qed  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
347  | 
|
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
348  | 
end  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
349  | 
|
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
350  | 
sublocale second_countable_topology <  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
351  | 
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
 | 
352  | 
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
 | 
353  | 
by unfold_locales safe  | 
| 
50094
 
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
 
immler 
parents: 
50087 
diff
changeset
 | 
354  | 
|
| 
50882
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
355  | 
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
 | 
356  | 
proof  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
357  | 
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
 | 
358  | 
using ex_countable_basis by auto  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
359  | 
moreover  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
360  | 
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
 | 
361  | 
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
 | 
362  | 
  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
 | 
363  | 
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
 | 
364  | 
topological_basis_imp_subbasis)  | 
| 
50882
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
365  | 
qed  | 
| 
 
a382bf90867e
move prod instantiation of second_countable_topology to its definition
 
hoelzl 
parents: 
50881 
diff
changeset
 | 
366  | 
|
| 50883 | 367  | 
instance second_countable_topology \<subseteq> first_countable_topology  | 
368  | 
proof  | 
|
369  | 
fix x :: 'a  | 
|
370  | 
def B \<equiv> "SOME B::'a set set. countable B \<and> topological_basis B"  | 
|
371  | 
then have B: "countable B" "topological_basis B"  | 
|
372  | 
using countable_basis is_basis  | 
|
373  | 
by (auto simp: countable_basis is_basis)  | 
|
| 53282 | 374  | 
then show "\<exists>A::nat \<Rightarrow> 'a set.  | 
375  | 
(\<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 | 376  | 
    by (intro first_countableI[of "{b\<in>B. x \<in> b}"])
 | 
377  | 
(fastforce simp: topological_space_class.topological_basis_def)+  | 
|
| 50883 | 378  | 
qed  | 
379  | 
||
| 53255 | 380  | 
|
| 50087 | 381  | 
subsection {* Polish spaces *}
 | 
382  | 
||
383  | 
text {* Textbooks define Polish spaces as completely metrizable.
 | 
|
384  | 
We assume the topology to be complete for a given metric. *}  | 
|
385  | 
||
| 
50881
 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
386  | 
class polish_space = complete_space + second_countable_topology  | 
| 50087 | 387  | 
|
| 44517 | 388  | 
subsection {* General notion of a topology as a value *}
 | 
| 33175 | 389  | 
|
| 53255 | 390  | 
definition "istopology L \<longleftrightarrow>  | 
391  | 
  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))"
 | 
|
392  | 
||
| 49834 | 393  | 
typedef 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
 | 
| 33175 | 394  | 
morphisms "openin" "topology"  | 
395  | 
unfolding istopology_def by blast  | 
|
396  | 
||
397  | 
lemma istopology_open_in[intro]: "istopology(openin U)"  | 
|
398  | 
using openin[of U] by blast  | 
|
399  | 
||
400  | 
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
 | 
401  | 
using topology_inverse[unfolded mem_Collect_eq] .  | 
| 33175 | 402  | 
|
403  | 
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U"  | 
|
404  | 
using topology_inverse[of U] istopology_open_in[of "topology U"] by auto  | 
|
405  | 
||
406  | 
lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)"  | 
|
| 53255 | 407  | 
proof  | 
408  | 
assume "T1 = T2"  | 
|
409  | 
then show "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp  | 
|
410  | 
next  | 
|
411  | 
assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S"  | 
|
412  | 
then have "openin T1 = openin T2" by (simp add: fun_eq_iff)  | 
|
413  | 
then have "topology (openin T1) = topology (openin T2)" by simp  | 
|
414  | 
then show "T1 = T2" unfolding openin_inverse .  | 
|
| 33175 | 415  | 
qed  | 
416  | 
||
417  | 
text{* Infer the "universe" from union of all sets in the topology. *}
 | 
|
418  | 
||
| 53640 | 419  | 
definition "topspace T = \<Union>{S. openin T S}"
 | 
| 33175 | 420  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
421  | 
subsubsection {* Main properties of open sets *}
 | 
| 33175 | 422  | 
|
423  | 
lemma openin_clauses:  | 
|
424  | 
fixes U :: "'a topology"  | 
|
| 53282 | 425  | 
shows  | 
426  | 
    "openin U {}"
 | 
|
427  | 
"\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)"  | 
|
428  | 
"\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)"  | 
|
429  | 
using openin[of U] unfolding istopology_def mem_Collect_eq by fast+  | 
|
| 33175 | 430  | 
|
431  | 
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U"  | 
|
432  | 
unfolding topspace_def by blast  | 
|
| 53255 | 433  | 
|
434  | 
lemma openin_empty[simp]: "openin U {}"
 | 
|
435  | 
by (simp add: openin_clauses)  | 
|
| 33175 | 436  | 
|
437  | 
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
 | 
438  | 
using openin_clauses by simp  | 
| 
 
06475a1547cb
fix lots of looping simp calls and other warnings
 
huffman 
parents: 
36360 
diff
changeset
 | 
439  | 
|
| 
 
06475a1547cb
fix lots of looping simp calls and other warnings
 
huffman 
parents: 
36360 
diff
changeset
 | 
440  | 
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
 | 
441  | 
using openin_clauses by simp  | 
| 33175 | 442  | 
|
443  | 
lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)"  | 
|
444  | 
  using openin_Union[of "{S,T}" U] by auto
 | 
|
445  | 
||
| 53255 | 446  | 
lemma openin_topspace[intro, simp]: "openin U (topspace U)"  | 
447  | 
by (simp add: openin_Union topspace_def)  | 
|
| 33175 | 448  | 
|
| 49711 | 449  | 
lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)"  | 
450  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 36584 | 451  | 
proof  | 
| 49711 | 452  | 
assume ?lhs  | 
453  | 
then show ?rhs by auto  | 
|
| 36584 | 454  | 
next  | 
455  | 
assume H: ?rhs  | 
|
456  | 
  let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
 | 
|
457  | 
have "openin U ?t" by (simp add: openin_Union)  | 
|
458  | 
also have "?t = S" using H by auto  | 
|
459  | 
finally show "openin U S" .  | 
|
| 33175 | 460  | 
qed  | 
461  | 
||
| 49711 | 462  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
463  | 
subsubsection {* Closed sets *}
 | 
| 33175 | 464  | 
|
465  | 
definition "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)"  | 
|
466  | 
||
| 53255 | 467  | 
lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U"  | 
468  | 
by (metis closedin_def)  | 
|
469  | 
||
470  | 
lemma closedin_empty[simp]: "closedin U {}"
 | 
|
471  | 
by (simp add: closedin_def)  | 
|
472  | 
||
473  | 
lemma closedin_topspace[intro, simp]: "closedin U (topspace U)"  | 
|
474  | 
by (simp add: closedin_def)  | 
|
475  | 
||
| 33175 | 476  | 
lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)"  | 
477  | 
by (auto simp add: Diff_Un closedin_def)  | 
|
478  | 
||
| 53255 | 479  | 
lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union> {A - s|s. s\<in>S}"
 | 
480  | 
by auto  | 
|
481  | 
||
482  | 
lemma closedin_Inter[intro]:  | 
|
483  | 
  assumes Ke: "K \<noteq> {}"
 | 
|
484  | 
and Kc: "\<forall>S \<in>K. closedin U S"  | 
|
485  | 
shows "closedin U (\<Inter> K)"  | 
|
486  | 
using Ke Kc unfolding closedin_def Diff_Inter by auto  | 
|
| 33175 | 487  | 
|
488  | 
lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)"  | 
|
489  | 
  using closedin_Inter[of "{S,T}" U] by auto
 | 
|
490  | 
||
| 53255 | 491  | 
lemma Diff_Diff_Int: "A - (A - B) = A \<inter> B"  | 
492  | 
by blast  | 
|
493  | 
||
| 33175 | 494  | 
lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)"  | 
495  | 
apply (auto simp add: closedin_def Diff_Diff_Int inf_absorb2)  | 
|
496  | 
apply (metis openin_subset subset_eq)  | 
|
497  | 
done  | 
|
498  | 
||
| 53255 | 499  | 
lemma openin_closedin: "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))"  | 
| 33175 | 500  | 
by (simp add: openin_closedin_eq)  | 
501  | 
||
| 53255 | 502  | 
lemma openin_diff[intro]:  | 
503  | 
assumes oS: "openin U S"  | 
|
504  | 
and cT: "closedin U T"  | 
|
505  | 
shows "openin U (S - T)"  | 
|
506  | 
proof -  | 
|
| 33175 | 507  | 
have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S] oS cT  | 
508  | 
by (auto simp add: topspace_def openin_subset)  | 
|
| 53282 | 509  | 
then show ?thesis using oS cT  | 
510  | 
by (auto simp add: closedin_def)  | 
|
| 33175 | 511  | 
qed  | 
512  | 
||
| 53255 | 513  | 
lemma closedin_diff[intro]:  | 
514  | 
assumes oS: "closedin U S"  | 
|
515  | 
and cT: "openin U T"  | 
|
516  | 
shows "closedin U (S - T)"  | 
|
517  | 
proof -  | 
|
518  | 
have "S - T = S \<inter> (topspace U - T)"  | 
|
| 53282 | 519  | 
using closedin_subset[of U S] oS cT by (auto simp add: topspace_def)  | 
| 53255 | 520  | 
then show ?thesis  | 
521  | 
using oS cT by (auto simp add: openin_closedin_eq)  | 
|
522  | 
qed  | 
|
523  | 
||
| 33175 | 524  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
525  | 
subsubsection {* Subspace topology *}
 | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
526  | 
|
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
527  | 
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
 | 
528  | 
|
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
529  | 
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
 | 
530  | 
(is "istopology ?L")  | 
| 53255 | 531  | 
proof -  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
532  | 
  have "?L {}" by blast
 | 
| 53255 | 533  | 
  {
 | 
534  | 
fix A B  | 
|
535  | 
assume A: "?L A" and B: "?L B"  | 
|
536  | 
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"  | 
|
537  | 
by blast  | 
|
538  | 
have "A \<inter> B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)"  | 
|
539  | 
using Sa Sb by blast+  | 
|
540  | 
then have "?L (A \<inter> B)" by blast  | 
|
541  | 
}  | 
|
| 33175 | 542  | 
moreover  | 
| 53255 | 543  | 
  {
 | 
| 53282 | 544  | 
fix K  | 
545  | 
assume K: "K \<subseteq> Collect ?L"  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
546  | 
have th0: "Collect ?L = (\<lambda>S. S \<inter> V) ` Collect (openin U)"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
547  | 
apply (rule set_eqI)  | 
| 33175 | 548  | 
apply (simp add: Ball_def image_iff)  | 
| 53255 | 549  | 
apply metis  | 
550  | 
done  | 
|
| 33175 | 551  | 
from K[unfolded th0 subset_image_iff]  | 
| 53255 | 552  | 
obtain Sk where Sk: "Sk \<subseteq> Collect (openin U)" "K = (\<lambda>S. S \<inter> V) ` Sk"  | 
553  | 
by blast  | 
|
554  | 
have "\<Union>K = (\<Union>Sk) \<inter> V"  | 
|
555  | 
using Sk by auto  | 
|
556  | 
moreover have "openin U (\<Union> Sk)"  | 
|
557  | 
using Sk by (auto simp add: subset_eq)  | 
|
558  | 
ultimately have "?L (\<Union>K)" by blast  | 
|
559  | 
}  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
560  | 
ultimately show ?thesis  | 
| 
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
561  | 
unfolding subset_eq mem_Collect_eq istopology_def by blast  | 
| 33175 | 562  | 
qed  | 
563  | 
||
| 53255 | 564  | 
lemma openin_subtopology: "openin (subtopology U V) S \<longleftrightarrow> (\<exists>T. openin U T \<and> S = T \<inter> V)"  | 
| 33175 | 565  | 
unfolding subtopology_def topology_inverse'[OF istopology_subtopology]  | 
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
566  | 
by auto  | 
| 33175 | 567  | 
|
| 53255 | 568  | 
lemma topspace_subtopology: "topspace (subtopology U V) = topspace U \<inter> V"  | 
| 33175 | 569  | 
by (auto simp add: topspace_def openin_subtopology)  | 
570  | 
||
| 53255 | 571  | 
lemma closedin_subtopology: "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)"  | 
| 33175 | 572  | 
unfolding closedin_def topspace_subtopology  | 
573  | 
apply (simp add: openin_subtopology)  | 
|
574  | 
apply (rule iffI)  | 
|
575  | 
apply clarify  | 
|
576  | 
apply (rule_tac x="topspace U - T" in exI)  | 
|
| 53255 | 577  | 
apply auto  | 
578  | 
done  | 
|
| 33175 | 579  | 
|
580  | 
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U"  | 
|
581  | 
unfolding openin_subtopology  | 
|
582  | 
apply (rule iffI, clarify)  | 
|
| 53255 | 583  | 
apply (frule openin_subset[of U])  | 
584  | 
apply blast  | 
|
| 33175 | 585  | 
apply (rule exI[where x="topspace U"])  | 
| 49711 | 586  | 
apply auto  | 
587  | 
done  | 
|
588  | 
||
589  | 
lemma subtopology_superset:  | 
|
590  | 
assumes UV: "topspace U \<subseteq> V"  | 
|
| 33175 | 591  | 
shows "subtopology U V = U"  | 
| 53255 | 592  | 
proof -  | 
593  | 
  {
 | 
|
594  | 
fix S  | 
|
595  | 
    {
 | 
|
596  | 
fix T  | 
|
597  | 
assume T: "openin U T" "S = T \<inter> V"  | 
|
598  | 
from T openin_subset[OF T(1)] UV have eq: "S = T"  | 
|
599  | 
by blast  | 
|
600  | 
have "openin U S"  | 
|
601  | 
unfolding eq using T by blast  | 
|
602  | 
}  | 
|
| 33175 | 603  | 
moreover  | 
| 53255 | 604  | 
    {
 | 
605  | 
assume S: "openin U S"  | 
|
606  | 
then have "\<exists>T. openin U T \<and> S = T \<inter> V"  | 
|
607  | 
using openin_subset[OF S] UV by auto  | 
|
608  | 
}  | 
|
609  | 
ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S"  | 
|
610  | 
by blast  | 
|
611  | 
}  | 
|
612  | 
then show ?thesis  | 
|
613  | 
unfolding topology_eq openin_subtopology by blast  | 
|
| 33175 | 614  | 
qed  | 
615  | 
||
616  | 
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U"  | 
|
617  | 
by (simp add: subtopology_superset)  | 
|
618  | 
||
619  | 
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U"  | 
|
620  | 
by (simp add: subtopology_superset)  | 
|
621  | 
||
| 53255 | 622  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
623  | 
subsubsection {* The standard Euclidean topology *}
 | 
| 33175 | 624  | 
|
| 53255 | 625  | 
definition euclidean :: "'a::topological_space topology"  | 
626  | 
where "euclidean = topology open"  | 
|
| 33175 | 627  | 
|
628  | 
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S"  | 
|
629  | 
unfolding euclidean_def  | 
|
630  | 
apply (rule cong[where x=S and y=S])  | 
|
631  | 
apply (rule topology_inverse[symmetric])  | 
|
632  | 
apply (auto simp add: istopology_def)  | 
|
| 
44170
 
510ac30f44c0
make Multivariate_Analysis work with separate set type
 
huffman 
parents: 
44167 
diff
changeset
 | 
633  | 
done  | 
| 33175 | 634  | 
|
635  | 
lemma topspace_euclidean: "topspace euclidean = UNIV"  | 
|
636  | 
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
 | 
637  | 
apply (rule set_eqI)  | 
| 53255 | 638  | 
apply (auto simp add: open_openin[symmetric])  | 
639  | 
done  | 
|
| 33175 | 640  | 
|
641  | 
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S"  | 
|
642  | 
by (simp add: topspace_euclidean topspace_subtopology)  | 
|
643  | 
||
644  | 
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S"  | 
|
645  | 
by (simp add: closed_def closedin_def topspace_euclidean open_openin Compl_eq_Diff_UNIV)  | 
|
646  | 
||
647  | 
lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)"  | 
|
648  | 
by (simp add: open_openin openin_subopen[symmetric])  | 
|
649  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
650  | 
text {* Basic "localization" results are handy for connectedness. *}
 | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
651  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
652  | 
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
 | 
653  | 
by (auto simp add: openin_subtopology open_openin[symmetric])  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
654  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
655  | 
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
 | 
656  | 
by (auto simp add: openin_open)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
657  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
658  | 
lemma open_openin_trans[trans]:  | 
| 53255 | 659  | 
"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
 | 
660  | 
by (metis Int_absorb1 openin_open_Int)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
661  | 
|
| 53255 | 662  | 
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
 | 
663  | 
by (auto simp add: openin_open)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
664  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
665  | 
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
 | 
666  | 
by (simp add: closedin_subtopology closed_closedin Int_ac)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
667  | 
|
| 53291 | 668  | 
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
 | 
669  | 
by (metis closedin_closed)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
670  | 
|
| 53282 | 671  | 
lemma closed_closedin_trans:  | 
672  | 
"closed S \<Longrightarrow> closed T \<Longrightarrow> T \<subseteq> S \<Longrightarrow> closedin (subtopology euclidean S) T"  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
673  | 
apply (subgoal_tac "S \<inter> T = T" )  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
674  | 
apply auto  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
675  | 
apply (frule closedin_closed_Int[of T S])  | 
| 52624 | 676  | 
apply simp  | 
677  | 
done  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
678  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
679  | 
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
 | 
680  | 
by (auto simp add: closedin_closed)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
681  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
682  | 
lemma openin_euclidean_subtopology_iff:  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
683  | 
fixes S U :: "'a::metric_space set"  | 
| 53255 | 684  | 
shows "openin (subtopology euclidean U) S \<longleftrightarrow>  | 
685  | 
S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)"  | 
|
686  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
687  | 
proof  | 
| 53255 | 688  | 
assume ?lhs  | 
| 53282 | 689  | 
then show ?rhs  | 
690  | 
unfolding openin_open open_dist by blast  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
691  | 
next  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
692  | 
  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
 | 
693  | 
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
 | 
694  | 
unfolding T_def  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
695  | 
apply clarsimp  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
696  | 
apply (rule_tac x="d - dist x a" in exI)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
697  | 
apply (clarsimp simp add: less_diff_eq)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
698  | 
apply (erule rev_bexI)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
699  | 
apply (rule_tac x=d in exI, clarify)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
700  | 
apply (erule le_less_trans [OF dist_triangle])  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
701  | 
done  | 
| 53282 | 702  | 
assume ?rhs then have 2: "S = U \<inter> T"  | 
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
703  | 
unfolding T_def  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
704  | 
apply auto  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
705  | 
apply (drule (1) bspec, erule rev_bexI)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
706  | 
apply auto  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
707  | 
done  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
708  | 
from 1 2 show ?lhs  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
709  | 
unfolding openin_open open_dist by fast  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
710  | 
qed  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
711  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
712  | 
text {* These "transitivity" results are handy too *}
 | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
713  | 
|
| 53255 | 714  | 
lemma openin_trans[trans]:  | 
715  | 
"openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T \<Longrightarrow>  | 
|
716  | 
openin (subtopology euclidean U) S"  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
717  | 
unfolding open_openin openin_open by blast  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
718  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
719  | 
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
 | 
720  | 
by (auto simp add: openin_open intro: openin_trans)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
721  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
722  | 
lemma closedin_trans[trans]:  | 
| 53255 | 723  | 
"closedin (subtopology euclidean T) S \<Longrightarrow> closedin (subtopology euclidean U) T \<Longrightarrow>  | 
724  | 
closedin (subtopology euclidean U) S"  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
725  | 
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
 | 
726  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
727  | 
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
 | 
728  | 
by (auto simp add: closedin_closed intro: closedin_trans)  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
729  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
730  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
731  | 
subsection {* Open and closed balls *}
 | 
| 33175 | 732  | 
|
| 53255 | 733  | 
definition ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"  | 
734  | 
  where "ball x e = {y. dist x y < e}"
 | 
|
735  | 
||
736  | 
definition cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set"  | 
|
737  | 
  where "cball x e = {y. dist x y \<le> e}"
 | 
|
| 33175 | 738  | 
|
| 
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
 | 
739  | 
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
 | 
740  | 
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
 | 
741  | 
|
| 
 
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  | 
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
 | 
743  | 
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
 | 
744  | 
|
| 
 
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  | 
lemma mem_ball_0:  | 
| 33175 | 746  | 
fixes x :: "'a::real_normed_vector"  | 
747  | 
shows "x \<in> ball 0 e \<longleftrightarrow> norm x < e"  | 
|
748  | 
by (simp add: dist_norm)  | 
|
749  | 
||
| 
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
 | 
750  | 
lemma mem_cball_0:  | 
| 33175 | 751  | 
fixes x :: "'a::real_normed_vector"  | 
752  | 
shows "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e"  | 
|
753  | 
by (simp add: dist_norm)  | 
|
754  | 
||
| 
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
 | 
755  | 
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
 | 
756  | 
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
 | 
757  | 
|
| 
 
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  | 
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
 | 
759  | 
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
 | 
760  | 
|
| 53255 | 761  | 
lemma ball_subset_cball[simp,intro]: "ball x e \<subseteq> cball x e"  | 
762  | 
by (simp add: subset_eq)  | 
|
763  | 
||
| 53282 | 764  | 
lemma subset_ball[intro]: "d \<le> e \<Longrightarrow> ball x d \<subseteq> ball x e"  | 
| 53255 | 765  | 
by (simp add: subset_eq)  | 
766  | 
||
| 53282 | 767  | 
lemma subset_cball[intro]: "d \<le> e \<Longrightarrow> cball x d \<subseteq> cball x e"  | 
| 53255 | 768  | 
by (simp add: subset_eq)  | 
769  | 
||
| 33175 | 770  | 
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
 | 
771  | 
by (simp add: set_eq_iff) arith  | 
| 33175 | 772  | 
|
773  | 
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
 | 
774  | 
by (simp add: set_eq_iff)  | 
| 33175 | 775  | 
|
| 53255 | 776  | 
lemma diff_less_iff:  | 
777  | 
"(a::real) - b > 0 \<longleftrightarrow> a > b"  | 
|
| 33175 | 778  | 
"(a::real) - b < 0 \<longleftrightarrow> a < b"  | 
| 53255 | 779  | 
"a - b < c \<longleftrightarrow> a < c + b" "a - b > c \<longleftrightarrow> a > c + b"  | 
780  | 
by arith+  | 
|
781  | 
||
782  | 
lemma diff_le_iff:  | 
|
783  | 
"(a::real) - b \<ge> 0 \<longleftrightarrow> a \<ge> b"  | 
|
784  | 
"(a::real) - b \<le> 0 \<longleftrightarrow> a \<le> b"  | 
|
785  | 
"a - b \<le> c \<longleftrightarrow> a \<le> c + b"  | 
|
786  | 
"a - b \<ge> c \<longleftrightarrow> a \<ge> c + b"  | 
|
787  | 
by arith+  | 
|
| 33175 | 788  | 
|
| 54070 | 789  | 
lemma open_vimage: (* TODO: move to Topological_Spaces.thy *)  | 
790  | 
assumes "open s" and "continuous_on UNIV f"  | 
|
791  | 
shows "open (vimage f s)"  | 
|
792  | 
using assms unfolding continuous_on_open_vimage [OF open_UNIV]  | 
|
793  | 
by simp  | 
|
794  | 
||
795  | 
lemma open_ball [intro, simp]: "open (ball x e)"  | 
|
796  | 
proof -  | 
|
797  | 
  have "open (dist x -` {..<e})"
 | 
|
798  | 
by (intro open_vimage open_lessThan continuous_on_intros)  | 
|
799  | 
  also have "dist x -` {..<e} = ball x e"
 | 
|
800  | 
by auto  | 
|
801  | 
finally show ?thesis .  | 
|
802  | 
qed  | 
|
| 33175 | 803  | 
|
804  | 
lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)"  | 
|
805  | 
unfolding open_dist subset_eq mem_ball Ball_def dist_commute ..  | 
|
806  | 
||
| 
33714
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33324 
diff
changeset
 | 
807  | 
lemma openE[elim?]:  | 
| 53282 | 808  | 
assumes "open S" "x\<in>S"  | 
| 
33714
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33324 
diff
changeset
 | 
809  | 
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
 | 
810  | 
using assms unfolding open_contains_ball by auto  | 
| 
 
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
 
hoelzl 
parents: 
33324 
diff
changeset
 | 
811  | 
|
| 33175 | 812  | 
lemma open_contains_ball_eq: "open S \<Longrightarrow> \<forall>x. x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"  | 
813  | 
by (metis open_contains_ball subset_eq centre_in_ball)  | 
|
814  | 
||
815  | 
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
 | 
816  | 
unfolding mem_ball set_eq_iff  | 
| 33175 | 817  | 
apply (simp add: not_less)  | 
| 52624 | 818  | 
apply (metis zero_le_dist order_trans dist_self)  | 
819  | 
done  | 
|
| 33175 | 820  | 
|
| 53291 | 821  | 
lemma ball_empty[intro]: "e \<le> 0 \<Longrightarrow> ball x e = {}" by simp
 | 
| 33175 | 822  | 
|
| 
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
 | 
823  | 
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
 | 
824  | 
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
 | 
825  | 
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
 | 
826  | 
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
 | 
827  | 
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
 | 
828  | 
|
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
829  | 
definition (in euclidean_space) eucl_less (infix "<e" 50)  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
830  | 
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
 | 
831  | 
|
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
832  | 
definition box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
 | 
| 
 
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"  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
836  | 
by (auto simp: box_eucl_less eucl_less_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
 | 
837  | 
|
| 50087 | 838  | 
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
 | 
839  | 
fixes x :: "'a\<Colon>euclidean_space"  | 
| 53291 | 840  | 
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
 | 
841  | 
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 | 842  | 
proof -  | 
843  | 
  def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
 | 
|
| 53291 | 844  | 
then have e: "e' > 0"  | 
| 53255 | 845  | 
using assms by (auto intro!: divide_pos_pos 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
 | 
846  | 
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 | 847  | 
proof  | 
| 53255 | 848  | 
fix i  | 
849  | 
from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e  | 
|
850  | 
show "?th i" by auto  | 
|
| 50087 | 851  | 
qed  | 
852  | 
from choice[OF this] guess a .. note a = this  | 
|
| 
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  | 
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 | 854  | 
proof  | 
| 53255 | 855  | 
fix i  | 
856  | 
from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e  | 
|
857  | 
show "?th i" by auto  | 
|
| 50087 | 858  | 
qed  | 
859  | 
from choice[OF this] guess b .. note b = this  | 
|
| 
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
 | 
860  | 
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
 | 
861  | 
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
 | 
862  | 
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe)  | 
| 53255 | 863  | 
fix y :: 'a  | 
864  | 
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
 | 
865  | 
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)"  | 
| 50087 | 866  | 
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
 | 
867  | 
    also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
 | 
| 50087 | 868  | 
proof (rule real_sqrt_less_mono, rule setsum_strict_mono)  | 
| 53255 | 869  | 
fix i :: "'a"  | 
870  | 
assume i: "i \<in> Basis"  | 
|
871  | 
have "a i < y\<bullet>i \<and> y\<bullet>i < b i"  | 
|
872  | 
using * i by (auto simp: box_def)  | 
|
873  | 
moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'"  | 
|
874  | 
using a by auto  | 
|
875  | 
moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'"  | 
|
876  | 
using b by auto  | 
|
877  | 
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'"  | 
|
878  | 
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
 | 
879  | 
      then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
 | 
| 50087 | 880  | 
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
 | 
881  | 
      then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
 | 
| 50087 | 882  | 
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
 | 
883  | 
      then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
 | 
| 50087 | 884  | 
by (simp add: power_divide)  | 
885  | 
qed auto  | 
|
| 53255 | 886  | 
also have "\<dots> = e"  | 
887  | 
using `0 < e` by (simp add: real_eq_of_nat)  | 
|
888  | 
finally show "y \<in> ball x e"  | 
|
889  | 
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
 | 
890  | 
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
 | 
891  | 
qed  | 
| 51103 | 892  | 
|
| 
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  | 
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
 | 
894  | 
fixes M :: "'a\<Colon>euclidean_space set"  | 
| 53282 | 895  | 
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
 | 
896  | 
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
 | 
897  | 
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
 | 
898  | 
  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
 | 
899  | 
shows "M = (\<Union>f\<in>I. box (a' f) (b' f))"  | 
| 52624 | 900  | 
proof -  | 
901  | 
  {
 | 
|
902  | 
fix x assume "x \<in> M"  | 
|
903  | 
obtain e where e: "e > 0" "ball x e \<subseteq> M"  | 
|
904  | 
using openE[OF `open M` `x \<in> M`] by auto  | 
|
| 53282 | 905  | 
moreover obtain a b where ab:  | 
906  | 
"x \<in> box a b"  | 
|
907  | 
"\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>"  | 
|
908  | 
"\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>"  | 
|
909  | 
"box a b \<subseteq> ball x e"  | 
|
| 52624 | 910  | 
using rational_boxes[OF e(1)] by metis  | 
911  | 
ultimately have "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))"  | 
|
912  | 
by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"])  | 
|
913  | 
(auto simp: euclidean_representation I_def a'_def b'_def)  | 
|
914  | 
}  | 
|
915  | 
then show ?thesis by (auto simp: I_def)  | 
|
916  | 
qed  | 
|
917  | 
||
| 33175 | 918  | 
|
919  | 
subsection{* Connectedness *}
 | 
|
920  | 
||
921  | 
lemma connected_local:  | 
|
| 53255 | 922  | 
"connected S \<longleftrightarrow>  | 
923  | 
\<not> (\<exists>e1 e2.  | 
|
924  | 
openin (subtopology euclidean S) e1 \<and>  | 
|
925  | 
openin (subtopology euclidean S) e2 \<and>  | 
|
926  | 
S \<subseteq> e1 \<union> e2 \<and>  | 
|
927  | 
      e1 \<inter> e2 = {} \<and>
 | 
|
928  | 
      e1 \<noteq> {} \<and>
 | 
|
929  | 
      e2 \<noteq> {})"
 | 
|
| 53282 | 930  | 
unfolding connected_def openin_open  | 
931  | 
apply safe  | 
|
932  | 
apply blast+  | 
|
933  | 
done  | 
|
| 33175 | 934  | 
|
| 34105 | 935  | 
lemma exists_diff:  | 
936  | 
fixes P :: "'a set \<Rightarrow> bool"  | 
|
937  | 
shows "(\<exists>S. P(- S)) \<longleftrightarrow> (\<exists>S. P S)" (is "?lhs \<longleftrightarrow> ?rhs")  | 
|
| 53255 | 938  | 
proof -  | 
939  | 
  {
 | 
|
940  | 
assume "?lhs"  | 
|
941  | 
then have ?rhs by blast  | 
|
942  | 
}  | 
|
| 33175 | 943  | 
moreover  | 
| 53255 | 944  | 
  {
 | 
945  | 
fix S  | 
|
946  | 
assume H: "P S"  | 
|
| 34105 | 947  | 
have "S = - (- S)" by auto  | 
| 53255 | 948  | 
with H have "P (- (- S))" by metis  | 
949  | 
}  | 
|
| 33175 | 950  | 
ultimately show ?thesis by metis  | 
951  | 
qed  | 
|
952  | 
||
953  | 
lemma connected_clopen: "connected S \<longleftrightarrow>  | 
|
| 53255 | 954  | 
(\<forall>T. openin (subtopology euclidean S) T \<and>  | 
955  | 
     closedin (subtopology euclidean S) T \<longrightarrow> T = {} \<or> T = S)" (is "?lhs \<longleftrightarrow> ?rhs")
 | 
|
956  | 
proof -  | 
|
957  | 
have "\<not> connected S \<longleftrightarrow>  | 
|
958  | 
    (\<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 | 959  | 
unfolding connected_def openin_open closedin_closed  | 
| 52624 | 960  | 
apply (subst exists_diff)  | 
961  | 
apply blast  | 
|
962  | 
done  | 
|
| 53282 | 963  | 
then have th0: "connected S \<longleftrightarrow>  | 
| 53255 | 964  | 
    \<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 | 965  | 
(is " _ \<longleftrightarrow> \<not> (\<exists>e2 e1. ?P e2 e1)")  | 
966  | 
apply (simp add: closed_def)  | 
|
967  | 
apply metis  | 
|
968  | 
done  | 
|
| 33175 | 969  | 
  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'))"
 | 
970  | 
(is "_ \<longleftrightarrow> \<not> (\<exists>t' t. ?Q t' t)")  | 
|
971  | 
unfolding connected_def openin_open closedin_closed by auto  | 
|
| 53255 | 972  | 
  {
 | 
973  | 
fix e2  | 
|
974  | 
    {
 | 
|
975  | 
fix e1  | 
|
| 53282 | 976  | 
      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 | 977  | 
by auto  | 
978  | 
}  | 
|
979  | 
then have "(\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)"  | 
|
980  | 
by metis  | 
|
981  | 
}  | 
|
982  | 
then have "\<forall>e2. (\<exists>e1. ?P e2 e1) \<longleftrightarrow> (\<exists>t. ?Q e2 t)"  | 
|
983  | 
by blast  | 
|
984  | 
then show ?thesis  | 
|
985  | 
unfolding th0 th1 by simp  | 
|
| 33175 | 986  | 
qed  | 
987  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
988  | 
|
| 33175 | 989  | 
subsection{* Limit points *}
 | 
990  | 
||
| 53282 | 991  | 
definition (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60)  | 
| 53255 | 992  | 
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 | 993  | 
|
994  | 
lemma islimptI:  | 
|
995  | 
assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"  | 
|
996  | 
shows "x islimpt S"  | 
|
997  | 
using assms unfolding islimpt_def by auto  | 
|
998  | 
||
999  | 
lemma islimptE:  | 
|
1000  | 
assumes "x islimpt S" and "x \<in> T" and "open T"  | 
|
1001  | 
obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x"  | 
|
1002  | 
using assms unfolding islimpt_def by auto  | 
|
1003  | 
||
| 44584 | 1004  | 
lemma islimpt_iff_eventually: "x islimpt S \<longleftrightarrow> \<not> eventually (\<lambda>y. y \<notin> S) (at x)"  | 
1005  | 
unfolding islimpt_def eventually_at_topological by auto  | 
|
1006  | 
||
| 53255 | 1007  | 
lemma islimpt_subset: "x islimpt S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> x islimpt T"  | 
| 44584 | 1008  | 
unfolding islimpt_def by fast  | 
| 33175 | 1009  | 
|
1010  | 
lemma islimpt_approachable:  | 
|
1011  | 
fixes x :: "'a::metric_space"  | 
|
1012  | 
shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e)"  | 
|
| 44584 | 1013  | 
unfolding islimpt_iff_eventually eventually_at by fast  | 
| 33175 | 1014  | 
|
1015  | 
lemma islimpt_approachable_le:  | 
|
1016  | 
fixes x :: "'a::metric_space"  | 
|
| 53640 | 1017  | 
shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x \<le> e)"  | 
| 33175 | 1018  | 
unfolding islimpt_approachable  | 
| 44584 | 1019  | 
using approachable_lt_le [where f="\<lambda>y. dist y x" and P="\<lambda>y. y \<notin> S \<or> y = x",  | 
1020  | 
THEN arg_cong [where f=Not]]  | 
|
1021  | 
by (simp add: Bex_def conj_commute conj_left_commute)  | 
|
| 33175 | 1022  | 
|
| 44571 | 1023  | 
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
 | 
1024  | 
  unfolding islimpt_def by (safe, fast, case_tac "T = {x}", fast, fast)
 | 
|
1025  | 
||
| 51351 | 1026  | 
lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})"
 | 
1027  | 
unfolding islimpt_def by blast  | 
|
1028  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1029  | 
text {* A perfect space has no isolated points. *}
 | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1030  | 
|
| 44571 | 1031  | 
lemma islimpt_UNIV [simp, intro]: "(x::'a::perfect_space) islimpt UNIV"  | 
1032  | 
unfolding islimpt_UNIV_iff by (rule not_open_singleton)  | 
|
| 33175 | 1033  | 
|
1034  | 
lemma perfect_choose_dist:  | 
|
| 
44072
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1035  | 
  fixes x :: "'a::{perfect_space, metric_space}"
 | 
| 33175 | 1036  | 
shows "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r"  | 
| 53255 | 1037  | 
using islimpt_UNIV [of x]  | 
1038  | 
by (simp add: islimpt_approachable)  | 
|
| 33175 | 1039  | 
|
1040  | 
lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)"  | 
|
1041  | 
unfolding closed_def  | 
|
1042  | 
apply (subst open_subopen)  | 
|
| 34105 | 1043  | 
apply (simp add: islimpt_def subset_eq)  | 
| 52624 | 1044  | 
apply (metis ComplE ComplI)  | 
1045  | 
done  | 
|
| 33175 | 1046  | 
|
1047  | 
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
 | 
|
1048  | 
unfolding islimpt_def by auto  | 
|
1049  | 
||
1050  | 
lemma finite_set_avoid:  | 
|
1051  | 
fixes a :: "'a::metric_space"  | 
|
| 53255 | 1052  | 
assumes fS: "finite S"  | 
| 53640 | 1053  | 
shows "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x"  | 
| 53255 | 1054  | 
proof (induct rule: finite_induct[OF fS])  | 
1055  | 
case 1  | 
|
1056  | 
then show ?case by (auto intro: zero_less_one)  | 
|
| 33175 | 1057  | 
next  | 
1058  | 
case (2 x F)  | 
|
| 53255 | 1059  | 
from 2 obtain d where d: "d >0" "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> d \<le> dist a x"  | 
1060  | 
by blast  | 
|
1061  | 
show ?case  | 
|
1062  | 
proof (cases "x = a")  | 
|
1063  | 
case True  | 
|
1064  | 
then show ?thesis using d by auto  | 
|
1065  | 
next  | 
|
1066  | 
case False  | 
|
| 33175 | 1067  | 
let ?d = "min d (dist a x)"  | 
| 53255 | 1068  | 
have dp: "?d > 0"  | 
1069  | 
using False d(1) using dist_nz by auto  | 
|
1070  | 
from d have d': "\<forall>x\<in>F. x\<noteq>a \<longrightarrow> ?d \<le> dist a x"  | 
|
1071  | 
by auto  | 
|
1072  | 
with dp False show ?thesis  | 
|
1073  | 
by (auto intro!: exI[where x="?d"])  | 
|
1074  | 
qed  | 
|
| 33175 | 1075  | 
qed  | 
1076  | 
||
1077  | 
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
 | 
1078  | 
by (simp add: islimpt_iff_eventually eventually_conj_iff)  | 
| 33175 | 1079  | 
|
1080  | 
lemma discrete_imp_closed:  | 
|
1081  | 
fixes S :: "'a::metric_space set"  | 
|
| 53255 | 1082  | 
assumes e: "0 < e"  | 
1083  | 
and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x"  | 
|
| 33175 | 1084  | 
shows "closed S"  | 
| 53255 | 1085  | 
proof -  | 
1086  | 
  {
 | 
|
1087  | 
fix x  | 
|
1088  | 
assume C: "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e"  | 
|
| 33175 | 1089  | 
from e have e2: "e/2 > 0" by arith  | 
| 53282 | 1090  | 
from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y \<noteq> x" "dist y x < e/2"  | 
| 53255 | 1091  | 
by blast  | 
| 33175 | 1092  | 
let ?m = "min (e/2) (dist x y) "  | 
| 53255 | 1093  | 
from e2 y(2) have mp: "?m > 0"  | 
| 53291 | 1094  | 
by (simp add: dist_nz[symmetric])  | 
| 53282 | 1095  | 
from C[rule_format, OF mp] obtain z where z: "z \<in> S" "z \<noteq> x" "dist z x < ?m"  | 
| 53255 | 1096  | 
by blast  | 
| 33175 | 1097  | 
have th: "dist z y < e" using z y  | 
1098  | 
by (intro dist_triangle_lt [where z=x], simp)  | 
|
1099  | 
from d[rule_format, OF y(1) z(1) th] y z  | 
|
1100  | 
have False by (auto simp add: dist_commute)}  | 
|
| 53255 | 1101  | 
then show ?thesis  | 
1102  | 
by (metis islimpt_approachable closed_limpt [where 'a='a])  | 
|
| 33175 | 1103  | 
qed  | 
1104  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1105  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1106  | 
subsection {* Interior of a Set *}
 | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1107  | 
|
| 44519 | 1108  | 
definition "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
 | 
1109  | 
||
1110  | 
lemma interiorI [intro?]:  | 
|
1111  | 
assumes "open T" and "x \<in> T" and "T \<subseteq> S"  | 
|
1112  | 
shows "x \<in> interior S"  | 
|
1113  | 
using assms unfolding interior_def by fast  | 
|
1114  | 
||
1115  | 
lemma interiorE [elim?]:  | 
|
1116  | 
assumes "x \<in> interior S"  | 
|
1117  | 
obtains T where "open T" and "x \<in> T" and "T \<subseteq> S"  | 
|
1118  | 
using assms unfolding interior_def by fast  | 
|
1119  | 
||
1120  | 
lemma open_interior [simp, intro]: "open (interior S)"  | 
|
1121  | 
by (simp add: interior_def open_Union)  | 
|
1122  | 
||
1123  | 
lemma interior_subset: "interior S \<subseteq> S"  | 
|
1124  | 
by (auto simp add: interior_def)  | 
|
1125  | 
||
1126  | 
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S"  | 
|
1127  | 
by (auto simp add: interior_def)  | 
|
1128  | 
||
1129  | 
lemma interior_open: "open S \<Longrightarrow> interior S = S"  | 
|
1130  | 
by (intro equalityI interior_subset interior_maximal subset_refl)  | 
|
| 33175 | 1131  | 
|
1132  | 
lemma interior_eq: "interior S = S \<longleftrightarrow> open S"  | 
|
| 44519 | 1133  | 
by (metis open_interior interior_open)  | 
1134  | 
||
1135  | 
lemma open_subset_interior: "open S \<Longrightarrow> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T"  | 
|
| 33175 | 1136  | 
by (metis interior_maximal interior_subset subset_trans)  | 
1137  | 
||
| 44519 | 1138  | 
lemma interior_empty [simp]: "interior {} = {}"
 | 
1139  | 
using open_empty by (rule interior_open)  | 
|
1140  | 
||
| 
44522
 
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
 
huffman 
parents: 
44519 
diff
changeset
 | 
1141  | 
lemma interior_UNIV [simp]: "interior UNIV = UNIV"  | 
| 
 
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
 
huffman 
parents: 
44519 
diff
changeset
 | 
1142  | 
using open_UNIV by (rule interior_open)  | 
| 
 
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
 
huffman 
parents: 
44519 
diff
changeset
 | 
1143  | 
|
| 44519 | 1144  | 
lemma interior_interior [simp]: "interior (interior S) = interior S"  | 
1145  | 
using open_interior by (rule interior_open)  | 
|
1146  | 
||
| 
44522
 
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
 
huffman 
parents: 
44519 
diff
changeset
 | 
1147  | 
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
 | 
1148  | 
by (auto simp add: interior_def)  | 
| 44519 | 1149  | 
|
1150  | 
lemma interior_unique:  | 
|
1151  | 
assumes "T \<subseteq> S" and "open T"  | 
|
1152  | 
assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T"  | 
|
1153  | 
shows "interior S = T"  | 
|
1154  | 
by (intro equalityI assms interior_subset open_interior interior_maximal)  | 
|
1155  | 
||
1156  | 
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
 | 
1157  | 
by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1  | 
| 44519 | 1158  | 
Int_lower2 interior_maximal interior_subset open_Int open_interior)  | 
1159  | 
||
1160  | 
lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)"  | 
|
1161  | 
using open_contains_ball_eq [where S="interior S"]  | 
|
1162  | 
by (simp add: open_subset_interior)  | 
|
| 33175 | 1163  | 
|
1164  | 
lemma interior_limit_point [intro]:  | 
|
1165  | 
fixes x :: "'a::perfect_space"  | 
|
| 53255 | 1166  | 
assumes x: "x \<in> interior S"  | 
1167  | 
shows "x islimpt S"  | 
|
| 
44072
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1168  | 
using x islimpt_UNIV [of x]  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1169  | 
unfolding interior_def islimpt_def  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1170  | 
apply (clarsimp, rename_tac T T')  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1171  | 
apply (drule_tac x="T \<inter> T'" in spec)  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1172  | 
apply (auto simp add: open_Int)  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1173  | 
done  | 
| 33175 | 1174  | 
|
1175  | 
lemma interior_closed_Un_empty_interior:  | 
|
| 53255 | 1176  | 
assumes cS: "closed S"  | 
1177  | 
    and iT: "interior T = {}"
 | 
|
| 44519 | 1178  | 
shows "interior (S \<union> T) = interior S"  | 
| 33175 | 1179  | 
proof  | 
| 44519 | 1180  | 
show "interior S \<subseteq> interior (S \<union> T)"  | 
| 53255 | 1181  | 
by (rule interior_mono) (rule Un_upper1)  | 
| 33175 | 1182  | 
show "interior (S \<union> T) \<subseteq> interior S"  | 
1183  | 
proof  | 
|
| 53255 | 1184  | 
fix x  | 
1185  | 
assume "x \<in> interior (S \<union> T)"  | 
|
| 44519 | 1186  | 
then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T" ..  | 
| 33175 | 1187  | 
show "x \<in> interior S"  | 
1188  | 
proof (rule ccontr)  | 
|
1189  | 
assume "x \<notin> interior S"  | 
|
1190  | 
with `x \<in> R` `open R` obtain y where "y \<in> R - S"  | 
|
| 44519 | 1191  | 
unfolding interior_def by fast  | 
| 53282 | 1192  | 
from `open R` `closed S` have "open (R - S)"  | 
1193  | 
by (rule open_Diff)  | 
|
1194  | 
from `R \<subseteq> S \<union> T` have "R - S \<subseteq> T"  | 
|
1195  | 
by fast  | 
|
1196  | 
      from `y \<in> R - S` `open (R - S)` `R - S \<subseteq> T` `interior T = {}` show False
 | 
|
1197  | 
unfolding interior_def by fast  | 
|
| 33175 | 1198  | 
qed  | 
1199  | 
qed  | 
|
1200  | 
qed  | 
|
1201  | 
||
| 44365 | 1202  | 
lemma interior_Times: "interior (A \<times> B) = interior A \<times> interior B"  | 
1203  | 
proof (rule interior_unique)  | 
|
1204  | 
show "interior A \<times> interior B \<subseteq> A \<times> B"  | 
|
1205  | 
by (intro Sigma_mono interior_subset)  | 
|
1206  | 
show "open (interior A \<times> interior B)"  | 
|
1207  | 
by (intro open_Times open_interior)  | 
|
| 53255 | 1208  | 
fix T  | 
1209  | 
assume "T \<subseteq> A \<times> B" and "open T"  | 
|
1210  | 
then show "T \<subseteq> interior A \<times> interior B"  | 
|
| 53282 | 1211  | 
proof safe  | 
| 53255 | 1212  | 
fix x y  | 
1213  | 
assume "(x, y) \<in> T"  | 
|
| 44519 | 1214  | 
then obtain C D where "open C" "open D" "C \<times> D \<subseteq> T" "x \<in> C" "y \<in> D"  | 
1215  | 
using `open T` unfolding open_prod_def by fast  | 
|
| 53255 | 1216  | 
then have "open C" "open D" "C \<subseteq> A" "D \<subseteq> B" "x \<in> C" "y \<in> D"  | 
| 44519 | 1217  | 
using `T \<subseteq> A \<times> B` by auto  | 
| 53255 | 1218  | 
then show "x \<in> interior A" and "y \<in> interior B"  | 
| 44519 | 1219  | 
by (auto intro: interiorI)  | 
1220  | 
qed  | 
|
| 44365 | 1221  | 
qed  | 
1222  | 
||
| 33175 | 1223  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1224  | 
subsection {* Closure of a Set *}
 | 
| 33175 | 1225  | 
|
1226  | 
definition "closure S = S \<union> {x | x. x islimpt S}"
 | 
|
1227  | 
||
| 44518 | 1228  | 
lemma interior_closure: "interior S = - (closure (- S))"  | 
1229  | 
unfolding interior_def closure_def islimpt_def by auto  | 
|
1230  | 
||
| 34105 | 1231  | 
lemma closure_interior: "closure S = - interior (- S)"  | 
| 44518 | 1232  | 
unfolding interior_closure by simp  | 
| 33175 | 1233  | 
|
1234  | 
lemma closed_closure[simp, intro]: "closed (closure S)"  | 
|
| 44518 | 1235  | 
unfolding closure_interior by (simp add: closed_Compl)  | 
1236  | 
||
1237  | 
lemma closure_subset: "S \<subseteq> closure S"  | 
|
1238  | 
unfolding closure_def by simp  | 
|
| 33175 | 1239  | 
|
1240  | 
lemma closure_hull: "closure S = closed hull S"  | 
|
| 44519 | 1241  | 
unfolding hull_def closure_interior interior_def by auto  | 
| 33175 | 1242  | 
|
1243  | 
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S"  | 
|
| 44519 | 1244  | 
unfolding closure_hull using closed_Inter by (rule hull_eq)  | 
1245  | 
||
1246  | 
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S"  | 
|
1247  | 
unfolding closure_eq .  | 
|
1248  | 
||
1249  | 
lemma closure_closure [simp]: "closure (closure S) = closure S"  | 
|
| 44518 | 1250  | 
unfolding closure_hull by (rule hull_hull)  | 
| 33175 | 1251  | 
|
| 
44522
 
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
 
huffman 
parents: 
44519 
diff
changeset
 | 
1252  | 
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T"  | 
| 44518 | 1253  | 
unfolding closure_hull by (rule hull_mono)  | 
| 33175 | 1254  | 
|
| 44519 | 1255  | 
lemma closure_minimal: "S \<subseteq> T \<Longrightarrow> closed T \<Longrightarrow> closure S \<subseteq> T"  | 
| 44518 | 1256  | 
unfolding closure_hull by (rule hull_minimal)  | 
| 33175 | 1257  | 
|
| 44519 | 1258  | 
lemma closure_unique:  | 
| 53255 | 1259  | 
assumes "S \<subseteq> T"  | 
1260  | 
and "closed T"  | 
|
1261  | 
and "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'"  | 
|
| 44519 | 1262  | 
shows "closure S = T"  | 
1263  | 
using assms unfolding closure_hull by (rule hull_unique)  | 
|
1264  | 
||
1265  | 
lemma closure_empty [simp]: "closure {} = {}"
 | 
|
| 44518 | 1266  | 
using closed_empty by (rule closure_closed)  | 
| 33175 | 1267  | 
|
| 
44522
 
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
 
huffman 
parents: 
44519 
diff
changeset
 | 
1268  | 
lemma closure_UNIV [simp]: "closure UNIV = UNIV"  | 
| 44518 | 1269  | 
using closed_UNIV by (rule closure_closed)  | 
1270  | 
||
1271  | 
lemma closure_union [simp]: "closure (S \<union> T) = closure S \<union> closure T"  | 
|
1272  | 
unfolding closure_interior by simp  | 
|
| 33175 | 1273  | 
|
1274  | 
lemma closure_eq_empty: "closure S = {} \<longleftrightarrow> S = {}"
 | 
|
1275  | 
using closure_empty closure_subset[of S]  | 
|
1276  | 
by blast  | 
|
1277  | 
||
1278  | 
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S"  | 
|
1279  | 
using closure_eq[of S] closure_subset[of S]  | 
|
1280  | 
by simp  | 
|
1281  | 
||
1282  | 
lemma open_inter_closure_eq_empty:  | 
|
1283  | 
  "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
 | 
|
| 34105 | 1284  | 
using open_subset_interior[of S "- T"]  | 
1285  | 
using interior_subset[of "- T"]  | 
|
| 33175 | 1286  | 
unfolding closure_interior  | 
1287  | 
by auto  | 
|
1288  | 
||
1289  | 
lemma open_inter_closure_subset:  | 
|
1290  | 
"open S \<Longrightarrow> (S \<inter> (closure T)) \<subseteq> closure(S \<inter> T)"  | 
|
1291  | 
proof  | 
|
1292  | 
fix x  | 
|
1293  | 
assume as: "open S" "x \<in> S \<inter> closure T"  | 
|
| 53255 | 1294  | 
  {
 | 
| 53282 | 1295  | 
assume *: "x islimpt T"  | 
| 33175 | 1296  | 
have "x islimpt (S \<inter> T)"  | 
1297  | 
proof (rule islimptI)  | 
|
1298  | 
fix A  | 
|
1299  | 
assume "x \<in> A" "open A"  | 
|
1300  | 
with as have "x \<in> A \<inter> S" "open (A \<inter> S)"  | 
|
1301  | 
by (simp_all add: open_Int)  | 
|
1302  | 
with * obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x"  | 
|
1303  | 
by (rule islimptE)  | 
|
| 53255 | 1304  | 
then have "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x"  | 
| 33175 | 1305  | 
by simp_all  | 
| 53255 | 1306  | 
then show "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" ..  | 
| 33175 | 1307  | 
qed  | 
1308  | 
}  | 
|
1309  | 
then show "x \<in> closure (S \<inter> T)" using as  | 
|
1310  | 
unfolding closure_def  | 
|
1311  | 
by blast  | 
|
1312  | 
qed  | 
|
1313  | 
||
| 44519 | 1314  | 
lemma closure_complement: "closure (- S) = - interior S"  | 
| 44518 | 1315  | 
unfolding closure_interior by simp  | 
| 33175 | 1316  | 
|
| 44519 | 1317  | 
lemma interior_complement: "interior (- S) = - closure S"  | 
| 44518 | 1318  | 
unfolding closure_interior by simp  | 
| 33175 | 1319  | 
|
| 44365 | 1320  | 
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B"  | 
| 44519 | 1321  | 
proof (rule closure_unique)  | 
| 44365 | 1322  | 
show "A \<times> B \<subseteq> closure A \<times> closure B"  | 
1323  | 
by (intro Sigma_mono closure_subset)  | 
|
1324  | 
show "closed (closure A \<times> closure B)"  | 
|
1325  | 
by (intro closed_Times closed_closure)  | 
|
| 53282 | 1326  | 
fix T  | 
1327  | 
assume "A \<times> B \<subseteq> T" and "closed T"  | 
|
1328  | 
then show "closure A \<times> closure B \<subseteq> T"  | 
|
| 44365 | 1329  | 
apply (simp add: closed_def open_prod_def, clarify)  | 
1330  | 
apply (rule ccontr)  | 
|
1331  | 
apply (drule_tac x="(a, b)" in bspec, simp, clarify, rename_tac C D)  | 
|
1332  | 
apply (simp add: closure_interior interior_def)  | 
|
1333  | 
apply (drule_tac x=C in spec)  | 
|
1334  | 
apply (drule_tac x=D in spec)  | 
|
1335  | 
apply auto  | 
|
1336  | 
done  | 
|
1337  | 
qed  | 
|
1338  | 
||
| 51351 | 1339  | 
lemma islimpt_in_closure: "(x islimpt S) = (x:closure(S-{x}))"
 | 
1340  | 
unfolding closure_def using islimpt_punctured by blast  | 
|
1341  | 
||
1342  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1343  | 
subsection {* Frontier (aka boundary) *}
 | 
| 33175 | 1344  | 
|
1345  | 
definition "frontier S = closure S - interior S"  | 
|
1346  | 
||
| 53255 | 1347  | 
lemma frontier_closed: "closed (frontier S)"  | 
| 33175 | 1348  | 
by (simp add: frontier_def closed_Diff)  | 
1349  | 
||
| 34105 | 1350  | 
lemma frontier_closures: "frontier S = (closure S) \<inter> (closure(- S))"  | 
| 33175 | 1351  | 
by (auto simp add: frontier_def interior_closure)  | 
1352  | 
||
1353  | 
lemma frontier_straddle:  | 
|
1354  | 
fixes a :: "'a::metric_space"  | 
|
| 44909 | 1355  | 
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))"  | 
1356  | 
unfolding frontier_def closure_interior  | 
|
1357  | 
by (auto simp add: mem_interior subset_eq ball_def)  | 
|
| 33175 | 1358  | 
|
1359  | 
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S"  | 
|
1360  | 
by (metis frontier_def closure_closed Diff_subset)  | 
|
1361  | 
||
| 34964 | 1362  | 
lemma frontier_empty[simp]: "frontier {} = {}"
 | 
| 
36362
 
06475a1547cb
fix lots of looping simp calls and other warnings
 
huffman 
parents: 
36360 
diff
changeset
 | 
1363  | 
by (simp add: frontier_def)  | 
| 33175 | 1364  | 
|
1365  | 
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S"  | 
|
1366  | 
proof-  | 
|
| 53255 | 1367  | 
  {
 | 
1368  | 
assume "frontier S \<subseteq> S"  | 
|
1369  | 
then have "closure S \<subseteq> S"  | 
|
1370  | 
using interior_subset unfolding frontier_def by auto  | 
|
1371  | 
then have "closed S"  | 
|
1372  | 
using closure_subset_eq by auto  | 
|
| 33175 | 1373  | 
}  | 
| 53255 | 1374  | 
then show ?thesis using frontier_subset_closed[of S] ..  | 
| 33175 | 1375  | 
qed  | 
1376  | 
||
| 34105 | 1377  | 
lemma frontier_complement: "frontier(- S) = frontier S"  | 
| 33175 | 1378  | 
by (auto simp add: frontier_def closure_complement interior_complement)  | 
1379  | 
||
1380  | 
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
 | 
|
| 34105 | 1381  | 
using frontier_complement frontier_subset_eq[of "- S"]  | 
1382  | 
unfolding open_closed by auto  | 
|
| 33175 | 1383  | 
|
| 
44081
 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 
huffman 
parents: 
44076 
diff
changeset
 | 
1384  | 
subsection {* Filters and the ``eventually true'' quantifier *}
 | 
| 
 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 
huffman 
parents: 
44076 
diff
changeset
 | 
1385  | 
|
| 52624 | 1386  | 
definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter"  | 
1387  | 
(infixr "indirection" 70)  | 
|
1388  | 
  where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
 | 
|
| 33175 | 1389  | 
|
| 36437 | 1390  | 
text {* Identify Trivial limits, where we can't approach arbitrarily closely. *}
 | 
| 33175 | 1391  | 
|
| 52624 | 1392  | 
lemma trivial_limit_within: "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S"  | 
| 33175 | 1393  | 
proof  | 
1394  | 
assume "trivial_limit (at a within S)"  | 
|
| 53255 | 1395  | 
then show "\<not> a islimpt S"  | 
| 33175 | 1396  | 
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
 | 
1397  | 
unfolding eventually_at_topological  | 
| 33175 | 1398  | 
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
 | 
1399  | 
apply (clarsimp simp add: set_eq_iff)  | 
| 33175 | 1400  | 
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
 | 
1401  | 
apply (clarsimp, drule_tac x=y in bspec, simp_all)  | 
| 33175 | 1402  | 
done  | 
1403  | 
next  | 
|
1404  | 
assume "\<not> a islimpt S"  | 
|
| 53255 | 1405  | 
then show "trivial_limit (at a within S)"  | 
| 33175 | 1406  | 
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
 | 
1407  | 
unfolding eventually_at_topological  | 
| 33175 | 1408  | 
unfolding islimpt_def  | 
| 
36358
 
246493d61204
define nets directly as filters, instead of as filter bases
 
huffman 
parents: 
36336 
diff
changeset
 | 
1409  | 
apply clarsimp  | 
| 
 
246493d61204
define nets directly as filters, instead of as filter bases
 
huffman 
parents: 
36336 
diff
changeset
 | 
1410  | 
apply (rule_tac x=T in exI)  | 
| 
 
246493d61204
define nets directly as filters, instead of as filter bases
 
huffman 
parents: 
36336 
diff
changeset
 | 
1411  | 
apply auto  | 
| 33175 | 1412  | 
done  | 
1413  | 
qed  | 
|
1414  | 
||
1415  | 
lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV"  | 
|
| 45031 | 1416  | 
using trivial_limit_within [of a UNIV] by simp  | 
| 33175 | 1417  | 
|
1418  | 
lemma trivial_limit_at:  | 
|
1419  | 
fixes a :: "'a::perfect_space"  | 
|
1420  | 
shows "\<not> trivial_limit (at a)"  | 
|
| 44571 | 1421  | 
by (rule at_neq_bot)  | 
| 33175 | 1422  | 
|
1423  | 
lemma trivial_limit_at_infinity:  | 
|
| 
44081
 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 
huffman 
parents: 
44076 
diff
changeset
 | 
1424  | 
  "\<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
 | 
1425  | 
unfolding trivial_limit_def eventually_at_infinity  | 
| 
 
246493d61204
define nets directly as filters, instead of as filter bases
 
huffman 
parents: 
36336 
diff
changeset
 | 
1426  | 
apply clarsimp  | 
| 
44072
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1427  | 
apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify)  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1428  | 
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
 | 
1429  | 
apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def])  | 
| 
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1430  | 
apply (drule_tac x=UNIV in spec, simp)  | 
| 33175 | 1431  | 
done  | 
1432  | 
||
| 53640 | 1433  | 
lemma not_trivial_limit_within: "\<not> trivial_limit (at x within S) = (x \<in> closure (S - {x}))"
 | 
1434  | 
using islimpt_in_closure  | 
|
1435  | 
by (metis trivial_limit_within)  | 
|
| 51351 | 1436  | 
|
| 36437 | 1437  | 
text {* Some property holds "sufficiently close" to the limit point. *}
 | 
| 33175 | 1438  | 
|
| 
51530
 
609914f0934a
rename eventually_at / _within, to distinguish them from the lemmas in the HOL image
 
hoelzl 
parents: 
51518 
diff
changeset
 | 
1439  | 
lemma eventually_at2:  | 
| 33175 | 1440  | 
"eventually P (at a) \<longleftrightarrow> (\<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> P x)"  | 
| 53255 | 1441  | 
unfolding eventually_at dist_nz by auto  | 
1442  | 
||
1443  | 
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
 | 
1444  | 
unfolding trivial_limit_def  | 
| 
 
246493d61204
define nets directly as filters, instead of as filter bases
 
huffman 
parents: 
36336 
diff
changeset
 | 
1445  | 
by (auto elim: eventually_rev_mp)  | 
| 33175 | 1446  | 
|
1447  | 
lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net"  | 
|
| 45031 | 1448  | 
by simp  | 
| 33175 | 1449  | 
|
1450  | 
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
 | 
1451  | 
by (simp add: filter_eq_iff)  | 
| 33175 | 1452  | 
|
1453  | 
text{* Combining theorems for "eventually" *}
 | 
|
1454  | 
||
1455  | 
lemma eventually_rev_mono:  | 
|
1456  | 
"eventually P net \<Longrightarrow> (\<forall>x. P x \<longrightarrow> Q x) \<Longrightarrow> eventually Q net"  | 
|
| 53255 | 1457  | 
using eventually_mono [of P Q] by fast  | 
| 33175 | 1458  | 
|
| 53282 | 1459  | 
lemma not_eventually: "(\<forall>x. \<not> P x ) \<Longrightarrow> \<not> trivial_limit net \<Longrightarrow> \<not> eventually (\<lambda>x. P x) net"  | 
| 33175 | 1460  | 
by (simp add: eventually_False)  | 
1461  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1462  | 
|
| 36437 | 1463  | 
subsection {* Limits *}
 | 
| 33175 | 1464  | 
|
1465  | 
lemma Lim:  | 
|
| 53255 | 1466  | 
"(f ---> l) net \<longleftrightarrow>  | 
| 33175 | 1467  | 
trivial_limit net \<or>  | 
1468  | 
(\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)"  | 
|
1469  | 
unfolding tendsto_iff trivial_limit_eq by auto  | 
|
1470  | 
||
1471  | 
text{* Show that they yield usual definitions in the various cases. *}
 | 
|
1472  | 
||
1473  | 
lemma Lim_within_le: "(f ---> l)(at a within S) \<longleftrightarrow>  | 
|
| 53640 | 1474  | 
(\<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
 | 
1475  | 
by (auto simp add: tendsto_iff eventually_at_le dist_nz)  | 
| 33175 | 1476  | 
|
1477  | 
lemma Lim_within: "(f ---> l) (at a within S) \<longleftrightarrow>  | 
|
| 53640 | 1478  | 
(\<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
 | 
1479  | 
by (auto simp add: tendsto_iff eventually_at dist_nz)  | 
| 33175 | 1480  | 
|
1481  | 
lemma Lim_at: "(f ---> l) (at a) \<longleftrightarrow>  | 
|
| 53640 | 1482  | 
(\<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
 | 
1483  | 
by (auto simp add: tendsto_iff eventually_at2)  | 
| 33175 | 1484  | 
|
1485  | 
lemma Lim_at_infinity:  | 
|
| 53640 | 1486  | 
"(f ---> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x \<ge> b \<longrightarrow> dist (f x) l < e)"  | 
| 33175 | 1487  | 
by (auto simp add: tendsto_iff eventually_at_infinity)  | 
1488  | 
||
1489  | 
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f ---> l) net"  | 
|
1490  | 
by (rule topological_tendstoI, auto elim: eventually_rev_mono)  | 
|
1491  | 
||
1492  | 
text{* The expected monotonicity property. *}
 | 
|
1493  | 
||
| 53255 | 1494  | 
lemma Lim_Un:  | 
1495  | 
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
 | 
1496  | 
shows "(f ---> l) (at x within (S \<union> T))"  | 
| 53860 | 1497  | 
using assms unfolding at_within_union by (rule filterlim_sup)  | 
| 33175 | 1498  | 
|
1499  | 
lemma Lim_Un_univ:  | 
|
| 53282 | 1500  | 
"(f ---> l) (at x within S) \<Longrightarrow> (f ---> l) (at x within T) \<Longrightarrow>  | 
| 53255 | 1501  | 
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
 | 
1502  | 
by (metis Lim_Un)  | 
| 33175 | 1503  | 
|
1504  | 
text{* Interrelations between restricted and unrestricted limits. *}
 | 
|
1505  | 
||
| 
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
 | 
1506  | 
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
 | 
1507  | 
"(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
 | 
1508  | 
by (metis order_refl filterlim_mono subset_UNIV at_le)  | 
| 33175 | 1509  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1510  | 
lemma eventually_within_interior:  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1511  | 
assumes "x \<in> interior S"  | 
| 53255 | 1512  | 
shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)"  | 
1513  | 
(is "?lhs = ?rhs")  | 
|
1514  | 
proof  | 
|
| 44519 | 1515  | 
from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S" ..  | 
| 53255 | 1516  | 
  {
 | 
1517  | 
assume "?lhs"  | 
|
| 53640 | 1518  | 
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
 | 
1519  | 
unfolding eventually_at_topological  | 
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1520  | 
by auto  | 
| 53640 | 1521  | 
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
 | 
1522  | 
by auto  | 
| 53255 | 1523  | 
then show "?rhs"  | 
| 51471 | 1524  | 
unfolding eventually_at_topological by auto  | 
| 53255 | 1525  | 
next  | 
1526  | 
assume "?rhs"  | 
|
1527  | 
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
 | 
1528  | 
by (auto elim: eventually_elim1 simp: eventually_at_filter)  | 
| 52624 | 1529  | 
}  | 
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1530  | 
qed  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1531  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1532  | 
lemma at_within_interior:  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1533  | 
"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
 | 
1534  | 
unfolding filter_eq_iff by (intro allI eventually_within_interior)  | 
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1535  | 
|
| 43338 | 1536  | 
lemma Lim_within_LIMSEQ:  | 
| 53862 | 1537  | 
fixes a :: "'a::first_countable_topology"  | 
| 43338 | 1538  | 
assumes "\<forall>S. (\<forall>n. S n \<noteq> a \<and> S n \<in> T) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"  | 
1539  | 
shows "(X ---> L) (at a within T)"  | 
|
| 44584 | 1540  | 
using assms unfolding tendsto_def [where l=L]  | 
1541  | 
by (simp add: sequentially_imp_eventually_within)  | 
|
| 43338 | 1542  | 
|
1543  | 
lemma Lim_right_bound:  | 
|
| 51773 | 1544  | 
  fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder, no_top} \<Rightarrow>
 | 
1545  | 
    'b::{linorder_topology, conditionally_complete_linorder}"
 | 
|
| 43338 | 1546  | 
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 | 1547  | 
and bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a"  | 
| 43338 | 1548  | 
  shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
 | 
| 53640 | 1549  | 
proof (cases "{x<..} \<inter> I = {}")
 | 
1550  | 
case True  | 
|
| 53859 | 1551  | 
then show ?thesis by simp  | 
| 43338 | 1552  | 
next  | 
| 53640 | 1553  | 
case False  | 
| 43338 | 1554  | 
show ?thesis  | 
| 
51518
 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 
hoelzl 
parents: 
51481 
diff
changeset
 | 
1555  | 
proof (rule order_tendstoI)  | 
| 53282 | 1556  | 
fix a  | 
1557  | 
    assume a: "a < Inf (f ` ({x<..} \<inter> I))"
 | 
|
| 53255 | 1558  | 
    {
 | 
1559  | 
fix y  | 
|
1560  | 
      assume "y \<in> {x<..} \<inter> I"
 | 
|
| 53640 | 1561  | 
      with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
 | 
| 
54263
 
c4159fe6fa46
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
 
hoelzl 
parents: 
54260 
diff
changeset
 | 
1562  | 
by (auto intro!: cInf_lower bdd_belowI2)  | 
| 53255 | 1563  | 
with a have "a < f y"  | 
1564  | 
by (blast intro: less_le_trans)  | 
|
1565  | 
}  | 
|
| 
51518
 
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
 
hoelzl 
parents: 
51481 
diff
changeset
 | 
1566  | 
    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
 | 
1567  | 
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
 | 
1568  | 
next  | 
| 53255 | 1569  | 
fix a  | 
1570  | 
    assume "Inf (f ` ({x<..} \<inter> I)) < a"
 | 
|
| 53640 | 1571  | 
from cInf_lessD[OF _ this] False obtain y where y: "x < y" "y \<in> I" "f y < a"  | 
| 53255 | 1572  | 
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
 | 
1573  | 
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
 | 
1574  | 
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
 | 
1575  | 
    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
 | 
1576  | 
unfolding eventually_at_filter by eventually_elim simp  | 
| 43338 | 1577  | 
qed  | 
1578  | 
qed  | 
|
1579  | 
||
| 33175 | 1580  | 
text{* Another limit point characterization. *}
 | 
1581  | 
||
1582  | 
lemma islimpt_sequential:  | 
|
| 50883 | 1583  | 
fixes x :: "'a::first_countable_topology"  | 
1584  | 
  shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S - {x}) \<and> (f ---> x) sequentially)"
 | 
|
| 33175 | 1585  | 
(is "?lhs = ?rhs")  | 
1586  | 
proof  | 
|
1587  | 
assume ?lhs  | 
|
| 50883 | 1588  | 
from countable_basis_at_decseq[of x] guess A . note A = this  | 
1589  | 
def f \<equiv> "\<lambda>n. SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"  | 
|
| 53255 | 1590  | 
  {
 | 
1591  | 
fix n  | 
|
| 50883 | 1592  | 
from `?lhs` have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y"  | 
1593  | 
unfolding islimpt_def using A(1,2)[of n] by auto  | 
|
1594  | 
then have "f n \<in> S \<and> f n \<in> A n \<and> x \<noteq> f n"  | 
|
1595  | 
unfolding f_def by (rule someI_ex)  | 
|
| 53255 | 1596  | 
then have "f n \<in> S" "f n \<in> A n" "x \<noteq> f n" by auto  | 
1597  | 
}  | 
|
| 50883 | 1598  | 
  then have "\<forall>n. f n \<in> S - {x}" by auto
 | 
1599  | 
moreover have "(\<lambda>n. f n) ----> x"  | 
|
1600  | 
proof (rule topological_tendstoI)  | 
|
| 53255 | 1601  | 
fix S  | 
1602  | 
assume "open S" "x \<in> S"  | 
|
| 50883 | 1603  | 
from A(3)[OF this] `\<And>n. f n \<in> A n`  | 
| 53255 | 1604  | 
show "eventually (\<lambda>x. f x \<in> S) sequentially"  | 
1605  | 
by (auto elim!: eventually_elim1)  | 
|
| 44584 | 1606  | 
qed  | 
1607  | 
ultimately show ?rhs by fast  | 
|
| 33175 | 1608  | 
next  | 
1609  | 
assume ?rhs  | 
|
| 53255 | 1610  | 
  then obtain f :: "nat \<Rightarrow> 'a" where f: "\<And>n. f n \<in> S - {x}" and lim: "f ----> x"
 | 
1611  | 
by auto  | 
|
| 50883 | 1612  | 
show ?lhs  | 
1613  | 
unfolding islimpt_def  | 
|
1614  | 
proof safe  | 
|
| 53255 | 1615  | 
fix T  | 
1616  | 
assume "open T" "x \<in> T"  | 
|
| 50883 | 1617  | 
from lim[THEN topological_tendstoD, OF this] f  | 
1618  | 
show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> x"  | 
|
1619  | 
unfolding eventually_sequentially by auto  | 
|
1620  | 
qed  | 
|
| 33175 | 1621  | 
qed  | 
1622  | 
||
1623  | 
lemma Lim_null:  | 
|
1624  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
| 44125 | 1625  | 
shows "(f ---> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) ---> 0) net"  | 
| 33175 | 1626  | 
by (simp add: Lim dist_norm)  | 
1627  | 
||
1628  | 
lemma Lim_null_comparison:  | 
|
1629  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
1630  | 
assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g ---> 0) net"  | 
|
1631  | 
shows "(f ---> 0) net"  | 
|
| 53282 | 1632  | 
using assms(2)  | 
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1633  | 
proof (rule metric_tendsto_imp_tendsto)  | 
| 
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1634  | 
show "eventually (\<lambda>x. dist (f x) 0 \<le> dist (g x) 0) net"  | 
| 53255 | 1635  | 
using assms(1) by (rule eventually_elim1) (simp add: dist_norm)  | 
| 33175 | 1636  | 
qed  | 
1637  | 
||
1638  | 
lemma Lim_transform_bound:  | 
|
1639  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
| 53255 | 1640  | 
and g :: "'a \<Rightarrow> 'c::real_normed_vector"  | 
| 53640 | 1641  | 
assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) net"  | 
| 53255 | 1642  | 
and "(g ---> 0) net"  | 
| 33175 | 1643  | 
shows "(f ---> 0) net"  | 
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1644  | 
using assms(1) tendsto_norm_zero [OF assms(2)]  | 
| 
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1645  | 
by (rule Lim_null_comparison)  | 
| 33175 | 1646  | 
|
1647  | 
text{* Deducing things about the limit from the elements. *}
 | 
|
1648  | 
||
1649  | 
lemma Lim_in_closed_set:  | 
|
| 53255 | 1650  | 
assumes "closed S"  | 
1651  | 
and "eventually (\<lambda>x. f(x) \<in> S) net"  | 
|
| 53640 | 1652  | 
and "\<not> trivial_limit net" "(f ---> l) net"  | 
| 33175 | 1653  | 
shows "l \<in> S"  | 
1654  | 
proof (rule ccontr)  | 
|
1655  | 
assume "l \<notin> S"  | 
|
1656  | 
with `closed S` have "open (- S)" "l \<in> - S"  | 
|
1657  | 
by (simp_all add: open_Compl)  | 
|
1658  | 
with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net"  | 
|
1659  | 
by (rule topological_tendstoD)  | 
|
1660  | 
with assms(2) have "eventually (\<lambda>x. False) net"  | 
|
1661  | 
by (rule eventually_elim2) simp  | 
|
1662  | 
with assms(3) show "False"  | 
|
1663  | 
by (simp add: eventually_False)  | 
|
1664  | 
qed  | 
|
1665  | 
||
1666  | 
text{* Need to prove closed(cball(x,e)) before deducing this as a corollary. *}
 | 
|
1667  | 
||
1668  | 
lemma Lim_dist_ubound:  | 
|
| 53255 | 1669  | 
assumes "\<not>(trivial_limit net)"  | 
1670  | 
and "(f ---> l) net"  | 
|
| 53640 | 1671  | 
and "eventually (\<lambda>x. dist a (f x) \<le> e) net"  | 
1672  | 
shows "dist a l \<le> e"  | 
|
| 52624 | 1673  | 
proof -  | 
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1674  | 
  have "dist a l \<in> {..e}"
 | 
| 
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1675  | 
proof (rule Lim_in_closed_set)  | 
| 53255 | 1676  | 
    show "closed {..e}"
 | 
1677  | 
by simp  | 
|
1678  | 
    show "eventually (\<lambda>x. dist a (f x) \<in> {..e}) net"
 | 
|
1679  | 
by (simp add: assms)  | 
|
1680  | 
show "\<not> trivial_limit net"  | 
|
1681  | 
by fact  | 
|
1682  | 
show "((\<lambda>x. dist a (f x)) ---> dist a l) net"  | 
|
1683  | 
by (intro tendsto_intros assms)  | 
|
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1684  | 
qed  | 
| 53255 | 1685  | 
then show ?thesis by simp  | 
| 33175 | 1686  | 
qed  | 
1687  | 
||
1688  | 
lemma Lim_norm_ubound:  | 
|
1689  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
| 53255 | 1690  | 
assumes "\<not>(trivial_limit net)" "(f ---> l) net" "eventually (\<lambda>x. norm(f x) \<le> e) net"  | 
1691  | 
shows "norm(l) \<le> e"  | 
|
| 52624 | 1692  | 
proof -  | 
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1693  | 
  have "norm l \<in> {..e}"
 | 
| 
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1694  | 
proof (rule Lim_in_closed_set)  | 
| 53255 | 1695  | 
    show "closed {..e}"
 | 
1696  | 
by simp  | 
|
1697  | 
    show "eventually (\<lambda>x. norm (f x) \<in> {..e}) net"
 | 
|
1698  | 
by (simp add: assms)  | 
|
1699  | 
show "\<not> trivial_limit net"  | 
|
1700  | 
by fact  | 
|
1701  | 
show "((\<lambda>x. norm (f x)) ---> norm l) net"  | 
|
1702  | 
by (intro tendsto_intros assms)  | 
|
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1703  | 
qed  | 
| 53255 | 1704  | 
then show ?thesis by simp  | 
| 33175 | 1705  | 
qed  | 
1706  | 
||
1707  | 
lemma Lim_norm_lbound:  | 
|
1708  | 
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"  | 
|
| 53640 | 1709  | 
assumes "\<not> trivial_limit net"  | 
1710  | 
and "(f ---> l) net"  | 
|
1711  | 
and "eventually (\<lambda>x. e \<le> norm (f x)) net"  | 
|
| 33175 | 1712  | 
shows "e \<le> norm l"  | 
| 52624 | 1713  | 
proof -  | 
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1714  | 
  have "norm l \<in> {e..}"
 | 
| 
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1715  | 
proof (rule Lim_in_closed_set)  | 
| 53255 | 1716  | 
    show "closed {e..}"
 | 
1717  | 
by simp  | 
|
1718  | 
    show "eventually (\<lambda>x. norm (f x) \<in> {e..}) net"
 | 
|
1719  | 
by (simp add: assms)  | 
|
1720  | 
show "\<not> trivial_limit net"  | 
|
1721  | 
by fact  | 
|
1722  | 
show "((\<lambda>x. norm (f x)) ---> norm l) net"  | 
|
1723  | 
by (intro tendsto_intros assms)  | 
|
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1724  | 
qed  | 
| 53255 | 1725  | 
then show ?thesis by simp  | 
| 33175 | 1726  | 
qed  | 
1727  | 
||
1728  | 
text{* Limit under bilinear function *}
 | 
|
1729  | 
||
1730  | 
lemma Lim_bilinear:  | 
|
| 53282 | 1731  | 
assumes "(f ---> l) net"  | 
1732  | 
and "(g ---> m) net"  | 
|
1733  | 
and "bounded_bilinear h"  | 
|
| 33175 | 1734  | 
shows "((\<lambda>x. h (f x) (g x)) ---> (h l m)) net"  | 
| 52624 | 1735  | 
using `bounded_bilinear h` `(f ---> l) net` `(g ---> m) net`  | 
1736  | 
by (rule bounded_bilinear.tendsto)  | 
|
| 33175 | 1737  | 
|
1738  | 
text{* These are special for limits out of the same vector space. *}
 | 
|
1739  | 
||
1740  | 
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
 | 
1741  | 
unfolding id_def by (rule tendsto_ident_at)  | 
| 33175 | 1742  | 
|
1743  | 
lemma Lim_at_id: "(id ---> a) (at a)"  | 
|
| 45031 | 1744  | 
unfolding id_def by (rule tendsto_ident_at)  | 
| 33175 | 1745  | 
|
1746  | 
lemma Lim_at_zero:  | 
|
1747  | 
fixes a :: "'a::real_normed_vector"  | 
|
| 53291 | 1748  | 
and l :: "'b::topological_space"  | 
| 53282 | 1749  | 
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
 | 
1750  | 
using LIM_offset_zero LIM_offset_zero_cancel ..  | 
| 33175 | 1751  | 
|
| 
44081
 
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
 
huffman 
parents: 
44076 
diff
changeset
 | 
1752  | 
text{* It's also sometimes useful to extract the limit point from the filter. *}
 | 
| 33175 | 1753  | 
|
| 52624 | 1754  | 
abbreviation netlimit :: "'a::t2_space filter \<Rightarrow> 'a"  | 
1755  | 
where "netlimit F \<equiv> Lim F (\<lambda>x. x)"  | 
|
| 33175 | 1756  | 
|
| 53282 | 1757  | 
lemma netlimit_within: "\<not> trivial_limit (at a within S) \<Longrightarrow> netlimit (at a within S) = a"  | 
| 51365 | 1758  | 
by (rule tendsto_Lim) (auto intro: tendsto_intros)  | 
| 33175 | 1759  | 
|
1760  | 
lemma netlimit_at:  | 
|
| 
44072
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
1761  | 
  fixes a :: "'a::{perfect_space,t2_space}"
 | 
| 33175 | 1762  | 
shows "netlimit (at a) = a"  | 
| 45031 | 1763  | 
using netlimit_within [of a UNIV] by simp  | 
| 33175 | 1764  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1765  | 
lemma lim_within_interior:  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1766  | 
"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
 | 
1767  | 
by (metis at_within_interior)  | 
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1768  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1769  | 
lemma netlimit_within_interior:  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1770  | 
  fixes x :: "'a::{t2_space,perfect_space}"
 | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1771  | 
assumes "x \<in> interior S"  | 
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1772  | 
shows "netlimit (at x within S) = x"  | 
| 52624 | 1773  | 
using assms by (metis at_within_interior netlimit_at)  | 
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
1774  | 
|
| 33175 | 1775  | 
text{* Transformation of limit. *}
 | 
1776  | 
||
1777  | 
lemma Lim_transform:  | 
|
1778  | 
fixes f g :: "'a::type \<Rightarrow> 'b::real_normed_vector"  | 
|
1779  | 
assumes "((\<lambda>x. f x - g x) ---> 0) net" "(f ---> l) net"  | 
|
1780  | 
shows "(g ---> l) net"  | 
|
| 
44252
 
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
 
huffman 
parents: 
44250 
diff
changeset
 | 
1781  | 
using tendsto_diff [OF assms(2) assms(1)] by simp  | 
| 33175 | 1782  | 
|
1783  | 
lemma Lim_transform_eventually:  | 
|
| 36667 | 1784  | 
"eventually (\<lambda>x. f x = g x) net \<Longrightarrow> (f ---> l) net \<Longrightarrow> (g ---> l) net"  | 
| 33175 | 1785  | 
apply (rule topological_tendstoI)  | 
1786  | 
apply (drule (2) topological_tendstoD)  | 
|
1787  | 
apply (erule (1) eventually_elim2, simp)  | 
|
1788  | 
done  | 
|
1789  | 
||
1790  | 
lemma Lim_transform_within:  | 
|
| 53282 | 1791  | 
assumes "0 < d"  | 
1792  | 
and "\<forall>x'\<in>S. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"  | 
|
1793  | 
and "(f ---> l) (at x within S)"  | 
|
| 36667 | 1794  | 
shows "(g ---> l) (at x within S)"  | 
1795  | 
proof (rule Lim_transform_eventually)  | 
|
1796  | 
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
 | 
1797  | 
using assms(1,2) by (auto simp: dist_nz eventually_at)  | 
| 36667 | 1798  | 
show "(f ---> l) (at x within S)" by fact  | 
1799  | 
qed  | 
|
| 33175 | 1800  | 
|
1801  | 
lemma Lim_transform_at:  | 
|
| 53282 | 1802  | 
assumes "0 < d"  | 
1803  | 
and "\<forall>x'. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"  | 
|
1804  | 
and "(f ---> l) (at x)"  | 
|
| 36667 | 1805  | 
shows "(g ---> l) (at x)"  | 
| 53282 | 1806  | 
using _ assms(3)  | 
| 36667 | 1807  | 
proof (rule Lim_transform_eventually)  | 
1808  | 
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
 | 
1809  | 
unfolding eventually_at2  | 
| 36667 | 1810  | 
using assms(1,2) by auto  | 
1811  | 
qed  | 
|
| 33175 | 1812  | 
|
1813  | 
text{* Common case assuming being away from some crucial point like 0. *}
 | 
|
1814  | 
||
1815  | 
lemma Lim_transform_away_within:  | 
|
| 36669 | 1816  | 
fixes a b :: "'a::t1_space"  | 
| 53282 | 1817  | 
assumes "a \<noteq> b"  | 
1818  | 
and "\<forall>x\<in>S. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"  | 
|
1819  | 
and "(f ---> l) (at a within S)"  | 
|
| 33175 | 1820  | 
shows "(g ---> l) (at a within S)"  | 
| 36669 | 1821  | 
proof (rule Lim_transform_eventually)  | 
1822  | 
show "(f ---> l) (at a within S)" by fact  | 
|
1823  | 
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
 | 
1824  | 
unfolding eventually_at_topological  | 
| 36669 | 1825  | 
    by (rule exI [where x="- {b}"], simp add: open_Compl assms)
 | 
| 33175 | 1826  | 
qed  | 
1827  | 
||
1828  | 
lemma Lim_transform_away_at:  | 
|
| 36669 | 1829  | 
fixes a b :: "'a::t1_space"  | 
| 52624 | 1830  | 
assumes ab: "a\<noteq>b"  | 
1831  | 
and fg: "\<forall>x. x \<noteq> a \<and> x \<noteq> b \<longrightarrow> f x = g x"  | 
|
1832  | 
and fl: "(f ---> l) (at a)"  | 
|
| 33175 | 1833  | 
shows "(g ---> l) (at a)"  | 
| 52624 | 1834  | 
using Lim_transform_away_within[OF ab, of UNIV f g l] fg fl by simp  | 
| 33175 | 1835  | 
|
1836  | 
text{* Alternatively, within an open set. *}
 | 
|
1837  | 
||
1838  | 
lemma Lim_transform_within_open:  | 
|
| 53282 | 1839  | 
assumes "open S" and "a \<in> S"  | 
1840  | 
and "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> f x = g x"  | 
|
1841  | 
and "(f ---> l) (at a)"  | 
|
| 33175 | 1842  | 
shows "(g ---> l) (at a)"  | 
| 36667 | 1843  | 
proof (rule Lim_transform_eventually)  | 
1844  | 
show "eventually (\<lambda>x. f x = g x) (at a)"  | 
|
1845  | 
unfolding eventually_at_topological  | 
|
1846  | 
using assms(1,2,3) by auto  | 
|
1847  | 
show "(f ---> l) (at a)" by fact  | 
|
| 33175 | 1848  | 
qed  | 
1849  | 
||
1850  | 
text{* A congruence rule allowing us to transform limits assuming not at point. *}
 | 
|
1851  | 
||
1852  | 
(* FIXME: Only one congruence rule for tendsto can be used at a time! *)  | 
|
1853  | 
||
| 
36362
 
06475a1547cb
fix lots of looping simp calls and other warnings
 
huffman 
parents: 
36360 
diff
changeset
 | 
1854  | 
lemma Lim_cong_within(*[cong add]*):  | 
| 53282 | 1855  | 
assumes "a = b"  | 
1856  | 
and "x = y"  | 
|
1857  | 
and "S = T"  | 
|
1858  | 
and "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x"  | 
|
| 43338 | 1859  | 
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
 | 
1860  | 
unfolding tendsto_def eventually_at_topological  | 
| 36667 | 1861  | 
using assms by simp  | 
1862  | 
||
1863  | 
lemma Lim_cong_at(*[cong add]*):  | 
|
| 43338 | 1864  | 
assumes "a = b" "x = y"  | 
| 53282 | 1865  | 
and "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"  | 
| 43338 | 1866  | 
shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"  | 
| 36667 | 1867  | 
unfolding tendsto_def eventually_at_topological  | 
1868  | 
using assms by simp  | 
|
| 33175 | 1869  | 
|
1870  | 
text{* Useful lemmas on closure and set of possible sequential limits.*}
 | 
|
1871  | 
||
1872  | 
lemma closure_sequential:  | 
|
| 50883 | 1873  | 
fixes l :: "'a::first_countable_topology"  | 
| 53291 | 1874  | 
shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially)"  | 
1875  | 
(is "?lhs = ?rhs")  | 
|
| 33175 | 1876  | 
proof  | 
| 53282 | 1877  | 
assume "?lhs"  | 
1878  | 
moreover  | 
|
1879  | 
  {
 | 
|
1880  | 
assume "l \<in> S"  | 
|
1881  | 
then have "?rhs" using tendsto_const[of l sequentially] by auto  | 
|
| 52624 | 1882  | 
}  | 
1883  | 
moreover  | 
|
| 53282 | 1884  | 
  {
 | 
1885  | 
assume "l islimpt S"  | 
|
1886  | 
then have "?rhs" unfolding islimpt_sequential by auto  | 
|
| 52624 | 1887  | 
}  | 
1888  | 
ultimately show "?rhs"  | 
|
1889  | 
unfolding closure_def by auto  | 
|
| 33175 | 1890  | 
next  | 
1891  | 
assume "?rhs"  | 
|
| 53282 | 1892  | 
then show "?lhs" unfolding closure_def islimpt_sequential by auto  | 
| 33175 | 1893  | 
qed  | 
1894  | 
||
1895  | 
lemma closed_sequential_limits:  | 
|
| 50883 | 1896  | 
fixes S :: "'a::first_countable_topology set"  | 
| 33175 | 1897  | 
shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x ---> l) sequentially \<longrightarrow> l \<in> S)"  | 
| 52624 | 1898  | 
using closure_sequential [where 'a='a] closure_closed [where 'a='a]  | 
1899  | 
closed_limpt [where 'a='a] islimpt_sequential [where 'a='a] mem_delete [where 'a='a]  | 
|
| 33175 | 1900  | 
by metis  | 
1901  | 
||
1902  | 
lemma closure_approachable:  | 
|
1903  | 
fixes S :: "'a::metric_space set"  | 
|
1904  | 
shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)"  | 
|
1905  | 
apply (auto simp add: closure_def islimpt_approachable)  | 
|
| 52624 | 1906  | 
apply (metis dist_self)  | 
1907  | 
done  | 
|
| 33175 | 1908  | 
|
1909  | 
lemma closed_approachable:  | 
|
1910  | 
fixes S :: "'a::metric_space set"  | 
|
| 53291 | 1911  | 
shows "closed S \<Longrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S"  | 
| 33175 | 1912  | 
by (metis closure_closed closure_approachable)  | 
1913  | 
||
| 51351 | 1914  | 
lemma closure_contains_Inf:  | 
1915  | 
fixes S :: "real set"  | 
|
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1916  | 
  assumes "S \<noteq> {}" "bdd_below S"
 | 
| 51351 | 1917  | 
shows "Inf S \<in> closure S"  | 
| 52624 | 1918  | 
proof -  | 
| 51351 | 1919  | 
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
 | 
1920  | 
using cInf_lower[of _ S] assms by metis  | 
| 52624 | 1921  | 
  {
 | 
| 53282 | 1922  | 
fix e :: real  | 
1923  | 
assume "e > 0"  | 
|
| 52624 | 1924  | 
then have "Inf S < Inf S + e" by simp  | 
1925  | 
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
 | 
1926  | 
by (subst (asm) cInf_less_iff) auto  | 
| 52624 | 1927  | 
with * have "\<exists>x\<in>S. dist x (Inf S) < e"  | 
1928  | 
by (intro bexI[of _ x]) (auto simp add: dist_real_def)  | 
|
1929  | 
}  | 
|
1930  | 
then show ?thesis unfolding closure_approachable by auto  | 
|
| 51351 | 1931  | 
qed  | 
1932  | 
||
1933  | 
lemma closed_contains_Inf:  | 
|
1934  | 
fixes S :: "real set"  | 
|
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1935  | 
  shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
 | 
| 51351 | 1936  | 
by (metis closure_contains_Inf closure_closed assms)  | 
1937  | 
||
1938  | 
lemma not_trivial_limit_within_ball:  | 
|
| 53640 | 1939  | 
  "\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
 | 
| 51351 | 1940  | 
(is "?lhs = ?rhs")  | 
1941  | 
proof -  | 
|
| 53282 | 1942  | 
  {
 | 
1943  | 
assume "?lhs"  | 
|
1944  | 
    {
 | 
|
1945  | 
fix e :: real  | 
|
1946  | 
assume "e > 0"  | 
|
| 53640 | 1947  | 
      then obtain y where "y \<in> S - {x}" and "dist y x < e"
 | 
| 51351 | 1948  | 
        using `?lhs` not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
 | 
1949  | 
by auto  | 
|
| 53640 | 1950  | 
      then have "y \<in> S \<inter> ball x e - {x}"
 | 
| 51351 | 1951  | 
unfolding ball_def by (simp add: dist_commute)  | 
| 53640 | 1952  | 
      then have "S \<inter> ball x e - {x} \<noteq> {}" by blast
 | 
| 52624 | 1953  | 
}  | 
1954  | 
then have "?rhs" by auto  | 
|
| 51351 | 1955  | 
}  | 
1956  | 
moreover  | 
|
| 53282 | 1957  | 
  {
 | 
1958  | 
assume "?rhs"  | 
|
1959  | 
    {
 | 
|
1960  | 
fix e :: real  | 
|
1961  | 
assume "e > 0"  | 
|
| 53640 | 1962  | 
      then obtain y where "y \<in> S \<inter> ball x e - {x}"
 | 
| 53282 | 1963  | 
using `?rhs` by blast  | 
| 53640 | 1964  | 
      then have "y \<in> S - {x}" and "dist y x < e"
 | 
1965  | 
unfolding ball_def by (simp_all add: dist_commute)  | 
|
1966  | 
      then have "\<exists>y \<in> S - {x}. dist y x < e"
 | 
|
| 53282 | 1967  | 
by auto  | 
| 51351 | 1968  | 
}  | 
1969  | 
then have "?lhs"  | 
|
| 53282 | 1970  | 
      using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
 | 
1971  | 
by auto  | 
|
| 51351 | 1972  | 
}  | 
1973  | 
ultimately show ?thesis by auto  | 
|
1974  | 
qed  | 
|
1975  | 
||
| 52624 | 1976  | 
|
| 50087 | 1977  | 
subsection {* Infimum Distance *}
 | 
1978  | 
||
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
1979  | 
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
 | 
1980  | 
|
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
1981  | 
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
 | 
1982  | 
by (auto intro!: zero_le_dist)  | 
| 
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1983  | 
|
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
1984  | 
lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = (INF a:A. dist x a)"
 | 
| 50087 | 1985  | 
by (simp add: infdist_def)  | 
1986  | 
||
| 52624 | 1987  | 
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
 | 
1988  | 
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
 | 
1989  | 
|
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
1990  | 
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
 | 
1991  | 
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
 | 
1992  | 
|
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
1993  | 
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
 | 
1994  | 
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
 | 
1995  | 
|
| 
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
1996  | 
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
 | 
1997  | 
by (auto intro!: antisym infdist_nonneg infdist_le2)  | 
| 50087 | 1998  | 
|
| 52624 | 1999  | 
lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y"  | 
| 53640 | 2000  | 
proof (cases "A = {}")
 | 
2001  | 
case True  | 
|
| 53282 | 2002  | 
then show ?thesis by (simp add: infdist_def)  | 
| 50087 | 2003  | 
next  | 
| 53640 | 2004  | 
case False  | 
| 52624 | 2005  | 
then obtain a where "a \<in> A" by auto  | 
| 50087 | 2006  | 
  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
 | 
2007  | 
proof (rule cInf_greatest)  | 
| 53282 | 2008  | 
    from `A \<noteq> {}` show "{dist x y + dist y a |a. a \<in> A} \<noteq> {}"
 | 
2009  | 
by simp  | 
|
2010  | 
fix d  | 
|
2011  | 
    assume "d \<in> {dist x y + dist y a |a. a \<in> A}"
 | 
|
2012  | 
then obtain a where d: "d = dist x y + dist y a" "a \<in> A"  | 
|
2013  | 
by auto  | 
|
| 50087 | 2014  | 
show "infdist x A \<le> d"  | 
2015  | 
      unfolding infdist_notempty[OF `A \<noteq> {}`]
 | 
|
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
2016  | 
proof (rule cINF_lower2)  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
2017  | 
show "a \<in> A" by fact  | 
| 53282 | 2018  | 
show "dist x a \<le> d"  | 
2019  | 
unfolding d by (rule dist_triangle)  | 
|
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2020  | 
qed simp  | 
| 50087 | 2021  | 
qed  | 
2022  | 
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
 | 
2023  | 
proof (rule cInf_eq, safe)  | 
| 53282 | 2024  | 
fix a  | 
2025  | 
assume "a \<in> A"  | 
|
2026  | 
then show "dist x y + infdist y A \<le> dist x y + dist y a"  | 
|
2027  | 
by (auto intro: infdist_le)  | 
|
| 50087 | 2028  | 
next  | 
| 53282 | 2029  | 
fix i  | 
2030  | 
    assume inf: "\<And>d. d \<in> {dist x y + dist y a |a. a \<in> A} \<Longrightarrow> i \<le> d"
 | 
|
2031  | 
then have "i - dist x y \<le> infdist y A"  | 
|
2032  | 
      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
 | 
2033  | 
by (intro cINF_greatest) (auto simp: field_simps)  | 
| 53282 | 2034  | 
then show "i \<le> dist x y + infdist y A"  | 
2035  | 
by simp  | 
|
| 50087 | 2036  | 
qed  | 
2037  | 
finally show ?thesis by simp  | 
|
2038  | 
qed  | 
|
2039  | 
||
| 
51475
 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 
hoelzl 
parents: 
51473 
diff
changeset
 | 
2040  | 
lemma in_closure_iff_infdist_zero:  | 
| 50087 | 2041  | 
  assumes "A \<noteq> {}"
 | 
2042  | 
shows "x \<in> closure A \<longleftrightarrow> infdist x A = 0"  | 
|
2043  | 
proof  | 
|
2044  | 
assume "x \<in> closure A"  | 
|
2045  | 
show "infdist x A = 0"  | 
|
2046  | 
proof (rule ccontr)  | 
|
2047  | 
assume "infdist x A \<noteq> 0"  | 
|
| 53282 | 2048  | 
with infdist_nonneg[of x A] have "infdist x A > 0"  | 
2049  | 
by auto  | 
|
2050  | 
    then have "ball x (infdist x A) \<inter> closure A = {}"
 | 
|
| 52624 | 2051  | 
apply auto  | 
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
2052  | 
apply (metis `x \<in> closure A` closure_approachable dist_commute infdist_le not_less)  | 
| 52624 | 2053  | 
done  | 
| 53282 | 2054  | 
then have "x \<notin> closure A"  | 
| 52624 | 2055  | 
by (metis `0 < infdist x A` centre_in_ball disjoint_iff_not_equal)  | 
| 53282 | 2056  | 
then show False using `x \<in> closure A` by simp  | 
| 50087 | 2057  | 
qed  | 
2058  | 
next  | 
|
2059  | 
assume x: "infdist x A = 0"  | 
|
| 53282 | 2060  | 
then obtain a where "a \<in> A"  | 
2061  | 
by atomize_elim (metis all_not_in_conv assms)  | 
|
2062  | 
show "x \<in> closure A"  | 
|
2063  | 
unfolding closure_approachable  | 
|
2064  | 
apply safe  | 
|
2065  | 
proof (rule ccontr)  | 
|
2066  | 
fix e :: real  | 
|
2067  | 
assume "e > 0"  | 
|
| 50087 | 2068  | 
assume "\<not> (\<exists>y\<in>A. dist y x < e)"  | 
| 53282 | 2069  | 
then have "infdist x A \<ge> e" using `a \<in> A`  | 
| 50087 | 2070  | 
unfolding infdist_def  | 
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
2071  | 
by (force simp: dist_commute intro: cINF_greatest)  | 
| 53282 | 2072  | 
with x `e > 0` show False by auto  | 
| 50087 | 2073  | 
qed  | 
2074  | 
qed  | 
|
2075  | 
||
| 
51475
 
ebf9d4fd00ba
introduct the conditional_complete_lattice type class; generalize theorems about real Sup and Inf to it
 
hoelzl 
parents: 
51473 
diff
changeset
 | 
2076  | 
lemma in_closed_iff_infdist_zero:  | 
| 50087 | 2077  | 
  assumes "closed A" "A \<noteq> {}"
 | 
2078  | 
shows "x \<in> A \<longleftrightarrow> infdist x A = 0"  | 
|
2079  | 
proof -  | 
|
2080  | 
have "x \<in> closure A \<longleftrightarrow> infdist x A = 0"  | 
|
2081  | 
by (rule in_closure_iff_infdist_zero) fact  | 
|
2082  | 
with assms show ?thesis by simp  | 
|
2083  | 
qed  | 
|
2084  | 
||
2085  | 
lemma tendsto_infdist [tendsto_intros]:  | 
|
2086  | 
assumes f: "(f ---> l) F"  | 
|
2087  | 
shows "((\<lambda>x. infdist (f x) A) ---> infdist l A) F"  | 
|
2088  | 
proof (rule tendstoI)  | 
|
| 53282 | 2089  | 
fix e ::real  | 
2090  | 
assume "e > 0"  | 
|
| 50087 | 2091  | 
from tendstoD[OF f this]  | 
2092  | 
show "eventually (\<lambda>x. dist (infdist (f x) A) (infdist l A) < e) F"  | 
|
2093  | 
proof (eventually_elim)  | 
|
2094  | 
fix x  | 
|
2095  | 
from infdist_triangle[of l A "f x"] infdist_triangle[of "f x" A l]  | 
|
2096  | 
have "dist (infdist (f x) A) (infdist l A) \<le> dist (f x) l"  | 
|
2097  | 
by (simp add: dist_commute dist_real_def)  | 
|
2098  | 
also assume "dist (f x) l < e"  | 
|
2099  | 
finally show "dist (infdist (f x) A) (infdist l A) < e" .  | 
|
2100  | 
qed  | 
|
2101  | 
qed  | 
|
2102  | 
||
| 33175 | 2103  | 
text{* Some other lemmas about sequences. *}
 | 
2104  | 
||
| 53597 | 2105  | 
lemma sequentially_offset: (* TODO: move to Topological_Spaces.thy *)  | 
| 36441 | 2106  | 
assumes "eventually (\<lambda>i. P i) sequentially"  | 
2107  | 
shows "eventually (\<lambda>i. P (i + k)) sequentially"  | 
|
| 53597 | 2108  | 
using assms by (rule eventually_sequentially_seg [THEN iffD2])  | 
2109  | 
||
2110  | 
lemma seq_offset_neg: (* TODO: move to Topological_Spaces.thy *)  | 
|
| 53291 | 2111  | 
"(f ---> l) sequentially \<Longrightarrow> ((\<lambda>i. f(i - k)) ---> l) sequentially"  | 
| 53597 | 2112  | 
apply (erule filterlim_compose)  | 
2113  | 
apply (simp add: filterlim_def le_sequentially eventually_filtermap eventually_sequentially)  | 
|
| 52624 | 2114  | 
apply arith  | 
2115  | 
done  | 
|
| 33175 | 2116  | 
|
2117  | 
lemma seq_harmonic: "((\<lambda>n. inverse (real n)) ---> 0) sequentially"  | 
|
| 53597 | 2118  | 
using LIMSEQ_inverse_real_of_nat by (rule LIMSEQ_imp_Suc) (* TODO: move to Limits.thy *)  | 
| 33175 | 2119  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
2120  | 
subsection {* More properties of closed balls *}
 | 
| 33175 | 2121  | 
|
| 54070 | 2122  | 
lemma closed_vimage: (* TODO: move to Topological_Spaces.thy *)  | 
2123  | 
assumes "closed s" and "continuous_on UNIV f"  | 
|
2124  | 
shows "closed (vimage f s)"  | 
|
2125  | 
using assms unfolding continuous_on_closed_vimage [OF closed_UNIV]  | 
|
2126  | 
by simp  | 
|
2127  | 
||
| 33175 | 2128  | 
lemma closed_cball: "closed (cball x e)"  | 
| 54070 | 2129  | 
proof -  | 
2130  | 
  have "closed (dist x -` {..e})"
 | 
|
2131  | 
by (intro closed_vimage closed_atMost continuous_on_intros)  | 
|
2132  | 
  also have "dist x -` {..e} = cball x e"
 | 
|
2133  | 
by auto  | 
|
2134  | 
finally show ?thesis .  | 
|
2135  | 
qed  | 
|
| 33175 | 2136  | 
|
2137  | 
lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. cball x e \<subseteq> S)"  | 
|
| 52624 | 2138  | 
proof -  | 
2139  | 
  {
 | 
|
2140  | 
fix x and e::real  | 
|
2141  | 
assume "x\<in>S" "e>0" "ball x e \<subseteq> S"  | 
|
| 53282 | 2142  | 
then have "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto)  | 
| 52624 | 2143  | 
}  | 
2144  | 
moreover  | 
|
2145  | 
  {
 | 
|
2146  | 
fix x and e::real  | 
|
2147  | 
assume "x\<in>S" "e>0" "cball x e \<subseteq> S"  | 
|
| 53282 | 2148  | 
then have "\<exists>d>0. ball x d \<subseteq> S"  | 
| 52624 | 2149  | 
unfolding subset_eq  | 
2150  | 
apply(rule_tac x="e/2" in exI)  | 
|
2151  | 
apply auto  | 
|
2152  | 
done  | 
|
2153  | 
}  | 
|
2154  | 
ultimately show ?thesis  | 
|
2155  | 
unfolding open_contains_ball by auto  | 
|
| 33175 | 2156  | 
qed  | 
2157  | 
||
| 53291 | 2158  | 
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
 | 
2159  | 
by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball)  | 
| 33175 | 2160  | 
|
2161  | 
lemma mem_interior_cball: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S)"  | 
|
2162  | 
apply (simp add: interior_def, safe)  | 
|
2163  | 
apply (force simp add: open_contains_cball)  | 
|
2164  | 
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
 | 
2165  | 
apply (simp add: subset_trans [OF ball_subset_cball])  | 
| 33175 | 2166  | 
done  | 
2167  | 
||
2168  | 
lemma islimpt_ball:  | 
|
2169  | 
  fixes x y :: "'a::{real_normed_vector,perfect_space}"
 | 
|
| 53291 | 2170  | 
shows "y islimpt ball x e \<longleftrightarrow> 0 < e \<and> y \<in> cball x e"  | 
2171  | 
(is "?lhs = ?rhs")  | 
|
| 33175 | 2172  | 
proof  | 
2173  | 
assume "?lhs"  | 
|
| 53282 | 2174  | 
  {
 | 
2175  | 
assume "e \<le> 0"  | 
|
2176  | 
    then have *:"ball x e = {}"
 | 
|
2177  | 
using ball_eq_empty[of x e] by auto  | 
|
2178  | 
have False using `?lhs`  | 
|
2179  | 
unfolding * using islimpt_EMPTY[of y] by auto  | 
|
| 33175 | 2180  | 
}  | 
| 53282 | 2181  | 
then have "e > 0" by (metis not_less)  | 
| 33175 | 2182  | 
moreover  | 
| 52624 | 2183  | 
have "y \<in> cball x e"  | 
2184  | 
using closed_cball[of x e] islimpt_subset[of y "ball x e" "cball x e"]  | 
|
2185  | 
ball_subset_cball[of x e] `?lhs`  | 
|
2186  | 
unfolding closed_limpt by auto  | 
|
| 33175 | 2187  | 
ultimately show "?rhs" by auto  | 
2188  | 
next  | 
|
| 53282 | 2189  | 
assume "?rhs"  | 
| 53640 | 2190  | 
then have "e > 0" by auto  | 
| 53282 | 2191  | 
  {
 | 
2192  | 
fix d :: real  | 
|
2193  | 
assume "d > 0"  | 
|
| 33175 | 2194  | 
have "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
| 53282 | 2195  | 
proof (cases "d \<le> dist x y")  | 
2196  | 
case True  | 
|
2197  | 
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
|
2198  | 
proof (cases "x = y")  | 
|
2199  | 
case True  | 
|
2200  | 
then have False  | 
|
2201  | 
using `d \<le> dist x y` `d>0` by auto  | 
|
2202  | 
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
|
2203  | 
by auto  | 
|
| 33175 | 2204  | 
next  | 
2205  | 
case False  | 
|
| 53282 | 2206  | 
have "dist x (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) =  | 
2207  | 
norm (x - y + (d / (2 * norm (y - x))) *\<^sub>R (y - x))"  | 
|
| 53291 | 2208  | 
unfolding mem_cball mem_ball dist_norm diff_diff_eq2 diff_add_eq[symmetric]  | 
| 53282 | 2209  | 
by auto  | 
| 33175 | 2210  | 
also have "\<dots> = \<bar>- 1 + d / (2 * norm (x - y))\<bar> * norm (x - y)"  | 
| 53291 | 2211  | 
using scaleR_left_distrib[of "- 1" "d / (2 * norm (y - x))", symmetric, of "y - x"]  | 
| 33175 | 2212  | 
unfolding scaleR_minus_left scaleR_one  | 
2213  | 
by (auto simp add: norm_minus_commute)  | 
|
2214  | 
also have "\<dots> = \<bar>- norm (x - y) + d / 2\<bar>"  | 
|
2215  | 
unfolding abs_mult_pos[of "norm (x - y)", OF norm_ge_zero[of "x - y"]]  | 
|
| 53282 | 2216  | 
unfolding distrib_right using `x\<noteq>y`[unfolded dist_nz, unfolded dist_norm]  | 
2217  | 
by auto  | 
|
2218  | 
also have "\<dots> \<le> e - d/2" using `d \<le> dist x y` and `d>0` and `?rhs`  | 
|
2219  | 
by (auto simp add: dist_norm)  | 
|
2220  | 
finally have "y - (d / (2 * dist y x)) *\<^sub>R (y - x) \<in> ball x e" using `d>0`  | 
|
2221  | 
by auto  | 
|
| 33175 | 2222  | 
moreover  | 
2223  | 
have "(d / (2*dist y x)) *\<^sub>R (y - x) \<noteq> 0"  | 
|
| 53282 | 2224  | 
using `x\<noteq>y`[unfolded dist_nz] `d>0` unfolding scaleR_eq_0_iff  | 
2225  | 
by (auto simp add: dist_commute)  | 
|
| 33175 | 2226  | 
moreover  | 
| 53282 | 2227  | 
have "dist (y - (d / (2 * dist y x)) *\<^sub>R (y - x)) y < d"  | 
2228  | 
unfolding dist_norm  | 
|
2229  | 
apply simp  | 
|
2230  | 
unfolding norm_minus_cancel  | 
|
2231  | 
using `d > 0` `x\<noteq>y`[unfolded dist_nz] dist_commute[of x y]  | 
|
2232  | 
unfolding dist_norm  | 
|
2233  | 
apply auto  | 
|
2234  | 
done  | 
|
2235  | 
ultimately show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
|
2236  | 
apply (rule_tac x = "y - (d / (2*dist y x)) *\<^sub>R (y - x)" in bexI)  | 
|
2237  | 
apply auto  | 
|
2238  | 
done  | 
|
| 33175 | 2239  | 
qed  | 
2240  | 
next  | 
|
| 53282 | 2241  | 
case False  | 
2242  | 
then have "d > dist x y" by auto  | 
|
2243  | 
show "\<exists>x' \<in> ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
|
2244  | 
proof (cases "x = y")  | 
|
| 33175 | 2245  | 
case True  | 
2246  | 
obtain z where **: "z \<noteq> y" "dist z y < min e d"  | 
|
2247  | 
using perfect_choose_dist[of "min e d" y]  | 
|
2248  | 
using `d > 0` `e>0` by auto  | 
|
2249  | 
show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
|
2250  | 
unfolding `x = y`  | 
|
2251  | 
using `z \<noteq> y` **  | 
|
| 53282 | 2252  | 
apply (rule_tac x=z in bexI)  | 
2253  | 
apply (auto simp add: dist_commute)  | 
|
2254  | 
done  | 
|
| 33175 | 2255  | 
next  | 
| 53282 | 2256  | 
case False  | 
2257  | 
then show "\<exists>x'\<in>ball x e. x' \<noteq> y \<and> dist x' y < d"  | 
|
2258  | 
using `d>0` `d > dist x y` `?rhs`  | 
|
2259  | 
apply (rule_tac x=x in bexI)  | 
|
2260  | 
apply auto  | 
|
2261  | 
done  | 
|
| 33175 | 2262  | 
qed  | 
| 53282 | 2263  | 
qed  | 
2264  | 
}  | 
|
2265  | 
then show "?lhs"  | 
|
2266  | 
unfolding mem_cball islimpt_approachable mem_ball by auto  | 
|
| 33175 | 2267  | 
qed  | 
2268  | 
||
2269  | 
lemma closure_ball_lemma:  | 
|
2270  | 
fixes x y :: "'a::real_normed_vector"  | 
|
| 53282 | 2271  | 
assumes "x \<noteq> y"  | 
2272  | 
shows "y islimpt ball x (dist x y)"  | 
|
| 33175 | 2273  | 
proof (rule islimptI)  | 
| 53282 | 2274  | 
fix T  | 
2275  | 
assume "y \<in> T" "open T"  | 
|
| 33175 | 2276  | 
then obtain r where "0 < r" "\<forall>z. dist z y < r \<longrightarrow> z \<in> T"  | 
2277  | 
unfolding open_dist by fast  | 
|
2278  | 
(* choose point between x and y, within distance r of y. *)  | 
|
2279  | 
def k \<equiv> "min 1 (r / (2 * dist x y))"  | 
|
2280  | 
def z \<equiv> "y + scaleR k (x - y)"  | 
|
2281  | 
have z_def2: "z = x + scaleR (1 - k) (y - x)"  | 
|
2282  | 
unfolding z_def by (simp add: algebra_simps)  | 
|
2283  | 
have "dist z y < r"  | 
|
2284  | 
unfolding z_def k_def using `0 < r`  | 
|
2285  | 
by (simp add: dist_norm min_def)  | 
|
| 53282 | 2286  | 
then have "z \<in> T"  | 
2287  | 
using `\<forall>z. dist z y < r \<longrightarrow> z \<in> T` by simp  | 
|
| 33175 | 2288  | 
have "dist x z < dist x y"  | 
2289  | 
unfolding z_def2 dist_norm  | 
|
2290  | 
apply (simp add: norm_minus_commute)  | 
|
2291  | 
apply (simp only: dist_norm [symmetric])  | 
|
2292  | 
apply (subgoal_tac "\<bar>1 - k\<bar> * dist x y < 1 * dist x y", simp)  | 
|
2293  | 
apply (rule mult_strict_right_mono)  | 
|
2294  | 
apply (simp add: k_def divide_pos_pos zero_less_dist_iff `0 < r` `x \<noteq> y`)  | 
|
2295  | 
apply (simp add: zero_less_dist_iff `x \<noteq> y`)  | 
|
2296  | 
done  | 
|
| 53282 | 2297  | 
then have "z \<in> ball x (dist x y)"  | 
2298  | 
by simp  | 
|
| 33175 | 2299  | 
have "z \<noteq> y"  | 
2300  | 
unfolding z_def k_def using `x \<noteq> y` `0 < r`  | 
|
2301  | 
by (simp add: min_def)  | 
|
2302  | 
show "\<exists>z\<in>ball x (dist x y). z \<in> T \<and> z \<noteq> y"  | 
|
2303  | 
using `z \<in> ball x (dist x y)` `z \<in> T` `z \<noteq> y`  | 
|
2304  | 
by fast  | 
|
2305  | 
qed  | 
|
2306  | 
||
2307  | 
lemma closure_ball:  | 
|
2308  | 
fixes x :: "'a::real_normed_vector"  | 
|
2309  | 
shows "0 < e \<Longrightarrow> closure (ball x e) = cball x e"  | 
|
| 52624 | 2310  | 
apply (rule equalityI)  | 
2311  | 
apply (rule closure_minimal)  | 
|
2312  | 
apply (rule ball_subset_cball)  | 
|
2313  | 
apply (rule closed_cball)  | 
|
2314  | 
apply (rule subsetI, rename_tac y)  | 
|
2315  | 
apply (simp add: le_less [where 'a=real])  | 
|
2316  | 
apply (erule disjE)  | 
|
2317  | 
apply (rule subsetD [OF closure_subset], simp)  | 
|
2318  | 
apply (simp add: closure_def)  | 
|
2319  | 
apply clarify  | 
|
2320  | 
apply (rule closure_ball_lemma)  | 
|
2321  | 
apply (simp add: zero_less_dist_iff)  | 
|
2322  | 
done  | 
|
| 33175 | 2323  | 
|
2324  | 
(* In a trivial vector space, this fails for e = 0. *)  | 
|
2325  | 
lemma interior_cball:  | 
|
2326  | 
  fixes x :: "'a::{real_normed_vector, perfect_space}"
 | 
|
2327  | 
shows "interior (cball x e) = ball x e"  | 
|
| 53640 | 2328  | 
proof (cases "e \<ge> 0")  | 
| 33175 | 2329  | 
case False note cs = this  | 
| 53282 | 2330  | 
  from cs have "ball x e = {}"
 | 
2331  | 
using ball_empty[of e x] by auto  | 
|
2332  | 
moreover  | 
|
2333  | 
  {
 | 
|
2334  | 
fix y  | 
|
2335  | 
assume "y \<in> cball x e"  | 
|
2336  | 
then have False  | 
|
2337  | 
unfolding mem_cball using dist_nz[of x y] cs by auto  | 
|
2338  | 
}  | 
|
2339  | 
  then have "cball x e = {}" by auto
 | 
|
2340  | 
  then have "interior (cball x e) = {}"
 | 
|
2341  | 
using interior_empty by auto  | 
|
| 33175 | 2342  | 
ultimately show ?thesis by blast  | 
2343  | 
next  | 
|
2344  | 
case True note cs = this  | 
|
| 53282 | 2345  | 
have "ball x e \<subseteq> cball x e"  | 
2346  | 
using ball_subset_cball by auto  | 
|
2347  | 
moreover  | 
|
2348  | 
  {
 | 
|
2349  | 
fix S y  | 
|
2350  | 
assume as: "S \<subseteq> cball x e" "open S" "y\<in>S"  | 
|
2351  | 
then obtain d where "d>0" and d: "\<forall>x'. dist x' y < d \<longrightarrow> x' \<in> S"  | 
|
2352  | 
unfolding open_dist by blast  | 
|
| 33175 | 2353  | 
then obtain xa where xa_y: "xa \<noteq> y" and xa: "dist xa y < d"  | 
2354  | 
using perfect_choose_dist [of d] by auto  | 
|
| 53282 | 2355  | 
have "xa \<in> S"  | 
2356  | 
using d[THEN spec[where x = xa]]  | 
|
2357  | 
using xa by (auto simp add: dist_commute)  | 
|
2358  | 
then have xa_cball: "xa \<in> cball x e"  | 
|
2359  | 
using as(1) by auto  | 
|
2360  | 
then have "y \<in> ball x e"  | 
|
2361  | 
proof (cases "x = y")  | 
|
| 33175 | 2362  | 
case True  | 
| 53282 | 2363  | 
then have "e > 0"  | 
2364  | 
using xa_y[unfolded dist_nz] xa_cball[unfolded mem_cball]  | 
|
| 52624 | 2365  | 
by (auto simp add: dist_commute)  | 
| 53282 | 2366  | 
then show "y \<in> ball x e"  | 
2367  | 
using `x = y ` by simp  | 
|
| 33175 | 2368  | 
next  | 
2369  | 
case False  | 
|
| 53282 | 2370  | 
have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) y < d"  | 
2371  | 
unfolding dist_norm  | 
|
| 33175 | 2372  | 
using `d>0` norm_ge_zero[of "y - x"] `x \<noteq> y` by auto  | 
| 53282 | 2373  | 
then have *: "y + (d / 2 / dist y x) *\<^sub>R (y - x) \<in> cball x e"  | 
| 52624 | 2374  | 
using d as(1)[unfolded subset_eq] by blast  | 
| 33175 | 2375  | 
have "y - x \<noteq> 0" using `x \<noteq> y` by auto  | 
| 53282 | 2376  | 
then have **:"d / (2 * norm (y - x)) > 0"  | 
| 53291 | 2377  | 
unfolding zero_less_norm_iff[symmetric]  | 
| 33175 | 2378  | 
using `d>0` divide_pos_pos[of d "2*norm (y - x)"] by auto  | 
| 53282 | 2379  | 
have "dist (y + (d / 2 / dist y x) *\<^sub>R (y - x)) x =  | 
2380  | 
norm (y + (d / (2 * norm (y - x))) *\<^sub>R y - (d / (2 * norm (y - x))) *\<^sub>R x - x)"  | 
|
| 33175 | 2381  | 
by (auto simp add: dist_norm algebra_simps)  | 
2382  | 
also have "\<dots> = norm ((1 + d / (2 * norm (y - x))) *\<^sub>R (y - x))"  | 
|
2383  | 
by (auto simp add: algebra_simps)  | 
|
2384  | 
also have "\<dots> = \<bar>1 + d / (2 * norm (y - x))\<bar> * norm (y - x)"  | 
|
2385  | 
using ** by auto  | 
|
| 53282 | 2386  | 
also have "\<dots> = (dist y x) + d/2"  | 
2387  | 
using ** by (auto simp add: distrib_right dist_norm)  | 
|
2388  | 
finally have "e \<ge> dist x y +d/2"  | 
|
2389  | 
using *[unfolded mem_cball] by (auto simp add: dist_commute)  | 
|
2390  | 
then show "y \<in> ball x e"  | 
|
2391  | 
unfolding mem_ball using `d>0` by auto  | 
|
| 52624 | 2392  | 
qed  | 
2393  | 
}  | 
|
| 53282 | 2394  | 
then have "\<forall>S \<subseteq> cball x e. open S \<longrightarrow> S \<subseteq> ball x e"  | 
2395  | 
by auto  | 
|
| 52624 | 2396  | 
ultimately show ?thesis  | 
| 53640 | 2397  | 
using interior_unique[of "ball x e" "cball x e"]  | 
2398  | 
using open_ball[of x e]  | 
|
2399  | 
by auto  | 
|
| 33175 | 2400  | 
qed  | 
2401  | 
||
2402  | 
lemma frontier_ball:  | 
|
2403  | 
fixes a :: "'a::real_normed_vector"  | 
|
| 53291 | 2404  | 
  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
 | 
2405  | 
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
 | 
2406  | 
apply (simp add: set_eq_iff)  | 
| 52624 | 2407  | 
apply arith  | 
2408  | 
done  | 
|
| 33175 | 2409  | 
|
2410  | 
lemma frontier_cball:  | 
|
2411  | 
  fixes a :: "'a::{real_normed_vector, perfect_space}"
 | 
|
| 53640 | 2412  | 
  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
 | 
2413  | 
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
 | 
2414  | 
apply (simp add: set_eq_iff)  | 
| 52624 | 2415  | 
apply arith  | 
2416  | 
done  | 
|
| 33175 | 2417  | 
|
| 53640 | 2418  | 
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
 | 
2419  | 
apply (simp add: set_eq_iff not_le)  | 
| 52624 | 2420  | 
apply (metis zero_le_dist dist_self order_less_le_trans)  | 
2421  | 
done  | 
|
2422  | 
||
| 53282 | 2423  | 
lemma cball_empty: "e < 0 \<Longrightarrow> cball x e = {}"
 | 
| 52624 | 2424  | 
by (simp add: cball_eq_empty)  | 
| 33175 | 2425  | 
|
2426  | 
lemma cball_eq_sing:  | 
|
| 
44072
 
5b970711fb39
class perfect_space inherits from topological_space;
 
huffman 
parents: 
43338 
diff
changeset
 | 
2427  | 
  fixes x :: "'a::{metric_space,perfect_space}"
 | 
| 53640 | 2428  | 
  shows "cball x e = {x} \<longleftrightarrow> e = 0"
 | 
| 33175 | 2429  | 
proof (rule linorder_cases)  | 
2430  | 
assume e: "0 < e"  | 
|
2431  | 
obtain a where "a \<noteq> x" "dist a x < e"  | 
|
2432  | 
using perfect_choose_dist [OF e] by auto  | 
|
| 53282 | 2433  | 
then have "a \<noteq> x" "dist x a \<le> e"  | 
2434  | 
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
 | 
2435  | 
with e show ?thesis by (auto simp add: set_eq_iff)  | 
| 33175 | 2436  | 
qed auto  | 
2437  | 
||
2438  | 
lemma cball_sing:  | 
|
2439  | 
fixes x :: "'a::metric_space"  | 
|
| 53291 | 2440  | 
  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
 | 
2441  | 
by (auto simp add: set_eq_iff)  | 
| 33175 | 2442  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
2443  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
2444  | 
subsection {* Boundedness *}
 | 
| 33175 | 2445  | 
|
2446  | 
(* FIXME: This has to be unified with BSEQ!! *)  | 
|
| 52624 | 2447  | 
definition (in metric_space) bounded :: "'a set \<Rightarrow> bool"  | 
2448  | 
where "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)"  | 
|
| 33175 | 2449  | 
|
| 50998 | 2450  | 
lemma bounded_subset_cball: "bounded S \<longleftrightarrow> (\<exists>e x. S \<subseteq> cball x e)"  | 
2451  | 
unfolding bounded_def subset_eq by auto  | 
|
2452  | 
||
| 33175 | 2453  | 
lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)"  | 
| 52624 | 2454  | 
unfolding bounded_def  | 
2455  | 
apply safe  | 
|
| 53640 | 2456  | 
apply (rule_tac x="dist a x + e" in exI)  | 
2457  | 
apply clarify  | 
|
| 52624 | 2458  | 
apply (drule (1) bspec)  | 
2459  | 
apply (erule order_trans [OF dist_triangle add_left_mono])  | 
|
2460  | 
apply auto  | 
|
2461  | 
done  | 
|
| 33175 | 2462  | 
|
2463  | 
lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)"  | 
|
| 52624 | 2464  | 
unfolding bounded_any_center [where a=0]  | 
2465  | 
by (simp add: dist_norm)  | 
|
| 33175 | 2466  | 
|
| 53282 | 2467  | 
lemma bounded_realI:  | 
2468  | 
assumes "\<forall>x\<in>s. abs (x::real) \<le> B"  | 
|
2469  | 
shows "bounded s"  | 
|
2470  | 
unfolding bounded_def dist_real_def  | 
|
2471  | 
apply (rule_tac x=0 in exI)  | 
|
2472  | 
using assms  | 
|
2473  | 
apply auto  | 
|
2474  | 
done  | 
|
| 50104 | 2475  | 
|
| 50948 | 2476  | 
lemma bounded_empty [simp]: "bounded {}"
 | 
2477  | 
by (simp add: bounded_def)  | 
|
2478  | 
||
| 53291 | 2479  | 
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> bounded S"  | 
| 33175 | 2480  | 
by (metis bounded_def subset_eq)  | 
2481  | 
||
| 53291 | 2482  | 
lemma bounded_interior[intro]: "bounded S \<Longrightarrow> bounded(interior S)"  | 
| 33175 | 2483  | 
by (metis bounded_subset interior_subset)  | 
2484  | 
||
| 52624 | 2485  | 
lemma bounded_closure[intro]:  | 
2486  | 
assumes "bounded S"  | 
|
2487  | 
shows "bounded (closure S)"  | 
|
2488  | 
proof -  | 
|
2489  | 
from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a"  | 
|
2490  | 
unfolding bounded_def by auto  | 
|
2491  | 
  {
 | 
|
2492  | 
fix y  | 
|
2493  | 
assume "y \<in> closure S"  | 
|
| 33175 | 2494  | 
then obtain f where f: "\<forall>n. f n \<in> S" "(f ---> y) sequentially"  | 
2495  | 
unfolding closure_sequential by auto  | 
|
2496  | 
have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp  | 
|
| 53282 | 2497  | 
then have "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially"  | 
| 33175 | 2498  | 
by (rule eventually_mono, simp add: f(1))  | 
2499  | 
have "dist x y \<le> a"  | 
|
2500  | 
apply (rule Lim_dist_ubound [of sequentially f])  | 
|
2501  | 
apply (rule trivial_limit_sequentially)  | 
|
2502  | 
apply (rule f(2))  | 
|
2503  | 
apply fact  | 
|
2504  | 
done  | 
|
2505  | 
}  | 
|
| 53282 | 2506  | 
then show ?thesis  | 
2507  | 
unfolding bounded_def by auto  | 
|
| 33175 | 2508  | 
qed  | 
2509  | 
||
2510  | 
lemma bounded_cball[simp,intro]: "bounded (cball x e)"  | 
|
2511  | 
apply (simp add: bounded_def)  | 
|
2512  | 
apply (rule_tac x=x in exI)  | 
|
2513  | 
apply (rule_tac x=e in exI)  | 
|
2514  | 
apply auto  | 
|
2515  | 
done  | 
|
2516  | 
||
| 53640 | 2517  | 
lemma bounded_ball[simp,intro]: "bounded (ball x e)"  | 
| 33175 | 2518  | 
by (metis ball_subset_cball bounded_cball bounded_subset)  | 
2519  | 
||
2520  | 
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T"  | 
|
2521  | 
apply (auto simp add: bounded_def)  | 
|
2522  | 
apply (rename_tac x y r s)  | 
|
2523  | 
apply (rule_tac x=x in exI)  | 
|
2524  | 
apply (rule_tac x="max r (dist x y + s)" in exI)  | 
|
| 53640 | 2525  | 
apply (rule ballI)  | 
2526  | 
apply safe  | 
|
2527  | 
apply (drule (1) bspec)  | 
|
2528  | 
apply simp  | 
|
| 33175 | 2529  | 
apply (drule (1) bspec)  | 
| 
54863
 
82acc20ded73
prefer more canonical names for lemmas on min/max
 
haftmann 
parents: 
54797 
diff
changeset
 | 
2530  | 
apply (rule max.coboundedI2)  | 
| 33175 | 2531  | 
apply (erule order_trans [OF dist_triangle add_left_mono])  | 
2532  | 
done  | 
|
2533  | 
||
| 53640 | 2534  | 
lemma bounded_Union[intro]: "finite F \<Longrightarrow> \<forall>S\<in>F. bounded S \<Longrightarrow> bounded (\<Union>F)"  | 
| 52624 | 2535  | 
by (induct rule: finite_induct[of F]) auto  | 
| 33175 | 2536  | 
|
| 50955 | 2537  | 
lemma bounded_UN [intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. bounded (B x) \<Longrightarrow> bounded (\<Union>x\<in>A. B x)"  | 
| 52624 | 2538  | 
by (induct set: finite) auto  | 
| 50955 | 2539  | 
|
| 50948 | 2540  | 
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S"  | 
2541  | 
proof -  | 
|
| 53640 | 2542  | 
  have "\<forall>y\<in>{x}. dist x y \<le> 0"
 | 
2543  | 
by simp  | 
|
2544  | 
  then have "bounded {x}"
 | 
|
2545  | 
unfolding bounded_def by fast  | 
|
2546  | 
then show ?thesis  | 
|
2547  | 
by (metis insert_is_Un bounded_Un)  | 
|
| 50948 | 2548  | 
qed  | 
2549  | 
||
2550  | 
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S"  | 
|
| 52624 | 2551  | 
by (induct set: finite) simp_all  | 
| 50948 | 2552  | 
|
| 53640 | 2553  | 
lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x \<le> b)"  | 
| 33175 | 2554  | 
apply (simp add: bounded_iff)  | 
| 53640 | 2555  | 
apply (subgoal_tac "\<And>x (y::real). 0 < 1 + abs y \<and> (x \<le> y \<longrightarrow> x \<le> 1 + abs y)")  | 
| 52624 | 2556  | 
apply metis  | 
2557  | 
apply arith  | 
|
2558  | 
done  | 
|
| 33175 | 2559  | 
|
| 53640 | 2560  | 
lemma Bseq_eq_bounded:  | 
2561  | 
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector"  | 
|
2562  | 
shows "Bseq f \<longleftrightarrow> bounded (range f)"  | 
|
| 50972 | 2563  | 
unfolding Bseq_def bounded_pos by auto  | 
2564  | 
||
| 33175 | 2565  | 
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)"  | 
2566  | 
by (metis Int_lower1 Int_lower2 bounded_subset)  | 
|
2567  | 
||
| 53291 | 2568  | 
lemma bounded_diff[intro]: "bounded S \<Longrightarrow> bounded (S - T)"  | 
| 52624 | 2569  | 
by (metis Diff_subset bounded_subset)  | 
| 33175 | 2570  | 
|
2571  | 
lemma not_bounded_UNIV[simp, intro]:  | 
|
2572  | 
  "\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
 | 
|
| 53640 | 2573  | 
proof (auto simp add: bounded_pos not_le)  | 
| 33175 | 2574  | 
obtain x :: 'a where "x \<noteq> 0"  | 
2575  | 
using perfect_choose_dist [OF zero_less_one] by fast  | 
|
| 53640 | 2576  | 
fix b :: real  | 
2577  | 
assume b: "b >0"  | 
|
2578  | 
have b1: "b +1 \<ge> 0"  | 
|
2579  | 
using b by simp  | 
|
| 33175 | 2580  | 
with `x \<noteq> 0` have "b < norm (scaleR (b + 1) (sgn x))"  | 
2581  | 
by (simp add: norm_sgn)  | 
|
2582  | 
then show "\<exists>x::'a. b < norm x" ..  | 
|
2583  | 
qed  | 
|
2584  | 
||
2585  | 
lemma bounded_linear_image:  | 
|
| 53291 | 2586  | 
assumes "bounded S"  | 
2587  | 
and "bounded_linear f"  | 
|
2588  | 
shows "bounded (f ` S)"  | 
|
| 52624 | 2589  | 
proof -  | 
| 53640 | 2590  | 
from assms(1) obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b"  | 
| 52624 | 2591  | 
unfolding bounded_pos by auto  | 
| 53640 | 2592  | 
from assms(2) obtain B where B: "B > 0" "\<forall>x. norm (f x) \<le> B * norm x"  | 
| 52624 | 2593  | 
using bounded_linear.pos_bounded by (auto simp add: mult_ac)  | 
2594  | 
  {
 | 
|
| 53282 | 2595  | 
fix x  | 
| 53640 | 2596  | 
assume "x \<in> S"  | 
2597  | 
then have "norm x \<le> b"  | 
|
2598  | 
using b by auto  | 
|
2599  | 
then have "norm (f x) \<le> B * b"  | 
|
2600  | 
using B(2)  | 
|
| 52624 | 2601  | 
apply (erule_tac x=x in allE)  | 
2602  | 
apply (metis B(1) B(2) order_trans mult_le_cancel_left_pos)  | 
|
2603  | 
done  | 
|
| 33175 | 2604  | 
}  | 
| 53282 | 2605  | 
then show ?thesis  | 
2606  | 
unfolding bounded_pos  | 
|
| 52624 | 2607  | 
apply (rule_tac x="b*B" in exI)  | 
2608  | 
using b B mult_pos_pos [of b B]  | 
|
2609  | 
apply (auto simp add: mult_commute)  | 
|
2610  | 
done  | 
|
| 33175 | 2611  | 
qed  | 
2612  | 
||
2613  | 
lemma bounded_scaling:  | 
|
2614  | 
fixes S :: "'a::real_normed_vector set"  | 
|
2615  | 
shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)"  | 
|
| 53291 | 2616  | 
apply (rule bounded_linear_image)  | 
2617  | 
apply assumption  | 
|
| 
44282
 
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
 
huffman 
parents: 
44252 
diff
changeset
 | 
2618  | 
apply (rule bounded_linear_scaleR_right)  | 
| 33175 | 2619  | 
done  | 
2620  | 
||
2621  | 
lemma bounded_translation:  | 
|
2622  | 
fixes S :: "'a::real_normed_vector set"  | 
|
| 52624 | 2623  | 
assumes "bounded S"  | 
2624  | 
shows "bounded ((\<lambda>x. a + x) ` S)"  | 
|
| 53282 | 2625  | 
proof -  | 
| 53640 | 2626  | 
from assms obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b"  | 
| 52624 | 2627  | 
unfolding bounded_pos by auto  | 
2628  | 
  {
 | 
|
2629  | 
fix x  | 
|
| 53640 | 2630  | 
assume "x \<in> S"  | 
| 53282 | 2631  | 
then have "norm (a + x) \<le> b + norm a"  | 
| 52624 | 2632  | 
using norm_triangle_ineq[of a x] b by auto  | 
| 33175 | 2633  | 
}  | 
| 53282 | 2634  | 
then show ?thesis  | 
| 52624 | 2635  | 
unfolding bounded_pos  | 
2636  | 
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
 | 
2637  | 
by (auto intro!: exI[of _ "b + norm a"])  | 
| 33175 | 2638  | 
qed  | 
2639  | 
||
2640  | 
||
2641  | 
text{* Some theorems on sups and infs using the notion "bounded". *}
 | 
|
2642  | 
||
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2643  | 
lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)"  | 
| 33175 | 2644  | 
by (simp add: bounded_iff)  | 
2645  | 
||
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2646  | 
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
 | 
2647  | 
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
 | 
2648  | 
(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
 | 
2649  | 
|
| 
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2650  | 
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
 | 
2651  | 
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
 | 
2652  | 
(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
 | 
2653  | 
|
| 
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2654  | 
(* 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
 | 
2655  | 
|
| 33270 | 2656  | 
lemma bounded_has_Sup:  | 
2657  | 
fixes S :: "real set"  | 
|
| 53640 | 2658  | 
assumes "bounded S"  | 
2659  | 
    and "S \<noteq> {}"
 | 
|
| 53282 | 2660  | 
shows "\<forall>x\<in>S. x \<le> Sup S"  | 
2661  | 
and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"  | 
|
| 33270 | 2662  | 
proof  | 
| 53282 | 2663  | 
show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"  | 
2664  | 
using assms by (metis cSup_least)  | 
|
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2665  | 
qed (metis cSup_upper assms(1) bounded_imp_bdd_above)  | 
| 33270 | 2666  | 
|
2667  | 
lemma Sup_insert:  | 
|
2668  | 
fixes S :: "real set"  | 
|
| 53291 | 2669  | 
  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
 | 
2670  | 
by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If)  | 
| 33270 | 2671  | 
|
2672  | 
lemma Sup_insert_finite:  | 
|
2673  | 
fixes S :: "real set"  | 
|
| 53291 | 2674  | 
  shows "finite S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
 | 
| 33270 | 2675  | 
apply (rule Sup_insert)  | 
2676  | 
apply (rule finite_imp_bounded)  | 
|
| 52624 | 2677  | 
apply simp  | 
2678  | 
done  | 
|
| 33270 | 2679  | 
|
2680  | 
lemma bounded_has_Inf:  | 
|
2681  | 
fixes S :: "real set"  | 
|
| 53640 | 2682  | 
assumes "bounded S"  | 
2683  | 
    and "S \<noteq> {}"
 | 
|
| 53282 | 2684  | 
shows "\<forall>x\<in>S. x \<ge> Inf S"  | 
2685  | 
and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"  | 
|
| 33175 | 2686  | 
proof  | 
| 53640 | 2687  | 
show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"  | 
| 53282 | 2688  | 
using assms by (metis cInf_greatest)  | 
| 
54258
 
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
 
hoelzl 
parents: 
54230 
diff
changeset
 | 
2689  | 
qed (metis cInf_lower assms(1) bounded_imp_bdd_below)  | 
| 33270 | 2690  | 
|
2691  | 
lemma Inf_insert:  | 
|
2692  | 
fixes S :: "real set"  | 
|
| 53291 | 2693  | 
  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
 | 
2694  | 
by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If)  | 
| 50944 | 2695  | 
|
| 33270 | 2696  | 
lemma Inf_insert_finite:  | 
2697  | 
fixes S :: "real set"  | 
|
| 53291 | 2698  | 
  shows "finite S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
 | 
| 53282 | 2699  | 
apply (rule Inf_insert)  | 
2700  | 
apply (rule finite_imp_bounded)  | 
|
2701  | 
apply simp  | 
|
2702  | 
done  | 
|
| 33270 | 2703  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2704  | 
subsection {* Compactness *}
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2705  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2706  | 
subsubsection {* Bolzano-Weierstrass property *}
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2707  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2708  | 
lemma heine_borel_imp_bolzano_weierstrass:  | 
| 53640 | 2709  | 
assumes "compact s"  | 
2710  | 
and "infinite t"  | 
|
2711  | 
and "t \<subseteq> s"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2712  | 
shows "\<exists>x \<in> s. x islimpt t"  | 
| 53291 | 2713  | 
proof (rule ccontr)  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2714  | 
assume "\<not> (\<exists>x \<in> s. x islimpt t)"  | 
| 53640 | 2715  | 
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 | 2716  | 
unfolding islimpt_def  | 
| 53282 | 2717  | 
using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"]  | 
2718  | 
by auto  | 
|
| 53640 | 2719  | 
  obtain g where g: "g \<subseteq> {t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
 | 
| 52624 | 2720  | 
    using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]]
 | 
2721  | 
using f by auto  | 
|
| 53640 | 2722  | 
from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa"  | 
2723  | 
by auto  | 
|
| 52624 | 2724  | 
  {
 | 
2725  | 
fix x y  | 
|
| 53640 | 2726  | 
assume "x \<in> t" "y \<in> t" "f x = f y"  | 
| 53282 | 2727  | 
then have "x \<in> f x" "y \<in> f x \<longrightarrow> y = x"  | 
| 53640 | 2728  | 
using f[THEN bspec[where x=x]] and `t \<subseteq> s` by auto  | 
| 53282 | 2729  | 
then have "x = y"  | 
| 53640 | 2730  | 
using `f x = f y` and f[THEN bspec[where x=y]] and `y \<in> t` and `t \<subseteq> s`  | 
2731  | 
by auto  | 
|
| 52624 | 2732  | 
}  | 
| 53282 | 2733  | 
then have "inj_on f t"  | 
| 52624 | 2734  | 
unfolding inj_on_def by simp  | 
| 53282 | 2735  | 
then have "infinite (f ` t)"  | 
| 52624 | 2736  | 
using assms(2) using finite_imageD by auto  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2737  | 
moreover  | 
| 52624 | 2738  | 
  {
 | 
2739  | 
fix x  | 
|
| 53640 | 2740  | 
assume "x \<in> t" "f x \<notin> g"  | 
2741  | 
from g(3) assms(3) `x \<in> t` obtain h where "h \<in> g" and "x \<in> h"  | 
|
2742  | 
by auto  | 
|
2743  | 
then obtain y where "y \<in> s" "h = f y"  | 
|
| 52624 | 2744  | 
using g'[THEN bspec[where x=h]] by auto  | 
| 53282 | 2745  | 
then have "y = x"  | 
| 53640 | 2746  | 
using f[THEN bspec[where x=y]] and `x\<in>t` and `x\<in>h`[unfolded `h = f y`]  | 
2747  | 
by auto  | 
|
| 53282 | 2748  | 
then have False  | 
| 53640 | 2749  | 
using `f x \<notin> g` `h \<in> g` unfolding `h = f y`  | 
2750  | 
by auto  | 
|
| 52624 | 2751  | 
}  | 
| 53282 | 2752  | 
then have "f ` t \<subseteq> g" by auto  | 
| 52624 | 2753  | 
ultimately show False  | 
2754  | 
using g(2) using finite_subset by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2755  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2756  | 
|
| 
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
 | 
2757  | 
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
 | 
2758  | 
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
 | 
2759  | 
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
 | 
2760  | 
shows "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2761  | 
proof -  | 
| 
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
 | 
2762  | 
from countable_basis_at_decseq[of l] guess A . note A = this  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2763  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2764  | 
def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> f j \<in> A (Suc n)"  | 
| 52624 | 2765  | 
  {
 | 
2766  | 
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
 | 
2767  | 
    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
 | 
2768  | 
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
 | 
2769  | 
    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
 | 
2770  | 
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
 | 
2771  | 
    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
 | 
2772  | 
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
 | 
2773  | 
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
 | 
2774  | 
by (auto simp: not_le)  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2775  | 
then have "i < s n i" "f (s n i) \<in> A (Suc n)"  | 
| 52624 | 2776  | 
unfolding s_def by (auto intro: someI2_ex)  | 
2777  | 
}  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2778  | 
note s = this  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2779  | 
def r \<equiv> "nat_rec (s 0 0) s"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2780  | 
have "subseq r"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2781  | 
by (auto simp: r_def s subseq_Suc_iff)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2782  | 
moreover  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2783  | 
have "(\<lambda>n. f (r n)) ----> l"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2784  | 
proof (rule topological_tendstoI)  | 
| 53282 | 2785  | 
fix S  | 
2786  | 
assume "open S" "l \<in> S"  | 
|
| 53640 | 2787  | 
with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially"  | 
2788  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2789  | 
moreover  | 
| 52624 | 2790  | 
    {
 | 
2791  | 
fix i  | 
|
| 53282 | 2792  | 
assume "Suc 0 \<le> i"  | 
2793  | 
then have "f (r i) \<in> A i"  | 
|
| 52624 | 2794  | 
by (cases i) (simp_all add: r_def s)  | 
2795  | 
}  | 
|
2796  | 
then have "eventually (\<lambda>i. f (r i) \<in> A i) sequentially"  | 
|
2797  | 
by (auto simp: eventually_sequentially)  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2798  | 
ultimately show "eventually (\<lambda>i. f (r i) \<in> S) sequentially"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2799  | 
by eventually_elim auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2800  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2801  | 
ultimately show "\<exists>r. subseq r \<and> (f \<circ> r) ----> l"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2802  | 
by (auto simp: convergent_def comp_def)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2803  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2804  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2805  | 
lemma sequence_infinite_lemma:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2806  | 
fixes f :: "nat \<Rightarrow> 'a::t1_space"  | 
| 53282 | 2807  | 
assumes "\<forall>n. f n \<noteq> l"  | 
2808  | 
and "(f ---> l) sequentially"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2809  | 
shows "infinite (range f)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2810  | 
proof  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2811  | 
assume "finite (range f)"  | 
| 53640 | 2812  | 
then have "closed (range f)"  | 
2813  | 
by (rule finite_imp_closed)  | 
|
2814  | 
then have "open (- range f)"  | 
|
2815  | 
by (rule open_Compl)  | 
|
2816  | 
from assms(1) have "l \<in> - range f"  | 
|
2817  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2818  | 
from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially"  | 
| 53640 | 2819  | 
using `open (- range f)` `l \<in> - range f`  | 
2820  | 
by (rule topological_tendstoD)  | 
|
2821  | 
then show False  | 
|
2822  | 
unfolding eventually_sequentially  | 
|
2823  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2824  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2825  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2826  | 
lemma closure_insert:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2827  | 
fixes x :: "'a::t1_space"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2828  | 
shows "closure (insert x s) = insert x (closure s)"  | 
| 52624 | 2829  | 
apply (rule closure_unique)  | 
2830  | 
apply (rule insert_mono [OF closure_subset])  | 
|
2831  | 
apply (rule closed_insert [OF closed_closure])  | 
|
2832  | 
apply (simp add: closure_minimal)  | 
|
2833  | 
done  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2834  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2835  | 
lemma islimpt_insert:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2836  | 
fixes x :: "'a::t1_space"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2837  | 
shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2838  | 
proof  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2839  | 
assume *: "x islimpt (insert a s)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2840  | 
show "x islimpt s"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2841  | 
proof (rule islimptI)  | 
| 53282 | 2842  | 
fix t  | 
2843  | 
assume t: "x \<in> t" "open t"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2844  | 
show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2845  | 
proof (cases "x = a")  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2846  | 
case True  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2847  | 
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
 | 
2848  | 
using * t by (rule islimptE)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2849  | 
with `x = a` show ?thesis by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2850  | 
next  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2851  | 
case False  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2852  | 
      with t have t': "x \<in> t - {a}" "open (t - {a})"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2853  | 
by (simp_all add: open_Diff)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2854  | 
      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
 | 
2855  | 
using * t' by (rule islimptE)  | 
| 53282 | 2856  | 
then show ?thesis by auto  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2857  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2858  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2859  | 
next  | 
| 53282 | 2860  | 
assume "x islimpt s"  | 
2861  | 
then show "x islimpt (insert a s)"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2862  | 
by (rule islimpt_subset) auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2863  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2864  | 
|
| 
50897
 
078590669527
generalize lemma islimpt_finite to class t1_space
 
huffman 
parents: 
50884 
diff
changeset
 | 
2865  | 
lemma islimpt_finite:  | 
| 
 
078590669527
generalize lemma islimpt_finite to class t1_space
 
huffman 
parents: 
50884 
diff
changeset
 | 
2866  | 
fixes x :: "'a::t1_space"  | 
| 
 
078590669527
generalize lemma islimpt_finite to class t1_space
 
huffman 
parents: 
50884 
diff
changeset
 | 
2867  | 
shows "finite s \<Longrightarrow> \<not> x islimpt s"  | 
| 52624 | 2868  | 
by (induct set: finite) (simp_all add: islimpt_insert)  | 
| 
50897
 
078590669527
generalize lemma islimpt_finite to class t1_space
 
huffman 
parents: 
50884 
diff
changeset
 | 
2869  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2870  | 
lemma islimpt_union_finite:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2871  | 
fixes x :: "'a::t1_space"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2872  | 
shows "finite s \<Longrightarrow> x islimpt (s \<union> t) \<longleftrightarrow> x islimpt t"  | 
| 52624 | 2873  | 
by (simp add: islimpt_Un islimpt_finite)  | 
| 
50897
 
078590669527
generalize lemma islimpt_finite to class t1_space
 
huffman 
parents: 
50884 
diff
changeset
 | 
2874  | 
|
| 
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
 | 
2875  | 
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
 | 
2876  | 
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
 | 
2877  | 
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
 | 
2878  | 
proof (safe intro!: islimptI)  | 
| 53282 | 2879  | 
fix U  | 
2880  | 
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
 | 
2881  | 
  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
 | 
2882  | 
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
 | 
2883  | 
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
 | 
2884  | 
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
 | 
2885  | 
next  | 
| 53282 | 2886  | 
fix T  | 
2887  | 
assume *: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S)" "l \<in> T" "open T"  | 
|
2888  | 
  then have "infinite (T \<inter> S - {l})"
 | 
|
2889  | 
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
 | 
2890  | 
  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
 | 
2891  | 
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
 | 
2892  | 
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
 | 
2893  | 
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
 | 
2894  | 
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
 | 
2895  | 
|
| 
 
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
 | 
2896  | 
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
 | 
2897  | 
  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
 | 
2898  | 
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
 | 
2899  | 
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
 | 
2900  | 
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
 | 
2901  | 
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
 | 
2902  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2903  | 
lemma sequence_unique_limpt:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2904  | 
fixes f :: "nat \<Rightarrow> 'a::t2_space"  | 
| 53282 | 2905  | 
assumes "(f ---> l) sequentially"  | 
2906  | 
and "l' islimpt (range f)"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2907  | 
shows "l' = l"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2908  | 
proof (rule ccontr)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2909  | 
assume "l' \<noteq> l"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2910  | 
  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
 | 
2911  | 
using hausdorff [OF `l' \<noteq> l`] by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2912  | 
have "eventually (\<lambda>n. f n \<in> t) sequentially"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2913  | 
using assms(1) `open t` `l \<in> t` by (rule topological_tendstoD)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2914  | 
then obtain N where "\<forall>n\<ge>N. f n \<in> t"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2915  | 
unfolding eventually_sequentially by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2916  | 
|
| 53282 | 2917  | 
  have "UNIV = {..<N} \<union> {N..}"
 | 
2918  | 
by auto  | 
|
2919  | 
  then have "l' islimpt (f ` ({..<N} \<union> {N..}))"
 | 
|
2920  | 
using assms(2) by simp  | 
|
2921  | 
  then have "l' islimpt (f ` {..<N} \<union> f ` {N..})"
 | 
|
2922  | 
by (simp add: image_Un)  | 
|
2923  | 
  then have "l' islimpt (f ` {N..})"
 | 
|
2924  | 
by (simp add: islimpt_union_finite)  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2925  | 
  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
 | 
2926  | 
using `l' \<in> s` `open s` by (rule islimptE)  | 
| 53282 | 2927  | 
then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'"  | 
2928  | 
by auto  | 
|
2929  | 
with `\<forall>n\<ge>N. f n \<in> t` have "f n \<in> s \<inter> t"  | 
|
2930  | 
by simp  | 
|
2931  | 
  with `s \<inter> t = {}` show False
 | 
|
2932  | 
by simp  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2933  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2934  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2935  | 
lemma bolzano_weierstrass_imp_closed:  | 
| 53640 | 2936  | 
  fixes s :: "'a::{first_countable_topology,t2_space} set"
 | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2937  | 
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
 | 
2938  | 
shows "closed s"  | 
| 52624 | 2939  | 
proof -  | 
2940  | 
  {
 | 
|
2941  | 
fix x l  | 
|
2942  | 
assume as: "\<forall>n::nat. x n \<in> s" "(x ---> l) sequentially"  | 
|
| 53282 | 2943  | 
then have "l \<in> s"  | 
| 52624 | 2944  | 
proof (cases "\<forall>n. x n \<noteq> l")  | 
2945  | 
case False  | 
|
| 53282 | 2946  | 
then show "l\<in>s" using as(1) by auto  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2947  | 
next  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2948  | 
case True note cas = this  | 
| 52624 | 2949  | 
with as(2) have "infinite (range x)"  | 
2950  | 
using sequence_infinite_lemma[of x l] by auto  | 
|
2951  | 
then obtain l' where "l'\<in>s" "l' islimpt (range x)"  | 
|
2952  | 
using assms[THEN spec[where x="range x"]] as(1) by auto  | 
|
| 53282 | 2953  | 
then show "l\<in>s" using sequence_unique_limpt[of x l l']  | 
| 52624 | 2954  | 
using as cas by auto  | 
2955  | 
qed  | 
|
2956  | 
}  | 
|
| 53282 | 2957  | 
then show ?thesis  | 
2958  | 
unfolding closed_sequential_limits by fast  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2959  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2960  | 
|
| 50944 | 2961  | 
lemma compact_imp_bounded:  | 
| 52624 | 2962  | 
assumes "compact U"  | 
2963  | 
shows "bounded U"  | 
|
| 50944 | 2964  | 
proof -  | 
| 52624 | 2965  | 
have "compact U" "\<forall>x\<in>U. open (ball x 1)" "U \<subseteq> (\<Union>x\<in>U. ball x 1)"  | 
2966  | 
using assms by auto  | 
|
| 50944 | 2967  | 
then obtain D where D: "D \<subseteq> U" "finite D" "U \<subseteq> (\<Union>x\<in>D. ball x 1)"  | 
| 52624 | 2968  | 
by (rule compactE_image)  | 
| 50955 | 2969  | 
from `finite D` have "bounded (\<Union>x\<in>D. ball x 1)"  | 
2970  | 
by (simp add: bounded_UN)  | 
|
| 53282 | 2971  | 
then show "bounded U" using `U \<subseteq> (\<Union>x\<in>D. ball x 1)`  | 
| 50955 | 2972  | 
by (rule bounded_subset)  | 
| 50944 | 2973  | 
qed  | 
2974  | 
||
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2975  | 
text{* In particular, some common special cases. *}
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2976  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2977  | 
lemma compact_union [intro]:  | 
| 53291 | 2978  | 
assumes "compact s"  | 
2979  | 
and "compact t"  | 
|
| 53282 | 2980  | 
shows " compact (s \<union> t)"  | 
| 50898 | 2981  | 
proof (rule compactI)  | 
| 52624 | 2982  | 
fix f  | 
2983  | 
assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2984  | 
from * `compact s` obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2985  | 
unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis  | 
| 52624 | 2986  | 
moreover  | 
2987  | 
from * `compact t` obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2988  | 
unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2989  | 
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
 | 
2990  | 
by (auto intro!: exI[of _ "s' \<union> t'"])  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2991  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2992  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2993  | 
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
 | 
2994  | 
by (induct set: finite) auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2995  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2996  | 
lemma compact_UN [intro]:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2997  | 
"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
 | 
2998  | 
unfolding SUP_def by (rule compact_Union) auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
2999  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3000  | 
lemma closed_inter_compact [intro]:  | 
| 53282 | 3001  | 
assumes "closed s"  | 
3002  | 
and "compact t"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3003  | 
shows "compact (s \<inter> t)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3004  | 
using compact_inter_closed [of t s] assms  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3005  | 
by (simp add: Int_commute)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3006  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3007  | 
lemma compact_inter [intro]:  | 
| 50898 | 3008  | 
fixes s t :: "'a :: t2_space set"  | 
| 53282 | 3009  | 
assumes "compact s"  | 
3010  | 
and "compact t"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3011  | 
shows "compact (s \<inter> t)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3012  | 
using assms by (intro compact_inter_closed compact_imp_closed)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3013  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3014  | 
lemma compact_sing [simp]: "compact {a}"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3015  | 
unfolding compact_eq_heine_borel by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3016  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3017  | 
lemma compact_insert [simp]:  | 
| 53282 | 3018  | 
assumes "compact s"  | 
3019  | 
shows "compact (insert x s)"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3020  | 
proof -  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3021  | 
  have "compact ({x} \<union> s)"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3022  | 
using compact_sing assms by (rule compact_union)  | 
| 53282 | 3023  | 
then show ?thesis by simp  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3024  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3025  | 
|
| 52624 | 3026  | 
lemma finite_imp_compact: "finite s \<Longrightarrow> compact s"  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3027  | 
by (induct set: finite) simp_all  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3028  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3029  | 
lemma open_delete:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3030  | 
fixes s :: "'a::t1_space set"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3031  | 
  shows "open s \<Longrightarrow> open (s - {x})"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3032  | 
by (simp add: open_Diff)  | 
| 
 
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  | 
text{*Compactness expressed with filters*}
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3035  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3036  | 
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
 | 
3037  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3038  | 
lemma eventually_filter_from_subbase:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3039  | 
"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
 | 
3040  | 
(is "_ \<longleftrightarrow> ?R P")  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3041  | 
unfolding filter_from_subbase_def  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3042  | 
proof (rule eventually_Abs_filter is_filter.intro)+  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3043  | 
show "?R (\<lambda>x. True)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3044  | 
    by (rule exI[of _ "{}"]) (simp add: le_fun_def)
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3045  | 
next  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3046  | 
fix P Q assume "?R P" then guess X ..  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3047  | 
moreover assume "?R Q" then guess Y ..  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3048  | 
ultimately show "?R (\<lambda>x. P x \<and> Q x)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3049  | 
by (intro exI[of _ "X \<union> Y"]) auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3050  | 
next  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3051  | 
fix P Q  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3052  | 
assume "?R P" then guess X ..  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3053  | 
moreover assume "\<forall>x. P x \<longrightarrow> Q x"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3054  | 
ultimately show "?R Q"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3055  | 
by (intro exI[of _ X]) auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3056  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3057  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3058  | 
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
 | 
3059  | 
  by (subst eventually_filter_from_subbase) (auto intro!: exI[of _ "{P}"])
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3060  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3061  | 
lemma filter_from_subbase_not_bot:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3062  | 
"\<forall>X \<subseteq> B. finite X \<longrightarrow> Inf X \<noteq> bot \<Longrightarrow> filter_from_subbase B \<noteq> bot"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3063  | 
unfolding trivial_limit_def eventually_filter_from_subbase by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3064  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3065  | 
lemma closure_iff_nhds_not_empty:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3066  | 
  "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
 | 
3067  | 
proof safe  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3068  | 
assume x: "x \<in> closure X"  | 
| 53282 | 3069  | 
fix S A  | 
3070  | 
  assume "open S" "x \<in> S" "X \<inter> A = {}" "S \<subseteq> A"
 | 
|
3071  | 
then have "x \<notin> closure (-S)"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3072  | 
by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3073  | 
with x have "x \<in> closure X - closure (-S)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3074  | 
by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3075  | 
also have "\<dots> \<subseteq> closure (X \<inter> S)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3076  | 
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
 | 
3077  | 
  finally have "X \<inter> S \<noteq> {}" by auto
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3078  | 
  then show False using `X \<inter> A = {}` `S \<subseteq> A` 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  | 
  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
 | 
3081  | 
from this[THEN spec, of "- X", THEN spec, of "- closure X"]  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3082  | 
show "x \<in> closure X"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3083  | 
by (simp add: closure_subset open_Compl)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3084  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3085  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3086  | 
lemma compact_filter:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3087  | 
"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
 | 
3088  | 
proof (intro allI iffI impI compact_fip[THEN iffD2] notI)  | 
| 53282 | 3089  | 
fix F  | 
3090  | 
assume "compact U"  | 
|
3091  | 
assume F: "F \<noteq> bot" "eventually (\<lambda>x. x \<in> U) F"  | 
|
3092  | 
  then have "U \<noteq> {}"
 | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3093  | 
by (auto simp: eventually_False)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3094  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3095  | 
  def Z \<equiv> "closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3096  | 
then have "\<forall>z\<in>Z. closed z"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3097  | 
by auto  | 
| 53282 | 3098  | 
moreover  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3099  | 
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
 | 
3100  | 
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
 | 
3101  | 
  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
 | 
3102  | 
proof (intro allI impI)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3103  | 
fix B assume "finite B" "B \<subseteq> Z"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3104  | 
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
 | 
3105  | 
by (auto intro!: eventually_Ball_finite)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3106  | 
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
 | 
3107  | 
by eventually_elim auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3108  | 
    with F show "U \<inter> \<Inter>B \<noteq> {}"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3109  | 
by (intro notI) (simp add: eventually_False)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3110  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3111  | 
  ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3112  | 
using `compact U` unfolding compact_fip by blast  | 
| 53282 | 3113  | 
then obtain x where "x \<in> U" and x: "\<And>z. z \<in> Z \<Longrightarrow> x \<in> z"  | 
3114  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3115  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3116  | 
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
 | 
3117  | 
unfolding eventually_inf eventually_nhds  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3118  | 
proof safe  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3119  | 
fix P Q R S  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3120  | 
assume "eventually R F" "open S" "x \<in> S"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3121  | 
    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
 | 
3122  | 
    have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3123  | 
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
 | 
3124  | 
ultimately show False by (auto simp: set_eq_iff)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3125  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3126  | 
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
 | 
3127  | 
by (metis eventually_bot)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3128  | 
next  | 
| 53282 | 3129  | 
fix A  | 
3130  | 
  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
 | 
3131  | 
def P' \<equiv> "(\<lambda>a (x::'a). x \<in> a)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3132  | 
then have inj_P': "\<And>A. inj_on P' A"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3133  | 
by (auto intro!: inj_onI simp: fun_eq_iff)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3134  | 
def F \<equiv> "filter_from_subbase (P' ` insert U A)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3135  | 
have "F \<noteq> bot"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3136  | 
unfolding F_def  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3137  | 
proof (safe intro!: filter_from_subbase_not_bot)  | 
| 53282 | 3138  | 
fix X  | 
3139  | 
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
 | 
3140  | 
then obtain B where "B \<subseteq> insert U A" "finite B" and B: "Inf (P' ` B) = bot"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3141  | 
unfolding subset_image_iff by (auto intro: inj_P' finite_imageD)  | 
| 53282 | 3142  | 
    with A(2)[THEN spec, of "B - {U}"] have "U \<inter> \<Inter>(B - {U}) \<noteq> {}"
 | 
3143  | 
by auto  | 
|
3144  | 
with B show False  | 
|
3145  | 
by (auto simp: P'_def fun_eq_iff)  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3146  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3147  | 
moreover have "eventually (\<lambda>x. x \<in> U) F"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3148  | 
unfolding F_def by (rule eventually_filter_from_subbaseI) (auto simp: P'_def)  | 
| 53282 | 3149  | 
moreover  | 
3150  | 
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
 | 
3151  | 
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
 | 
3152  | 
by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3153  | 
|
| 53282 | 3154  | 
  {
 | 
3155  | 
fix V  | 
|
3156  | 
assume "V \<in> A"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3157  | 
then have V: "eventually (\<lambda>x. x \<in> V) F"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3158  | 
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
 | 
3159  | 
have "x \<in> closure V"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3160  | 
unfolding closure_iff_nhds_not_empty  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3161  | 
proof (intro impI allI)  | 
| 53282 | 3162  | 
fix S A  | 
3163  | 
assume "open S" "x \<in> S" "S \<subseteq> A"  | 
|
3164  | 
then have "eventually (\<lambda>x. x \<in> A) (nhds x)"  | 
|
3165  | 
by (auto simp: eventually_nhds)  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3166  | 
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
 | 
3167  | 
by (auto simp: eventually_inf)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3168  | 
      with x show "V \<inter> A \<noteq> {}"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3169  | 
by (auto simp del: Int_iff simp add: trivial_limit_def)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3170  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3171  | 
then have "x \<in> V"  | 
| 53282 | 3172  | 
using `V \<in> A` A(1) by simp  | 
3173  | 
}  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3174  | 
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
 | 
3175  | 
  with `U \<inter> \<Inter>A = {}` show False by auto
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3176  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3177  | 
|
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3178  | 
definition "countably_compact U \<longleftrightarrow>  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3179  | 
(\<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
 | 
3180  | 
|
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3181  | 
lemma countably_compactE:  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3182  | 
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
 | 
3183  | 
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
 | 
3184  | 
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
 | 
3185  | 
|
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3186  | 
lemma countably_compactI:  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3187  | 
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
 | 
3188  | 
shows "countably_compact s"  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3189  | 
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
 | 
3190  | 
|
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3191  | 
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
 | 
3192  | 
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
 | 
3193  | 
|
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3194  | 
lemma countably_compact_imp_compact:  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3195  | 
assumes "countably_compact U"  | 
| 53282 | 3196  | 
and ccover: "countable B" "\<forall>b\<in>B. open b"  | 
3197  | 
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
 | 
3198  | 
shows "compact U"  | 
| 53282 | 3199  | 
using `countably_compact U`  | 
3200  | 
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
 | 
3201  | 
proof safe  | 
| 53282 | 3202  | 
fix A  | 
3203  | 
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
 | 
3204  | 
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
 | 
3205  | 
|
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3206  | 
  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
 | 
3207  | 
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
 | 
3208  | 
unfolding C_def using ccover by auto  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3209  | 
moreover  | 
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3210  | 
have "\<Union>A \<inter> U \<subseteq> \<Union>C"  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3211  | 
proof safe  | 
| 53282 | 3212  | 
fix x a  | 
3213  | 
assume "x \<in> U" "x \<in> a" "a \<in> A"  | 
|
3214  | 
with basis[of a x] A obtain b where "b \<in> B" "x \<in> b" "b \<inter> U \<subseteq> a"  | 
|
3215  | 
by blast  | 
|
3216  | 
with `a \<in> A` show "x \<in> \<Union>C"  | 
|
3217  | 
unfolding C_def by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3218  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3219  | 
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
 | 
3220  | 
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
 | 
3221  | 
using * by metis  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
3222  | 
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
 | 
3223  | 
by (auto simp: C_def)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3224  | 
then guess f unfolding bchoice_iff Bex_def ..  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
3225  | 
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
 | 
3226  | 
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
 | 
3227  | 
qed  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3228  | 
|
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3229  | 
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
 | 
3230  | 
"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
 | 
3231  | 
proof (rule countably_compact_imp_compact)  | 
| 53282 | 3232  | 
fix T and x :: 'a  | 
3233  | 
assume "open T" "x \<in> T"  | 
|
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3234  | 
from topological_basisE[OF is_basis this] guess b .  | 
| 53282 | 3235  | 
then show "\<exists>b\<in>SOME B. countable B \<and> topological_basis B. x \<in> b \<and> b \<inter> U \<subseteq> T"  | 
3236  | 
by auto  | 
|
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3237  | 
qed (insert countable_basis topological_basis_open[OF is_basis], auto)  | 
| 36437 | 3238  | 
|
| 
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
 | 
3239  | 
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
 | 
3240  | 
"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
 | 
3241  | 
using countably_compact_imp_compact_second_countable compact_imp_countably_compact by blast  | 
| 53282 | 3242  | 
|
| 36437 | 3243  | 
subsubsection{* Sequential compactness *}
 | 
| 33175 | 3244  | 
|
| 53282 | 3245  | 
definition seq_compact :: "'a::topological_space set \<Rightarrow> bool"  | 
3246  | 
where "seq_compact S \<longleftrightarrow>  | 
|
| 53640 | 3247  | 
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially))"  | 
| 33175 | 3248  | 
|
| 54070 | 3249  | 
lemma seq_compactI:  | 
3250  | 
assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"  | 
|
3251  | 
shows "seq_compact S"  | 
|
3252  | 
unfolding seq_compact_def using assms by fast  | 
|
3253  | 
||
3254  | 
lemma seq_compactE:  | 
|
3255  | 
assumes "seq_compact S" "\<forall>n. f n \<in> S"  | 
|
3256  | 
obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) ---> l) sequentially"  | 
|
3257  | 
using assms unfolding seq_compact_def by fast  | 
|
3258  | 
||
3259  | 
lemma closed_sequentially: (* TODO: move upwards *)  | 
|
3260  | 
assumes "closed s" and "\<forall>n. f n \<in> s" and "f ----> l"  | 
|
3261  | 
shows "l \<in> s"  | 
|
3262  | 
proof (rule ccontr)  | 
|
3263  | 
assume "l \<notin> s"  | 
|
3264  | 
with `closed s` and `f ----> l` have "eventually (\<lambda>n. f n \<in> - s) sequentially"  | 
|
3265  | 
by (fast intro: topological_tendstoD)  | 
|
3266  | 
with `\<forall>n. f n \<in> s` show "False"  | 
|
3267  | 
by simp  | 
|
3268  | 
qed  | 
|
3269  | 
||
3270  | 
lemma seq_compact_inter_closed:  | 
|
3271  | 
assumes "seq_compact s" and "closed t"  | 
|
3272  | 
shows "seq_compact (s \<inter> t)"  | 
|
3273  | 
proof (rule seq_compactI)  | 
|
3274  | 
fix f assume "\<forall>n::nat. f n \<in> s \<inter> t"  | 
|
3275  | 
hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"  | 
|
3276  | 
by simp_all  | 
|
3277  | 
from `seq_compact s` and `\<forall>n. f n \<in> s`  | 
|
3278  | 
obtain l r where "l \<in> s" and r: "subseq r" and l: "(f \<circ> r) ----> l"  | 
|
3279  | 
by (rule seq_compactE)  | 
|
3280  | 
from `\<forall>n. f n \<in> t` have "\<forall>n. (f \<circ> r) n \<in> t"  | 
|
3281  | 
by simp  | 
|
3282  | 
from `closed t` and this and l have "l \<in> t"  | 
|
3283  | 
by (rule closed_sequentially)  | 
|
3284  | 
with `l \<in> s` and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"  | 
|
3285  | 
by fast  | 
|
3286  | 
qed  | 
|
3287  | 
||
3288  | 
lemma seq_compact_closed_subset:  | 
|
3289  | 
assumes "closed s" and "s \<subseteq> t" and "seq_compact t"  | 
|
3290  | 
shows "seq_compact s"  | 
|
3291  | 
using assms seq_compact_inter_closed [of t s] by (simp add: Int_absorb1)  | 
|
3292  | 
||
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3293  | 
lemma seq_compact_imp_countably_compact:  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3294  | 
fixes U :: "'a :: first_countable_topology set"  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3295  | 
assumes "seq_compact U"  | 
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3296  | 
shows "countably_compact U"  | 
| 
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3297  | 
proof (safe intro!: countably_compactI)  | 
| 52624 | 3298  | 
fix A  | 
3299  | 
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
 | 
3300  | 
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
 | 
3301  | 
using `seq_compact U` by (fastforce simp: seq_compact_def subset_eq)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3302  | 
show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3303  | 
proof cases  | 
| 52624 | 3304  | 
assume "finite A"  | 
3305  | 
with A show ?thesis by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3306  | 
next  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3307  | 
assume "infinite A"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3308  | 
    then have "A \<noteq> {}" by auto
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3309  | 
show ?thesis  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3310  | 
proof (rule ccontr)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3311  | 
assume "\<not> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)"  | 
| 53282 | 3312  | 
then have "\<forall>T. \<exists>x. T \<subseteq> A \<and> finite T \<longrightarrow> (x \<in> U - \<Union>T)"  | 
3313  | 
by auto  | 
|
3314  | 
then obtain X' where T: "\<And>T. T \<subseteq> A \<Longrightarrow> finite T \<Longrightarrow> X' T \<in> U - \<Union>T"  | 
|
3315  | 
by metis  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3316  | 
      def X \<equiv> "\<lambda>n. X' (from_nat_into A ` {.. n})"
 | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3317  | 
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
 | 
3318  | 
        using `A \<noteq> {}` unfolding X_def SUP_def by (intro T) (auto intro: from_nat_into)
 | 
| 53282 | 3319  | 
then have "range X \<subseteq> U"  | 
3320  | 
by auto  | 
|
3321  | 
with subseq[of X] obtain r x where "x \<in> U" and r: "subseq r" "(X \<circ> r) ----> x"  | 
|
3322  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3323  | 
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
 | 
3324  | 
obtain n where "x \<in> from_nat_into A n" by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3325  | 
      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
 | 
3326  | 
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
 | 
3327  | 
unfolding tendsto_def by (auto simp: comp_def)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3328  | 
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
 | 
3329  | 
by (auto simp: eventually_sequentially)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3330  | 
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
 | 
3331  | 
by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3332  | 
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
 | 
3333  | 
by (auto intro!: exI[of _ "max n N"])  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3334  | 
ultimately show False  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3335  | 
by auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3336  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3337  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3338  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3339  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3340  | 
lemma compact_imp_seq_compact:  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3341  | 
fixes U :: "'a :: first_countable_topology set"  | 
| 53282 | 3342  | 
assumes "compact U"  | 
3343  | 
shows "seq_compact U"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3344  | 
unfolding seq_compact_def  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3345  | 
proof safe  | 
| 52624 | 3346  | 
fix X :: "nat \<Rightarrow> 'a"  | 
3347  | 
assume "\<forall>n. X n \<in> U"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3348  | 
then have "eventually (\<lambda>x. x \<in> U) (filtermap X sequentially)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3349  | 
by (auto simp: eventually_filtermap)  | 
| 52624 | 3350  | 
moreover  | 
3351  | 
have "filtermap X sequentially \<noteq> bot"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3352  | 
by (simp add: trivial_limit_def eventually_filtermap)  | 
| 52624 | 3353  | 
ultimately  | 
3354  | 
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
 | 
3355  | 
using `compact U` by (auto simp: compact_filter)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3356  | 
|
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3357  | 
from countable_basis_at_decseq[of x] guess A . note A = this  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3358  | 
def s \<equiv> "\<lambda>n i. SOME j. i < j \<and> X j \<in> A (Suc n)"  | 
| 52624 | 3359  | 
  {
 | 
3360  | 
fix n i  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3361  | 
have "\<exists>a. i < a \<and> X a \<in> A (Suc n)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3362  | 
proof (rule ccontr)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3363  | 
assume "\<not> (\<exists>a>i. X a \<in> A (Suc n))"  | 
| 53282 | 3364  | 
then have "\<And>a. Suc i \<le> a \<Longrightarrow> X a \<notin> A (Suc n)"  | 
3365  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3366  | 
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
 | 
3367  | 
by (auto simp: eventually_filtermap eventually_sequentially)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3368  | 
moreover have "eventually (\<lambda>x. x \<in> A (Suc n)) (nhds x)"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3369  | 
using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3370  | 
ultimately have "eventually (\<lambda>x. False) ?F"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3371  | 
by (auto simp add: eventually_inf)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3372  | 
with x show False  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3373  | 
by (simp add: eventually_False)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3374  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3375  | 
then have "i < s n i" "X (s n i) \<in> A (Suc n)"  | 
| 52624 | 3376  | 
unfolding s_def by (auto intro: someI2_ex)  | 
3377  | 
}  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3378  | 
note s = this  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3379  | 
def r \<equiv> "nat_rec (s 0 0) s"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3380  | 
have "subseq r"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3381  | 
by (auto simp: r_def s subseq_Suc_iff)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3382  | 
moreover  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3383  | 
have "(\<lambda>n. X (r n)) ----> x"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3384  | 
proof (rule topological_tendstoI)  | 
| 52624 | 3385  | 
fix S  | 
3386  | 
assume "open S" "x \<in> S"  | 
|
| 53282 | 3387  | 
with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially"  | 
3388  | 
by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3389  | 
moreover  | 
| 52624 | 3390  | 
    {
 | 
3391  | 
fix i  | 
|
3392  | 
assume "Suc 0 \<le> i"  | 
|
3393  | 
then have "X (r i) \<in> A i"  | 
|
3394  | 
by (cases i) (simp_all add: r_def s)  | 
|
3395  | 
}  | 
|
3396  | 
then have "eventually (\<lambda>i. X (r i) \<in> A i) sequentially"  | 
|
3397  | 
by (auto simp: eventually_sequentially)  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3398  | 
ultimately show "eventually (\<lambda>i. X (r i) \<in> S) sequentially"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3399  | 
by eventually_elim auto  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3400  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3401  | 
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
 | 
3402  | 
using `x \<in> U` by (auto simp: convergent_def comp_def)  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3403  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3404  | 
|
| 
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
 | 
3405  | 
lemma countably_compact_imp_acc_point:  | 
| 53291 | 3406  | 
assumes "countably_compact s"  | 
3407  | 
and "countable t"  | 
|
3408  | 
and "infinite t"  | 
|
3409  | 
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
 | 
3410  | 
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
 | 
3411  | 
proof (rule ccontr)  | 
| 53282 | 3412  | 
  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
 | 
3413  | 
note `countably_compact s`  | 
| 53282 | 3414  | 
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
 | 
3415  | 
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
 | 
3416  | 
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
 | 
3417  | 
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
 | 
3418  | 
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
 | 
3419  | 
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
 | 
3420  | 
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
 | 
3421  | 
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
 | 
3422  | 
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
 | 
3423  | 
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
 | 
3424  | 
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
 | 
3425  | 
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
 | 
3426  | 
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
 | 
3427  | 
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
 | 
3428  | 
unfolding C_def by (auto intro: countable_Collect_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
 | 
3429  | 
ultimately guess D by (rule countably_compactE)  | 
| 53282 | 3430  | 
  then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E"
 | 
3431  | 
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
 | 
3432  | 
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
 | 
3433  | 
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
 | 
3434  | 
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
 | 
3435  | 
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
 | 
3436  | 
using E by auto  | 
| 53282 | 3437  | 
ultimately show False using `infinite t`  | 
3438  | 
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
 | 
3439  | 
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
 | 
3440  | 
|
| 
 
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
 | 
3441  | 
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
 | 
3442  | 
fixes s :: "'a::first_countable_topology set"  | 
| 53291 | 3443  | 
assumes "\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s \<longrightarrow>  | 
3444  | 
(\<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
 | 
3445  | 
shows "seq_compact s"  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3446  | 
proof -  | 
| 52624 | 3447  | 
  {
 | 
3448  | 
fix f :: "nat \<Rightarrow> 'a"  | 
|
3449  | 
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
 | 
3450  | 
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
 | 
3451  | 
proof (cases "finite (range f)")  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3452  | 
case True  | 
| 
50941
 
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
 
hoelzl 
parents: 
50940 
diff
changeset
 | 
3453  | 
      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
 | 
3454  | 
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
 | 
3455  | 
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
 | 
3456  | 
using infinite_enumerate by blast  | 
| 53282 | 3457  | 
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
 | 
3458  | 
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
 | 
3459  | 
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
 | 
3460  | 
by auto  | 
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3461  | 
next  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3462  | 
case False  | 
| 53282 | 3463  | 
with f assms have "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)"  | 
3464  | 
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
 | 
3465  | 
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
 | 
3466  | 
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
 | 
3467  | 
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
 | 
3468  | 
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
 | 
3469  | 
qed  | 
| 
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3470  | 
}  | 
| 53282 | 3471  | 
then show ?thesis  | 
3472  | 
unfolding seq_compact_def by auto  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3473  | 
qed  | 
| 44075 | 3474  | 
|
| 
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
 | 
3475  | 
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
 | 
3476  | 
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
 | 
3477  | 
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
 | 
3478  | 
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
 | 
3479  | 
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
 | 
3480  | 
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
 | 
3481  | 
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
 | 
3482  | 
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
 | 
3483  | 
|
| 
 
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
 | 
3484  | 
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
 | 
3485  | 
fixes s :: "'a :: first_countable_topology set"  | 
| 53291 | 3486  | 
shows "seq_compact s \<longleftrightarrow>  | 
3487  | 
(\<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
 | 
3488  | 
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
 | 
3489  | 
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
 | 
3490  | 
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
 | 
3491  | 
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
 | 
3492  | 
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
 | 
3493  | 
|
| 
 
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
 | 
3494  | 
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
 | 
3495  | 
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
 | 
3496  | 
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
 | 
3497  | 
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
 | 
3498  | 
|
| 
 
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
 | 
3499  | 
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
 | 
3500  | 
  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
 | 
3501  | 
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
 | 
3502  | 
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
 | 
3503  | 
|
| 
50940
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3504  | 
subsubsection{* Total boundedness *}
 | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3505  | 
|
| 53282 | 3506  | 
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 | 3507  | 
unfolding Cauchy_def by metis  | 
3508  | 
||
| 53282 | 3509  | 
fun helper_1 :: "('a::metric_space set) \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> 'a"
 | 
3510  | 
where  | 
|
| 
50940
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3511  | 
"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
 | 
3512  | 
declare helper_1.simps[simp del]  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3513  | 
|
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3514  | 
lemma seq_compact_imp_totally_bounded:  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3515  | 
assumes "seq_compact s"  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3516  | 
shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k))"  | 
| 52624 | 3517  | 
proof (rule, rule, rule ccontr)  | 
3518  | 
fix e::real  | 
|
| 53282 | 3519  | 
assume "e > 0"  | 
3520  | 
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
 | 
3521  | 
def x \<equiv> "helper_1 s e"  | 
| 52624 | 3522  | 
  {
 | 
3523  | 
fix n  | 
|
| 
50940
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3524  | 
have "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"  | 
| 52624 | 3525  | 
proof (induct n rule: nat_less_induct)  | 
3526  | 
fix n  | 
|
3527  | 
def Q \<equiv> "(\<lambda>y. y \<in> s \<and> (\<forall>m<n. \<not> dist (x m) y < e))"  | 
|
3528  | 
assume as: "\<forall>m<n. x m \<in> s \<and> (\<forall>ma<m. \<not> dist (x ma) (x m) < e)"  | 
|
3529  | 
      have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)"
 | 
|
3530  | 
using assm  | 
|
3531  | 
apply simp  | 
|
3532  | 
        apply (erule_tac x="x ` {0 ..< n}" in allE)
 | 
|
3533  | 
using as  | 
|
3534  | 
apply auto  | 
|
3535  | 
done  | 
|
3536  | 
      then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)"
 | 
|
3537  | 
unfolding subset_eq by auto  | 
|
3538  | 
have "Q (x n)"  | 
|
3539  | 
unfolding x_def and helper_1.simps[of s e n]  | 
|
3540  | 
apply (rule someI2[where a=z])  | 
|
3541  | 
unfolding x_def[symmetric] and Q_def  | 
|
3542  | 
using z  | 
|
3543  | 
apply auto  | 
|
3544  | 
done  | 
|
| 53282 | 3545  | 
then show "x n \<in> s \<and> (\<forall>m<n. \<not> dist (x m) (x n) < e)"  | 
| 52624 | 3546  | 
unfolding Q_def by auto  | 
3547  | 
qed  | 
|
3548  | 
}  | 
|
| 53282 | 3549  | 
then have "\<forall>n::nat. x n \<in> s" and x:"\<forall>n. \<forall>m < n. \<not> (dist (x m) (x n) < e)"  | 
| 52624 | 3550  | 
by blast+  | 
3551  | 
then obtain l r where "l\<in>s" and r:"subseq r" and "((x \<circ> r) ---> l) sequentially"  | 
|
3552  | 
using assms(1)[unfolded seq_compact_def, THEN spec[where x=x]] by auto  | 
|
3553  | 
from this(3) have "Cauchy (x \<circ> r)"  | 
|
3554  | 
using LIMSEQ_imp_Cauchy by auto  | 
|
3555  | 
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"  | 
|
3556  | 
unfolding cauchy_def using `e>0` by auto  | 
|
| 
50940
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3557  | 
show False  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3558  | 
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
 | 
3559  | 
using r[unfolded subseq_def, THEN spec[where x=N], THEN spec[where x="N+1"]]  | 
| 52624 | 3560  | 
using x[THEN spec[where x="r (N+1)"], THEN spec[where x="r (N)"]]  | 
3561  | 
by auto  | 
|
| 
50940
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3562  | 
qed  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3563  | 
|
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3564  | 
subsubsection{* Heine-Borel theorem *}
 | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3565  | 
|
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3566  | 
lemma seq_compact_imp_heine_borel:  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3567  | 
fixes s :: "'a :: metric_space set"  | 
| 53282 | 3568  | 
assumes "seq_compact s"  | 
3569  | 
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
 | 
3570  | 
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
 | 
3571  | 
from seq_compact_imp_totally_bounded[OF `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
 | 
3572  | 
guess f unfolding choice_iff' .. note f = this  | 
| 
 
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
 | 
3573  | 
  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
 | 
3574  | 
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
 | 
3575  | 
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
 | 
3576  | 
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
 | 
3577  | 
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
 | 
3578  | 
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
 | 
3579  | 
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
 | 
3580  | 
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
 | 
3581  | 
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
 | 
3582  | 
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
 | 
3583  | 
next  | 
| 53282 | 3584  | 
fix T x  | 
3585  | 
assume T: "open T" "x \<in> T" and x: "x \<in> s"  | 
|
3586  | 
from openE[OF T] obtain e where "0 < e" "ball x e \<subseteq> T"  | 
|
3587  | 
by auto  | 
|
3588  | 
then have "0 < e / 2" "ball x (e / 2) \<subseteq> T"  | 
|
3589  | 
by auto  | 
|
3590  | 
from Rats_dense_in_real[OF `0 < e / 2`] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2"  | 
|
3591  | 
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
 | 
3592  | 
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
 | 
3593  | 
unfolding Union_image_eq by auto  | 
| 53282 | 3594  | 
from `r \<in> \<rat>` `0 < r` `k \<in> f r` have "ball k r \<in> K"  | 
3595  | 
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
 | 
3596  | 
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
 | 
3597  | 
proof (rule bexI[rotated], safe)  | 
| 53282 | 3598  | 
fix y  | 
3599  | 
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
 | 
3600  | 
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
 | 
3601  | 
by (intro dist_double[where x = k and d=e]) (auto simp: dist_commute)  | 
| 53282 | 3602  | 
with `ball x e \<subseteq> T` show "y \<in> T"  | 
3603  | 
by auto  | 
|
3604  | 
next  | 
|
3605  | 
show "x \<in> ball k r" by fact  | 
|
3606  | 
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
 | 
3607  | 
qed  | 
| 
50940
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3608  | 
qed  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3609  | 
|
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3610  | 
lemma compact_eq_seq_compact_metric:  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3611  | 
"compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s"  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3612  | 
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
 | 
3613  | 
|
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3614  | 
lemma compact_def:  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3615  | 
"compact (S :: 'a::metric_space set) \<longleftrightarrow>  | 
| 53640 | 3616  | 
(\<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
 | 
3617  | 
unfolding compact_eq_seq_compact_metric seq_compact_def by auto  | 
| 
 
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
 
hoelzl 
parents: 
50939 
diff
changeset
 | 
3618  | 
|
| 50944 | 3619  | 
subsubsection {* Complete the chain of compactness variants *}
 | 
3620  | 
||
3621  | 
lemma compact_eq_bolzano_weierstrass:  | 
|
3622  | 
fixes s :: "'a::metric_space set"  | 
|
| 53282 | 3623  | 
shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))"  | 
3624  | 
(is "?lhs = ?rhs")  | 
|
| 50944 | 3625  | 
proof  | 
| 52624 | 3626  | 
assume ?lhs  | 
| 53282 | 3627  | 
then show ?rhs  | 
3628  | 
using heine_borel_imp_bolzano_weierstrass[of s] by auto  | 
|
| 50944 | 3629  | 
next  | 
| 52624 | 3630  | 
assume ?rhs  | 
| 53282 | 3631  | 
then show ?lhs  | 
| 50944 | 3632  | 
unfolding compact_eq_seq_compact_metric by (rule bolzano_weierstrass_imp_seq_compact)  | 
3633  | 
qed  | 
|
3634  | 
||
3635  | 
lemma bolzano_weierstrass_imp_bounded:  | 
|
| 53282 | 3636  | 
"\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s"  | 
| 50944 | 3637  | 
using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass .  | 
3638  | 
||
| 54070 | 3639  | 
subsection {* Metric spaces with the Heine-Borel property *}
 | 
3640  | 
||
| 33175 | 3641  | 
text {*
 | 
3642  | 
A metric space (or topological vector space) is said to have the  | 
|
3643  | 
Heine-Borel property if every closed and bounded subset is compact.  | 
|
3644  | 
*}  | 
|
3645  | 
||
| 
44207
 
ea99698c2070
locale-ize some definitions, so perfect_space and heine_borel can inherit from the proper superclasses
 
huffman 
parents: 
44170 
diff
changeset
 | 
3646  | 
class heine_borel = metric_space +  | 
| 33175 | 3647  | 
assumes bounded_imp_convergent_subsequence:  | 
| 50998 | 3648  | 
"bounded (range f) \<Longrightarrow> \<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"  | 
| 33175 | 3649  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3650  | 
lemma bounded_closed_imp_seq_compact:  | 
| 33175 | 3651  | 
fixes s::"'a::heine_borel set"  | 
| 53282 | 3652  | 
assumes "bounded s"  | 
3653  | 
and "closed s"  | 
|
3654  | 
shows "seq_compact s"  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
3655  | 
proof (unfold seq_compact_def, clarify)  | 
| 53282 | 3656  | 
fix f :: "nat \<Rightarrow> 'a"  | 
3657  | 
assume f: "\<forall>n. f n \<in> s"  | 
|
3658  | 
with `bounded s` have "bounded (range f)"  | 
|
3659  | 
by (auto intro: bounded_subset)  | 
|
| 33175 | 3660  | 
obtain l r where r: "subseq r" and l: "((f \<circ> r) ---> l) sequentially"  | 
| 50998 | 3661  | 
using bounded_imp_convergent_subsequence [OF `bounded (range f)`] by auto  | 
| 53282 | 3662  | 
from f have fr: "\<forall>n. (f \<circ> r) n \<in> s"  | 
3663  | 
by simp  | 
|
| 33175 | 3664  | 
have "l \<in> s" using `closed s` fr l  | 
| 54070 | 3665  | 
by (rule closed_sequentially)  | 
| 33175 | 3666  | 
show "\<exists>l\<in>s. \<exists>r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"  | 
3667  | 
using `l \<in> s` r l by blast  | 
|
3668  | 
qed  | 
|
3669  | 
||
| 50944 | 3670  | 
lemma compact_eq_bounded_closed:  | 
3671  | 
fixes s :: "'a::heine_borel set"  | 
|
| 53291 | 3672  | 
shows "compact s \<longleftrightarrow> bounded s \<and> closed s"  | 
3673  | 
(is "?lhs = ?rhs")  | 
|
| 50944 | 3674  | 
proof  | 
| 52624 | 3675  | 
assume ?lhs  | 
| 53282 | 3676  | 
then show ?rhs  | 
| 52624 | 3677  | 
using compact_imp_closed compact_imp_bounded  | 
3678  | 
by blast  | 
|
| 50944 | 3679  | 
next  | 
| 52624 | 3680  | 
assume ?rhs  | 
| 53282 | 3681  | 
then show ?lhs  | 
| 52624 | 3682  | 
using bounded_closed_imp_seq_compact[of s]  | 
3683  | 
unfolding compact_eq_seq_compact_metric  | 
|
3684  | 
by auto  | 
|
| 50944 | 3685  | 
qed  | 
3686  | 
||
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
3687  | 
(* TODO: is this lemma necessary? *)  | 
| 50972 | 3688  | 
lemma bounded_increasing_convergent:  | 
3689  | 
fixes s :: "nat \<Rightarrow> real"  | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
3690  | 
  shows "bounded {s n| n. True} \<Longrightarrow> \<forall>n. s n \<le> s (Suc n) \<Longrightarrow> \<exists>l. s ----> l"
 | 
| 50972 | 3691  | 
using Bseq_mono_convergent[of s] incseq_Suc_iff[of s]  | 
3692  | 
by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def)  | 
|
| 33175 | 3693  | 
|
3694  | 
instance real :: heine_borel  | 
|
3695  | 
proof  | 
|
| 50998 | 3696  | 
fix f :: "nat \<Rightarrow> real"  | 
3697  | 
assume f: "bounded (range f)"  | 
|
| 50972 | 3698  | 
obtain r where r: "subseq r" "monoseq (f \<circ> r)"  | 
3699  | 
unfolding comp_def by (metis seq_monosub)  | 
|
3700  | 
then have "Bseq (f \<circ> r)"  | 
|
| 50998 | 3701  | 
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
 | 
3702  | 
with r show "\<exists>l r. subseq r \<and> (f \<circ> r) ----> l"  | 
| 50972 | 3703  | 
using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def)  | 
| 33175 | 3704  | 
qed  | 
3705  | 
||
3706  | 
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
 | 
3707  | 
fixes f :: "nat \<Rightarrow> 'a::euclidean_space"  | 
| 50998 | 3708  | 
assumes "bounded (range f)"  | 
| 53291 | 3709  | 
shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r.  | 
3710  | 
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
 | 
3711  | 
proof safe  | 
| 52624 | 3712  | 
fix d :: "'a set"  | 
| 53282 | 3713  | 
assume d: "d \<subseteq> Basis"  | 
3714  | 
with finite_Basis have "finite d"  | 
|
3715  | 
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
 | 
3716  | 
from this d show "\<exists>l::'a. \<exists>r. subseq r \<and>  | 
| 52624 | 3717  | 
(\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)"  | 
3718  | 
proof (induct d)  | 
|
3719  | 
case empty  | 
|
| 53282 | 3720  | 
then show ?case  | 
3721  | 
unfolding subseq_def by auto  | 
|
| 52624 | 3722  | 
next  | 
3723  | 
case (insert k d)  | 
|
| 53282 | 3724  | 
have k[intro]: "k \<in> Basis"  | 
3725  | 
using insert by auto  | 
|
3726  | 
have s': "bounded ((\<lambda>x. x \<bullet> k) ` range f)"  | 
|
3727  | 
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
 | 
3728  | 
by (auto intro!: bounded_linear_image bounded_linear_inner_left)  | 
| 53282 | 3729  | 
obtain l1::"'a" and r1 where r1: "subseq r1"  | 
3730  | 
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
 | 
3731  | 
using insert(3) using insert(4) by auto  | 
| 53282 | 3732  | 
have f': "\<forall>n. f (r1 n) \<bullet> k \<in> (\<lambda>x. x \<bullet> k) ` range f"  | 
3733  | 
by simp  | 
|
| 50998 | 3734  | 
have "bounded (range (\<lambda>i. f (r1 i) \<bullet> k))"  | 
3735  | 
by (metis (lifting) bounded_subset f' image_subsetI s')  | 
|
3736  | 
then obtain l2 r2 where r2:"subseq r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) \<bullet> k) ---> l2) sequentially"  | 
|
| 53282 | 3737  | 
using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) \<bullet> k"]  | 
3738  | 
by (auto simp: o_def)  | 
|
3739  | 
def r \<equiv> "r1 \<circ> r2"  | 
|
3740  | 
have r:"subseq r"  | 
|
| 33175 | 3741  | 
using r1 and r2 unfolding r_def o_def subseq_def by auto  | 
3742  | 
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
 | 
3743  | 
def l \<equiv> "(\<Sum>i\<in>Basis. (if i = k then l2 else l1\<bullet>i) *\<^sub>R i)::'a"  | 
| 52624 | 3744  | 
    {
 | 
3745  | 
fix e::real  | 
|
| 53282 | 3746  | 
assume "e > 0"  | 
3747  | 
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 | 3748  | 
by blast  | 
| 53282 | 3749  | 
from lr2 `e > 0` have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) \<bullet> k) l2 < e) sequentially"  | 
| 52624 | 3750  | 
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
 | 
3751  | 
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 | 3752  | 
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
 | 
3753  | 
have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"  | 
| 53282 | 3754  | 
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
 | 
3755  | 
by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def)  | 
| 33175 | 3756  | 
}  | 
3757  | 
ultimately show ?case by auto  | 
|
3758  | 
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
 | 
3759  | 
qed  | 
| 
 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 
hoelzl 
parents: 
37452 
diff
changeset
 | 
3760  | 
|
| 
 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 
hoelzl 
parents: 
37452 
diff
changeset
 | 
3761  | 
instance euclidean_space \<subseteq> heine_borel  | 
| 33175 | 3762  | 
proof  | 
| 50998 | 3763  | 
fix f :: "nat \<Rightarrow> 'a"  | 
3764  | 
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
 | 
3765  | 
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
 | 
3766  | 
and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially"  | 
| 50998 | 3767  | 
using compact_lemma [OF f] by blast  | 
| 52624 | 3768  | 
  {
 | 
3769  | 
fix e::real  | 
|
| 53282 | 3770  | 
assume "e > 0"  | 
3771  | 
    then have "e / real_of_nat DIM('a) > 0"
 | 
|
| 52624 | 3772  | 
by (auto intro!: divide_pos_pos 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
 | 
3773  | 
    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 | 3774  | 
by simp  | 
3775  | 
moreover  | 
|
| 52624 | 3776  | 
    {
 | 
3777  | 
fix n  | 
|
3778  | 
      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
 | 
3779  | 
have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))"  | 
| 52624 | 3780  | 
apply (subst euclidean_dist_l2)  | 
3781  | 
using zero_le_dist  | 
|
| 53282 | 3782  | 
apply (rule setL2_le_setsum)  | 
3783  | 
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
 | 
3784  | 
      also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
 | 
| 52624 | 3785  | 
apply (rule setsum_strict_mono)  | 
3786  | 
using n  | 
|
| 53282 | 3787  | 
apply auto  | 
3788  | 
done  | 
|
3789  | 
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
 | 
3790  | 
by auto  | 
| 33175 | 3791  | 
}  | 
3792  | 
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially"  | 
|
3793  | 
by (rule eventually_elim1)  | 
|
3794  | 
}  | 
|
| 53282 | 3795  | 
then have *: "((f \<circ> r) ---> l) sequentially"  | 
| 52624 | 3796  | 
unfolding o_def tendsto_iff by simp  | 
3797  | 
with r show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"  | 
|
3798  | 
by auto  | 
|
| 33175 | 3799  | 
qed  | 
3800  | 
||
3801  | 
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)"  | 
|
| 52624 | 3802  | 
unfolding bounded_def  | 
3803  | 
apply clarify  | 
|
3804  | 
apply (rule_tac x="a" in exI)  | 
|
3805  | 
apply (rule_tac x="e" in exI)  | 
|
3806  | 
apply clarsimp  | 
|
3807  | 
apply (drule (1) bspec)  | 
|
3808  | 
apply (simp add: dist_Pair_Pair)  | 
|
3809  | 
apply (erule order_trans [OF real_sqrt_sum_squares_ge1])  | 
|
3810  | 
done  | 
|
| 33175 | 3811  | 
|
3812  | 
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)"  | 
|
| 52624 | 3813  | 
unfolding bounded_def  | 
3814  | 
apply clarify  | 
|
3815  | 
apply (rule_tac x="b" in exI)  | 
|
3816  | 
apply (rule_tac x="e" in exI)  | 
|
3817  | 
apply clarsimp  | 
|
3818  | 
apply (drule (1) bspec)  | 
|
3819  | 
apply (simp add: dist_Pair_Pair)  | 
|
3820  | 
apply (erule order_trans [OF real_sqrt_sum_squares_ge2])  | 
|
3821  | 
done  | 
|
| 33175 | 3822  | 
|
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
37649 
diff
changeset
 | 
3823  | 
instance prod :: (heine_borel, heine_borel) heine_borel  | 
| 33175 | 3824  | 
proof  | 
| 50998 | 3825  | 
fix f :: "nat \<Rightarrow> 'a \<times> 'b"  | 
3826  | 
assume f: "bounded (range f)"  | 
|
| 53282 | 3827  | 
from f have s1: "bounded (range (fst \<circ> f))"  | 
3828  | 
unfolding image_comp by (rule bounded_fst)  | 
|
| 50998 | 3829  | 
obtain l1 r1 where r1: "subseq r1" and l1: "(\<lambda>n. fst (f (r1 n))) ----> l1"  | 
3830  | 
using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast  | 
|
3831  | 
from f have s2: "bounded (range (snd \<circ> f \<circ> r1))"  | 
|
3832  | 
by (auto simp add: image_comp intro: bounded_snd bounded_subset)  | 
|
| 53282 | 3833  | 
obtain l2 r2 where r2: "subseq r2" and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) ---> l2) sequentially"  | 
| 50998 | 3834  | 
using bounded_imp_convergent_subsequence [OF s2]  | 
| 33175 | 3835  | 
unfolding o_def by fast  | 
3836  | 
have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) ---> l1) sequentially"  | 
|
| 50972 | 3837  | 
using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def .  | 
| 33175 | 3838  | 
have l: "((f \<circ> (r1 \<circ> r2)) ---> (l1, l2)) sequentially"  | 
3839  | 
using tendsto_Pair [OF l1' l2] unfolding o_def by simp  | 
|
3840  | 
have r: "subseq (r1 \<circ> r2)"  | 
|
3841  | 
using r1 r2 unfolding subseq_def by simp  | 
|
3842  | 
show "\<exists>l r. subseq r \<and> ((f \<circ> r) ---> l) sequentially"  | 
|
3843  | 
using l r by fast  | 
|
3844  | 
qed  | 
|
3845  | 
||
| 54070 | 3846  | 
subsubsection {* Completeness *}
 | 
| 33175 | 3847  | 
|
| 52624 | 3848  | 
definition complete :: "'a::metric_space set \<Rightarrow> bool"  | 
3849  | 
where "complete s \<longleftrightarrow> (\<forall>f. (\<forall>n. f n \<in> s) \<and> Cauchy f \<longrightarrow> (\<exists>l\<in>s. f ----> l))"  | 
|
3850  | 
||
| 54070 | 3851  | 
lemma completeI:  | 
3852  | 
assumes "\<And>f. \<forall>n. f n \<in> s \<Longrightarrow> Cauchy f \<Longrightarrow> \<exists>l\<in>s. f ----> l"  | 
|
3853  | 
shows "complete s"  | 
|
3854  | 
using assms unfolding complete_def by fast  | 
|
3855  | 
||
3856  | 
lemma completeE:  | 
|
3857  | 
assumes "complete s" and "\<forall>n. f n \<in> s" and "Cauchy f"  | 
|
3858  | 
obtains l where "l \<in> s" and "f ----> l"  | 
|
3859  | 
using assms unfolding complete_def by fast  | 
|
3860  | 
||
| 52624 | 3861  | 
lemma compact_imp_complete:  | 
3862  | 
assumes "compact s"  | 
|
3863  | 
shows "complete s"  | 
|
3864  | 
proof -  | 
|
3865  | 
  {
 | 
|
3866  | 
fix f  | 
|
3867  | 
assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f"  | 
|
| 50971 | 3868  | 
from as(1) obtain l r where lr: "l\<in>s" "subseq r" "(f \<circ> r) ----> l"  | 
3869  | 
using assms unfolding compact_def by blast  | 
|
3870  | 
||
3871  | 
note lr' = seq_suble [OF lr(2)]  | 
|
3872  | 
||
| 52624 | 3873  | 
    {
 | 
| 53282 | 3874  | 
fix e :: real  | 
3875  | 
assume "e > 0"  | 
|
| 52624 | 3876  | 
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"  | 
3877  | 
unfolding cauchy_def  | 
|
| 53282 | 3878  | 
using `e > 0`  | 
3879  | 
apply (erule_tac x="e/2" in allE)  | 
|
| 52624 | 3880  | 
apply auto  | 
3881  | 
done  | 
|
3882  | 
from lr(3)[unfolded LIMSEQ_def, THEN spec[where x="e/2"]]  | 
|
| 53282 | 3883  | 
obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2"  | 
3884  | 
using `e > 0` by auto  | 
|
| 52624 | 3885  | 
      {
 | 
| 53282 | 3886  | 
fix n :: nat  | 
3887  | 
assume n: "n \<ge> max N M"  | 
|
3888  | 
have "dist ((f \<circ> r) n) l < e/2"  | 
|
3889  | 
using n M by auto  | 
|
3890  | 
moreover have "r n \<ge> N"  | 
|
3891  | 
using lr'[of n] n by auto  | 
|
3892  | 
then have "dist (f n) ((f \<circ> r) n) < e / 2"  | 
|
3893  | 
using N and n by auto  | 
|
| 52624 | 3894  | 
ultimately have "dist (f n) l < e"  | 
| 53282 | 3895  | 
using dist_triangle_half_r[of "f (r n)" "f n" e l]  | 
3896  | 
by (auto simp add: dist_commute)  | 
|
| 52624 | 3897  | 
}  | 
| 53282 | 3898  | 
then have "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast  | 
| 52624 | 3899  | 
}  | 
| 53282 | 3900  | 
then have "\<exists>l\<in>s. (f ---> l) sequentially" using `l\<in>s`  | 
| 52624 | 3901  | 
unfolding LIMSEQ_def by auto  | 
3902  | 
}  | 
|
| 53282 | 3903  | 
then show ?thesis unfolding complete_def by auto  | 
| 50971 | 3904  | 
qed  | 
3905  | 
||
3906  | 
lemma nat_approx_posE:  | 
|
3907  | 
fixes e::real  | 
|
3908  | 
assumes "0 < e"  | 
|
| 53282 | 3909  | 
obtains n :: nat where "1 / (Suc n) < e"  | 
| 50971 | 3910  | 
proof atomize_elim  | 
3911  | 
have " 1 / real (Suc (nat (ceiling (1/e)))) < 1 / (ceiling (1/e))"  | 
|
3912  | 
by (rule divide_strict_left_mono) (auto intro!: mult_pos_pos simp: `0 < e`)  | 
|
3913  | 
also have "1 / (ceiling (1/e)) \<le> 1 / (1/e)"  | 
|
3914  | 
by (rule divide_left_mono) (auto intro!: divide_pos_pos simp: `0 < e`)  | 
|
3915  | 
also have "\<dots> = e" by simp  | 
|
3916  | 
finally show "\<exists>n. 1 / real (Suc n) < e" ..  | 
|
3917  | 
qed  | 
|
3918  | 
||
3919  | 
lemma compact_eq_totally_bounded:  | 
|
3920  | 
"compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>((\<lambda>x. ball x e) ` k)))"  | 
|
3921  | 
(is "_ \<longleftrightarrow> ?rhs")  | 
|
3922  | 
proof  | 
|
3923  | 
assume assms: "?rhs"  | 
|
3924  | 
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)"  | 
|
3925  | 
by (auto simp: choice_iff')  | 
|
3926  | 
||
3927  | 
show "compact s"  | 
|
3928  | 
proof cases  | 
|
| 53282 | 3929  | 
    assume "s = {}"
 | 
3930  | 
then show "compact s" by (simp add: compact_def)  | 
|
| 50971 | 3931  | 
next  | 
3932  | 
    assume "s \<noteq> {}"
 | 
|
3933  | 
show ?thesis  | 
|
3934  | 
unfolding compact_def  | 
|
3935  | 
proof safe  | 
|
| 53282 | 3936  | 
fix f :: "nat \<Rightarrow> 'a"  | 
3937  | 
assume f: "\<forall>n. f n \<in> s"  | 
|
3938  | 
||
| 50971 | 3939  | 
def e \<equiv> "\<lambda>n. 1 / (2 * Suc n)"  | 
3940  | 
then have [simp]: "\<And>n. 0 < e n" by auto  | 
|
3941  | 
      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 | 3942  | 
      {
 | 
3943  | 
fix n U  | 
|
3944  | 
        assume "infinite {n. f n \<in> U}"
 | 
|
| 50971 | 3945  | 
        then have "\<exists>b\<in>k (e n). infinite {i\<in>{n. f n \<in> U}. f i \<in> ball b (e n)}"
 | 
3946  | 
using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq)  | 
|
3947  | 
then guess a ..  | 
|
3948  | 
        then have "\<exists>b. infinite {i. f i \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
 | 
|
3949  | 
by (intro exI[of _ "ball a (e n) \<inter> U"] exI[of _ a]) (auto simp: ac_simps)  | 
|
3950  | 
from someI_ex[OF this]  | 
|
3951  | 
        have "infinite {i. f i \<in> B n U}" "\<exists>x. B n U \<subseteq> ball x (e n) \<inter> U"
 | 
|
| 53282 | 3952  | 
unfolding B_def by auto  | 
3953  | 
}  | 
|
| 50971 | 3954  | 
note B = this  | 
3955  | 
||
3956  | 
def F \<equiv> "nat_rec (B 0 UNIV) B"  | 
|
| 53282 | 3957  | 
      {
 | 
3958  | 
fix n  | 
|
3959  | 
        have "infinite {i. f i \<in> F n}"
 | 
|
3960  | 
by (induct n) (auto simp: F_def B)  | 
|
3961  | 
}  | 
|
| 50971 | 3962  | 
then have F: "\<And>n. \<exists>x. F (Suc n) \<subseteq> ball x (e n) \<inter> F n"  | 
3963  | 
using B by (simp add: F_def)  | 
|
3964  | 
then have F_dec: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m"  | 
|
3965  | 
using decseq_SucI[of F] by (auto simp: decseq_def)  | 
|
3966  | 
||
3967  | 
obtain sel where sel: "\<And>k i. i < sel k i" "\<And>k i. f (sel k i) \<in> F k"  | 
|
3968  | 
proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI)  | 
|
3969  | 
fix k i  | 
|
3970  | 
        have "infinite ({n. f n \<in> F k} - {.. i})"
 | 
|
3971  | 
          using `infinite {n. f n \<in> F k}` by auto
 | 
|
3972  | 
from infinite_imp_nonempty[OF this]  | 
|
3973  | 
show "\<exists>x>i. f x \<in> F k"  | 
|
3974  | 
by (simp add: set_eq_iff not_le conj_commute)  | 
|
3975  | 
qed  | 
|
3976  | 
||
3977  | 
def t \<equiv> "nat_rec (sel 0 0) (\<lambda>n i. sel (Suc n) i)"  | 
|
3978  | 
have "subseq t"  | 
|
3979  | 
unfolding subseq_Suc_iff by (simp add: t_def sel)  | 
|
3980  | 
moreover have "\<forall>i. (f \<circ> t) i \<in> s"  | 
|
3981  | 
using f by auto  | 
|
3982  | 
moreover  | 
|
| 53282 | 3983  | 
      {
 | 
3984  | 
fix n  | 
|
3985  | 
have "(f \<circ> t) n \<in> F n"  | 
|
3986  | 
by (cases n) (simp_all add: t_def sel)  | 
|
3987  | 
}  | 
|
| 50971 | 3988  | 
note t = this  | 
3989  | 
||
3990  | 
have "Cauchy (f \<circ> t)"  | 
|
3991  | 
proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE)  | 
|
| 53282 | 3992  | 
fix r :: real and N n m  | 
3993  | 
assume "1 / Suc N < r" "Suc N \<le> n" "Suc N \<le> m"  | 
|
| 50971 | 3994  | 
then have "(f \<circ> t) n \<in> F (Suc N)" "(f \<circ> t) m \<in> F (Suc N)" "2 * e N < r"  | 
3995  | 
using F_dec t by (auto simp: e_def field_simps real_of_nat_Suc)  | 
|
3996  | 
with F[of N] obtain x where "dist x ((f \<circ> t) n) < e N" "dist x ((f \<circ> t) m) < e N"  | 
|
3997  | 
by (auto simp: subset_eq)  | 
|
3998  | 
with dist_triangle[of "(f \<circ> t) m" "(f \<circ> t) n" x] `2 * e N < r`  | 
|
3999  | 
show "dist ((f \<circ> t) m) ((f \<circ> t) n) < r"  | 
|
4000  | 
by (simp add: dist_commute)  | 
|
4001  | 
qed  | 
|
4002  | 
||
4003  | 
ultimately show "\<exists>l\<in>s. \<exists>r. subseq r \<and> (f \<circ> r) ----> l"  | 
|
4004  | 
using assms unfolding complete_def by blast  | 
|
4005  | 
qed  | 
|
4006  | 
qed  | 
|
4007  | 
qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded)  | 
|
| 33175 | 4008  | 
|
4009  | 
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 | 4010  | 
proof -  | 
4011  | 
  {
 | 
|
4012  | 
assume ?rhs  | 
|
4013  | 
    {
 | 
|
4014  | 
fix e::real  | 
|
| 33175 | 4015  | 
assume "e>0"  | 
4016  | 
with `?rhs` obtain N where N:"\<forall>n\<ge>N. dist (s n) (s N) < e/2"  | 
|
4017  | 
by (erule_tac x="e/2" in allE) auto  | 
|
| 53282 | 4018  | 
      {
 | 
4019  | 
fix n m  | 
|
| 33175 | 4020  | 
assume nm:"N \<le> m \<and> N \<le> n"  | 
| 53282 | 4021  | 
then have "dist (s m) (s n) < e" using N  | 
| 33175 | 4022  | 
using dist_triangle_half_l[of "s m" "s N" "e" "s n"]  | 
4023  | 
by blast  | 
|
4024  | 
}  | 
|
| 53282 | 4025  | 
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < e"  | 
| 33175 | 4026  | 
by blast  | 
4027  | 
}  | 
|
| 53282 | 4028  | 
then have ?lhs  | 
| 33175 | 4029  | 
unfolding cauchy_def  | 
4030  | 
by blast  | 
|
4031  | 
}  | 
|
| 53282 | 4032  | 
then show ?thesis  | 
| 33175 | 4033  | 
unfolding cauchy_def  | 
4034  | 
using dist_triangle_half_l  | 
|
4035  | 
by blast  | 
|
4036  | 
qed  | 
|
4037  | 
||
| 53282 | 4038  | 
lemma cauchy_imp_bounded:  | 
4039  | 
assumes "Cauchy s"  | 
|
4040  | 
shows "bounded (range s)"  | 
|
4041  | 
proof -  | 
|
4042  | 
from assms obtain N :: nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1"  | 
|
| 52624 | 4043  | 
unfolding cauchy_def  | 
4044  | 
apply (erule_tac x= 1 in allE)  | 
|
4045  | 
apply auto  | 
|
4046  | 
done  | 
|
| 53282 | 4047  | 
then have N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto  | 
| 33175 | 4048  | 
moreover  | 
| 52624 | 4049  | 
  have "bounded (s ` {0..N})"
 | 
4050  | 
    using finite_imp_bounded[of "s ` {1..N}"] by auto
 | 
|
| 33175 | 4051  | 
  then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
 | 
4052  | 
unfolding bounded_any_center [where a="s N"] by auto  | 
|
4053  | 
ultimately show "?thesis"  | 
|
4054  | 
unfolding bounded_any_center [where a="s N"]  | 
|
| 52624 | 4055  | 
apply (rule_tac x="max a 1" in exI)  | 
4056  | 
apply auto  | 
|
4057  | 
apply (erule_tac x=y in allE)  | 
|
4058  | 
apply (erule_tac x=y in ballE)  | 
|
4059  | 
apply auto  | 
|
4060  | 
done  | 
|
| 33175 | 4061  | 
qed  | 
4062  | 
||
4063  | 
instance heine_borel < complete_space  | 
|
4064  | 
proof  | 
|
4065  | 
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"  | 
|
| 53282 | 4066  | 
then have "bounded (range f)"  | 
| 34104 | 4067  | 
by (rule cauchy_imp_bounded)  | 
| 53282 | 4068  | 
then have "compact (closure (range f))"  | 
| 50971 | 4069  | 
unfolding compact_eq_bounded_closed by auto  | 
| 53282 | 4070  | 
then have "complete (closure (range f))"  | 
| 50971 | 4071  | 
by (rule compact_imp_complete)  | 
| 33175 | 4072  | 
moreover have "\<forall>n. f n \<in> closure (range f)"  | 
4073  | 
using closure_subset [of "range f"] by auto  | 
|
4074  | 
ultimately have "\<exists>l\<in>closure (range f). (f ---> l) sequentially"  | 
|
4075  | 
using `Cauchy f` unfolding complete_def by auto  | 
|
4076  | 
then show "convergent f"  | 
|
| 
36660
 
1cc4ab4b7ff7
make (X ----> L) an abbreviation for (X ---> L) sequentially
 
huffman 
parents: 
36659 
diff
changeset
 | 
4077  | 
unfolding convergent_def by auto  | 
| 33175 | 4078  | 
qed  | 
4079  | 
||
| 44632 | 4080  | 
instance euclidean_space \<subseteq> banach ..  | 
4081  | 
||
| 54070 | 4082  | 
lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
 | 
4083  | 
proof (rule completeI)  | 
|
4084  | 
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f"  | 
|
| 53282 | 4085  | 
then have "convergent f" by (rule Cauchy_convergent)  | 
| 54070 | 4086  | 
then show "\<exists>l\<in>UNIV. f ----> l" unfolding convergent_def by simp  | 
| 53282 | 4087  | 
qed  | 
4088  | 
||
4089  | 
lemma complete_imp_closed:  | 
|
4090  | 
assumes "complete s"  | 
|
4091  | 
shows "closed s"  | 
|
| 54070 | 4092  | 
proof (unfold closed_sequential_limits, clarify)  | 
4093  | 
fix f x assume "\<forall>n. f n \<in> s" and "f ----> x"  | 
|
4094  | 
from `f ----> x` have "Cauchy f"  | 
|
4095  | 
by (rule LIMSEQ_imp_Cauchy)  | 
|
4096  | 
with `complete s` and `\<forall>n. f n \<in> s` obtain l where "l \<in> s" and "f ----> l"  | 
|
4097  | 
by (rule completeE)  | 
|
4098  | 
from `f ----> x` and `f ----> l` have "x = l"  | 
|
4099  | 
by (rule LIMSEQ_unique)  | 
|
4100  | 
with `l \<in> s` show "x \<in> s"  | 
|
4101  | 
by simp  | 
|
4102  | 
qed  | 
|
4103  | 
||
4104  | 
lemma complete_inter_closed:  | 
|
4105  | 
assumes "complete s" and "closed t"  | 
|
4106  | 
shows "complete (s \<inter> t)"  | 
|
4107  | 
proof (rule completeI)  | 
|
4108  | 
fix f assume "\<forall>n. f n \<in> s \<inter> t" and "Cauchy f"  | 
|
4109  | 
then have "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t"  | 
|
4110  | 
by simp_all  | 
|
4111  | 
from `complete s` obtain l where "l \<in> s" and "f ----> l"  | 
|
4112  | 
using `\<forall>n. f n \<in> s` and `Cauchy f` by (rule completeE)  | 
|
4113  | 
from `closed t` and `\<forall>n. f n \<in> t` and `f ----> l` have "l \<in> t"  | 
|
4114  | 
by (rule closed_sequentially)  | 
|
4115  | 
with `l \<in> s` and `f ----> l` show "\<exists>l\<in>s \<inter> t. f ----> l"  | 
|
4116  | 
by fast  | 
|
4117  | 
qed  | 
|
4118  | 
||
4119  | 
lemma complete_closed_subset:  | 
|
4120  | 
assumes "closed s" and "s \<subseteq> t" and "complete t"  | 
|
4121  | 
shows "complete s"  | 
|
4122  | 
using assms complete_inter_closed [of t s] by (simp add: Int_absorb1)  | 
|
| 33175 | 4123  | 
|
4124  | 
lemma complete_eq_closed:  | 
|
| 54070 | 4125  | 
  fixes s :: "('a::complete_space) set"
 | 
4126  | 
shows "complete s \<longleftrightarrow> closed s"  | 
|
| 33175 | 4127  | 
proof  | 
| 54070 | 4128  | 
assume "closed s" then show "complete s"  | 
4129  | 
using subset_UNIV complete_UNIV by (rule complete_closed_subset)  | 
|
| 33175 | 4130  | 
next  | 
| 54070 | 4131  | 
assume "complete s" then show "closed s"  | 
4132  | 
by (rule complete_imp_closed)  | 
|
| 33175 | 4133  | 
qed  | 
4134  | 
||
4135  | 
lemma convergent_eq_cauchy:  | 
|
4136  | 
fixes s :: "nat \<Rightarrow> 'a::complete_space"  | 
|
| 44632 | 4137  | 
shows "(\<exists>l. (s ---> l) sequentially) \<longleftrightarrow> Cauchy s"  | 
4138  | 
unfolding Cauchy_convergent_iff convergent_def ..  | 
|
| 33175 | 4139  | 
|
4140  | 
lemma convergent_imp_bounded:  | 
|
4141  | 
fixes s :: "nat \<Rightarrow> 'a::metric_space"  | 
|
| 44632 | 4142  | 
shows "(s ---> l) sequentially \<Longrightarrow> bounded (range s)"  | 
| 
50939
 
ae7cd20ed118
replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
 
hoelzl 
parents: 
50938 
diff
changeset
 | 
4143  | 
by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy)  | 
| 33175 | 4144  | 
|
4145  | 
lemma compact_cball[simp]:  | 
|
4146  | 
fixes x :: "'a::heine_borel"  | 
|
| 54070 | 4147  | 
shows "compact (cball x e)"  | 
| 33175 | 4148  | 
using compact_eq_bounded_closed bounded_cball closed_cball  | 
4149  | 
by blast  | 
|
4150  | 
||
4151  | 
lemma compact_frontier_bounded[intro]:  | 
|
4152  | 
fixes s :: "'a::heine_borel set"  | 
|
| 54070 | 4153  | 
shows "bounded s \<Longrightarrow> compact (frontier s)"  | 
| 33175 | 4154  | 
unfolding frontier_def  | 
4155  | 
using compact_eq_bounded_closed  | 
|
4156  | 
by blast  | 
|
4157  | 
||
4158  | 
lemma compact_frontier[intro]:  | 
|
4159  | 
fixes s :: "'a::heine_borel set"  | 
|
| 53291 | 4160  | 
shows "compact s \<Longrightarrow> compact (frontier s)"  | 
| 33175 | 4161  | 
using compact_eq_bounded_closed compact_frontier_bounded  | 
4162  | 
by blast  | 
|
4163  | 
||
4164  | 
lemma frontier_subset_compact:  | 
|
4165  | 
fixes s :: "'a::heine_borel set"  | 
|
| 53291 | 4166  | 
shows "compact s \<Longrightarrow> frontier s \<subseteq> s"  | 
| 33175 | 4167  | 
using frontier_subset_closed compact_eq_bounded_closed  | 
4168  | 
by blast  | 
|
4169  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4170  | 
subsection {* Bounded closed nest property (proof does not use Heine-Borel) *}
 | 
| 33175 | 4171  | 
|
4172  | 
lemma bounded_closed_nest:  | 
|
| 54070 | 4173  | 
  fixes s :: "nat \<Rightarrow> ('a::heine_borel) set"
 | 
4174  | 
assumes "\<forall>n. closed (s n)"  | 
|
4175  | 
    and "\<forall>n. s n \<noteq> {}"
 | 
|
4176  | 
and "\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"  | 
|
4177  | 
and "bounded (s 0)"  | 
|
4178  | 
shows "\<exists>a. \<forall>n. a \<in> s n"  | 
|
| 52624 | 4179  | 
proof -  | 
| 54070 | 4180  | 
from assms(2) obtain x where x: "\<forall>n. x n \<in> s n"  | 
4181  | 
using choice[of "\<lambda>n x. x \<in> s n"] by auto  | 
|
4182  | 
from assms(4,1) have "seq_compact (s 0)"  | 
|
4183  | 
by (simp add: bounded_closed_imp_seq_compact)  | 
|
4184  | 
then obtain l r where lr: "l \<in> s 0" "subseq r" "(x \<circ> r) ----> l"  | 
|
4185  | 
using x and assms(3) unfolding seq_compact_def by blast  | 
|
4186  | 
have "\<forall>n. l \<in> s n"  | 
|
4187  | 
proof  | 
|
| 53282 | 4188  | 
fix n :: nat  | 
| 54070 | 4189  | 
have "closed (s n)"  | 
4190  | 
using assms(1) by simp  | 
|
4191  | 
moreover have "\<forall>i. (x \<circ> r) i \<in> s i"  | 
|
4192  | 
using x and assms(3) and lr(2) [THEN seq_suble] by auto  | 
|
4193  | 
then have "\<forall>i. (x \<circ> r) (i + n) \<in> s n"  | 
|
4194  | 
using assms(3) by (fast intro!: le_add2)  | 
|
4195  | 
moreover have "(\<lambda>i. (x \<circ> r) (i + n)) ----> l"  | 
|
4196  | 
using lr(3) by (rule LIMSEQ_ignore_initial_segment)  | 
|
4197  | 
ultimately show "l \<in> s n"  | 
|
4198  | 
by (rule closed_sequentially)  | 
|
4199  | 
qed  | 
|
4200  | 
then show ?thesis ..  | 
|
| 33175 | 4201  | 
qed  | 
4202  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4203  | 
text {* Decreasing case does not even need compactness, just completeness. *}
 | 
| 33175 | 4204  | 
|
4205  | 
lemma decreasing_closed_nest:  | 
|
| 54070 | 4206  | 
  fixes s :: "nat \<Rightarrow> ('a::complete_space) set"
 | 
| 53282 | 4207  | 
assumes  | 
| 54070 | 4208  | 
"\<forall>n. closed (s n)"  | 
4209  | 
    "\<forall>n. s n \<noteq> {}"
 | 
|
4210  | 
"\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"  | 
|
4211  | 
"\<forall>e>0. \<exists>n. \<forall>x\<in>s n. \<forall>y\<in>s n. dist x y < e"  | 
|
4212  | 
shows "\<exists>a. \<forall>n. a \<in> s n"  | 
|
4213  | 
proof -  | 
|
4214  | 
have "\<forall>n. \<exists>x. x \<in> s n"  | 
|
| 53282 | 4215  | 
using assms(2) by auto  | 
4216  | 
then have "\<exists>t. \<forall>n. t n \<in> s n"  | 
|
| 54070 | 4217  | 
using choice[of "\<lambda>n x. x \<in> s n"] by auto  | 
| 33175 | 4218  | 
then obtain t where t: "\<forall>n. t n \<in> s n" by auto  | 
| 53282 | 4219  | 
  {
 | 
4220  | 
fix e :: real  | 
|
4221  | 
assume "e > 0"  | 
|
4222  | 
then obtain N where N:"\<forall>x\<in>s N. \<forall>y\<in>s N. dist x y < e"  | 
|
4223  | 
using assms(4) by auto  | 
|
4224  | 
    {
 | 
|
4225  | 
fix m n :: nat  | 
|
4226  | 
assume "N \<le> m \<and> N \<le> n"  | 
|
4227  | 
then have "t m \<in> s N" "t n \<in> s N"  | 
|
4228  | 
using assms(3) t unfolding subset_eq t by blast+  | 
|
4229  | 
then have "dist (t m) (t n) < e"  | 
|
4230  | 
using N by auto  | 
|
| 33175 | 4231  | 
}  | 
| 53282 | 4232  | 
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (t m) (t n) < e"  | 
4233  | 
by auto  | 
|
| 33175 | 4234  | 
}  | 
| 53282 | 4235  | 
then have "Cauchy t"  | 
4236  | 
unfolding cauchy_def by auto  | 
|
4237  | 
then obtain l where l:"(t ---> l) sequentially"  | 
|
| 54070 | 4238  | 
using complete_UNIV unfolding complete_def by auto  | 
| 53282 | 4239  | 
  {
 | 
4240  | 
fix n :: nat  | 
|
4241  | 
    {
 | 
|
4242  | 
fix e :: real  | 
|
4243  | 
assume "e > 0"  | 
|
4244  | 
then obtain N :: nat where N: "\<forall>n\<ge>N. dist (t n) l < e"  | 
|
4245  | 
using l[unfolded LIMSEQ_def] by auto  | 
|
4246  | 
have "t (max n N) \<in> s n"  | 
|
4247  | 
using assms(3)  | 
|
4248  | 
unfolding subset_eq  | 
|
4249  | 
apply (erule_tac x=n in allE)  | 
|
4250  | 
apply (erule_tac x="max n N" in allE)  | 
|
4251  | 
using t  | 
|
4252  | 
apply auto  | 
|
4253  | 
done  | 
|
4254  | 
then have "\<exists>y\<in>s n. dist y l < e"  | 
|
4255  | 
apply (rule_tac x="t (max n N)" in bexI)  | 
|
4256  | 
using N  | 
|
4257  | 
apply auto  | 
|
4258  | 
done  | 
|
| 33175 | 4259  | 
}  | 
| 53282 | 4260  | 
then have "l \<in> s n"  | 
4261  | 
using closed_approachable[of "s n" l] assms(1) by auto  | 
|
| 33175 | 4262  | 
}  | 
4263  | 
then show ?thesis by auto  | 
|
4264  | 
qed  | 
|
4265  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4266  | 
text {* Strengthen it to the intersection actually being a singleton. *}
 | 
| 33175 | 4267  | 
|
4268  | 
lemma decreasing_closed_nest_sing:  | 
|
| 44632 | 4269  | 
fixes s :: "nat \<Rightarrow> 'a::complete_space set"  | 
| 53282 | 4270  | 
assumes  | 
4271  | 
"\<forall>n. closed(s n)"  | 
|
4272  | 
    "\<forall>n. s n \<noteq> {}"
 | 
|
| 54070 | 4273  | 
"\<forall>m n. m \<le> n \<longrightarrow> s n \<subseteq> s m"  | 
| 53282 | 4274  | 
"\<forall>e>0. \<exists>n. \<forall>x \<in> (s n). \<forall> y\<in>(s n). dist x y < e"  | 
| 34104 | 4275  | 
  shows "\<exists>a. \<Inter>(range s) = {a}"
 | 
| 53282 | 4276  | 
proof -  | 
4277  | 
obtain a where a: "\<forall>n. a \<in> s n"  | 
|
4278  | 
using decreasing_closed_nest[of s] using assms by auto  | 
|
4279  | 
  {
 | 
|
4280  | 
fix b  | 
|
4281  | 
assume b: "b \<in> \<Inter>(range s)"  | 
|
4282  | 
    {
 | 
|
4283  | 
fix e :: real  | 
|
4284  | 
assume "e > 0"  | 
|
4285  | 
then have "dist a b < e"  | 
|
4286  | 
using assms(4) and b and a by blast  | 
|
| 33175 | 4287  | 
}  | 
| 53282 | 4288  | 
then have "dist a b = 0"  | 
4289  | 
by (metis dist_eq_0_iff dist_nz less_le)  | 
|
| 33175 | 4290  | 
}  | 
| 53282 | 4291  | 
  with a have "\<Inter>(range s) = {a}"
 | 
4292  | 
unfolding image_def by auto  | 
|
4293  | 
then show ?thesis ..  | 
|
| 33175 | 4294  | 
qed  | 
4295  | 
||
4296  | 
text{* Cauchy-type criteria for uniform convergence. *}
 | 
|
4297  | 
||
| 53282 | 4298  | 
lemma uniformly_convergent_eq_cauchy:  | 
4299  | 
fixes s::"nat \<Rightarrow> 'b \<Rightarrow> 'a::complete_space"  | 
|
4300  | 
shows  | 
|
| 53291 | 4301  | 
"(\<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>  | 
4302  | 
(\<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 | 4303  | 
(is "?lhs = ?rhs")  | 
4304  | 
proof  | 
|
| 33175 | 4305  | 
assume ?lhs  | 
| 53282 | 4306  | 
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"  | 
4307  | 
by auto  | 
|
4308  | 
  {
 | 
|
4309  | 
fix e :: real  | 
|
4310  | 
assume "e > 0"  | 
|
4311  | 
then obtain N :: nat where N: "\<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist (s n x) (l x) < e / 2"  | 
|
4312  | 
using l[THEN spec[where x="e/2"]] by auto  | 
|
4313  | 
    {
 | 
|
4314  | 
fix n m :: nat and x :: "'b"  | 
|
4315  | 
assume "N \<le> m \<and> N \<le> n \<and> P x"  | 
|
4316  | 
then have "dist (s m x) (s n x) < e"  | 
|
| 33175 | 4317  | 
using N[THEN spec[where x=m], THEN spec[where x=x]]  | 
4318  | 
using N[THEN spec[where x=n], THEN spec[where x=x]]  | 
|
| 53282 | 4319  | 
using dist_triangle_half_l[of "s m x" "l x" e "s n x"] by auto  | 
4320  | 
}  | 
|
4321  | 
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  | 
|
4322  | 
}  | 
|
4323  | 
then show ?rhs by auto  | 
|
| 33175 | 4324  | 
next  | 
4325  | 
assume ?rhs  | 
|
| 53282 | 4326  | 
then have "\<forall>x. P x \<longrightarrow> Cauchy (\<lambda>n. s n x)"  | 
4327  | 
unfolding cauchy_def  | 
|
4328  | 
apply auto  | 
|
4329  | 
apply (erule_tac x=e in allE)  | 
|
4330  | 
apply auto  | 
|
4331  | 
done  | 
|
4332  | 
then obtain l where l: "\<forall>x. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l x) sequentially"  | 
|
| 53291 | 4333  | 
unfolding convergent_eq_cauchy[symmetric]  | 
| 53282 | 4334  | 
using choice[of "\<lambda>x l. P x \<longrightarrow> ((\<lambda>n. s n x) ---> l) sequentially"]  | 
4335  | 
by auto  | 
|
4336  | 
  {
 | 
|
4337  | 
fix e :: real  | 
|
4338  | 
assume "e > 0"  | 
|
| 33175 | 4339  | 
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"  | 
4340  | 
using `?rhs`[THEN spec[where x="e/2"]] by auto  | 
|
| 53282 | 4341  | 
    {
 | 
4342  | 
fix x  | 
|
4343  | 
assume "P x"  | 
|
| 33175 | 4344  | 
then obtain M where M:"\<forall>n\<ge>M. dist (s n x) (l x) < e/2"  | 
| 53282 | 4345  | 
using l[THEN spec[where x=x], unfolded LIMSEQ_def] and `e > 0`  | 
4346  | 
by (auto elim!: allE[where x="e/2"])  | 
|
4347  | 
fix n :: nat  | 
|
4348  | 
assume "n \<ge> N"  | 
|
4349  | 
then have "dist(s n x)(l x) < e"  | 
|
4350  | 
using `P x`and N[THEN spec[where x=n], THEN spec[where x="N+M"], THEN spec[where x=x]]  | 
|
4351  | 
using M[THEN spec[where x="N+M"]] and dist_triangle_half_l[of "s n x" "s (N+M) x" e "l x"]  | 
|
4352  | 
by (auto simp add: dist_commute)  | 
|
4353  | 
}  | 
|
4354  | 
then have "\<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e"  | 
|
4355  | 
by auto  | 
|
4356  | 
}  | 
|
4357  | 
then show ?lhs by auto  | 
|
| 33175 | 4358  | 
qed  | 
4359  | 
||
4360  | 
lemma uniformly_cauchy_imp_uniformly_convergent:  | 
|
| 51102 | 4361  | 
fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::complete_space"  | 
| 33175 | 4362  | 
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 | 4363  | 
and "\<forall>x. P x --> (\<forall>e>0. \<exists>N. \<forall>n. N \<le> n \<longrightarrow> dist(s n x)(l x) < e)"  | 
4364  | 
shows "\<forall>e>0. \<exists>N. \<forall>n x. N \<le> n \<and> P x \<longrightarrow> dist(s n x)(l x) < e"  | 
|
| 53282 | 4365  | 
proof -  | 
| 33175 | 4366  | 
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 | 4367  | 
using assms(1) unfolding uniformly_convergent_eq_cauchy[symmetric] by auto  | 
| 33175 | 4368  | 
moreover  | 
| 53282 | 4369  | 
  {
 | 
4370  | 
fix x  | 
|
4371  | 
assume "P x"  | 
|
4372  | 
then have "l x = l' x"  | 
|
4373  | 
using tendsto_unique[OF trivial_limit_sequentially, of "\<lambda>n. s n x" "l x" "l' x"]  | 
|
4374  | 
using l and assms(2) unfolding LIMSEQ_def by blast  | 
|
4375  | 
}  | 
|
| 33175 | 4376  | 
ultimately show ?thesis by auto  | 
4377  | 
qed  | 
|
4378  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4379  | 
|
| 36437 | 4380  | 
subsection {* Continuity *}
 | 
4381  | 
||
| 33175 | 4382  | 
text{* Derive the epsilon-delta forms, which we often use as "definitions" *}
 | 
4383  | 
||
4384  | 
lemma continuous_within_eps_delta:  | 
|
4385  | 
"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)"  | 
|
4386  | 
unfolding continuous_within and Lim_within  | 
|
| 53282 | 4387  | 
apply auto  | 
| 53291 | 4388  | 
unfolding dist_nz[symmetric]  | 
| 53282 | 4389  | 
apply (auto del: allE elim!:allE)  | 
4390  | 
apply(rule_tac x=d in exI)  | 
|
4391  | 
apply auto  | 
|
4392  | 
done  | 
|
4393  | 
||
4394  | 
lemma continuous_at_eps_delta:  | 
|
4395  | 
"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 | 4396  | 
using continuous_within_eps_delta [of x UNIV f] by simp  | 
| 33175 | 4397  | 
|
4398  | 
text{* Versions in terms of open balls. *}
 | 
|
4399  | 
||
4400  | 
lemma continuous_within_ball:  | 
|
| 53282 | 4401  | 
"continuous (at x within s) f \<longleftrightarrow>  | 
4402  | 
(\<forall>e > 0. \<exists>d > 0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)"  | 
|
4403  | 
(is "?lhs = ?rhs")  | 
|
| 33175 | 4404  | 
proof  | 
4405  | 
assume ?lhs  | 
|
| 53282 | 4406  | 
  {
 | 
4407  | 
fix e :: real  | 
|
4408  | 
assume "e > 0"  | 
|
| 33175 | 4409  | 
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"  | 
4410  | 
using `?lhs`[unfolded continuous_within Lim_within] by auto  | 
|
| 53282 | 4411  | 
    {
 | 
4412  | 
fix y  | 
|
4413  | 
assume "y \<in> f ` (ball x d \<inter> s)"  | 
|
4414  | 
then have "y \<in> ball (f x) e"  | 
|
4415  | 
using d(2)  | 
|
| 53291 | 4416  | 
unfolding dist_nz[symmetric]  | 
| 53282 | 4417  | 
apply (auto simp add: dist_commute)  | 
4418  | 
apply (erule_tac x=xa in ballE)  | 
|
4419  | 
apply auto  | 
|
4420  | 
using `e > 0`  | 
|
4421  | 
apply auto  | 
|
4422  | 
done  | 
|
| 33175 | 4423  | 
}  | 
| 53282 | 4424  | 
then have "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e"  | 
4425  | 
using `d > 0`  | 
|
4426  | 
unfolding subset_eq ball_def by (auto simp add: dist_commute)  | 
|
4427  | 
}  | 
|
4428  | 
then show ?rhs by auto  | 
|
| 33175 | 4429  | 
next  | 
| 53282 | 4430  | 
assume ?rhs  | 
4431  | 
then show ?lhs  | 
|
4432  | 
unfolding continuous_within Lim_within ball_def subset_eq  | 
|
4433  | 
apply (auto simp add: dist_commute)  | 
|
4434  | 
apply (erule_tac x=e in allE)  | 
|
4435  | 
apply auto  | 
|
4436  | 
done  | 
|
| 33175 | 4437  | 
qed  | 
4438  | 
||
4439  | 
lemma continuous_at_ball:  | 
|
4440  | 
"continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs")  | 
|
4441  | 
proof  | 
|
| 53282 | 4442  | 
assume ?lhs  | 
4443  | 
then show ?rhs  | 
|
4444  | 
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball  | 
|
4445  | 
apply auto  | 
|
4446  | 
apply (erule_tac x=e in allE)  | 
|
4447  | 
apply auto  | 
|
4448  | 
apply (rule_tac x=d in exI)  | 
|
4449  | 
apply auto  | 
|
4450  | 
apply (erule_tac x=xa in allE)  | 
|
4451  | 
apply (auto simp add: dist_commute dist_nz)  | 
|
| 53291 | 4452  | 
unfolding dist_nz[symmetric]  | 
| 53282 | 4453  | 
apply auto  | 
4454  | 
done  | 
|
| 33175 | 4455  | 
next  | 
| 53282 | 4456  | 
assume ?rhs  | 
4457  | 
then show ?lhs  | 
|
4458  | 
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball  | 
|
4459  | 
apply auto  | 
|
4460  | 
apply (erule_tac x=e in allE)  | 
|
4461  | 
apply auto  | 
|
4462  | 
apply (rule_tac x=d in exI)  | 
|
4463  | 
apply auto  | 
|
4464  | 
apply (erule_tac x="f xa" in allE)  | 
|
4465  | 
apply (auto simp add: dist_commute dist_nz)  | 
|
4466  | 
done  | 
|
| 33175 | 4467  | 
qed  | 
4468  | 
||
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4469  | 
text{* Define setwise continuity in terms of limits within the set. *}
 | 
| 33175 | 4470  | 
|
| 36359 | 4471  | 
lemma continuous_on_iff:  | 
4472  | 
"continuous_on s f \<longleftrightarrow>  | 
|
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4473  | 
(\<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 | 4474  | 
unfolding continuous_on_def Lim_within  | 
4475  | 
apply (intro ball_cong [OF refl] all_cong ex_cong)  | 
|
4476  | 
apply (rename_tac y, case_tac "y = x")  | 
|
4477  | 
apply simp  | 
|
4478  | 
apply (simp add: dist_nz)  | 
|
4479  | 
done  | 
|
4480  | 
||
4481  | 
definition uniformly_continuous_on :: "'a set \<Rightarrow> ('a::metric_space \<Rightarrow> 'b::metric_space) \<Rightarrow> bool"
 | 
|
4482  | 
where "uniformly_continuous_on s f \<longleftrightarrow>  | 
|
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4483  | 
(\<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
 | 
4484  | 
|
| 33175 | 4485  | 
text{* Some simple consequential lemmas. *}
 | 
4486  | 
||
4487  | 
lemma uniformly_continuous_imp_continuous:  | 
|
| 53282 | 4488  | 
"uniformly_continuous_on s f \<Longrightarrow> continuous_on s f"  | 
| 36359 | 4489  | 
unfolding uniformly_continuous_on_def continuous_on_iff by blast  | 
| 33175 | 4490  | 
|
4491  | 
lemma continuous_at_imp_continuous_within:  | 
|
| 53282 | 4492  | 
"continuous (at x) f \<Longrightarrow> continuous (at x within s) f"  | 
| 33175 | 4493  | 
unfolding continuous_within continuous_at using Lim_at_within by auto  | 
4494  | 
||
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4495  | 
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
 | 
4496  | 
by simp  | 
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4497  | 
|
| 
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4498  | 
lemmas continuous_on = continuous_on_def -- "legacy theorem name"  | 
| 33175 | 4499  | 
|
4500  | 
lemma continuous_within_subset:  | 
|
| 53282 | 4501  | 
"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
 | 
4502  | 
unfolding continuous_within by(metis tendsto_within_subset)  | 
| 33175 | 4503  | 
|
4504  | 
lemma continuous_on_interior:  | 
|
| 53282 | 4505  | 
"continuous_on s f \<Longrightarrow> x \<in> interior s \<Longrightarrow> continuous (at x) f"  | 
4506  | 
apply (erule interiorE)  | 
|
4507  | 
apply (drule (1) continuous_on_subset)  | 
|
4508  | 
apply (simp add: continuous_on_eq_continuous_at)  | 
|
4509  | 
done  | 
|
| 33175 | 4510  | 
|
4511  | 
lemma continuous_on_eq:  | 
|
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4512  | 
"(\<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
 | 
4513  | 
unfolding continuous_on_def tendsto_def eventually_at_topological  | 
| 
36440
 
89a70297564d
simplify definition of continuous_on; generalize some lemmas
 
huffman 
parents: 
36439 
diff
changeset
 | 
4514  | 
by simp  | 
| 33175 | 4515  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4516  | 
text {* Characterization of various kinds of continuity in terms of sequences. *}
 | 
| 33175 | 4517  | 
|
4518  | 
lemma continuous_within_sequentially:  | 
|
| 
44533
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4519  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"  | 
| 33175 | 4520  | 
shows "continuous (at a within s) f \<longleftrightarrow>  | 
| 53282 | 4521  | 
(\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x ---> a) sequentially  | 
| 53640 | 4522  | 
\<longrightarrow> ((f \<circ> x) ---> f a) sequentially)"  | 
| 53282 | 4523  | 
(is "?lhs = ?rhs")  | 
| 33175 | 4524  | 
proof  | 
4525  | 
assume ?lhs  | 
|
| 53282 | 4526  | 
  {
 | 
4527  | 
fix x :: "nat \<Rightarrow> 'a"  | 
|
4528  | 
assume x: "\<forall>n. x n \<in> s" "\<forall>e>0. eventually (\<lambda>n. dist (x n) a < e) sequentially"  | 
|
4529  | 
fix T :: "'b set"  | 
|
4530  | 
assume "open T" and "f a \<in> T"  | 
|
| 
44533
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4531  | 
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
 | 
4532  | 
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
 | 
4533  | 
have "eventually (\<lambda>n. dist (x n) a < d) sequentially"  | 
| 
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4534  | 
using x(2) `d>0` by simp  | 
| 53282 | 4535  | 
then have "eventually (\<lambda>n. (f \<circ> x) n \<in> T) sequentially"  | 
| 46887 | 4536  | 
proof eventually_elim  | 
| 53282 | 4537  | 
case (elim n)  | 
4538  | 
then show ?case  | 
|
| 53291 | 4539  | 
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
 | 
4540  | 
qed  | 
| 33175 | 4541  | 
}  | 
| 53282 | 4542  | 
then show ?rhs  | 
4543  | 
unfolding tendsto_iff tendsto_def by simp  | 
|
| 33175 | 4544  | 
next  | 
| 53282 | 4545  | 
assume ?rhs  | 
4546  | 
then show ?lhs  | 
|
| 
44533
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4547  | 
unfolding continuous_within tendsto_def [where l="f a"]  | 
| 
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4548  | 
by (simp add: sequentially_imp_eventually_within)  | 
| 33175 | 4549  | 
qed  | 
4550  | 
||
4551  | 
lemma continuous_at_sequentially:  | 
|
| 
44533
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4552  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"  | 
| 53291 | 4553  | 
shows "continuous (at a) f \<longleftrightarrow>  | 
| 53640 | 4554  | 
(\<forall>x. (x ---> a) sequentially --> ((f \<circ> x) ---> f a) sequentially)"  | 
| 45031 | 4555  | 
using continuous_within_sequentially[of a UNIV f] by simp  | 
| 33175 | 4556  | 
|
4557  | 
lemma continuous_on_sequentially:  | 
|
| 
44533
 
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
 
huffman 
parents: 
44531 
diff
changeset
 | 
4558  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space"  | 
| 36359 | 4559  | 
shows "continuous_on s f \<longleftrightarrow>  | 
4560  | 
(\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x ---> a) sequentially  | 
|
| 53640 | 4561  | 
--> ((f \<circ> x) ---> f a) sequentially)"  | 
| 53291 | 4562  | 
(is "?lhs = ?rhs")  | 
| 33175 | 4563  | 
proof  | 
| 53282 | 4564  | 
assume ?rhs  | 
4565  | 
then show ?lhs  | 
|
4566  | 
using continuous_within_sequentially[of _ s f]  | 
|
4567  | 
unfolding continuous_on_eq_continuous_within  | 
|
4568  | 
by auto  | 
|
| 33175 | 4569  | 
next  | 
| 53282 | 4570  | 
assume ?lhs  | 
4571  | 
then show ?rhs  | 
|
4572  | 
unfolding continuous_on_eq_continuous_within  | 
|
4573  | 
using continuous_within_sequentially[of _ s f]  | 
|
4574  | 
by auto  | 
|
| 33175 | 4575  | 
qed  | 
4576  | 
||
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4577  | 
lemma uniformly_continuous_on_sequentially:  | 
| 36441 | 4578  | 
"uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and>  | 
4579  | 
((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially  | 
|
4580  | 
\<longrightarrow> ((\<lambda>n. dist (f(x n)) (f(y n))) ---> 0) sequentially)" (is "?lhs = ?rhs")  | 
|
| 33175 | 4581  | 
proof  | 
4582  | 
assume ?lhs  | 
|
| 53282 | 4583  | 
  {
 | 
4584  | 
fix x y  | 
|
4585  | 
assume x: "\<forall>n. x n \<in> s"  | 
|
4586  | 
and y: "\<forall>n. y n \<in> s"  | 
|
4587  | 
and xy: "((\<lambda>n. dist (x n) (y n)) ---> 0) sequentially"  | 
|
4588  | 
    {
 | 
|
4589  | 
fix e :: real  | 
|
4590  | 
assume "e > 0"  | 
|
4591  | 
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 | 4592  | 
using `?lhs`[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto  | 
| 53282 | 4593  | 
obtain N where N: "\<forall>n\<ge>N. dist (x n) (y n) < d"  | 
4594  | 
using xy[unfolded LIMSEQ_def dist_norm] and `d>0` by auto  | 
|
4595  | 
      {
 | 
|
4596  | 
fix n  | 
|
4597  | 
assume "n\<ge>N"  | 
|
4598  | 
then have "dist (f (x n)) (f (y n)) < e"  | 
|
4599  | 
using N[THEN spec[where x=n]]  | 
|
4600  | 
using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]]  | 
|
4601  | 
using x and y  | 
|
4602  | 
unfolding dist_commute  | 
|
4603  | 
by simp  | 
|
4604  | 
}  | 
|
4605  | 
then have "\<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"  | 
|
4606  | 
by auto  | 
|
4607  | 
}  | 
|
4608  | 
then have "((\<lambda>n. dist (f(x n)) (f(y n))) ---> 0) sequentially"  | 
|
4609  | 
unfolding LIMSEQ_def and dist_real_def by auto  | 
|
4610  | 
}  | 
|
4611  | 
then show ?rhs by auto  | 
|
| 33175 | 4612  | 
next  | 
4613  | 
assume ?rhs  | 
|
| 53282 | 4614  | 
  {
 | 
4615  | 
assume "\<not> ?lhs"  | 
|
4616  | 
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"  | 
|
4617  | 
unfolding uniformly_continuous_on_def by auto  | 
|
4618  | 
then obtain fa where fa:  | 
|
4619  | 
"\<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"  | 
|
4620  | 
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"]  | 
|
4621  | 
unfolding Bex_def  | 
|
| 33175 | 4622  | 
by (auto simp add: dist_commute)  | 
4623  | 
def x \<equiv> "\<lambda>n::nat. fst (fa (inverse (real n + 1)))"  | 
|
4624  | 
def y \<equiv> "\<lambda>n::nat. snd (fa (inverse (real n + 1)))"  | 
|
| 53282 | 4625  | 
have xyn: "\<forall>n. x n \<in> s \<and> y n \<in> s"  | 
4626  | 
and xy0: "\<forall>n. dist (x n) (y n) < inverse (real n + 1)"  | 
|
4627  | 
and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e"  | 
|
4628  | 
unfolding x_def and y_def using fa  | 
|
4629  | 
by auto  | 
|
4630  | 
    {
 | 
|
4631  | 
fix e :: real  | 
|
4632  | 
assume "e > 0"  | 
|
4633  | 
then obtain N :: nat where "N \<noteq> 0" and N: "0 < inverse (real N) \<and> inverse (real N) < e"  | 
|
4634  | 
unfolding real_arch_inv[of e] by auto  | 
|
4635  | 
      {
 | 
|
4636  | 
fix n :: nat  | 
|
4637  | 
assume "n \<ge> N"  | 
|
4638  | 
then have "inverse (real n + 1) < inverse (real N)"  | 
|
4639  | 
using real_of_nat_ge_zero and `N\<noteq>0` by auto  | 
|
| 33175 | 4640  | 
also have "\<dots> < e" using N by auto  | 
4641  | 
finally have "inverse (real n + 1) < e" by auto  | 
|
| 53282 | 4642  | 
then have "dist (x n) (y n) < e"  | 
4643  | 
using xy0[THEN spec[where x=n]] by auto  | 
|
4644  | 
}  | 
|
4645  | 
then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto  | 
|
4646  | 
}  | 
|
4647  | 
then have "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e"  | 
|
4648  | 
using `?rhs`[THEN spec[where x=x], THEN spec[where x=y]] and xyn  | 
|
4649  | 
unfolding LIMSEQ_def dist_real_def by auto  | 
|
4650  | 
then have False using fxy and `e>0` by auto  | 
|
4651  | 
}  | 
|
4652  | 
then show ?lhs  | 
|
4653  | 
unfolding uniformly_continuous_on_def by blast  | 
|
| 33175 | 4654  | 
qed  | 
4655  | 
||
4656  | 
text{* The usual transformation theorems. *}
 | 
|
4657  | 
||
4658  | 
lemma continuous_transform_within:  | 
|
| 36667 | 4659  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"  | 
| 53282 | 4660  | 
assumes "0 < d"  | 
4661  | 
and "x \<in> s"  | 
|
4662  | 
and "\<forall>x' \<in> s. dist x' x < d --> f x' = g x'"  | 
|
4663  | 
and "continuous (at x within s) f"  | 
|
| 33175 | 4664  | 
shows "continuous (at x within s) g"  | 
| 53282 | 4665  | 
unfolding continuous_within  | 
| 36667 | 4666  | 
proof (rule Lim_transform_within)  | 
4667  | 
show "0 < d" by fact  | 
|
4668  | 
show "\<forall>x'\<in>s. 0 < dist x' x \<and> dist x' x < d \<longrightarrow> f x' = g x'"  | 
|
4669  | 
using assms(3) by auto  | 
|
4670  | 
have "f x = g x"  | 
|
4671  | 
using assms(1,2,3) by auto  | 
|
| 53282 | 4672  | 
then show "(f ---> g x) (at x within s)"  | 
| 36667 | 4673  | 
using assms(4) unfolding continuous_within by simp  | 
| 33175 | 4674  | 
qed  | 
4675  | 
||
4676  | 
lemma continuous_transform_at:  | 
|
| 36667 | 4677  | 
fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space"  | 
| 53282 | 4678  | 
assumes "0 < d"  | 
4679  | 
and "\<forall>x'. dist x' x < d --> f x' = g x'"  | 
|
4680  | 
and "continuous (at x) f"  | 
|
| 33175 | 4681  | 
shows "continuous (at x) g"  | 
| 45031 | 4682  | 
using continuous_transform_within [of d x UNIV f g] assms by simp  | 
| 33175 | 4683  | 
|
| 53282 | 4684  | 
|
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4685  | 
subsubsection {* Structural rules for pointwise continuity *}
 | 
| 33175 | 4686  | 
|
| 
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
 | 
4687  | 
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
 | 
4688  | 
|
| 
 
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
 | 
4689  | 
lemmas continuous_at_id = isCont_ident  | 
| 
44647
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4690  | 
|
| 
51361
 
21e5b6efb317
changed continuous_intros into a named theorems collection
 
hoelzl 
parents: 
51351 
diff
changeset
 | 
4691  | 
lemma continuous_infdist[continuous_intros]:  | 
| 50087 | 4692  | 
assumes "continuous F f"  | 
4693  | 
shows "continuous F (\<lambda>x. infdist (f x) A)"  | 
|
4694  | 
using assms unfolding continuous_def by (rule tendsto_infdist)  | 
|
4695  | 
||
| 
51361
 
21e5b6efb317
changed continuous_intros into a named theorems collection
 
hoelzl 
parents: 
51351 
diff
changeset
 | 
4696  | 
lemma continuous_infnorm[continuous_intros]:  | 
| 53282 | 4697  | 
"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
 | 
4698  | 
unfolding continuous_def by (rule tendsto_infnorm)  | 
| 33175 | 4699  | 
|
| 
51361
 
21e5b6efb317
changed continuous_intros into a named theorems collection
 
hoelzl 
parents: 
51351 
diff
changeset
 | 
4700  | 
lemma continuous_inner[continuous_intros]:  | 
| 53282 | 4701  | 
assumes "continuous F f"  | 
4702  | 
and "continuous F g"  | 
|
| 
44647
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4703  | 
shows "continuous F (\<lambda>x. inner (f x) (g x))"  | 
| 
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4704  | 
using assms unfolding continuous_def by (rule tendsto_inner)  | 
| 
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4705  | 
|
| 
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
 | 
4706  | 
lemmas continuous_at_inverse = isCont_inverse  | 
| 
44647
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4707  | 
|
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4708  | 
subsubsection {* Structural rules for setwise continuity *}
 | 
| 33175 | 4709  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4710  | 
lemma continuous_on_infnorm[continuous_on_intros]:  | 
| 53282 | 4711  | 
"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
 | 
4712  | 
unfolding continuous_on by (fast intro: tendsto_infnorm)  | 
| 
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4713  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4714  | 
lemma continuous_on_inner[continuous_on_intros]:  | 
| 
44531
 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 
huffman 
parents: 
44530 
diff
changeset
 | 
4715  | 
fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner"  | 
| 53282 | 4716  | 
assumes "continuous_on s f"  | 
4717  | 
and "continuous_on s g"  | 
|
| 
44531
 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 
huffman 
parents: 
44530 
diff
changeset
 | 
4718  | 
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
 | 
4719  | 
using bounded_bilinear_inner assms  | 
| 
 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 
huffman 
parents: 
44530 
diff
changeset
 | 
4720  | 
by (rule bounded_bilinear.continuous_on)  | 
| 
 
1d477a2b1572
replace some continuous_on lemmas with more general versions
 
huffman 
parents: 
44530 
diff
changeset
 | 
4721  | 
|
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4722  | 
subsubsection {* Structural rules for uniform continuity *}
 | 
| 33175 | 4723  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4724  | 
lemma uniformly_continuous_on_id[continuous_on_intros]:  | 
| 53282 | 4725  | 
"uniformly_continuous_on s (\<lambda>x. x)"  | 
| 
44647
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4726  | 
unfolding uniformly_continuous_on_def by auto  | 
| 
 
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
 
huffman 
parents: 
44632 
diff
changeset
 | 
4727  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4728  | 
lemma uniformly_continuous_on_const[continuous_on_intros]:  | 
| 53282 | 4729  | 
"uniformly_continuous_on s (\<lambda>x. c)"  | 
| 33175 | 4730  | 
unfolding uniformly_continuous_on_def by simp  | 
4731  | 
||
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4732  | 
lemma uniformly_continuous_on_dist[continuous_on_intros]:  | 
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4733  | 
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
 | 
4734  | 
assumes "uniformly_continuous_on s f"  | 
| 53282 | 4735  | 
and "uniformly_continuous_on s g"  | 
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4736  | 
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
 | 
4737  | 
proof -  | 
| 53282 | 4738  | 
  {
 | 
4739  | 
fix a b c d :: 'b  | 
|
4740  | 
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
 | 
4741  | 
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
 | 
4742  | 
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
 | 
4743  | 
by arith  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4744  | 
} note le = this  | 
| 53282 | 4745  | 
  {
 | 
4746  | 
fix x y  | 
|
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4747  | 
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
 | 
4748  | 
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
 | 
4749  | 
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
 | 
4750  | 
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
 | 
4751  | 
simp add: le)  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4752  | 
}  | 
| 53282 | 4753  | 
then show ?thesis  | 
4754  | 
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
 | 
4755  | 
unfolding dist_real_def by simp  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4756  | 
qed  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4757  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4758  | 
lemma uniformly_continuous_on_norm[continuous_on_intros]:  | 
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4759  | 
assumes "uniformly_continuous_on s f"  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4760  | 
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
 | 
4761  | 
unfolding norm_conv_dist using assms  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4762  | 
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
 | 
4763  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4764  | 
lemma (in bounded_linear) uniformly_continuous_on[continuous_on_intros]:  | 
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4765  | 
assumes "uniformly_continuous_on s g"  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4766  | 
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
 | 
4767  | 
using assms unfolding uniformly_continuous_on_sequentially  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4768  | 
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
 | 
4769  | 
by (auto intro: tendsto_zero)  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4770  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4771  | 
lemma uniformly_continuous_on_cmul[continuous_on_intros]:  | 
| 36441 | 4772  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 33175 | 4773  | 
assumes "uniformly_continuous_on s f"  | 
4774  | 
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
 | 
4775  | 
using bounded_linear_scaleR_right assms  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4776  | 
by (rule bounded_linear.uniformly_continuous_on)  | 
| 33175 | 4777  | 
|
4778  | 
lemma dist_minus:  | 
|
4779  | 
fixes x y :: "'a::real_normed_vector"  | 
|
4780  | 
shows "dist (- x) (- y) = dist x y"  | 
|
4781  | 
unfolding dist_norm minus_diff_minus norm_minus_cancel ..  | 
|
4782  | 
||
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4783  | 
lemma uniformly_continuous_on_minus[continuous_on_intros]:  | 
| 33175 | 4784  | 
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
 | 
4785  | 
shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. - f x)"  | 
| 33175 | 4786  | 
unfolding uniformly_continuous_on_def dist_minus .  | 
4787  | 
||
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4788  | 
lemma uniformly_continuous_on_add[continuous_on_intros]:  | 
| 36441 | 4789  | 
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
 | 
4790  | 
assumes "uniformly_continuous_on s f"  | 
| 53282 | 4791  | 
and "uniformly_continuous_on s g"  | 
| 33175 | 4792  | 
shows "uniformly_continuous_on s (\<lambda>x. f x + g x)"  | 
| 53282 | 4793  | 
using assms  | 
4794  | 
unfolding uniformly_continuous_on_sequentially  | 
|
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4795  | 
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
 | 
4796  | 
by (auto intro: tendsto_add_zero)  | 
| 
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4797  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4798  | 
lemma uniformly_continuous_on_diff[continuous_on_intros]:  | 
| 36441 | 4799  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
| 53282 | 4800  | 
assumes "uniformly_continuous_on s f"  | 
4801  | 
and "uniformly_continuous_on s g"  | 
|
| 
44648
 
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
 
huffman 
parents: 
44647 
diff
changeset
 | 
4802  | 
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
 | 
4803  | 
using assms uniformly_continuous_on_add [of s f "- g"]  | 
| 
 
b1d955791529
more simplification rules on unary and binary minus
 
haftmann 
parents: 
54070 
diff
changeset
 | 
4804  | 
by (simp add: fun_Compl_def uniformly_continuous_on_minus)  | 
| 33175 | 4805  | 
|
4806  | 
text{* Continuity of all kinds is preserved under composition. *}
 | 
|
4807  | 
||
| 
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
 | 
4808  | 
lemmas continuous_at_compose = isCont_o  | 
| 33175 | 4809  | 
|
| 
51362
 
dac3f564a98d
changed continuous_on_intros into a named theorems collection
 
hoelzl 
parents: 
51361 
diff
changeset
 | 
4810  | 
lemma uniformly_continuous_on_compose[continuous_on_intros]:  | 
| 33175 | 4811  | 
assumes "uniformly_continuous_on s f" "uniformly_continuous_on (f ` s) g"  | 
| 53640 | 4812  | 
shows "uniformly_continuous_on s (g \<circ> f)"  | 
4813  | 
proof -  | 
|
| 53282 | 4814  | 
  {
 | 
4815  | 
fix e :: real  | 
|
4816  | 
assume "e > 0"  | 
|
4817  | 
then obtain d where "d > 0"  | 
|
4818  | 
and d: "\<forall>x\<in>f ` s. \<forall>x'\<in>f ` s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e"  | 
|
4819  | 
using assms(2) unfolding uniformly_continuous_on_def by auto  | 
|
4820  | 
obtain d' where "d'>0" "\<forall>x\<in>s. \<forall>x'\<in>s. dist x' x < d' \<longrightarrow> dist (f x') (f x) < d"  | 
|
4821  | 
using `d > 0` using assms(1) unfolding uniformly_continuous_on_def by auto  | 
|
4822  | 
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"  | 
|
4823  | 
using `d>0` using d by auto  | 
|
4824  | 
}  | 
|
4825  | 
then show ?thesis  | 
|
4826  | 
using assms unfolding uniformly_continuous_on_def by auto  | 
|
| 33175 | 4827  | 
qed  | 
4828  | 
||
4829  | 
text{* Continuity in terms of open preimages. *}
 | 
|
4830  | 
||
4831  | 
lemma continuous_at_open:  | 
|
| 53282 | 4832  | 
"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)))"  | 
4833  | 
unfolding continuous_within_topological [of x UNIV f]  | 
|
4834  | 
unfolding imp_conjL  | 
|
4835  | 
by (intro all_cong imp_cong ex_cong conj_cong refl) auto  | 
|
| 33175 | 4836  | 
|
| 51351 | 4837  | 
lemma continuous_imp_tendsto:  | 
| 53282 | 4838  | 
assumes "continuous (at x0) f"  | 
4839  | 
and "x ----> x0"  | 
|
| 51351 | 4840  | 
shows "(f \<circ> x) ----> (f x0)"  | 
4841  | 
proof (rule topological_tendstoI)  | 
|
4842  | 
fix S  | 
|
4843  | 
assume "open S" "f x0 \<in> S"  | 
|
4844  | 
then obtain T where T_def: "open T" "x0 \<in> T" "\<forall>x\<in>T. f x \<in> S"  | 
|
4845  | 
using assms continuous_at_open by metis  | 
|
4846  | 
then have "eventually (\<lambda>n. x n \<in> T) sequentially"  | 
|
4847  | 
using assms T_def by (auto simp: tendsto_def)  | 
|
4848  | 
then show "eventually (\<lambda>n. (f \<circ> x) n \<in> S) sequentially"  | 
|
4849  | 
using T_def by (auto elim!: eventually_elim1)  | 
|
4850  | 
qed  | 
|
4851  | 
||
| 33175 | 4852  | 
lemma continuous_on_open:  | 
| 
51481
 
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
 
hoelzl 
parents: 
51480 
diff
changeset
 | 
4853  | 
"continuous_on s f \<longleftrightarrow>  | 
| 53282 | 4854  | 
(\<forall>t. openin (subtopology euclidean (f ` s)) t \<longrightarrow>  | 
4855  | 
      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
 | 
4856  | 
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
 | 
4857  | 
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong)  | 
| 36441 | 4858  | 
|
4859  | 
text {* Similarly in terms of closed sets. *}
 | 
|
| 33175 | 4860  | 
|
4861  | 
lemma continuous_on_closed:  | 
|
| 53282 | 4862  | 
"continuous_on s f \<longleftrightarrow>  | 
4863  | 
(\<forall>t. closedin (subtopology euclidean (f ` s)) t \<longrightarrow>  | 
|
4864  | 
      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
 | 
4865  | 
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
 | 
4866  | 
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong)  | 
| 33175 | 4867  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4868  | 
text {* Half-global and completely global cases. *}
 | 
| 33175 | 4869  | 
|
4870  | 
lemma continuous_open_in_preimage:  | 
|
4871  | 
assumes "continuous_on s f" "open t"  | 
|
4872  | 
  shows "openin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
 | 
|
| 53282 | 4873  | 
proof -  | 
4874  | 
have *: "\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)"  | 
|
4875  | 
by auto  | 
|
| 33175 | 4876  | 
have "openin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"  | 
4877  | 
using openin_open_Int[of t "f ` s", OF assms(2)] unfolding openin_open by auto  | 
|
| 53282 | 4878  | 
then show ?thesis  | 
4879  | 
using assms(1)[unfolded continuous_on_open, THEN spec[where x="t \<inter> f ` s"]]  | 
|
4880  | 
using * by auto  | 
|
| 33175 | 4881  | 
qed  | 
4882  | 
||
4883  | 
lemma continuous_closed_in_preimage:  | 
|
| 53291 | 4884  | 
assumes "continuous_on s f" and "closed t"  | 
| 33175 | 4885  | 
  shows "closedin (subtopology euclidean s) {x \<in> s. f x \<in> t}"
 | 
| 53282 | 4886  | 
proof -  | 
4887  | 
have *: "\<forall>x. x \<in> s \<and> f x \<in> t \<longleftrightarrow> x \<in> s \<and> f x \<in> (t \<inter> f ` s)"  | 
|
4888  | 
by auto  | 
|
| 33175 | 4889  | 
have "closedin (subtopology euclidean (f ` s)) (t \<inter> f ` s)"  | 
| 53282 | 4890  | 
using closedin_closed_Int[of t "f ` s", OF assms(2)] unfolding Int_commute  | 
4891  | 
by auto  | 
|
4892  | 
then show ?thesis  | 
|
4893  | 
using assms(1)[unfolded continuous_on_closed, THEN spec[where x="t \<inter> f ` s"]]  | 
|
4894  | 
using * by auto  | 
|
| 33175 | 4895  | 
qed  | 
4896  | 
||
4897  | 
lemma continuous_open_preimage:  | 
|
| 53291 | 4898  | 
assumes "continuous_on s f"  | 
4899  | 
and "open s"  | 
|
4900  | 
and "open t"  | 
|
| 33175 | 4901  | 
  shows "open {x \<in> s. f x \<in> t}"
 | 
4902  | 
proof-  | 
|
4903  | 
  obtain T where T: "open T" "{x \<in> s. f x \<in> t} = s \<inter> T"
 | 
|
4904  | 
using continuous_open_in_preimage[OF assms(1,3)] unfolding openin_open by auto  | 
|
| 53282 | 4905  | 
then show ?thesis  | 
4906  | 
using open_Int[of s T, OF assms(2)] by auto  | 
|
| 33175 | 4907  | 
qed  | 
4908  | 
||
4909  | 
lemma continuous_closed_preimage:  | 
|
| 53291 | 4910  | 
assumes "continuous_on s f"  | 
4911  | 
and "closed s"  | 
|
4912  | 
and "closed t"  | 
|
| 33175 | 4913  | 
  shows "closed {x \<in> s. f x \<in> t}"
 | 
4914  | 
proof-  | 
|
| 53282 | 4915  | 
  obtain T where "closed T" "{x \<in> s. f x \<in> t} = s \<inter> T"
 | 
4916  | 
using continuous_closed_in_preimage[OF assms(1,3)]  | 
|
4917  | 
unfolding closedin_closed by auto  | 
|
4918  | 
then show ?thesis using closed_Int[of s T, OF assms(2)] by auto  | 
|
| 33175 | 4919  | 
qed  | 
4920  | 
||
4921  | 
lemma continuous_open_preimage_univ:  | 
|
| 53282 | 4922  | 
  "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open {x. f x \<in> s}"
 | 
| 33175 | 4923  | 
using continuous_open_preimage[of UNIV f s] open_UNIV continuous_at_imp_continuous_on by auto  | 
4924  | 
||
4925  | 
lemma continuous_closed_preimage_univ:  | 
|
| 53291 | 4926  | 
  "(\<forall>x. continuous (at x) f) \<Longrightarrow> closed s \<Longrightarrow> closed {x. f x \<in> s}"
 | 
| 33175 | 4927  | 
using continuous_closed_preimage[of UNIV f s] closed_UNIV continuous_at_imp_continuous_on by auto  | 
4928  | 
||
| 53282 | 4929  | 
lemma continuous_open_vimage: "\<forall>x. continuous (at x) f \<Longrightarrow> open s \<Longrightarrow> open (f -` s)"  | 
| 33175 | 4930  | 
unfolding vimage_def by (rule continuous_open_preimage_univ)  | 
4931  | 
||
| 53282 | 4932  | 
lemma continuous_closed_vimage: "\<forall>x. continuous (at x) f \<Longrightarrow> closed s \<Longrightarrow> closed (f -` s)"  | 
| 33175 | 4933  | 
unfolding vimage_def by (rule continuous_closed_preimage_univ)  | 
4934  | 
||
| 36441 | 4935  | 
lemma interior_image_subset:  | 
| 53291 | 4936  | 
assumes "\<forall>x. continuous (at x) f"  | 
4937  | 
and "inj f"  | 
|
| 
35172
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35028 
diff
changeset
 | 
4938  | 
shows "interior (f ` s) \<subseteq> f ` (interior s)"  | 
| 44519 | 4939  | 
proof  | 
4940  | 
fix x assume "x \<in> interior (f ` s)"  | 
|
4941  | 
then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` s" ..  | 
|
| 53282 | 4942  | 
then have "x \<in> f ` s" by auto  | 
| 44519 | 4943  | 
then obtain y where y: "y \<in> s" "x = f y" by auto  | 
4944  | 
have "open (vimage f T)"  | 
|
4945  | 
using assms(1) `open T` by (rule continuous_open_vimage)  | 
|
4946  | 
moreover have "y \<in> vimage f T"  | 
|
4947  | 
using `x = f y` `x \<in> T` by simp  | 
|
4948  | 
moreover have "vimage f T \<subseteq> s"  | 
|
4949  | 
using `T \<subseteq> image f s` `inj f` unfolding inj_on_def subset_eq by auto  | 
|
4950  | 
ultimately have "y \<in> interior s" ..  | 
|
4951  | 
with `x = f y` show "x \<in> f ` interior s" ..  | 
|
4952  | 
qed  | 
|
| 
35172
 
579dd5570f96
Added integration to Multivariate-Analysis (upto FTC)
 
himmelma 
parents: 
35028 
diff
changeset
 | 
4953  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
4954  | 
text {* Equality of continuous functions on closure and related results. *}
 | 
| 33175 | 4955  | 
|
4956  | 
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
 | 
4957  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 53291 | 4958  | 
  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
 | 
4959  | 
  using continuous_closed_in_preimage[of s f "{a}"] by auto
 | 
| 33175 | 4960  | 
|
4961  | 
lemma continuous_closed_preimage_constant:  | 
|
| 
36668
 
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
 
huffman 
parents: 
36667 
diff
changeset
 | 
4962  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 53291 | 4963  | 
  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
 | 
4964  | 
  using continuous_closed_preimage[of s f "{a}"] by auto
 | 
| 33175 | 4965  | 
|
4966  | 
lemma continuous_constant_on_closure:  | 
|
| 
36668
 
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
 
huffman 
parents: 
36667 
diff
changeset
 | 
4967  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 33175 | 4968  | 
assumes "continuous_on (closure s) f"  | 
| 53282 | 4969  | 
and "\<forall>x \<in> s. f x = a"  | 
| 33175 | 4970  | 
shows "\<forall>x \<in> (closure s). f x = a"  | 
4971  | 
using continuous_closed_preimage_constant[of "closure s" f a]  | 
|
| 53282 | 4972  | 
      assms closure_minimal[of s "{x \<in> closure s. f x = a}"] closure_subset
 | 
4973  | 
unfolding subset_eq  | 
|
4974  | 
by auto  | 
|
| 33175 | 4975  | 
|
4976  | 
lemma image_closure_subset:  | 
|
| 53291 | 4977  | 
assumes "continuous_on (closure s) f"  | 
4978  | 
and "closed t"  | 
|
4979  | 
and "(f ` s) \<subseteq> t"  | 
|
| 33175 | 4980  | 
shows "f ` (closure s) \<subseteq> t"  | 
| 53282 | 4981  | 
proof -  | 
4982  | 
  have "s \<subseteq> {x \<in> closure s. f x \<in> t}"
 | 
|
4983  | 
using assms(3) closure_subset by auto  | 
|
| 33175 | 4984  | 
  moreover have "closed {x \<in> closure s. f x \<in> t}"
 | 
4985  | 
using continuous_closed_preimage[OF assms(1)] and assms(2) by auto  | 
|
4986  | 
  ultimately have "closure s = {x \<in> closure s . f x \<in> t}"
 | 
|
4987  | 
    using closure_minimal[of s "{x \<in> closure s. f x \<in> t}"] by auto
 | 
|
| 53282 | 4988  | 
then show ?thesis by auto  | 
| 33175 | 4989  | 
qed  | 
4990  | 
||
4991  | 
lemma continuous_on_closure_norm_le:  | 
|
4992  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector"  | 
|
| 53282 | 4993  | 
assumes "continuous_on (closure s) f"  | 
4994  | 
and "\<forall>y \<in> s. norm(f y) \<le> b"  | 
|
4995  | 
and "x \<in> (closure s)"  | 
|
| 53291 | 4996  | 
shows "norm (f x) \<le> b"  | 
| 53282 | 4997  | 
proof -  | 
4998  | 
have *: "f ` s \<subseteq> cball 0 b"  | 
|
| 53291 | 4999  | 
using assms(2)[unfolded mem_cball_0[symmetric]] by auto  | 
| 33175 | 5000  | 
show ?thesis  | 
5001  | 
using image_closure_subset[OF assms(1) closed_cball[of 0 b] *] assms(3)  | 
|
| 53282 | 5002  | 
unfolding subset_eq  | 
5003  | 
apply (erule_tac x="f x" in ballE)  | 
|
5004  | 
apply (auto simp add: dist_norm)  | 
|
5005  | 
done  | 
|
| 33175 | 5006  | 
qed  | 
5007  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5008  | 
text {* Making a continuous function avoid some value in a neighbourhood. *}
 | 
| 33175 | 5009  | 
|
5010  | 
lemma continuous_within_avoid:  | 
|
| 50898 | 5011  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"  | 
| 53282 | 5012  | 
assumes "continuous (at x within s) f"  | 
5013  | 
and "f x \<noteq> a"  | 
|
| 33175 | 5014  | 
shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e --> f y \<noteq> a"  | 
| 53291 | 5015  | 
proof -  | 
| 50898 | 5016  | 
obtain U where "open U" and "f x \<in> U" and "a \<notin> U"  | 
5017  | 
using t1_space [OF `f x \<noteq> a`] by fast  | 
|
5018  | 
have "(f ---> f x) (at x within s)"  | 
|
5019  | 
using assms(1) by (simp add: continuous_within)  | 
|
| 53282 | 5020  | 
then have "eventually (\<lambda>y. f y \<in> U) (at x within s)"  | 
| 50898 | 5021  | 
using `open U` and `f x \<in> U`  | 
5022  | 
unfolding tendsto_def by fast  | 
|
| 53282 | 5023  | 
then have "eventually (\<lambda>y. f y \<noteq> a) (at x within s)"  | 
| 50898 | 5024  | 
using `a \<notin> U` by (fast elim: eventually_mono [rotated])  | 
| 53282 | 5025  | 
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
 | 
5026  | 
using `f x \<noteq> a` by (auto simp: dist_commute zero_less_dist_iff eventually_at)  | 
| 33175 | 5027  | 
qed  | 
5028  | 
||
5029  | 
lemma continuous_at_avoid:  | 
|
| 50898 | 5030  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"  | 
| 53282 | 5031  | 
assumes "continuous (at x) f"  | 
5032  | 
and "f x \<noteq> a"  | 
|
| 33175 | 5033  | 
shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"  | 
| 45031 | 5034  | 
using assms continuous_within_avoid[of x UNIV f a] by simp  | 
| 33175 | 5035  | 
|
5036  | 
lemma continuous_on_avoid:  | 
|
| 50898 | 5037  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"  | 
| 53282 | 5038  | 
assumes "continuous_on s f"  | 
5039  | 
and "x \<in> s"  | 
|
5040  | 
and "f x \<noteq> a"  | 
|
| 33175 | 5041  | 
shows "\<exists>e>0. \<forall>y \<in> s. dist x y < e \<longrightarrow> f y \<noteq> a"  | 
| 53282 | 5042  | 
using assms(1)[unfolded continuous_on_eq_continuous_within, THEN bspec[where x=x],  | 
5043  | 
OF assms(2)] continuous_within_avoid[of x s f a]  | 
|
5044  | 
using assms(3)  | 
|
5045  | 
by auto  | 
|
| 33175 | 5046  | 
|
5047  | 
lemma continuous_on_open_avoid:  | 
|
| 50898 | 5048  | 
fixes f :: "'a::metric_space \<Rightarrow> 'b::t1_space"  | 
| 53291 | 5049  | 
assumes "continuous_on s f"  | 
5050  | 
and "open s"  | 
|
5051  | 
and "x \<in> s"  | 
|
5052  | 
and "f x \<noteq> a"  | 
|
| 33175 | 5053  | 
shows "\<exists>e>0. \<forall>y. dist x y < e \<longrightarrow> f y \<noteq> a"  | 
| 53282 | 5054  | 
using assms(1)[unfolded continuous_on_eq_continuous_at[OF assms(2)], THEN bspec[where x=x], OF assms(3)]  | 
5055  | 
using continuous_at_avoid[of x f a] assms(4)  | 
|
5056  | 
by auto  | 
|
| 33175 | 5057  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5058  | 
text {* Proving a function is constant by proving open-ness of level set. *}
 | 
| 33175 | 5059  | 
|
5060  | 
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
 | 
5061  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 36359 | 5062  | 
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>  | 
| 33175 | 5063  | 
        openin (subtopology euclidean s) {x \<in> s. f x = a}
 | 
| 53282 | 5064  | 
\<Longrightarrow> (\<forall>x \<in> s. f x \<noteq> a) \<or> (\<forall>x \<in> s. f x = a)"  | 
5065  | 
unfolding connected_clopen  | 
|
5066  | 
using continuous_closed_in_preimage_constant by auto  | 
|
| 33175 | 5067  | 
|
5068  | 
lemma continuous_levelset_open_in:  | 
|
| 
36668
 
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
 
huffman 
parents: 
36667 
diff
changeset
 | 
5069  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 36359 | 5070  | 
shows "connected s \<Longrightarrow> continuous_on s f \<Longrightarrow>  | 
| 33175 | 5071  | 
        openin (subtopology euclidean s) {x \<in> s. f x = a} \<Longrightarrow>
 | 
| 53291 | 5072  | 
(\<exists>x \<in> s. f x = a) \<Longrightarrow> (\<forall>x \<in> s. f x = a)"  | 
| 53282 | 5073  | 
using continuous_levelset_open_in_cases[of s f ]  | 
5074  | 
by meson  | 
|
| 33175 | 5075  | 
|
5076  | 
lemma continuous_levelset_open:  | 
|
| 
36668
 
941ba2da372e
simplify definition of t1_space; generalize lemma closed_sing and related lemmas
 
huffman 
parents: 
36667 
diff
changeset
 | 
5077  | 
fixes f :: "_ \<Rightarrow> 'b::t1_space"  | 
| 53282 | 5078  | 
assumes "connected s"  | 
5079  | 
and "continuous_on s f"  | 
|
5080  | 
    and "open {x \<in> s. f x = a}"
 | 
|
5081  | 
and "\<exists>x \<in> s. f x = a"  | 
|
| 33175 | 5082  | 
shows "\<forall>x \<in> s. f x = a"  | 
| 53282 | 5083  | 
using continuous_levelset_open_in[OF assms(1,2), of a, unfolded openin_open]  | 
5084  | 
using assms (3,4)  | 
|
5085  | 
by fast  | 
|
| 33175 | 5086  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5087  | 
text {* Some arithmetical combinations (more to prove). *}
 | 
| 33175 | 5088  | 
|
5089  | 
lemma open_scaling[intro]:  | 
|
5090  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53291 | 5091  | 
assumes "c \<noteq> 0"  | 
5092  | 
and "open s"  | 
|
| 33175 | 5093  | 
shows "open((\<lambda>x. c *\<^sub>R x) ` s)"  | 
| 53282 | 5094  | 
proof -  | 
5095  | 
  {
 | 
|
5096  | 
fix x  | 
|
5097  | 
assume "x \<in> s"  | 
|
5098  | 
then obtain e where "e>0"  | 
|
5099  | 
and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> s" using assms(2)[unfolded open_dist, THEN bspec[where x=x]]  | 
|
5100  | 
by auto  | 
|
5101  | 
have "e * abs c > 0"  | 
|
| 53291 | 5102  | 
using assms(1)[unfolded zero_less_abs_iff[symmetric]]  | 
| 53282 | 5103  | 
using mult_pos_pos[OF `e>0`]  | 
5104  | 
by auto  | 
|
| 33175 | 5105  | 
moreover  | 
| 53282 | 5106  | 
    {
 | 
5107  | 
fix y  | 
|
5108  | 
assume "dist y (c *\<^sub>R x) < e * \<bar>c\<bar>"  | 
|
5109  | 
then have "norm ((1 / c) *\<^sub>R y - x) < e"  | 
|
5110  | 
unfolding dist_norm  | 
|
| 33175 | 5111  | 
using norm_scaleR[of c "(1 / c) *\<^sub>R y - x", unfolded scaleR_right_diff_distrib, unfolded scaleR_scaleR] assms(1)  | 
| 53291 | 5112  | 
assms(1)[unfolded zero_less_abs_iff[symmetric]] by (simp del:zero_less_abs_iff)  | 
| 53282 | 5113  | 
then have "y \<in> op *\<^sub>R c ` s"  | 
5114  | 
using rev_image_eqI[of "(1 / c) *\<^sub>R y" s y "op *\<^sub>R c"]  | 
|
5115  | 
using e[THEN spec[where x="(1 / c) *\<^sub>R y"]]  | 
|
5116  | 
using assms(1)  | 
|
5117  | 
unfolding dist_norm scaleR_scaleR  | 
|
5118  | 
by auto  | 
|
5119  | 
}  | 
|
5120  | 
ultimately have "\<exists>e>0. \<forall>x'. dist x' (c *\<^sub>R x) < e \<longrightarrow> x' \<in> op *\<^sub>R c ` s"  | 
|
5121  | 
apply (rule_tac x="e * abs c" in exI)  | 
|
5122  | 
apply auto  | 
|
5123  | 
done  | 
|
5124  | 
}  | 
|
5125  | 
then show ?thesis unfolding open_dist by auto  | 
|
| 33175 | 5126  | 
qed  | 
5127  | 
||
5128  | 
lemma minus_image_eq_vimage:  | 
|
5129  | 
fixes A :: "'a::ab_group_add set"  | 
|
5130  | 
shows "(\<lambda>x. - x) ` A = (\<lambda>x. - x) -` A"  | 
|
5131  | 
by (auto intro!: image_eqI [where f="\<lambda>x. - x"])  | 
|
5132  | 
||
5133  | 
lemma open_negations:  | 
|
5134  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 
54489
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
5135  | 
shows "open s \<Longrightarrow> open ((\<lambda>x. - x) ` s)"  | 
| 
 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 
haftmann 
parents: 
54263 
diff
changeset
 | 
5136  | 
using open_scaling [of "- 1" s] by simp  | 
| 33175 | 5137  | 
|
5138  | 
lemma open_translation:  | 
|
5139  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53291 | 5140  | 
assumes "open s"  | 
5141  | 
shows "open((\<lambda>x. a + x) ` s)"  | 
|
| 53282 | 5142  | 
proof -  | 
5143  | 
  {
 | 
|
5144  | 
fix x  | 
|
5145  | 
have "continuous (at x) (\<lambda>x. x - a)"  | 
|
5146  | 
by (intro continuous_diff continuous_at_id continuous_const)  | 
|
5147  | 
}  | 
|
5148  | 
  moreover have "{x. x - a \<in> s} = op + a ` s"
 | 
|
5149  | 
by force  | 
|
5150  | 
ultimately show ?thesis using continuous_open_preimage_univ[of "\<lambda>x. x - a" s]  | 
|
5151  | 
using assms by auto  | 
|
| 33175 | 5152  | 
qed  | 
5153  | 
||
5154  | 
lemma open_affinity:  | 
|
5155  | 
fixes s :: "'a::real_normed_vector set"  | 
|
5156  | 
assumes "open s" "c \<noteq> 0"  | 
|
5157  | 
shows "open ((\<lambda>x. a + c *\<^sub>R x) ` s)"  | 
|
| 53282 | 5158  | 
proof -  | 
5159  | 
have *: "(\<lambda>x. a + c *\<^sub>R x) = (\<lambda>x. a + x) \<circ> (\<lambda>x. c *\<^sub>R x)"  | 
|
5160  | 
unfolding o_def ..  | 
|
5161  | 
have "op + a ` op *\<^sub>R c ` s = (op + a \<circ> op *\<^sub>R c) ` s"  | 
|
5162  | 
by auto  | 
|
5163  | 
then show ?thesis  | 
|
5164  | 
using assms open_translation[of "op *\<^sub>R c ` s" a]  | 
|
5165  | 
unfolding *  | 
|
5166  | 
by auto  | 
|
| 33175 | 5167  | 
qed  | 
5168  | 
||
5169  | 
lemma interior_translation:  | 
|
5170  | 
fixes s :: "'a::real_normed_vector set"  | 
|
5171  | 
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
 | 
5172  | 
proof (rule set_eqI, rule)  | 
| 53282 | 5173  | 
fix x  | 
5174  | 
assume "x \<in> interior (op + a ` s)"  | 
|
5175  | 
then obtain e where "e > 0" and e: "ball x e \<subseteq> op + a ` s"  | 
|
5176  | 
unfolding mem_interior by auto  | 
|
5177  | 
then have "ball (x - a) e \<subseteq> s"  | 
|
5178  | 
unfolding subset_eq Ball_def mem_ball dist_norm  | 
|
5179  | 
apply auto  | 
|
5180  | 
apply (erule_tac x="a + xa" in allE)  | 
|
| 53291 | 5181  | 
unfolding ab_group_add_class.diff_diff_eq[symmetric]  | 
| 53282 | 5182  | 
apply auto  | 
5183  | 
done  | 
|
5184  | 
then show "x \<in> op + a ` interior s"  | 
|
5185  | 
unfolding image_iff  | 
|
5186  | 
apply (rule_tac x="x - a" in bexI)  | 
|
5187  | 
unfolding mem_interior  | 
|
5188  | 
using `e > 0`  | 
|
5189  | 
apply auto  | 
|
5190  | 
done  | 
|
| 33175 | 5191  | 
next  | 
| 53282 | 5192  | 
fix x  | 
5193  | 
assume "x \<in> op + a ` interior s"  | 
|
5194  | 
then obtain y e where "e > 0" and e: "ball y e \<subseteq> s" and y: "x = a + y"  | 
|
5195  | 
unfolding image_iff Bex_def mem_interior by auto  | 
|
5196  | 
  {
 | 
|
5197  | 
fix z  | 
|
5198  | 
have *: "a + y - z = y + a - z" by auto  | 
|
5199  | 
assume "z \<in> ball x e"  | 
|
5200  | 
then have "z - a \<in> s"  | 
|
5201  | 
using e[unfolded subset_eq, THEN bspec[where x="z - a"]]  | 
|
5202  | 
unfolding mem_ball dist_norm y group_add_class.diff_diff_eq2 *  | 
|
5203  | 
by auto  | 
|
5204  | 
then have "z \<in> op + a ` s"  | 
|
5205  | 
unfolding image_iff by (auto intro!: bexI[where x="z - a"])  | 
|
5206  | 
}  | 
|
5207  | 
then have "ball x e \<subseteq> op + a ` s"  | 
|
5208  | 
unfolding subset_eq by auto  | 
|
5209  | 
then show "x \<in> interior (op + a ` s)"  | 
|
5210  | 
unfolding mem_interior using `e > 0` by auto  | 
|
| 33175 | 5211  | 
qed  | 
5212  | 
||
| 36437 | 5213  | 
text {* Topological properties of linear functions. *}
 | 
5214  | 
||
5215  | 
lemma linear_lim_0:  | 
|
| 53282 | 5216  | 
assumes "bounded_linear f"  | 
5217  | 
shows "(f ---> 0) (at (0))"  | 
|
5218  | 
proof -  | 
|
| 36437 | 5219  | 
interpret f: bounded_linear f by fact  | 
5220  | 
have "(f ---> f 0) (at 0)"  | 
|
5221  | 
using tendsto_ident_at by (rule f.tendsto)  | 
|
| 53282 | 5222  | 
then show ?thesis unfolding f.zero .  | 
| 36437 | 5223  | 
qed  | 
5224  | 
||
5225  | 
lemma linear_continuous_at:  | 
|
| 53282 | 5226  | 
assumes "bounded_linear f"  | 
5227  | 
shows "continuous (at a) f"  | 
|
| 36437 | 5228  | 
unfolding continuous_at using assms  | 
5229  | 
apply (rule bounded_linear.tendsto)  | 
|
5230  | 
apply (rule tendsto_ident_at)  | 
|
5231  | 
done  | 
|
5232  | 
||
5233  | 
lemma linear_continuous_within:  | 
|
| 53291 | 5234  | 
"bounded_linear f \<Longrightarrow> continuous (at x within s) f"  | 
| 36437 | 5235  | 
using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto  | 
5236  | 
||
5237  | 
lemma linear_continuous_on:  | 
|
| 53291 | 5238  | 
"bounded_linear f \<Longrightarrow> continuous_on s f"  | 
| 36437 | 5239  | 
using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto  | 
5240  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5241  | 
text {* Also bilinear functions, in composition form. *}
 | 
| 36437 | 5242  | 
|
5243  | 
lemma bilinear_continuous_at_compose:  | 
|
| 53282 | 5244  | 
"continuous (at x) f \<Longrightarrow> continuous (at x) g \<Longrightarrow> bounded_bilinear h \<Longrightarrow>  | 
5245  | 
continuous (at x) (\<lambda>x. h (f x) (g x))"  | 
|
5246  | 
unfolding continuous_at  | 
|
5247  | 
using Lim_bilinear[of f "f x" "(at x)" g "g x" h]  | 
|
5248  | 
by auto  | 
|
| 36437 | 5249  | 
|
5250  | 
lemma bilinear_continuous_within_compose:  | 
|
| 53282 | 5251  | 
"continuous (at x within s) f \<Longrightarrow> continuous (at x within s) g \<Longrightarrow> bounded_bilinear h \<Longrightarrow>  | 
5252  | 
continuous (at x within s) (\<lambda>x. h (f x) (g x))"  | 
|
5253  | 
unfolding continuous_within  | 
|
5254  | 
using Lim_bilinear[of f "f x"]  | 
|
5255  | 
by auto  | 
|
| 36437 | 5256  | 
|
5257  | 
lemma bilinear_continuous_on_compose:  | 
|
| 53282 | 5258  | 
"continuous_on s f \<Longrightarrow> continuous_on s g \<Longrightarrow> bounded_bilinear h \<Longrightarrow>  | 
5259  | 
continuous_on s (\<lambda>x. h (f x) (g x))"  | 
|
| 36441 | 5260  | 
unfolding continuous_on_def  | 
5261  | 
by (fast elim: bounded_bilinear.tendsto)  | 
|
| 36437 | 5262  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5263  | 
text {* Preservation of compactness and connectedness under continuous function. *}
 | 
| 33175 | 5264  | 
|
| 50898 | 5265  | 
lemma compact_eq_openin_cover:  | 
5266  | 
"compact S \<longleftrightarrow>  | 
|
5267  | 
(\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>  | 
|
5268  | 
(\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D))"  | 
|
5269  | 
proof safe  | 
|
5270  | 
fix C  | 
|
5271  | 
assume "compact S" and "\<forall>c\<in>C. openin (subtopology euclidean S) c" and "S \<subseteq> \<Union>C"  | 
|
| 53282 | 5272  | 
  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 | 5273  | 
unfolding openin_open by force+  | 
5274  | 
  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"
 | 
|
5275  | 
by (rule compactE)  | 
|
| 53282 | 5276  | 
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 | 5277  | 
by auto  | 
| 53282 | 5278  | 
then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..  | 
| 50898 | 5279  | 
next  | 
5280  | 
assume 1: "\<forall>C. (\<forall>c\<in>C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>C \<longrightarrow>  | 
|
5281  | 
(\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D)"  | 
|
5282  | 
show "compact S"  | 
|
5283  | 
proof (rule compactI)  | 
|
5284  | 
fix C  | 
|
5285  | 
let ?C = "image (\<lambda>T. S \<inter> T) C"  | 
|
5286  | 
assume "\<forall>t\<in>C. open t" and "S \<subseteq> \<Union>C"  | 
|
| 53282 | 5287  | 
then have "(\<forall>c\<in>?C. openin (subtopology euclidean S) c) \<and> S \<subseteq> \<Union>?C"  | 
| 50898 | 5288  | 
unfolding openin_open by auto  | 
5289  | 
with 1 obtain D where "D \<subseteq> ?C" and "finite D" and "S \<subseteq> \<Union>D"  | 
|
5290  | 
by metis  | 
|
5291  | 
let ?D = "inv_into C (\<lambda>T. S \<inter> T) ` D"  | 
|
5292  | 
have "?D \<subseteq> C \<and> finite ?D \<and> S \<subseteq> \<Union>?D"  | 
|
5293  | 
proof (intro conjI)  | 
|
5294  | 
from `D \<subseteq> ?C` show "?D \<subseteq> C"  | 
|
5295  | 
by (fast intro: inv_into_into)  | 
|
5296  | 
from `finite D` show "finite ?D"  | 
|
5297  | 
by (rule finite_imageI)  | 
|
5298  | 
from `S \<subseteq> \<Union>D` show "S \<subseteq> \<Union>?D"  | 
|
5299  | 
apply (rule subset_trans)  | 
|
5300  | 
apply clarsimp  | 
|
5301  | 
apply (frule subsetD [OF `D \<subseteq> ?C`, THEN f_inv_into_f])  | 
|
5302  | 
apply (erule rev_bexI, fast)  | 
|
5303  | 
done  | 
|
5304  | 
qed  | 
|
| 53282 | 5305  | 
then show "\<exists>D\<subseteq>C. finite D \<and> S \<subseteq> \<Union>D" ..  | 
| 50898 | 5306  | 
qed  | 
5307  | 
qed  | 
|
5308  | 
||
| 33175 | 5309  | 
lemma connected_continuous_image:  | 
| 53291 | 5310  | 
assumes "continuous_on s f"  | 
5311  | 
and "connected s"  | 
|
| 33175 | 5312  | 
shows "connected(f ` s)"  | 
| 53282 | 5313  | 
proof -  | 
5314  | 
  {
 | 
|
5315  | 
fix T  | 
|
| 53291 | 5316  | 
assume as:  | 
5317  | 
      "T \<noteq> {}"
 | 
|
5318  | 
"T \<noteq> f ` s"  | 
|
5319  | 
"openin (subtopology euclidean (f ` s)) T"  | 
|
5320  | 
"closedin (subtopology euclidean (f ` s)) T"  | 
|
| 33175 | 5321  | 
    have "{x \<in> s. f x \<in> T} = {} \<or> {x \<in> s. f x \<in> T} = s"
 | 
5322  | 
using assms(1)[unfolded continuous_on_open, THEN spec[where x=T]]  | 
|
5323  | 
using assms(1)[unfolded continuous_on_closed, THEN spec[where x=T]]  | 
|
5324  | 
      using assms(2)[unfolded connected_clopen, THEN spec[where x="{x \<in> s. f x \<in> T}"]] as(3,4) by auto
 | 
|
| 53282 | 5325  | 
then have False using as(1,2)  | 
5326  | 
using as(4)[unfolded closedin_def topspace_euclidean_subtopology] by auto  | 
|
5327  | 
}  | 
|
5328  | 
then show ?thesis  | 
|
5329  | 
unfolding connected_clopen by auto  | 
|
| 33175 | 5330  | 
qed  | 
5331  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5332  | 
text {* Continuity implies uniform continuity on a compact domain. *}
 | 
| 53282 | 5333  | 
|
| 33175 | 5334  | 
lemma compact_uniformly_continuous:  | 
| 53291 | 5335  | 
assumes f: "continuous_on s f"  | 
5336  | 
and s: "compact s"  | 
|
| 33175 | 5337  | 
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
 | 
5338  | 
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
 | 
5339  | 
proof (cases, safe)  | 
| 53282 | 5340  | 
fix e :: real  | 
5341  | 
  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
 | 
5342  | 
  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 | 5343  | 
let ?b = "(\<lambda>(y, d). ball y (d/2))"  | 
5344  | 
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
 | 
5345  | 
proof safe  | 
| 53282 | 5346  | 
fix y  | 
5347  | 
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
 | 
5348  | 
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
 | 
5349  | 
    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
 | 
5350  | 
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
 | 
5351  | 
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
 | 
5352  | 
using `0 < e` `y \<in> s` by (auto elim!: openE)  | 
| 50944 | 5353  | 
with T `y \<in> s` show "y \<in> (\<Union>r\<in>R. ?b r)"  | 
5354  | 
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
 | 
5355  | 
qed auto  | 
| 50944 | 5356  | 
with s obtain D where D: "finite D" "D \<subseteq> R" "s \<subseteq> (\<Union>(y, d)\<in>D. ball y (d/2))"  | 
5357  | 
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
 | 
5358  | 
  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
 | 
5359  | 
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
 | 
5360  | 
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
 | 
5361  | 
proof (rule, safe)  | 
| 53282 | 5362  | 
fix x x'  | 
5363  | 
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
 | 
5364  | 
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
 | 
5365  | 
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
 | 
5366  | 
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
 | 
5367  | 
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
 | 
5368  | 
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
 | 
5369  | 
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
 | 
5370  | 
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
 | 
5371  | 
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
 | 
5372  | 
qed auto  | 
| 33175 | 5373  | 
|
| 36437 | 5374  | 
text {* A uniformly convergent limit of continuous functions is continuous. *}
 | 
| 33175 | 5375  | 
|
5376  | 
lemma continuous_uniform_limit:  | 
|
| 
44212
 
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
 
huffman 
parents: 
44211 
diff
changeset
 | 
5377  | 
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
 | 
5378  | 
assumes "\<not> trivial_limit F"  | 
| 53282 | 5379  | 
and "eventually (\<lambda>n. continuous_on s (f n)) F"  | 
5380  | 
and "\<forall>e>0. eventually (\<lambda>n. \<forall>x\<in>s. dist (f n x) (g x) < e) F"  | 
|
| 33175 | 5381  | 
shows "continuous_on s g"  | 
| 53282 | 5382  | 
proof -  | 
5383  | 
  {
 | 
|
5384  | 
fix x and e :: real  | 
|
5385  | 
assume "x\<in>s" "e>0"  | 
|
| 
44212
 
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
 
huffman 
parents: 
44211 
diff
changeset
 | 
5386  | 
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
 | 
5387  | 
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
 | 
5388  | 
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
 | 
5389  | 
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
 | 
5390  | 
using assms(1) by blast  | 
| 33175 | 5391  | 
have "e / 3 > 0" using `e>0` by auto  | 
5392  | 
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 | 5393  | 
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 | 5394  | 
    {
 | 
5395  | 
fix y  | 
|
5396  | 
assume "y \<in> s" and "dist y x < d"  | 
|
5397  | 
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
 | 
5398  | 
by (rule d [rule_format])  | 
| 53282 | 5399  | 
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
 | 
5400  | 
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
 | 
5401  | 
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
 | 
5402  | 
by auto  | 
| 53282 | 5403  | 
then have "dist (g y) (g x) < e"  | 
| 
44212
 
4d10e7f342b1
generalize lemma continuous_uniform_limit to class metric_space
 
huffman 
parents: 
44211 
diff
changeset
 | 
5404  | 
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
 | 
5405  | 
using dist_triangle3 [of "g y" "g x" "f n y"]  | 
| 53282 | 5406  | 
by auto  | 
5407  | 
}  | 
|
5408  | 
then have "\<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (g x') (g x) < e"  | 
|
5409  | 
using `d>0` by auto  | 
|
5410  | 
}  | 
|
5411  | 
then show ?thesis  | 
|
5412  | 
unfolding continuous_on_iff by auto  | 
|
| 33175 | 5413  | 
qed  | 
5414  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5415  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5416  | 
subsection {* Topological stuff lifted from and dropped to R *}
 | 
| 33175 | 5417  | 
|
5418  | 
lemma open_real:  | 
|
| 53282 | 5419  | 
fixes s :: "real set"  | 
5420  | 
shows "open s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e>0. \<forall>x'. abs(x' - x) < e --> x' \<in> s)"  | 
|
| 33175 | 5421  | 
unfolding open_dist dist_norm by simp  | 
5422  | 
||
5423  | 
lemma islimpt_approachable_real:  | 
|
5424  | 
fixes s :: "real set"  | 
|
5425  | 
shows "x islimpt s \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> s. x' \<noteq> x \<and> abs(x' - x) < e)"  | 
|
5426  | 
unfolding islimpt_approachable dist_norm by simp  | 
|
5427  | 
||
5428  | 
lemma closed_real:  | 
|
5429  | 
fixes s :: "real set"  | 
|
| 53282 | 5430  | 
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 | 5431  | 
unfolding closed_limpt islimpt_approachable dist_norm by simp  | 
5432  | 
||
5433  | 
lemma continuous_at_real_range:  | 
|
5434  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> real"  | 
|
| 53282 | 5435  | 
shows "continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. \<forall>x'. norm(x' - x) < d --> abs(f x' - f x) < e)"  | 
5436  | 
unfolding continuous_at  | 
|
5437  | 
unfolding Lim_at  | 
|
| 53291 | 5438  | 
unfolding dist_nz[symmetric]  | 
| 53282 | 5439  | 
unfolding dist_norm  | 
5440  | 
apply auto  | 
|
5441  | 
apply (erule_tac x=e in allE)  | 
|
5442  | 
apply auto  | 
|
5443  | 
apply (rule_tac x=d in exI)  | 
|
5444  | 
apply auto  | 
|
5445  | 
apply (erule_tac x=x' in allE)  | 
|
5446  | 
apply auto  | 
|
5447  | 
apply (erule_tac x=e in allE)  | 
|
5448  | 
apply auto  | 
|
5449  | 
done  | 
|
| 33175 | 5450  | 
|
5451  | 
lemma continuous_on_real_range:  | 
|
5452  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> real"  | 
|
| 53282 | 5453  | 
shows "continuous_on s f \<longleftrightarrow>  | 
5454  | 
(\<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 | 5455  | 
unfolding continuous_on_iff dist_norm by simp  | 
| 33175 | 5456  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5457  | 
text {* Hence some handy theorems on distance, diameter etc. of/from a set. *}
 | 
| 33175 | 5458  | 
|
5459  | 
lemma distance_attains_sup:  | 
|
5460  | 
  assumes "compact s" "s \<noteq> {}"
 | 
|
| 
51346
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5461  | 
shows "\<exists>x\<in>s. \<forall>y\<in>s. dist a y \<le> dist a x"  | 
| 33175 | 5462  | 
proof (rule continuous_attains_sup [OF assms])  | 
| 53282 | 5463  | 
  {
 | 
5464  | 
fix x  | 
|
5465  | 
assume "x\<in>s"  | 
|
| 33175 | 5466  | 
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
 | 
5467  | 
by (intro tendsto_dist tendsto_const tendsto_ident_at)  | 
| 33175 | 5468  | 
}  | 
| 53282 | 5469  | 
then show "continuous_on s (dist a)"  | 
| 33175 | 5470  | 
unfolding continuous_on ..  | 
5471  | 
qed  | 
|
5472  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5473  | 
text {* For \emph{minimal} distance, we only need closure, not compactness. *}
 | 
| 33175 | 5474  | 
|
5475  | 
lemma distance_attains_inf:  | 
|
5476  | 
fixes a :: "'a::heine_borel"  | 
|
| 53291 | 5477  | 
assumes "closed s"  | 
5478  | 
    and "s \<noteq> {}"
 | 
|
| 
51346
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5479  | 
shows "\<exists>x\<in>s. \<forall>y\<in>s. dist a x \<le> dist a y"  | 
| 53282 | 5480  | 
proof -  | 
| 
51346
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5481  | 
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
 | 
5482  | 
let ?B = "s \<inter> cball a (dist b a)"  | 
| 53282 | 5483  | 
  have "?B \<noteq> {}" using `b \<in> s`
 | 
5484  | 
by (auto simp add: dist_commute)  | 
|
| 
51346
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5485  | 
moreover have "continuous_on ?B (dist a)"  | 
| 
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5486  | 
by (auto intro!: continuous_at_imp_continuous_on continuous_dist continuous_at_id continuous_const)  | 
| 33175 | 5487  | 
moreover have "compact ?B"  | 
| 
51346
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5488  | 
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
 | 
5489  | 
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
 | 
5490  | 
by (metis continuous_attains_inf)  | 
| 53282 | 5491  | 
then show ?thesis by fastforce  | 
| 33175 | 5492  | 
qed  | 
5493  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5494  | 
|
| 36437 | 5495  | 
subsection {* Pasted sets *}
 | 
| 33175 | 5496  | 
|
5497  | 
lemma bounded_Times:  | 
|
| 53282 | 5498  | 
assumes "bounded s" "bounded t"  | 
5499  | 
shows "bounded (s \<times> t)"  | 
|
5500  | 
proof -  | 
|
| 33175 | 5501  | 
obtain x y a b where "\<forall>z\<in>s. dist x z \<le> a" "\<forall>z\<in>t. dist y z \<le> b"  | 
5502  | 
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
 | 
5503  | 
then have "\<forall>z\<in>s \<times> t. dist (x, y) z \<le> sqrt (a\<^sup>2 + b\<^sup>2)"  | 
| 33175 | 5504  | 
by (auto simp add: dist_Pair_Pair real_sqrt_le_mono add_mono power_mono)  | 
| 53282 | 5505  | 
then show ?thesis unfolding bounded_any_center [where a="(x, y)"] by auto  | 
| 33175 | 5506  | 
qed  | 
5507  | 
||
5508  | 
lemma mem_Times_iff: "x \<in> A \<times> B \<longleftrightarrow> fst x \<in> A \<and> snd x \<in> B"  | 
|
| 53282 | 5509  | 
by (induct x) simp  | 
| 33175 | 5510  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
5511  | 
lemma seq_compact_Times: "seq_compact s \<Longrightarrow> seq_compact t \<Longrightarrow> seq_compact (s \<times> t)"  | 
| 53282 | 5512  | 
unfolding seq_compact_def  | 
5513  | 
apply clarify  | 
|
5514  | 
apply (drule_tac x="fst \<circ> f" in spec)  | 
|
5515  | 
apply (drule mp, simp add: mem_Times_iff)  | 
|
5516  | 
apply (clarify, rename_tac l1 r1)  | 
|
5517  | 
apply (drule_tac x="snd \<circ> f \<circ> r1" in spec)  | 
|
5518  | 
apply (drule mp, simp add: mem_Times_iff)  | 
|
5519  | 
apply (clarify, rename_tac l2 r2)  | 
|
5520  | 
apply (rule_tac x="(l1, l2)" in rev_bexI, simp)  | 
|
5521  | 
apply (rule_tac x="r1 \<circ> r2" in exI)  | 
|
5522  | 
apply (rule conjI, simp add: subseq_def)  | 
|
5523  | 
apply (drule_tac f=r2 in LIMSEQ_subseq_LIMSEQ, assumption)  | 
|
5524  | 
apply (drule (1) tendsto_Pair) back  | 
|
5525  | 
apply (simp add: o_def)  | 
|
5526  | 
done  | 
|
5527  | 
||
5528  | 
lemma compact_Times:  | 
|
| 51349 | 5529  | 
assumes "compact s" "compact t"  | 
5530  | 
shows "compact (s \<times> t)"  | 
|
5531  | 
proof (rule compactI)  | 
|
| 53282 | 5532  | 
fix C  | 
5533  | 
assume C: "\<forall>t\<in>C. open t" "s \<times> t \<subseteq> \<Union>C"  | 
|
| 51349 | 5534  | 
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)"  | 
5535  | 
proof  | 
|
| 53282 | 5536  | 
fix x  | 
5537  | 
assume "x \<in> s"  | 
|
| 51349 | 5538  | 
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 | 5539  | 
proof  | 
5540  | 
fix y  | 
|
5541  | 
assume "y \<in> t"  | 
|
| 51349 | 5542  | 
with `x \<in> s` C obtain c where "c \<in> C" "(x, y) \<in> c" "open c" by auto  | 
5543  | 
then show "?P y" by (auto elim!: open_prod_elim)  | 
|
5544  | 
qed  | 
|
5545  | 
then obtain a b c where b: "\<And>y. y \<in> t \<Longrightarrow> open (b y)"  | 
|
5546  | 
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"  | 
|
5547  | 
by metis  | 
|
5548  | 
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
 | 
5549  | 
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 | 5550  | 
by auto  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
5551  | 
moreover from D c have "(\<Inter>y\<in>D. a y) \<times> t \<subseteq> (\<Union>y\<in>D. c y)"  | 
| 51349 | 5552  | 
by (fastforce simp: subset_eq)  | 
5553  | 
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
 | 
5554  | 
using c by (intro exI[of _ "c`D"] exI[of _ "\<Inter>(a`D)"] conjI) (auto intro!: open_INT)  | 
| 51349 | 5555  | 
qed  | 
5556  | 
then obtain a d where a: "\<forall>x\<in>s. open (a x)" "s \<subseteq> (\<Union>x\<in>s. a x)"  | 
|
5557  | 
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"  | 
|
5558  | 
unfolding subset_eq UN_iff by metis  | 
|
| 53282 | 5559  | 
moreover  | 
5560  | 
from compactE_image[OF `compact s` a]  | 
|
5561  | 
obtain e where e: "e \<subseteq> s" "finite e" and s: "s \<subseteq> (\<Union>x\<in>e. a x)"  | 
|
5562  | 
by auto  | 
|
| 51349 | 5563  | 
moreover  | 
| 53282 | 5564  | 
  {
 | 
5565  | 
from s have "s \<times> t \<subseteq> (\<Union>x\<in>e. a x \<times> t)"  | 
|
5566  | 
by auto  | 
|
5567  | 
also have "\<dots> \<subseteq> (\<Union>x\<in>e. \<Union>d x)"  | 
|
5568  | 
using d `e \<subseteq> s` by (intro UN_mono) auto  | 
|
5569  | 
finally have "s \<times> t \<subseteq> (\<Union>x\<in>e. \<Union>d x)" .  | 
|
5570  | 
}  | 
|
| 51349 | 5571  | 
ultimately show "\<exists>C'\<subseteq>C. finite C' \<and> s \<times> t \<subseteq> \<Union>C'"  | 
5572  | 
by (intro exI[of _ "(\<Union>x\<in>e. d x)"]) (auto simp add: subset_eq)  | 
|
5573  | 
qed  | 
|
| 
50884
 
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
 
hoelzl 
parents: 
50883 
diff
changeset
 | 
5574  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5575  | 
text{* Hence some useful properties follow quite easily. *}
 | 
| 33175 | 5576  | 
|
5577  | 
lemma compact_scaling:  | 
|
5578  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5579  | 
assumes "compact s"  | 
5580  | 
shows "compact ((\<lambda>x. c *\<^sub>R x) ` s)"  | 
|
5581  | 
proof -  | 
|
| 33175 | 5582  | 
let ?f = "\<lambda>x. scaleR c x"  | 
| 53282 | 5583  | 
have *: "bounded_linear ?f" by (rule bounded_linear_scaleR_right)  | 
5584  | 
show ?thesis  | 
|
5585  | 
using compact_continuous_image[of s ?f] continuous_at_imp_continuous_on[of s ?f]  | 
|
5586  | 
using linear_continuous_at[OF *] assms  | 
|
5587  | 
by auto  | 
|
| 33175 | 5588  | 
qed  | 
5589  | 
||
5590  | 
lemma compact_negations:  | 
|
5591  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5592  | 
assumes "compact s"  | 
| 53291 | 5593  | 
shows "compact ((\<lambda>x. - x) ` s)"  | 
| 33175 | 5594  | 
using compact_scaling [OF assms, of "- 1"] by auto  | 
5595  | 
||
5596  | 
lemma compact_sums:  | 
|
5597  | 
fixes s t :: "'a::real_normed_vector set"  | 
|
| 53291 | 5598  | 
assumes "compact s"  | 
5599  | 
and "compact t"  | 
|
| 53282 | 5600  | 
  shows "compact {x + y | x y. x \<in> s \<and> y \<in> t}"
 | 
5601  | 
proof -  | 
|
5602  | 
  have *: "{x + y | x y. x \<in> s \<and> y \<in> t} = (\<lambda>z. fst z + snd z) ` (s \<times> t)"
 | 
|
5603  | 
apply auto  | 
|
5604  | 
unfolding image_iff  | 
|
5605  | 
apply (rule_tac x="(xa, y)" in bexI)  | 
|
5606  | 
apply auto  | 
|
5607  | 
done  | 
|
| 33175 | 5608  | 
have "continuous_on (s \<times> t) (\<lambda>z. fst z + snd z)"  | 
5609  | 
unfolding continuous_on by (rule ballI) (intro tendsto_intros)  | 
|
| 53282 | 5610  | 
then show ?thesis  | 
5611  | 
unfolding * using compact_continuous_image compact_Times [OF assms] by auto  | 
|
| 33175 | 5612  | 
qed  | 
5613  | 
||
5614  | 
lemma compact_differences:  | 
|
5615  | 
fixes s t :: "'a::real_normed_vector set"  | 
|
| 53291 | 5616  | 
assumes "compact s"  | 
5617  | 
and "compact t"  | 
|
5618  | 
  shows "compact {x - y | x y. x \<in> s \<and> y \<in> t}"
 | 
|
| 33175 | 5619  | 
proof-  | 
5620  | 
  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 | 5621  | 
apply auto  | 
5622  | 
apply (rule_tac x= xa in exI)  | 
|
5623  | 
apply auto  | 
|
5624  | 
done  | 
|
5625  | 
then show ?thesis  | 
|
5626  | 
using compact_sums[OF assms(1) compact_negations[OF assms(2)]] by auto  | 
|
| 33175 | 5627  | 
qed  | 
5628  | 
||
5629  | 
lemma compact_translation:  | 
|
5630  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5631  | 
assumes "compact s"  | 
5632  | 
shows "compact ((\<lambda>x. a + x) ` s)"  | 
|
5633  | 
proof -  | 
|
5634  | 
  have "{x + y |x y. x \<in> s \<and> y \<in> {a}} = (\<lambda>x. a + x) ` s"
 | 
|
5635  | 
by auto  | 
|
5636  | 
then show ?thesis  | 
|
5637  | 
using compact_sums[OF assms compact_sing[of a]] by auto  | 
|
| 33175 | 5638  | 
qed  | 
5639  | 
||
5640  | 
lemma compact_affinity:  | 
|
5641  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5642  | 
assumes "compact s"  | 
5643  | 
shows "compact ((\<lambda>x. a + c *\<^sub>R x) ` s)"  | 
|
5644  | 
proof -  | 
|
5645  | 
have "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s"  | 
|
5646  | 
by auto  | 
|
5647  | 
then show ?thesis  | 
|
5648  | 
using compact_translation[OF compact_scaling[OF assms], of a c] by auto  | 
|
| 33175 | 5649  | 
qed  | 
5650  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5651  | 
text {* Hence we get the following. *}
 | 
| 33175 | 5652  | 
|
5653  | 
lemma compact_sup_maxdistance:  | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5654  | 
fixes s :: "'a::metric_space set"  | 
| 53291 | 5655  | 
assumes "compact s"  | 
5656  | 
    and "s \<noteq> {}"
 | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5657  | 
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 | 5658  | 
proof -  | 
5659  | 
have "compact (s \<times> s)"  | 
|
5660  | 
using `compact s` by (intro compact_Times)  | 
|
5661  | 
  moreover have "s \<times> s \<noteq> {}"
 | 
|
5662  | 
    using `s \<noteq> {}` by auto
 | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5663  | 
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
 | 
5664  | 
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
 | 
5665  | 
ultimately show ?thesis  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5666  | 
using continuous_attains_sup[of "s \<times> s" "\<lambda>x. dist (fst x) (snd x)"] by auto  | 
| 33175 | 5667  | 
qed  | 
5668  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5669  | 
text {* We can state this in terms of diameter of a set. *}
 | 
| 33175 | 5670  | 
|
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5671  | 
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
 | 
5672  | 
  "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
 | 
5673  | 
|
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5674  | 
lemma diameter_bounded_bound:  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5675  | 
fixes s :: "'a :: metric_space set"  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5676  | 
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
 | 
5677  | 
shows "dist x y \<le> diameter s"  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5678  | 
proof -  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5679  | 
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
 | 
5680  | 
unfolding bounded_def by auto  | 
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5681  | 
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
 | 
5682  | 
proof (intro bdd_aboveI, safe)  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5683  | 
fix a b  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5684  | 
assume "a \<in> s" "b \<in> s"  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5685  | 
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
 | 
5686  | 
show "dist a b \<le> 2 * d"  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5687  | 
by (simp add: dist_commute)  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5688  | 
qed  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5689  | 
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
 | 
5690  | 
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
 | 
5691  | 
by (rule cSUP_upper2) simp  | 
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5692  | 
with `x \<in> s` show ?thesis  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5693  | 
by (auto simp add: diameter_def)  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5694  | 
qed  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5695  | 
|
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5696  | 
lemma diameter_lower_bounded:  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5697  | 
fixes s :: "'a :: metric_space set"  | 
| 53282 | 5698  | 
assumes s: "bounded s"  | 
5699  | 
and d: "0 < d" "d < diameter s"  | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5700  | 
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
 | 
5701  | 
proof (rule ccontr)  | 
| 
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5702  | 
assume contr: "\<not> ?thesis"  | 
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5703  | 
  moreover have "s \<noteq> {}"
 | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5704  | 
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
 | 
5705  | 
ultimately have "diameter s \<le> d"  | 
| 
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5706  | 
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
 | 
5707  | 
with `d < diameter s` show False by auto  | 
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5708  | 
qed  | 
| 33175 | 5709  | 
|
5710  | 
lemma diameter_bounded:  | 
|
5711  | 
assumes "bounded s"  | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5712  | 
shows "\<forall>x\<in>s. \<forall>y\<in>s. dist x y \<le> diameter s"  | 
| 53291 | 5713  | 
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
 | 
5714  | 
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
 | 
5715  | 
by auto  | 
| 33175 | 5716  | 
|
5717  | 
lemma diameter_compact_attained:  | 
|
| 53291 | 5718  | 
assumes "compact s"  | 
5719  | 
    and "s \<noteq> {}"
 | 
|
| 
50973
 
4a2c82644889
generalized diameter from real_normed_vector to metric_space
 
hoelzl 
parents: 
50972 
diff
changeset
 | 
5720  | 
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
 | 
5721  | 
proof -  | 
| 53282 | 5722  | 
have b: "bounded s" using assms(1)  | 
5723  | 
by (rule compact_imp_bounded)  | 
|
| 53291 | 5724  | 
then obtain x y where xys: "x\<in>s" "y\<in>s"  | 
5725  | 
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
 | 
5726  | 
using compact_sup_maxdistance[OF assms] by auto  | 
| 53282 | 5727  | 
then have "diameter s \<le> dist x y"  | 
5728  | 
unfolding diameter_def  | 
|
5729  | 
apply clarsimp  | 
|
| 
54260
 
6a967667fd45
use INF and SUP on conditionally complete lattices in multivariate analysis
 
hoelzl 
parents: 
54259 
diff
changeset
 | 
5730  | 
apply (rule cSUP_least)  | 
| 53282 | 5731  | 
apply fast+  | 
5732  | 
done  | 
|
5733  | 
then show ?thesis  | 
|
| 
36362
 
06475a1547cb
fix lots of looping simp calls and other warnings
 
huffman 
parents: 
36360 
diff
changeset
 | 
5734  | 
by (metis b diameter_bounded_bound order_antisym xys)  | 
| 33175 | 5735  | 
qed  | 
5736  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5737  | 
text {* Related results with closure as the conclusion. *}
 | 
| 33175 | 5738  | 
|
5739  | 
lemma closed_scaling:  | 
|
5740  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5741  | 
assumes "closed s"  | 
5742  | 
shows "closed ((\<lambda>x. c *\<^sub>R x) ` s)"  | 
|
| 53813 | 5743  | 
proof (cases "c = 0")  | 
5744  | 
case True then show ?thesis  | 
|
5745  | 
by (auto simp add: image_constant_conv)  | 
|
| 33175 | 5746  | 
next  | 
5747  | 
case False  | 
|
| 53813 | 5748  | 
from assms have "closed ((\<lambda>x. inverse c *\<^sub>R x) -` s)"  | 
5749  | 
by (simp add: continuous_closed_vimage)  | 
|
5750  | 
also have "(\<lambda>x. inverse c *\<^sub>R x) -` s = (\<lambda>x. c *\<^sub>R x) ` s"  | 
|
5751  | 
using `c \<noteq> 0` by (auto elim: image_eqI [rotated])  | 
|
5752  | 
finally show ?thesis .  | 
|
| 33175 | 5753  | 
qed  | 
5754  | 
||
5755  | 
lemma closed_negations:  | 
|
5756  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5757  | 
assumes "closed s"  | 
5758  | 
shows "closed ((\<lambda>x. -x) ` s)"  | 
|
| 33175 | 5759  | 
using closed_scaling[OF assms, of "- 1"] by simp  | 
5760  | 
||
5761  | 
lemma compact_closed_sums:  | 
|
5762  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 5763  | 
assumes "compact s" and "closed t"  | 
5764  | 
  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
 | 
|
5765  | 
proof -  | 
|
| 33175 | 5766  | 
  let ?S = "{x + y |x y. x \<in> s \<and> y \<in> t}"
 | 
| 53282 | 5767  | 
  {
 | 
5768  | 
fix x l  | 
|
5769  | 
assume as: "\<forall>n. x n \<in> ?S" "(x ---> l) sequentially"  | 
|
5770  | 
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 | 5771  | 
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 | 5772  | 
obtain l' r where "l'\<in>s" and r: "subseq r" and lr: "(((\<lambda>n. fst (f n)) \<circ> r) ---> l') sequentially"  | 
| 33175 | 5773  | 
using assms(1)[unfolded compact_def, THEN spec[where x="\<lambda> n. fst (f n)"]] using f(2) by auto  | 
5774  | 
have "((\<lambda>n. snd (f (r n))) ---> l - l') sequentially"  | 
|
| 53282 | 5775  | 
using tendsto_diff[OF LIMSEQ_subseq_LIMSEQ[OF as(2) r] lr] and f(1)  | 
5776  | 
unfolding o_def  | 
|
5777  | 
by auto  | 
|
5778  | 
then have "l - l' \<in> t"  | 
|
| 53291 | 5779  | 
using assms(2)[unfolded closed_sequential_limits,  | 
5780  | 
THEN spec[where x="\<lambda> n. snd (f (r n))"],  | 
|
5781  | 
THEN spec[where x="l - l'"]]  | 
|
| 53282 | 5782  | 
using f(3)  | 
5783  | 
by auto  | 
|
5784  | 
then have "l \<in> ?S"  | 
|
5785  | 
using `l' \<in> s`  | 
|
5786  | 
apply auto  | 
|
5787  | 
apply (rule_tac x=l' in exI)  | 
|
5788  | 
apply (rule_tac x="l - l'" in exI)  | 
|
5789  | 
apply auto  | 
|
5790  | 
done  | 
|
| 33175 | 5791  | 
}  | 
| 53282 | 5792  | 
then show ?thesis  | 
5793  | 
unfolding closed_sequential_limits by fast  | 
|
| 33175 | 5794  | 
qed  | 
5795  | 
||
5796  | 
lemma closed_compact_sums:  | 
|
5797  | 
fixes s t :: "'a::real_normed_vector set"  | 
|
| 53291 | 5798  | 
assumes "closed s"  | 
5799  | 
and "compact t"  | 
|
| 33175 | 5800  | 
  shows "closed {x + y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 53282 | 5801  | 
proof -  | 
5802  | 
  have "{x + y |x y. x \<in> t \<and> y \<in> s} = {x + y |x y. x \<in> s \<and> y \<in> t}"
 | 
|
5803  | 
apply auto  | 
|
5804  | 
apply (rule_tac x=y in exI)  | 
|
5805  | 
apply auto  | 
|
5806  | 
apply (rule_tac x=y in exI)  | 
|
5807  | 
apply auto  | 
|
5808  | 
done  | 
|
5809  | 
then show ?thesis  | 
|
5810  | 
using compact_closed_sums[OF assms(2,1)] by simp  | 
|
| 33175 | 5811  | 
qed  | 
5812  | 
||
5813  | 
lemma compact_closed_differences:  | 
|
5814  | 
fixes s t :: "'a::real_normed_vector set"  | 
|
| 53291 | 5815  | 
assumes "compact s"  | 
5816  | 
and "closed t"  | 
|
| 33175 | 5817  | 
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 53282 | 5818  | 
proof -  | 
| 33175 | 5819  | 
  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 | 5820  | 
apply auto  | 
5821  | 
apply (rule_tac x=xa in exI)  | 
|
5822  | 
apply auto  | 
|
5823  | 
apply (rule_tac x=xa in exI)  | 
|
5824  | 
apply auto  | 
|
5825  | 
done  | 
|
5826  | 
then show ?thesis  | 
|
5827  | 
using compact_closed_sums[OF assms(1) closed_negations[OF assms(2)]] by auto  | 
|
| 33175 | 5828  | 
qed  | 
5829  | 
||
5830  | 
lemma closed_compact_differences:  | 
|
5831  | 
fixes s t :: "'a::real_normed_vector set"  | 
|
| 53291 | 5832  | 
assumes "closed s"  | 
5833  | 
and "compact t"  | 
|
| 33175 | 5834  | 
  shows "closed {x - y | x y. x \<in> s \<and> y \<in> t}"
 | 
| 53282 | 5835  | 
proof -  | 
| 33175 | 5836  | 
  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 | 5837  | 
apply auto  | 
5838  | 
apply (rule_tac x=xa in exI)  | 
|
5839  | 
apply auto  | 
|
5840  | 
apply (rule_tac x=xa in exI)  | 
|
5841  | 
apply auto  | 
|
5842  | 
done  | 
|
5843  | 
then show ?thesis  | 
|
5844  | 
using closed_compact_sums[OF assms(1) compact_negations[OF assms(2)]] by simp  | 
|
| 33175 | 5845  | 
qed  | 
5846  | 
||
5847  | 
lemma closed_translation:  | 
|
5848  | 
fixes a :: "'a::real_normed_vector"  | 
|
| 53282 | 5849  | 
assumes "closed s"  | 
5850  | 
shows "closed ((\<lambda>x. a + x) ` s)"  | 
|
5851  | 
proof -  | 
|
| 33175 | 5852  | 
  have "{a + y |y. y \<in> s} = (op + a ` s)" by auto
 | 
| 53282 | 5853  | 
then show ?thesis  | 
5854  | 
using compact_closed_sums[OF compact_sing[of a] assms] by auto  | 
|
| 33175 | 5855  | 
qed  | 
5856  | 
||
| 34105 | 5857  | 
lemma translation_Compl:  | 
5858  | 
fixes a :: "'a::ab_group_add"  | 
|
5859  | 
shows "(\<lambda>x. a + x) ` (- t) = - ((\<lambda>x. a + x) ` t)"  | 
|
| 53282 | 5860  | 
apply (auto simp add: image_iff)  | 
5861  | 
apply (rule_tac x="x - a" in bexI)  | 
|
5862  | 
apply auto  | 
|
5863  | 
done  | 
|
| 34105 | 5864  | 
|
| 33175 | 5865  | 
lemma translation_UNIV:  | 
| 53282 | 5866  | 
fixes a :: "'a::ab_group_add"  | 
5867  | 
shows "range (\<lambda>x. a + x) = UNIV"  | 
|
5868  | 
apply (auto simp add: image_iff)  | 
|
5869  | 
apply (rule_tac x="x - a" in exI)  | 
|
5870  | 
apply auto  | 
|
5871  | 
done  | 
|
| 33175 | 5872  | 
|
5873  | 
lemma translation_diff:  | 
|
5874  | 
fixes a :: "'a::ab_group_add"  | 
|
5875  | 
shows "(\<lambda>x. a + x) ` (s - t) = ((\<lambda>x. a + x) ` s) - ((\<lambda>x. a + x) ` t)"  | 
|
5876  | 
by auto  | 
|
5877  | 
||
5878  | 
lemma closure_translation:  | 
|
5879  | 
fixes a :: "'a::real_normed_vector"  | 
|
5880  | 
shows "closure ((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (closure s)"  | 
|
| 53282 | 5881  | 
proof -  | 
5882  | 
have *: "op + a ` (- s) = - op + a ` s"  | 
|
5883  | 
apply auto  | 
|
5884  | 
unfolding image_iff  | 
|
5885  | 
apply (rule_tac x="x - a" in bexI)  | 
|
5886  | 
apply auto  | 
|
5887  | 
done  | 
|
5888  | 
show ?thesis  | 
|
5889  | 
unfolding closure_interior translation_Compl  | 
|
5890  | 
using interior_translation[of a "- s"]  | 
|
5891  | 
unfolding *  | 
|
5892  | 
by auto  | 
|
| 33175 | 5893  | 
qed  | 
5894  | 
||
5895  | 
lemma frontier_translation:  | 
|
5896  | 
fixes a :: "'a::real_normed_vector"  | 
|
5897  | 
shows "frontier((\<lambda>x. a + x) ` s) = (\<lambda>x. a + x) ` (frontier s)"  | 
|
| 53282 | 5898  | 
unfolding frontier_def translation_diff interior_translation closure_translation  | 
5899  | 
by auto  | 
|
| 33175 | 5900  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5901  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
5902  | 
subsection {* Separation between points and sets *}
 | 
| 33175 | 5903  | 
|
5904  | 
lemma separate_point_closed:  | 
|
5905  | 
fixes s :: "'a::heine_borel set"  | 
|
| 53291 | 5906  | 
assumes "closed s"  | 
5907  | 
and "a \<notin> s"  | 
|
| 53282 | 5908  | 
shows "\<exists>d>0. \<forall>x\<in>s. d \<le> dist a x"  | 
5909  | 
proof (cases "s = {}")
 | 
|
| 33175 | 5910  | 
case True  | 
| 53282 | 5911  | 
then show ?thesis by(auto intro!: exI[where x=1])  | 
| 33175 | 5912  | 
next  | 
5913  | 
case False  | 
|
| 53282 | 5914  | 
from assms obtain x where "x\<in>s" "\<forall>y\<in>s. dist a x \<le> dist a y"  | 
5915  | 
    using `s \<noteq> {}` distance_attains_inf [of s a] by blast
 | 
|
5916  | 
with `x\<in>s` show ?thesis using dist_pos_lt[of a x] and`a \<notin> s`  | 
|
5917  | 
by blast  | 
|
| 33175 | 5918  | 
qed  | 
5919  | 
||
5920  | 
lemma separate_compact_closed:  | 
|
| 50949 | 5921  | 
fixes s t :: "'a::heine_borel set"  | 
| 53282 | 5922  | 
assumes "compact s"  | 
5923  | 
    and t: "closed t" "s \<inter> t = {}"
 | 
|
| 33175 | 5924  | 
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
 | 
5925  | 
proof cases  | 
| 
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5926  | 
  assume "s \<noteq> {} \<and> t \<noteq> {}"
 | 
| 
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5927  | 
  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
 | 
5928  | 
let ?inf = "\<lambda>x. infdist x t"  | 
| 
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5929  | 
have "continuous_on s ?inf"  | 
| 
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5930  | 
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
 | 
5931  | 
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
 | 
5932  | 
    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
 | 
5933  | 
then have "0 < ?inf x"  | 
| 
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5934  | 
    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
 | 
5935  | 
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
 | 
5936  | 
using x by (auto intro: order_trans infdist_le)  | 
| 53282 | 5937  | 
ultimately show ?thesis by auto  | 
| 
51346
 
d33de22432e2
tuned proofs (used continuity of infdist, dist and continuous_attains_sup)
 
hoelzl 
parents: 
51345 
diff
changeset
 | 
5938  | 
qed (auto intro!: exI[of _ 1])  | 
| 33175 | 5939  | 
|
5940  | 
lemma separate_closed_compact:  | 
|
| 50949 | 5941  | 
fixes s t :: "'a::heine_borel set"  | 
| 53282 | 5942  | 
assumes "closed s"  | 
5943  | 
and "compact t"  | 
|
5944  | 
    and "s \<inter> t = {}"
 | 
|
| 33175 | 5945  | 
shows "\<exists>d>0. \<forall>x\<in>s. \<forall>y\<in>t. d \<le> dist x y"  | 
| 53282 | 5946  | 
proof -  | 
5947  | 
  have *: "t \<inter> s = {}"
 | 
|
5948  | 
using assms(3) by auto  | 
|
5949  | 
show ?thesis  | 
|
5950  | 
using separate_compact_closed[OF assms(2,1) *]  | 
|
5951  | 
apply auto  | 
|
5952  | 
apply (rule_tac x=d in exI)  | 
|
5953  | 
apply auto  | 
|
5954  | 
apply (erule_tac x=y in ballE)  | 
|
5955  | 
apply (auto simp add: dist_commute)  | 
|
5956  | 
done  | 
|
| 33175 | 5957  | 
qed  | 
5958  | 
||
| 36439 | 5959  | 
subsection {* Intervals *}
 | 
| 53282 | 5960  | 
|
| 
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
 | 
5961  | 
lemma open_box: "open (box a 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
 | 
5962  | 
proof -  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5963  | 
  have "open (\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<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
 | 
5964  | 
by (auto intro!: continuous_open_vimage continuous_inner continuous_at_id continuous_const)  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5965  | 
  also have "(\<Inter>i\<in>Basis. (op \<bullet> i) -` {a \<bullet> i <..< b \<bullet> i}) = box a 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
 | 
5966  | 
by (auto simp add: box_def inner_commute)  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5967  | 
finally 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
 | 
5968  | 
qed  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5969  | 
|
| 
50881
 
ae630bab13da
renamed countable_basis_space to second_countable_topology
 
hoelzl 
parents: 
50526 
diff
changeset
 | 
5970  | 
instance euclidean_space \<subseteq> second_countable_topology  | 
| 50087 | 5971  | 
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
 | 
5972  | 
def a \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. fst (f i) *\<^sub>R i"  | 
| 53282 | 5973  | 
then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f"  | 
5974  | 
by simp  | 
|
| 
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
 | 
5975  | 
def b \<equiv> "\<lambda>f :: 'a \<Rightarrow> (real \<times> real). \<Sum>i\<in>Basis. snd (f i) *\<^sub>R i"  | 
| 53282 | 5976  | 
then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f"  | 
5977  | 
by simp  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52625 
diff
changeset
 | 
5978  | 
def B \<equiv> "(\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))"  | 
| 
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
 | 
5979  | 
|
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5980  | 
have "Ball B open" by (simp add: B_def open_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
 | 
5981  | 
moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))"  | 
| 50087 | 5982  | 
proof safe  | 
| 53282 | 5983  | 
fix A::"'a set"  | 
5984  | 
assume "open 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
 | 
5985  | 
show "\<exists>B'\<subseteq>B. \<Union>B' = A"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5986  | 
      apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
 | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5987  | 
apply (subst (3) open_UNION_box[OF `open A`])  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50324 
diff
changeset
 | 
5988  | 
apply (auto simp add: a b B_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
 | 
5989  | 
done  | 
| 50087 | 5990  | 
qed  | 
5991  | 
ultimately  | 
|
| 53282 | 5992  | 
have "topological_basis B"  | 
5993  | 
unfolding topological_basis_def by blast  | 
|
| 
51343
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
5994  | 
moreover  | 
| 53282 | 5995  | 
have "countable B"  | 
5996  | 
unfolding B_def  | 
|
| 
51343
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
5997  | 
by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat)  | 
| 
 
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
 
hoelzl 
parents: 
51342 
diff
changeset
 | 
5998  | 
ultimately show "\<exists>B::'a 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
 | 
5999  | 
by (blast intro: topological_basis_imp_subbasis)  | 
| 50087 | 6000  | 
qed  | 
6001  | 
||
| 51103 | 6002  | 
instance euclidean_space \<subseteq> polish_space ..  | 
| 50087 | 6003  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6004  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6005  | 
subsection {* Closure of halfspaces and hyperplanes *}
 | 
| 33175 | 6006  | 
|
| 
44219
 
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
 
huffman 
parents: 
44216 
diff
changeset
 | 
6007  | 
lemma isCont_open_vimage:  | 
| 53282 | 6008  | 
assumes "\<And>x. isCont f x"  | 
6009  | 
and "open s"  | 
|
6010  | 
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
 | 
6011  | 
proof -  | 
| 
 
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
 
huffman 
parents: 
44216 
diff
changeset
 | 
6012  | 
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
 | 
6013  | 
unfolding isCont_def continuous_on_def by simp  | 
| 53282 | 6014  | 
  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
 | 
6015  | 
using open_UNIV `open s` by (rule continuous_open_preimage)  | 
| 53282 | 6016  | 
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
 | 
6017  | 
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
 | 
6018  | 
qed  | 
| 
 
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
 
huffman 
parents: 
44216 
diff
changeset
 | 
6019  | 
|
| 
 
f738e3200e24
generalize lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq to class topological_space
 
huffman 
parents: 
44216 
diff
changeset
 | 
6020  | 
lemma isCont_closed_vimage:  | 
| 53282 | 6021  | 
assumes "\<And>x. isCont f x"  | 
6022  | 
and "closed s"  | 
|
6023  | 
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
 | 
6024  | 
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
 | 
6025  | 
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
 | 
6026  | 
|
| 
44213
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6027  | 
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
 | 
6028  | 
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
 | 
6029  | 
assumes f: "\<And>x. isCont f x"  | 
| 53282 | 6030  | 
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
 | 
6031  | 
  shows "open {x. f x < g x}"
 | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6032  | 
proof -  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6033  | 
  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
 | 
6034  | 
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
 | 
6035  | 
by (rule isCont_open_vimage)  | 
| 
44213
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6036  | 
  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
 | 
6037  | 
by auto  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6038  | 
finally show ?thesis .  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6039  | 
qed  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6040  | 
|
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6041  | 
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
 | 
6042  | 
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
 | 
6043  | 
assumes f: "\<And>x. isCont f x"  | 
| 53282 | 6044  | 
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
 | 
6045  | 
  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
 | 
6046  | 
proof -  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6047  | 
  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
 | 
6048  | 
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
 | 
6049  | 
by (rule isCont_closed_vimage)  | 
| 
44213
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6050  | 
  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
 | 
6051  | 
by auto  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6052  | 
finally show ?thesis .  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6053  | 
qed  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6054  | 
|
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6055  | 
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
 | 
6056  | 
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
 | 
6057  | 
assumes f: "\<And>x. isCont f x"  | 
| 53282 | 6058  | 
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
 | 
6059  | 
  shows "closed {x. f x = g x}"
 | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6060  | 
proof -  | 
| 44216 | 6061  | 
  have "open {(x::'b, y::'b). x \<noteq> y}"
 | 
6062  | 
unfolding open_prod_def by (auto dest!: hausdorff)  | 
|
| 53282 | 6063  | 
  then have "closed {(x::'b, y::'b). x = y}"
 | 
| 44216 | 6064  | 
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
 | 
6065  | 
with isCont_Pair [OF f g]  | 
| 44216 | 6066  | 
  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
 | 
6067  | 
by (rule isCont_closed_vimage)  | 
| 44216 | 6068  | 
  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
 | 
6069  | 
finally show ?thesis .  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6070  | 
qed  | 
| 
 
6fb54701a11b
add lemmas open_Collect_less, closed_Collect_le, closed_Collect_eq;
 
huffman 
parents: 
44212 
diff
changeset
 | 
6071  | 
|
| 33175 | 6072  | 
lemma continuous_at_inner: "continuous (at x) (inner a)"  | 
6073  | 
unfolding continuous_at by (intro tendsto_intros)  | 
|
6074  | 
||
6075  | 
lemma closed_halfspace_le: "closed {x. inner a x \<le> b}"
 | 
|
| 44233 | 6076  | 
by (simp add: closed_Collect_le)  | 
| 33175 | 6077  | 
|
6078  | 
lemma closed_halfspace_ge: "closed {x. inner a x \<ge> b}"
 | 
|
| 44233 | 6079  | 
by (simp add: closed_Collect_le)  | 
| 33175 | 6080  | 
|
6081  | 
lemma closed_hyperplane: "closed {x. inner a x = b}"
 | 
|
| 44233 | 6082  | 
by (simp add: closed_Collect_eq)  | 
| 33175 | 6083  | 
|
| 53282 | 6084  | 
lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\<bullet>i \<le> a}"
 | 
| 44233 | 6085  | 
by (simp add: closed_Collect_le)  | 
| 33175 | 6086  | 
|
| 53282 | 6087  | 
lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\<bullet>i \<ge> a}"
 | 
| 44233 | 6088  | 
by (simp add: closed_Collect_le)  | 
| 33175 | 6089  | 
|
| 53813 | 6090  | 
lemma closed_interval_left:  | 
6091  | 
fixes b :: "'a::euclidean_space"  | 
|
6092  | 
  shows "closed {x::'a. \<forall>i\<in>Basis. x\<bullet>i \<le> b\<bullet>i}"
 | 
|
6093  | 
by (simp add: Collect_ball_eq closed_INT closed_Collect_le)  | 
|
6094  | 
||
6095  | 
lemma closed_interval_right:  | 
|
6096  | 
fixes a :: "'a::euclidean_space"  | 
|
6097  | 
  shows "closed {x::'a. \<forall>i\<in>Basis. a\<bullet>i \<le> x\<bullet>i}"
 | 
|
6098  | 
by (simp add: Collect_ball_eq closed_INT closed_Collect_le)  | 
|
6099  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6100  | 
text {* Openness of halfspaces. *}
 | 
| 33175 | 6101  | 
|
6102  | 
lemma open_halfspace_lt: "open {x. inner a x < b}"
 | 
|
| 44233 | 6103  | 
by (simp add: open_Collect_less)  | 
| 33175 | 6104  | 
|
6105  | 
lemma open_halfspace_gt: "open {x. inner a x > b}"
 | 
|
| 44233 | 6106  | 
by (simp add: open_Collect_less)  | 
| 33175 | 6107  | 
|
| 53282 | 6108  | 
lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\<bullet>i < a}"
 | 
| 44233 | 6109  | 
by (simp add: open_Collect_less)  | 
| 33175 | 6110  | 
|
| 53282 | 6111  | 
lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\<bullet>i > a}"
 | 
| 44233 | 6112  | 
by (simp add: open_Collect_less)  | 
| 33175 | 6113  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6114  | 
text {* This gives a simple derivation of limit component bounds. *}
 | 
| 33175 | 6115  | 
|
| 53282 | 6116  | 
lemma Lim_component_le:  | 
6117  | 
fixes f :: "'a \<Rightarrow> 'b::euclidean_space"  | 
|
6118  | 
assumes "(f ---> l) net"  | 
|
6119  | 
and "\<not> (trivial_limit net)"  | 
|
6120  | 
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
 | 
6121  | 
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
 | 
6122  | 
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
 | 
6123  | 
|
| 53282 | 6124  | 
lemma Lim_component_ge:  | 
6125  | 
fixes f :: "'a \<Rightarrow> 'b::euclidean_space"  | 
|
6126  | 
assumes "(f ---> l) net"  | 
|
6127  | 
and "\<not> (trivial_limit net)"  | 
|
6128  | 
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
 | 
6129  | 
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
 | 
6130  | 
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
 | 
6131  | 
|
| 53282 | 6132  | 
lemma Lim_component_eq:  | 
6133  | 
fixes f :: "'a \<Rightarrow> 'b::euclidean_space"  | 
|
| 53640 | 6134  | 
assumes net: "(f ---> l) net" "\<not> trivial_limit net"  | 
| 53282 | 6135  | 
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
 | 
6136  | 
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
 | 
6137  | 
using ev[unfolded order_eq_iff eventually_conj_iff]  | 
| 53282 | 6138  | 
using Lim_component_ge[OF net, of b i]  | 
6139  | 
using Lim_component_le[OF net, of i b]  | 
|
6140  | 
by auto  | 
|
6141  | 
||
6142  | 
text {* Limits relative to a union. *}
 | 
|
| 33175 | 6143  | 
|
6144  | 
lemma eventually_within_Un:  | 
|
| 53282 | 6145  | 
"eventually P (at x within (s \<union> t)) \<longleftrightarrow>  | 
6146  | 
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
 | 
6147  | 
unfolding eventually_at_filter  | 
| 33175 | 6148  | 
by (auto elim!: eventually_rev_mp)  | 
6149  | 
||
6150  | 
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
 | 
6151  | 
"(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
 | 
6152  | 
(f ---> l) (at x within s) \<and> (f ---> l) (at x within t)"  | 
| 33175 | 6153  | 
unfolding tendsto_def  | 
6154  | 
by (auto simp add: eventually_within_Un)  | 
|
6155  | 
||
| 36442 | 6156  | 
lemma Lim_topological:  | 
| 53282 | 6157  | 
"(f ---> l) net \<longleftrightarrow>  | 
6158  | 
trivial_limit net \<or> (\<forall>S. open S \<longrightarrow> l \<in> S \<longrightarrow> eventually (\<lambda>x. f x \<in> S) net)"  | 
|
| 36442 | 6159  | 
unfolding tendsto_def trivial_limit_eq by auto  | 
6160  | 
||
| 53282 | 6161  | 
text{* Some more convenient intermediate-value theorem formulations. *}
 | 
| 33175 | 6162  | 
|
6163  | 
lemma connected_ivt_hyperplane:  | 
|
| 53291 | 6164  | 
assumes "connected s"  | 
6165  | 
and "x \<in> s"  | 
|
6166  | 
and "y \<in> s"  | 
|
6167  | 
and "inner a x \<le> b"  | 
|
6168  | 
and "b \<le> inner a y"  | 
|
| 33175 | 6169  | 
shows "\<exists>z \<in> s. inner a z = b"  | 
| 53282 | 6170  | 
proof (rule ccontr)  | 
| 33175 | 6171  | 
assume as:"\<not> (\<exists>z\<in>s. inner a z = b)"  | 
6172  | 
  let ?A = "{x. inner a x < b}"
 | 
|
6173  | 
  let ?B = "{x. inner a x > b}"
 | 
|
| 53282 | 6174  | 
have "open ?A" "open ?B"  | 
6175  | 
using open_halfspace_lt and open_halfspace_gt by auto  | 
|
| 53291 | 6176  | 
moreover  | 
6177  | 
  have "?A \<inter> ?B = {}" by auto
 | 
|
6178  | 
moreover  | 
|
6179  | 
have "s \<subseteq> ?A \<union> ?B" using as by auto  | 
|
6180  | 
ultimately  | 
|
6181  | 
show False  | 
|
| 53282 | 6182  | 
using assms(1)[unfolded connected_def not_ex,  | 
6183  | 
THEN spec[where x="?A"], THEN spec[where x="?B"]]  | 
|
6184  | 
using assms(2-5)  | 
|
| 52625 | 6185  | 
by auto  | 
6186  | 
qed  | 
|
6187  | 
||
6188  | 
lemma connected_ivt_component:  | 
|
6189  | 
fixes x::"'a::euclidean_space"  | 
|
6190  | 
shows "connected s \<Longrightarrow>  | 
|
6191  | 
x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow>  | 
|
6192  | 
x\<bullet>k \<le> a \<Longrightarrow> a \<le> y\<bullet>k \<Longrightarrow> (\<exists>z\<in>s. z\<bullet>k = a)"  | 
|
6193  | 
using connected_ivt_hyperplane[of s x y "k::'a" a]  | 
|
6194  | 
by (auto simp: inner_commute)  | 
|
| 33175 | 6195  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6196  | 
|
| 36437 | 6197  | 
subsection {* Homeomorphisms *}
 | 
| 33175 | 6198  | 
|
| 52625 | 6199  | 
definition "homeomorphism s t f g \<longleftrightarrow>  | 
6200  | 
(\<forall>x\<in>s. (g(f x) = x)) \<and> (f ` s = t) \<and> continuous_on s f \<and>  | 
|
6201  | 
(\<forall>y\<in>t. (f(g y) = y)) \<and> (g ` t = s) \<and> continuous_on t g"  | 
|
| 33175 | 6202  | 
|
| 53640 | 6203  | 
definition homeomorphic :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool"  | 
| 53282 | 6204  | 
(infixr "homeomorphic" 60)  | 
6205  | 
where "s homeomorphic t \<equiv> (\<exists>f g. homeomorphism s t f g)"  | 
|
| 33175 | 6206  | 
|
6207  | 
lemma homeomorphic_refl: "s homeomorphic s"  | 
|
6208  | 
unfolding homeomorphic_def  | 
|
6209  | 
unfolding homeomorphism_def  | 
|
6210  | 
using continuous_on_id  | 
|
| 53282 | 6211  | 
apply (rule_tac x = "(\<lambda>x. x)" in exI)  | 
6212  | 
apply (rule_tac x = "(\<lambda>x. x)" in exI)  | 
|
| 52625 | 6213  | 
apply blast  | 
6214  | 
done  | 
|
6215  | 
||
6216  | 
lemma homeomorphic_sym: "s homeomorphic t \<longleftrightarrow> t homeomorphic s"  | 
|
6217  | 
unfolding homeomorphic_def  | 
|
6218  | 
unfolding homeomorphism_def  | 
|
| 53282 | 6219  | 
by blast  | 
| 33175 | 6220  | 
|
6221  | 
lemma homeomorphic_trans:  | 
|
| 53282 | 6222  | 
assumes "s homeomorphic t"  | 
6223  | 
and "t homeomorphic u"  | 
|
| 52625 | 6224  | 
shows "s homeomorphic u"  | 
| 53282 | 6225  | 
proof -  | 
6226  | 
obtain f1 g1 where fg1: "\<forall>x\<in>s. g1 (f1 x) = x" "f1 ` s = t"  | 
|
6227  | 
"continuous_on s f1" "\<forall>y\<in>t. f1 (g1 y) = y" "g1 ` t = s" "continuous_on t g1"  | 
|
| 33175 | 6228  | 
using assms(1) unfolding homeomorphic_def homeomorphism_def by auto  | 
| 53282 | 6229  | 
obtain f2 g2 where fg2: "\<forall>x\<in>t. g2 (f2 x) = x" "f2 ` t = u" "continuous_on t f2"  | 
6230  | 
"\<forall>y\<in>u. f2 (g2 y) = y" "g2 ` u = t" "continuous_on u g2"  | 
|
| 33175 | 6231  | 
using assms(2) unfolding homeomorphic_def homeomorphism_def by auto  | 
| 52625 | 6232  | 
  {
 | 
6233  | 
fix x  | 
|
6234  | 
assume "x\<in>s"  | 
|
| 53282 | 6235  | 
then have "(g1 \<circ> g2) ((f2 \<circ> f1) x) = x"  | 
| 52625 | 6236  | 
using fg1(1)[THEN bspec[where x=x]] and fg2(1)[THEN bspec[where x="f1 x"]] and fg1(2)  | 
6237  | 
by auto  | 
|
6238  | 
}  | 
|
6239  | 
moreover have "(f2 \<circ> f1) ` s = u"  | 
|
6240  | 
using fg1(2) fg2(2) by auto  | 
|
6241  | 
moreover have "continuous_on s (f2 \<circ> f1)"  | 
|
6242  | 
using continuous_on_compose[OF fg1(3)] and fg2(3) unfolding fg1(2) by auto  | 
|
| 53282 | 6243  | 
moreover  | 
6244  | 
  {
 | 
|
| 52625 | 6245  | 
fix y  | 
6246  | 
assume "y\<in>u"  | 
|
| 53282 | 6247  | 
then have "(f2 \<circ> f1) ((g1 \<circ> g2) y) = y"  | 
| 52625 | 6248  | 
using fg2(4)[THEN bspec[where x=y]] and fg1(4)[THEN bspec[where x="g2 y"]] and fg2(5)  | 
6249  | 
by auto  | 
|
6250  | 
}  | 
|
| 33175 | 6251  | 
moreover have "(g1 \<circ> g2) ` u = s" using fg1(5) fg2(5) by auto  | 
| 52625 | 6252  | 
moreover have "continuous_on u (g1 \<circ> g2)"  | 
6253  | 
using continuous_on_compose[OF fg2(6)] and fg1(6)  | 
|
6254  | 
unfolding fg2(5)  | 
|
6255  | 
by auto  | 
|
6256  | 
ultimately show ?thesis  | 
|
6257  | 
unfolding homeomorphic_def homeomorphism_def  | 
|
6258  | 
apply (rule_tac x="f2 \<circ> f1" in exI)  | 
|
6259  | 
apply (rule_tac x="g1 \<circ> g2" in exI)  | 
|
6260  | 
apply auto  | 
|
6261  | 
done  | 
|
| 33175 | 6262  | 
qed  | 
6263  | 
||
6264  | 
lemma homeomorphic_minimal:  | 
|
| 52625 | 6265  | 
"s homeomorphic t \<longleftrightarrow>  | 
| 33175 | 6266  | 
(\<exists>f g. (\<forall>x\<in>s. f(x) \<in> t \<and> (g(f(x)) = x)) \<and>  | 
6267  | 
(\<forall>y\<in>t. g(y) \<in> s \<and> (f(g(y)) = y)) \<and>  | 
|
6268  | 
continuous_on s f \<and> continuous_on t g)"  | 
|
| 52625 | 6269  | 
unfolding homeomorphic_def homeomorphism_def  | 
6270  | 
apply auto  | 
|
6271  | 
apply (rule_tac x=f in exI)  | 
|
6272  | 
apply (rule_tac x=g in exI)  | 
|
6273  | 
apply auto  | 
|
6274  | 
apply (rule_tac x=f in exI)  | 
|
6275  | 
apply (rule_tac x=g in exI)  | 
|
6276  | 
apply auto  | 
|
6277  | 
unfolding image_iff  | 
|
6278  | 
apply (erule_tac x="g x" in ballE)  | 
|
6279  | 
apply (erule_tac x="x" in ballE)  | 
|
6280  | 
apply auto  | 
|
6281  | 
apply (rule_tac x="g x" in bexI)  | 
|
6282  | 
apply auto  | 
|
6283  | 
apply (erule_tac x="f x" in ballE)  | 
|
6284  | 
apply (erule_tac x="x" in ballE)  | 
|
6285  | 
apply auto  | 
|
6286  | 
apply (rule_tac x="f x" in bexI)  | 
|
6287  | 
apply auto  | 
|
6288  | 
done  | 
|
| 33175 | 6289  | 
|
| 36437 | 6290  | 
text {* Relatively weak hypotheses if a set is compact. *}
 | 
| 33175 | 6291  | 
|
6292  | 
lemma homeomorphism_compact:  | 
|
| 50898 | 6293  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"  | 
| 33175 | 6294  | 
assumes "compact s" "continuous_on s f" "f ` s = t" "inj_on f s"  | 
6295  | 
shows "\<exists>g. homeomorphism s t f g"  | 
|
| 53282 | 6296  | 
proof -  | 
| 33175 | 6297  | 
def g \<equiv> "\<lambda>x. SOME y. y\<in>s \<and> f y = x"  | 
| 52625 | 6298  | 
have g: "\<forall>x\<in>s. g (f x) = x"  | 
6299  | 
using assms(3) assms(4)[unfolded inj_on_def] unfolding g_def by auto  | 
|
6300  | 
  {
 | 
|
| 53282 | 6301  | 
fix y  | 
6302  | 
assume "y \<in> t"  | 
|
6303  | 
then obtain x where x:"f x = y" "x\<in>s"  | 
|
6304  | 
using assms(3) by auto  | 
|
6305  | 
then have "g (f x) = x" using g by auto  | 
|
| 53291 | 6306  | 
then have "f (g y) = y" unfolding x(1)[symmetric] by auto  | 
| 52625 | 6307  | 
}  | 
| 53282 | 6308  | 
then have g':"\<forall>x\<in>t. f (g x) = x" by auto  | 
| 33175 | 6309  | 
moreover  | 
| 52625 | 6310  | 
  {
 | 
6311  | 
fix x  | 
|
6312  | 
have "x\<in>s \<Longrightarrow> x \<in> g ` t"  | 
|
6313  | 
using g[THEN bspec[where x=x]]  | 
|
6314  | 
unfolding image_iff  | 
|
6315  | 
using assms(3)  | 
|
6316  | 
by (auto intro!: bexI[where x="f x"])  | 
|
| 33175 | 6317  | 
moreover  | 
| 52625 | 6318  | 
    {
 | 
6319  | 
assume "x\<in>g ` t"  | 
|
| 33175 | 6320  | 
then obtain y where y:"y\<in>t" "g y = x" by auto  | 
| 52625 | 6321  | 
then obtain x' where x':"x'\<in>s" "f x' = y"  | 
6322  | 
using assms(3) by auto  | 
|
| 53282 | 6323  | 
then have "x \<in> s"  | 
| 52625 | 6324  | 
unfolding g_def  | 
6325  | 
using someI2[of "\<lambda>b. b\<in>s \<and> f b = y" x' "\<lambda>x. x\<in>s"]  | 
|
| 53291 | 6326  | 
unfolding y(2)[symmetric] and g_def  | 
| 52625 | 6327  | 
by auto  | 
6328  | 
}  | 
|
6329  | 
ultimately have "x\<in>s \<longleftrightarrow> x \<in> g ` t" ..  | 
|
6330  | 
}  | 
|
| 53282 | 6331  | 
then have "g ` t = s" by auto  | 
| 52625 | 6332  | 
ultimately show ?thesis  | 
6333  | 
unfolding homeomorphism_def homeomorphic_def  | 
|
6334  | 
apply (rule_tac x=g in exI)  | 
|
6335  | 
using g and assms(3) and continuous_on_inv[OF assms(2,1), of g, unfolded assms(3)] and assms(2)  | 
|
6336  | 
apply auto  | 
|
6337  | 
done  | 
|
| 33175 | 6338  | 
qed  | 
6339  | 
||
6340  | 
lemma homeomorphic_compact:  | 
|
| 50898 | 6341  | 
fixes f :: "'a::topological_space \<Rightarrow> 'b::t2_space"  | 
| 53282 | 6342  | 
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
 | 
6343  | 
unfolding homeomorphic_def by (metis homeomorphism_compact)  | 
| 33175 | 6344  | 
|
| 53282 | 6345  | 
text{* Preservation of topological properties. *}
 | 
| 33175 | 6346  | 
|
| 52625 | 6347  | 
lemma homeomorphic_compactness: "s homeomorphic t \<Longrightarrow> (compact s \<longleftrightarrow> compact t)"  | 
6348  | 
unfolding homeomorphic_def homeomorphism_def  | 
|
6349  | 
by (metis compact_continuous_image)  | 
|
| 33175 | 6350  | 
|
| 53282 | 6351  | 
text{* Results on translation, scaling etc. *}
 | 
| 33175 | 6352  | 
|
6353  | 
lemma homeomorphic_scaling:  | 
|
6354  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 53282 | 6355  | 
assumes "c \<noteq> 0"  | 
6356  | 
shows "s homeomorphic ((\<lambda>x. c *\<^sub>R x) ` s)"  | 
|
| 33175 | 6357  | 
unfolding homeomorphic_minimal  | 
| 52625 | 6358  | 
apply (rule_tac x="\<lambda>x. c *\<^sub>R x" in exI)  | 
6359  | 
apply (rule_tac x="\<lambda>x. (1 / c) *\<^sub>R x" in exI)  | 
|
6360  | 
using assms  | 
|
6361  | 
apply (auto simp add: continuous_on_intros)  | 
|
6362  | 
done  | 
|
| 33175 | 6363  | 
|
6364  | 
lemma homeomorphic_translation:  | 
|
6365  | 
fixes s :: "'a::real_normed_vector set"  | 
|
6366  | 
shows "s homeomorphic ((\<lambda>x. a + x) ` s)"  | 
|
6367  | 
unfolding homeomorphic_minimal  | 
|
| 52625 | 6368  | 
apply (rule_tac x="\<lambda>x. a + x" in exI)  | 
6369  | 
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
 | 
6370  | 
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
 | 
6371  | 
continuous_on_add [OF continuous_on_const continuous_on_id, of "plus a ` s" "- a"]  | 
| 52625 | 6372  | 
apply auto  | 
6373  | 
done  | 
|
| 33175 | 6374  | 
|
6375  | 
lemma homeomorphic_affinity:  | 
|
6376  | 
fixes s :: "'a::real_normed_vector set"  | 
|
| 52625 | 6377  | 
assumes "c \<noteq> 0"  | 
6378  | 
shows "s homeomorphic ((\<lambda>x. a + c *\<^sub>R x) ` s)"  | 
|
| 53282 | 6379  | 
proof -  | 
| 52625 | 6380  | 
have *: "op + a ` op *\<^sub>R c ` s = (\<lambda>x. a + c *\<^sub>R x) ` s" by auto  | 
| 33175 | 6381  | 
show ?thesis  | 
6382  | 
using homeomorphic_trans  | 
|
6383  | 
using homeomorphic_scaling[OF assms, of s]  | 
|
| 52625 | 6384  | 
using homeomorphic_translation[of "(\<lambda>x. c *\<^sub>R x) ` s" a]  | 
6385  | 
unfolding *  | 
|
6386  | 
by auto  | 
|
| 33175 | 6387  | 
qed  | 
6388  | 
||
6389  | 
lemma homeomorphic_balls:  | 
|
| 50898 | 6390  | 
fixes a b ::"'a::real_normed_vector"  | 
| 33175 | 6391  | 
assumes "0 < d" "0 < e"  | 
6392  | 
shows "(ball a d) homeomorphic (ball b e)" (is ?th)  | 
|
| 53282 | 6393  | 
and "(cball a d) homeomorphic (cball b e)" (is ?cth)  | 
6394  | 
proof -  | 
|
| 33175 | 6395  | 
show ?th unfolding homeomorphic_minimal  | 
6396  | 
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)  | 
|
6397  | 
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)  | 
|
| 51364 | 6398  | 
using assms  | 
6399  | 
apply (auto intro!: continuous_on_intros  | 
|
| 52625 | 6400  | 
simp: dist_commute dist_norm pos_divide_less_eq mult_strict_left_mono)  | 
| 51364 | 6401  | 
done  | 
| 33175 | 6402  | 
show ?cth unfolding homeomorphic_minimal  | 
6403  | 
apply(rule_tac x="\<lambda>x. b + (e/d) *\<^sub>R (x - a)" in exI)  | 
|
6404  | 
apply(rule_tac x="\<lambda>x. a + (d/e) *\<^sub>R (x - b)" in exI)  | 
|
| 51364 | 6405  | 
using assms  | 
6406  | 
apply (auto intro!: continuous_on_intros  | 
|
| 52625 | 6407  | 
simp: dist_commute dist_norm pos_divide_le_eq mult_strict_left_mono)  | 
| 51364 | 6408  | 
done  | 
| 33175 | 6409  | 
qed  | 
6410  | 
||
6411  | 
text{* "Isometry" (up to constant bounds) of injective linear map etc.           *}
 | 
|
6412  | 
||
6413  | 
lemma cauchy_isometric:  | 
|
| 
37489
 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 
hoelzl 
parents: 
37452 
diff
changeset
 | 
6414  | 
fixes x :: "nat \<Rightarrow> 'a::euclidean_space"  | 
| 53640 | 6415  | 
assumes e: "e > 0"  | 
| 52625 | 6416  | 
and s: "subspace s"  | 
6417  | 
and f: "bounded_linear f"  | 
|
| 53640 | 6418  | 
and normf: "\<forall>x\<in>s. norm (f x) \<ge> e * norm x"  | 
6419  | 
and xs: "\<forall>n. x n \<in> s"  | 
|
6420  | 
and cf: "Cauchy (f \<circ> x)"  | 
|
| 33175 | 6421  | 
shows "Cauchy x"  | 
| 52625 | 6422  | 
proof -  | 
| 33175 | 6423  | 
interpret f: bounded_linear f by fact  | 
| 52625 | 6424  | 
  {
 | 
| 53291 | 6425  | 
fix d :: real  | 
6426  | 
assume "d > 0"  | 
|
| 33175 | 6427  | 
then obtain N where N:"\<forall>n\<ge>N. norm (f (x n) - f (x N)) < e * d"  | 
| 53291 | 6428  | 
using cf[unfolded cauchy o_def dist_norm, THEN spec[where x="e*d"]]  | 
6429  | 
and e and mult_pos_pos[of e d]  | 
|
| 52625 | 6430  | 
by auto  | 
6431  | 
    {
 | 
|
6432  | 
fix n  | 
|
6433  | 
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
 | 
6434  | 
have "e * norm (x n - x N) \<le> norm (f (x n - x N))"  | 
| 52625 | 6435  | 
using subspace_sub[OF s, of "x n" "x N"]  | 
6436  | 
using xs[THEN spec[where x=N]] and xs[THEN spec[where x=n]]  | 
|
6437  | 
using normf[THEN bspec[where x="x n - x N"]]  | 
|
6438  | 
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
 | 
6439  | 
also have "norm (f (x n - x N)) < e * d"  | 
| 53291 | 6440  | 
using `N \<le> n` N unfolding f.diff[symmetric] by auto  | 
| 52625 | 6441  | 
finally have "norm (x n - x N) < d" using `e>0` by simp  | 
6442  | 
}  | 
|
| 53282 | 6443  | 
then have "\<exists>N. \<forall>n\<ge>N. norm (x n - x N) < d" by auto  | 
| 52625 | 6444  | 
}  | 
| 53282 | 6445  | 
then show ?thesis unfolding cauchy and dist_norm by auto  | 
| 33175 | 6446  | 
qed  | 
6447  | 
||
6448  | 
lemma complete_isometric_image:  | 
|
| 53291 | 6449  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 52625 | 6450  | 
assumes "0 < e"  | 
6451  | 
and s: "subspace s"  | 
|
6452  | 
and f: "bounded_linear f"  | 
|
6453  | 
and normf: "\<forall>x\<in>s. norm(f x) \<ge> e * norm(x)"  | 
|
6454  | 
and cs: "complete s"  | 
|
| 53291 | 6455  | 
shows "complete (f ` s)"  | 
| 52625 | 6456  | 
proof -  | 
6457  | 
  {
 | 
|
6458  | 
fix g  | 
|
6459  | 
assume as:"\<forall>n::nat. g n \<in> f ` s" and cfg:"Cauchy g"  | 
|
| 53282 | 6460  | 
then obtain x where "\<forall>n. x n \<in> s \<and> g n = f (x n)"  | 
| 53640 | 6461  | 
using choice[of "\<lambda> n xa. xa \<in> s \<and> g n = f xa"]  | 
6462  | 
by auto  | 
|
6463  | 
then have x:"\<forall>n. x n \<in> s" "\<forall>n. g n = f (x n)"  | 
|
6464  | 
by auto  | 
|
6465  | 
then have "f \<circ> x = g"  | 
|
6466  | 
unfolding fun_eq_iff  | 
|
6467  | 
by auto  | 
|
| 33175 | 6468  | 
then obtain l where "l\<in>s" and l:"(x ---> l) sequentially"  | 
6469  | 
using cs[unfolded complete_def, THEN spec[where x="x"]]  | 
|
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
6470  | 
using cauchy_isometric[OF `0 < e` s f normf] and cfg and x(1)  | 
| 53640 | 6471  | 
by auto  | 
| 53282 | 6472  | 
then have "\<exists>l\<in>f ` s. (g ---> l) sequentially"  | 
| 33175 | 6473  | 
using linear_continuous_at[OF f, unfolded continuous_at_sequentially, THEN spec[where x=x], of l]  | 
| 53640 | 6474  | 
unfolding `f \<circ> x = g`  | 
6475  | 
by auto  | 
|
| 52625 | 6476  | 
}  | 
| 53640 | 6477  | 
then show ?thesis  | 
6478  | 
unfolding complete_def by auto  | 
|
| 33175 | 6479  | 
qed  | 
6480  | 
||
| 52625 | 6481  | 
lemma injective_imp_isometric:  | 
6482  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
|
6483  | 
assumes s: "closed s" "subspace s"  | 
|
| 53640 | 6484  | 
and f: "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0"  | 
6485  | 
shows "\<exists>e>0. \<forall>x\<in>s. norm (f x) \<ge> e * norm x"  | 
|
| 52625 | 6486  | 
proof (cases "s \<subseteq> {0::'a}")
 | 
| 33175 | 6487  | 
case True  | 
| 52625 | 6488  | 
  {
 | 
6489  | 
fix x  | 
|
6490  | 
assume "x \<in> s"  | 
|
| 53282 | 6491  | 
then have "x = 0" using True by auto  | 
6492  | 
then have "norm x \<le> norm (f x)" by auto  | 
|
| 52625 | 6493  | 
}  | 
| 53282 | 6494  | 
then show ?thesis by (auto intro!: exI[where x=1])  | 
| 33175 | 6495  | 
next  | 
6496  | 
interpret f: bounded_linear f by fact  | 
|
6497  | 
case False  | 
|
| 53640 | 6498  | 
then obtain a where a: "a \<noteq> 0" "a \<in> s"  | 
6499  | 
by auto  | 
|
6500  | 
  from False have "s \<noteq> {}"
 | 
|
6501  | 
by auto  | 
|
| 33175 | 6502  | 
  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
 | 
6503  | 
  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
 | 
6504  | 
  let ?S'' = "{x::'a. norm x = norm a}"
 | 
| 33175 | 6505  | 
|
| 52625 | 6506  | 
have "?S'' = frontier(cball 0 (norm a))"  | 
6507  | 
unfolding frontier_cball and dist_norm by auto  | 
|
| 53282 | 6508  | 
then have "compact ?S''"  | 
| 52625 | 6509  | 
using compact_frontier[OF compact_cball, of 0 "norm a"] by auto  | 
| 33175 | 6510  | 
moreover have "?S' = s \<inter> ?S''" by auto  | 
| 52625 | 6511  | 
ultimately have "compact ?S'"  | 
6512  | 
using closed_inter_compact[of s ?S''] using s(1) by auto  | 
|
| 33175 | 6513  | 
moreover have *:"f ` ?S' = ?S" by auto  | 
| 52625 | 6514  | 
ultimately have "compact ?S"  | 
6515  | 
using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto  | 
|
| 53282 | 6516  | 
then have "closed ?S" using compact_imp_closed by auto  | 
| 33175 | 6517  | 
  moreover have "?S \<noteq> {}" using a by auto
 | 
| 52625 | 6518  | 
ultimately obtain b' where "b'\<in>?S" "\<forall>y\<in>?S. norm b' \<le> norm y"  | 
6519  | 
using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto  | 
|
| 53282 | 6520  | 
then obtain b where "b\<in>s"  | 
6521  | 
and ba: "norm b = norm a"  | 
|
6522  | 
    and b: "\<forall>x\<in>{x \<in> s. norm x = norm a}. norm (f b) \<le> norm (f x)"
 | 
|
| 53291 | 6523  | 
unfolding *[symmetric] unfolding image_iff by auto  | 
| 33175 | 6524  | 
|
6525  | 
let ?e = "norm (f b) / norm b"  | 
|
6526  | 
have "norm b > 0" using ba and a and norm_ge_zero by auto  | 
|
| 52625 | 6527  | 
moreover have "norm (f b) > 0"  | 
6528  | 
using f(2)[THEN bspec[where x=b], OF `b\<in>s`]  | 
|
6529  | 
using `norm b >0`  | 
|
6530  | 
unfolding zero_less_norm_iff  | 
|
6531  | 
by auto  | 
|
6532  | 
ultimately have "0 < norm (f b) / norm b"  | 
|
6533  | 
by (simp only: divide_pos_pos)  | 
|
| 33175 | 6534  | 
moreover  | 
| 52625 | 6535  | 
  {
 | 
6536  | 
fix x  | 
|
6537  | 
assume "x\<in>s"  | 
|
| 53282 | 6538  | 
then have "norm (f b) / norm b * norm x \<le> norm (f x)"  | 
| 52625 | 6539  | 
proof (cases "x=0")  | 
6540  | 
case True  | 
|
| 53282 | 6541  | 
then show "norm (f b) / norm b * norm x \<le> norm (f x)" by auto  | 
| 33175 | 6542  | 
next  | 
6543  | 
case False  | 
|
| 53282 | 6544  | 
then have *: "0 < norm a / norm x"  | 
| 52625 | 6545  | 
using `a\<noteq>0`  | 
| 53291 | 6546  | 
unfolding zero_less_norm_iff[symmetric]  | 
| 52625 | 6547  | 
by (simp only: divide_pos_pos)  | 
6548  | 
have "\<forall>c. \<forall>x\<in>s. c *\<^sub>R x \<in> s"  | 
|
6549  | 
using s[unfolded subspace_def] by auto  | 
|
| 53282 | 6550  | 
      then have "(norm a / norm x) *\<^sub>R x \<in> {x \<in> s. norm x = norm a}"
 | 
| 52625 | 6551  | 
using `x\<in>s` and `x\<noteq>0` by auto  | 
| 53282 | 6552  | 
then show "norm (f b) / norm b * norm x \<le> norm (f x)"  | 
| 52625 | 6553  | 
using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]]  | 
| 33175 | 6554  | 
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
 | 
6555  | 
by (auto simp add: mult_commute pos_le_divide_eq pos_divide_le_eq)  | 
| 52625 | 6556  | 
qed  | 
6557  | 
}  | 
|
6558  | 
ultimately show ?thesis by auto  | 
|
| 33175 | 6559  | 
qed  | 
6560  | 
||
6561  | 
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
 | 
6562  | 
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"  | 
| 53282 | 6563  | 
assumes "subspace s" "bounded_linear f" "\<forall>x\<in>s. f x = 0 \<longrightarrow> x = 0" "closed s"  | 
| 33175 | 6564  | 
shows "closed(f ` s)"  | 
| 53282 | 6565  | 
proof -  | 
6566  | 
obtain e where "e > 0" and e: "\<forall>x\<in>s. e * norm x \<le> norm (f x)"  | 
|
| 52625 | 6567  | 
using injective_imp_isometric[OF assms(4,1,2,3)] by auto  | 
6568  | 
show ?thesis  | 
|
6569  | 
using complete_isometric_image[OF `e>0` assms(1,2) e] and assms(4)  | 
|
| 53291 | 6570  | 
unfolding complete_eq_closed[symmetric] by auto  | 
| 33175 | 6571  | 
qed  | 
6572  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6573  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6574  | 
subsection {* Some properties of a canonical subspace *}
 | 
| 33175 | 6575  | 
|
6576  | 
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
 | 
6577  | 
  "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
 | 
6578  | 
unfolding subspace_def by (auto simp: inner_add_left)  | 
| 33175 | 6579  | 
|
6580  | 
lemma closed_substandard:  | 
|
| 52625 | 6581  | 
  "closed {x::'a::euclidean_space. \<forall>i\<in>Basis. P i --> x\<bullet>i = 0}" (is "closed ?A")
 | 
6582  | 
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
 | 
6583  | 
  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
 | 
6584  | 
  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
 | 
6585  | 
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
 | 
6586  | 
  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
 | 
6587  | 
by auto  | 
| 
 
d366fa5551ef
declare euclidean_simps [simp] at the point they are proved;
 
huffman 
parents: 
44365 
diff
changeset
 | 
6588  | 
finally show "closed ?A" .  | 
| 33175 | 6589  | 
qed  | 
6590  | 
||
| 52625 | 6591  | 
lemma dim_substandard:  | 
6592  | 
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
 | 
6593  | 
  shows "dim {x::'a::euclidean_space. \<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x\<bullet>i = 0} = card d" (is "dim ?A = _")
 | 
| 53813 | 6594  | 
proof (rule dim_unique)  | 
6595  | 
show "d \<subseteq> ?A"  | 
|
6596  | 
using d by (auto simp: inner_Basis)  | 
|
6597  | 
show "independent d"  | 
|
6598  | 
using independent_mono [OF independent_Basis d] .  | 
|
6599  | 
show "?A \<subseteq> span d"  | 
|
6600  | 
proof (clarify)  | 
|
6601  | 
fix x assume x: "\<forall>i\<in>Basis. i \<notin> d \<longrightarrow> x \<bullet> i = 0"  | 
|
6602  | 
have "finite d"  | 
|
6603  | 
using finite_subset [OF d finite_Basis] .  | 
|
6604  | 
then have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) \<in> span d"  | 
|
6605  | 
by (simp add: span_setsum span_clauses)  | 
|
6606  | 
also have "(\<Sum>i\<in>d. (x \<bullet> i) *\<^sub>R i) = (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)"  | 
|
6607  | 
by (rule setsum_mono_zero_cong_left [OF finite_Basis d]) (auto simp add: x)  | 
|
6608  | 
finally show "x \<in> span d"  | 
|
6609  | 
unfolding euclidean_representation .  | 
|
6610  | 
qed  | 
|
6611  | 
qed simp  | 
|
| 33175 | 6612  | 
|
| 53282 | 6613  | 
text{* Hence closure and completeness of all subspaces. *}
 | 
6614  | 
||
6615  | 
lemma ex_card:  | 
|
6616  | 
assumes "n \<le> card A"  | 
|
6617  | 
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
 | 
6618  | 
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
 | 
6619  | 
assume "finite A"  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
6620  | 
from ex_bij_betw_nat_finite[OF this] guess f .. note f = this  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53291 
diff
changeset
 | 
6621  | 
  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
 | 
6622  | 
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
 | 
6623  | 
  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
 | 
6624  | 
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
 | 
6625  | 
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
 | 
6626  | 
next  | 
| 52625 | 6627  | 
assume "\<not> finite A"  | 
6628  | 
with `n \<le> card A` show ?thesis by force  | 
|
6629  | 
qed  | 
|
6630  | 
||
6631  | 
lemma closed_subspace:  | 
|
| 53291 | 6632  | 
fixes s :: "'a::euclidean_space set"  | 
| 52625 | 6633  | 
assumes "subspace s"  | 
6634  | 
shows "closed s"  | 
|
6635  | 
proof -  | 
|
6636  | 
have "dim s \<le> card (Basis :: 'a set)"  | 
|
6637  | 
using dim_subset_UNIV by auto  | 
|
6638  | 
with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \<subseteq> Basis"  | 
|
6639  | 
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
 | 
6640  | 
  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
 | 
6641  | 
  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
 | 
6642  | 
      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
 | 
6643  | 
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
 | 
6644  | 
by (intro subspace_isomorphism[OF subspace_substandard[of "\<lambda>i. i \<notin> d"]]) (auto simp: inner_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
 | 
6645  | 
then guess f by (elim exE conjE) note f = this  | 
| 52625 | 6646  | 
interpret f: bounded_linear f  | 
6647  | 
using f unfolding linear_conv_bounded_linear by auto  | 
|
6648  | 
  {
 | 
|
6649  | 
fix x  | 
|
6650  | 
have "x\<in>?t \<Longrightarrow> f x = 0 \<Longrightarrow> x = 0"  | 
|
6651  | 
using f.zero d f(3)[THEN inj_onD, of x 0] by auto  | 
|
6652  | 
}  | 
|
| 33175 | 6653  | 
moreover have "closed ?t" using closed_substandard .  | 
6654  | 
moreover have "subspace ?t" using subspace_substandard .  | 
|
| 52625 | 6655  | 
ultimately show ?thesis  | 
6656  | 
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
 | 
6657  | 
unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto  | 
| 33175 | 6658  | 
qed  | 
6659  | 
||
6660  | 
lemma complete_subspace:  | 
|
| 52625 | 6661  | 
  fixes s :: "('a::euclidean_space) set"
 | 
6662  | 
shows "subspace s \<Longrightarrow> complete s"  | 
|
6663  | 
using complete_eq_closed closed_subspace by auto  | 
|
| 33175 | 6664  | 
|
6665  | 
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
 | 
6666  | 
  fixes s :: "('a::euclidean_space) set"
 | 
| 33175 | 6667  | 
shows "dim(closure s) = dim s" (is "?dc = ?d")  | 
| 52625 | 6668  | 
proof -  | 
| 33175 | 6669  | 
have "?dc \<le> ?d" using closure_minimal[OF span_inc, of s]  | 
6670  | 
using closed_subspace[OF subspace_span, of s]  | 
|
| 52625 | 6671  | 
using dim_subset[of "closure s" "span s"]  | 
6672  | 
unfolding dim_span  | 
|
6673  | 
by auto  | 
|
| 53282 | 6674  | 
then show ?thesis using dim_subset[OF closure_subset, of s]  | 
| 52625 | 6675  | 
by auto  | 
| 33175 | 6676  | 
qed  | 
6677  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6678  | 
|
| 36437 | 6679  | 
subsection {* Affine transformations of intervals *}
 | 
| 33175 | 6680  | 
|
6681  | 
lemma real_affinity_le:  | 
|
| 53291 | 6682  | 
"0 < (m::'a::linordered_field) \<Longrightarrow> (m * x + c \<le> y \<longleftrightarrow> x \<le> inverse(m) * y + -(c / m))"  | 
| 33175 | 6683  | 
by (simp add: field_simps inverse_eq_divide)  | 
6684  | 
||
6685  | 
lemma real_le_affinity:  | 
|
| 53291 | 6686  | 
"0 < (m::'a::linordered_field) \<Longrightarrow> (y \<le> m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) \<le> x)"  | 
| 33175 | 6687  | 
by (simp add: field_simps inverse_eq_divide)  | 
6688  | 
||
6689  | 
lemma real_affinity_lt:  | 
|
| 53291 | 6690  | 
"0 < (m::'a::linordered_field) \<Longrightarrow> (m * x + c < y \<longleftrightarrow> x < inverse(m) * y + -(c / m))"  | 
| 33175 | 6691  | 
by (simp add: field_simps inverse_eq_divide)  | 
6692  | 
||
6693  | 
lemma real_lt_affinity:  | 
|
| 53291 | 6694  | 
"0 < (m::'a::linordered_field) \<Longrightarrow> (y < m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) < x)"  | 
| 33175 | 6695  | 
by (simp add: field_simps inverse_eq_divide)  | 
6696  | 
||
6697  | 
lemma real_affinity_eq:  | 
|
| 53291 | 6698  | 
"(m::'a::linordered_field) \<noteq> 0 \<Longrightarrow> (m * x + c = y \<longleftrightarrow> x = inverse(m) * y + -(c / m))"  | 
| 33175 | 6699  | 
by (simp add: field_simps inverse_eq_divide)  | 
6700  | 
||
6701  | 
lemma real_eq_affinity:  | 
|
| 53291 | 6702  | 
"(m::'a::linordered_field) \<noteq> 0 \<Longrightarrow> (y = m * x + c \<longleftrightarrow> inverse(m) * y + -(c / m) = x)"  | 
| 33175 | 6703  | 
by (simp add: field_simps inverse_eq_divide)  | 
6704  | 
||
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6705  | 
|
| 
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6706  | 
subsection {* Banach fixed point theorem (not really topological...) *}
 | 
| 33175 | 6707  | 
|
6708  | 
lemma banach_fix:  | 
|
| 53282 | 6709  | 
  assumes s: "complete s" "s \<noteq> {}"
 | 
6710  | 
and c: "0 \<le> c" "c < 1"  | 
|
6711  | 
and f: "(f ` s) \<subseteq> s"  | 
|
| 53291 | 6712  | 
and lipschitz: "\<forall>x\<in>s. \<forall>y\<in>s. dist (f x) (f y) \<le> c * dist x y"  | 
6713  | 
shows "\<exists>!x\<in>s. f x = x"  | 
|
| 53282 | 6714  | 
proof -  | 
| 33175 | 6715  | 
have "1 - c > 0" using c by auto  | 
6716  | 
||
6717  | 
from s(2) obtain z0 where "z0 \<in> s" by auto  | 
|
6718  | 
def z \<equiv> "\<lambda>n. (f ^^ n) z0"  | 
|
| 53282 | 6719  | 
  {
 | 
6720  | 
fix n :: nat  | 
|
| 33175 | 6721  | 
have "z n \<in> s" unfolding z_def  | 
| 52625 | 6722  | 
proof (induct n)  | 
6723  | 
case 0  | 
|
| 53282 | 6724  | 
then show ?case using `z0 \<in> s` by auto  | 
| 52625 | 6725  | 
next  | 
6726  | 
case Suc  | 
|
| 53282 | 6727  | 
then show ?case using f by auto qed  | 
| 52625 | 6728  | 
} note z_in_s = this  | 
| 33175 | 6729  | 
|
6730  | 
def d \<equiv> "dist (z 0) (z 1)"  | 
|
6731  | 
||
6732  | 
have fzn:"\<And>n. f (z n) = z (Suc n)" unfolding z_def by auto  | 
|
| 52625 | 6733  | 
  {
 | 
| 53282 | 6734  | 
fix n :: nat  | 
| 33175 | 6735  | 
have "dist (z n) (z (Suc n)) \<le> (c ^ n) * d"  | 
| 52625 | 6736  | 
proof (induct n)  | 
| 53282 | 6737  | 
case 0  | 
6738  | 
then show ?case  | 
|
| 52625 | 6739  | 
unfolding d_def by auto  | 
| 33175 | 6740  | 
next  | 
6741  | 
case (Suc m)  | 
|
| 53282 | 6742  | 
then have "c * dist (z m) (z (Suc m)) \<le> c ^ Suc m * d"  | 
| 52625 | 6743  | 
using `0 \<le> c`  | 
6744  | 
using mult_left_mono[of "dist (z m) (z (Suc m))" "c ^ m * d" c]  | 
|
6745  | 
by auto  | 
|
| 53282 | 6746  | 
then show ?case  | 
| 52625 | 6747  | 
using lipschitz[THEN bspec[where x="z m"], OF z_in_s, THEN bspec[where x="z (Suc m)"], OF z_in_s]  | 
6748  | 
unfolding fzn and mult_le_cancel_left  | 
|
6749  | 
by auto  | 
|
| 33175 | 6750  | 
qed  | 
6751  | 
} note cf_z = this  | 
|
6752  | 
||
| 52625 | 6753  | 
  {
 | 
| 53282 | 6754  | 
fix n m :: nat  | 
| 33175 | 6755  | 
have "(1 - c) * dist (z m) (z (m+n)) \<le> (c ^ m) * d * (1 - c ^ n)"  | 
| 52625 | 6756  | 
proof (induct n)  | 
| 53282 | 6757  | 
case 0  | 
6758  | 
show ?case by auto  | 
|
| 33175 | 6759  | 
next  | 
6760  | 
case (Suc k)  | 
|
| 52625 | 6761  | 
have "(1 - c) * dist (z m) (z (m + Suc k)) \<le>  | 
6762  | 
(1 - c) * (dist (z m) (z (m + k)) + dist (z (m + k)) (z (Suc (m + k))))"  | 
|
6763  | 
using dist_triangle and c by (auto simp add: dist_triangle)  | 
|
| 33175 | 6764  | 
also have "\<dots> \<le> (1 - c) * (dist (z m) (z (m + k)) + c ^ (m + k) * d)"  | 
6765  | 
using cf_z[of "m + k"] and c by auto  | 
|
6766  | 
also have "\<dots> \<le> c ^ m * d * (1 - c ^ k) + (1 - c) * c ^ (m + k) * d"  | 
|
| 36350 | 6767  | 
using Suc by (auto simp add: field_simps)  | 
| 33175 | 6768  | 
also have "\<dots> = (c ^ m) * (d * (1 - c ^ k) + (1 - c) * c ^ k * d)"  | 
| 36350 | 6769  | 
unfolding power_add by (auto simp add: field_simps)  | 
| 33175 | 6770  | 
also have "\<dots> \<le> (c ^ m) * d * (1 - c ^ Suc k)"  | 
| 36350 | 6771  | 
using c by (auto simp add: field_simps)  | 
| 33175 | 6772  | 
finally show ?case by auto  | 
6773  | 
qed  | 
|
6774  | 
} note cf_z2 = this  | 
|
| 52625 | 6775  | 
  {
 | 
| 53282 | 6776  | 
fix e :: real  | 
6777  | 
assume "e > 0"  | 
|
6778  | 
then have "\<exists>N. \<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (z m) (z n) < e"  | 
|
| 52625 | 6779  | 
proof (cases "d = 0")  | 
| 33175 | 6780  | 
case True  | 
| 41863 | 6781  | 
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
 | 
6782  | 
by (metis mult_zero_left mult_commute real_mult_le_cancel_iff1)  | 
| 41863 | 6783  | 
from True have "\<And>n. z n = z0" using cf_z2[of 0] and c unfolding z_def  | 
6784  | 
by (simp add: *)  | 
|
| 53282 | 6785  | 
then show ?thesis using `e>0` by auto  | 
| 33175 | 6786  | 
next  | 
| 52625 | 6787  | 
case False  | 
| 53282 | 6788  | 
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
 | 
6789  | 
by (metis False d_def less_le)  | 
| 53282 | 6790  | 
then have "0 < e * (1 - c) / d"  | 
| 52625 | 6791  | 
using `e>0` and `1-c>0`  | 
6792  | 
using divide_pos_pos[of "e * (1 - c)" d] and mult_pos_pos[of e "1 - c"]  | 
|
6793  | 
by auto  | 
|
6794  | 
then obtain N where N:"c ^ N < e * (1 - c) / d"  | 
|
6795  | 
using real_arch_pow_inv[of "e * (1 - c) / d" c] and c by auto  | 
|
6796  | 
      {
 | 
|
6797  | 
fix m n::nat  | 
|
6798  | 
assume "m>n" and as:"m\<ge>N" "n\<ge>N"  | 
|
6799  | 
have *:"c ^ n \<le> c ^ N" using `n\<ge>N` and c  | 
|
6800  | 
using power_decreasing[OF `n\<ge>N`, of c] by auto  | 
|
6801  | 
have "1 - c ^ (m - n) > 0"  | 
|
6802  | 
using c and power_strict_mono[of c 1 "m - n"] using `m>n` by auto  | 
|
| 53282 | 6803  | 
then have **: "d * (1 - c ^ (m - n)) / (1 - c) > 0"  | 
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36669 
diff
changeset
 | 
6804  | 
using mult_pos_pos[OF `d>0`, of "1 - c ^ (m - n)"]  | 
| 33175 | 6805  | 
using divide_pos_pos[of "d * (1 - c ^ (m - n))" "1 - c"]  | 
| 52625 | 6806  | 
using `0 < 1 - c`  | 
6807  | 
by auto  | 
|
| 33175 | 6808  | 
|
6809  | 
have "dist (z m) (z n) \<le> c ^ n * d * (1 - c ^ (m - n)) / (1 - c)"  | 
|
| 52625 | 6810  | 
using cf_z2[of n "m - n"] and `m>n`  | 
6811  | 
unfolding pos_le_divide_eq[OF `1-c>0`]  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36669 
diff
changeset
 | 
6812  | 
by (auto simp add: mult_commute dist_commute)  | 
| 33175 | 6813  | 
also have "\<dots> \<le> c ^ N * d * (1 - c ^ (m - n)) / (1 - c)"  | 
6814  | 
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
 | 
6815  | 
unfolding mult_assoc by auto  | 
| 33175 | 6816  | 
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
 | 
6817  | 
using mult_strict_right_mono[OF N **] unfolding mult_assoc by auto  | 
| 53282 | 6818  | 
also have "\<dots> = e * (1 - c ^ (m - n))"  | 
6819  | 
using c and `d>0` and `1 - c > 0` by auto  | 
|
6820  | 
also have "\<dots> \<le> e" using c and `1 - c ^ (m - n) > 0` and `e>0`  | 
|
6821  | 
using mult_right_le_one_le[of e "1 - c ^ (m - n)"] by auto  | 
|
| 33175 | 6822  | 
finally have "dist (z m) (z n) < e" by auto  | 
6823  | 
} note * = this  | 
|
| 52625 | 6824  | 
      {
 | 
| 53282 | 6825  | 
fix m n :: nat  | 
6826  | 
assume as: "N \<le> m" "N \<le> n"  | 
|
6827  | 
then have "dist (z n) (z m) < e"  | 
|
| 52625 | 6828  | 
proof (cases "n = m")  | 
6829  | 
case True  | 
|
| 53282 | 6830  | 
then show ?thesis using `e>0` by auto  | 
| 33175 | 6831  | 
next  | 
| 52625 | 6832  | 
case False  | 
| 53282 | 6833  | 
then show ?thesis using as and *[of n m] *[of m n]  | 
| 52625 | 6834  | 
unfolding nat_neq_iff by (auto simp add: dist_commute)  | 
6835  | 
qed  | 
|
6836  | 
}  | 
|
| 53282 | 6837  | 
then show ?thesis by auto  | 
| 33175 | 6838  | 
qed  | 
6839  | 
}  | 
|
| 53282 | 6840  | 
then have "Cauchy z"  | 
6841  | 
unfolding cauchy_def by auto  | 
|
| 52625 | 6842  | 
then obtain x where "x\<in>s" and x:"(z ---> x) sequentially"  | 
6843  | 
using s(1)[unfolded compact_def complete_def, THEN spec[where x=z]] and z_in_s by auto  | 
|
| 33175 | 6844  | 
|
6845  | 
def e \<equiv> "dist (f x) x"  | 
|
| 52625 | 6846  | 
have "e = 0"  | 
6847  | 
proof (rule ccontr)  | 
|
6848  | 
assume "e \<noteq> 0"  | 
|
| 53282 | 6849  | 
then have "e > 0"  | 
6850  | 
unfolding e_def using zero_le_dist[of "f x" x]  | 
|
| 33175 | 6851  | 
by (metis dist_eq_0_iff dist_nz e_def)  | 
6852  | 
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
 | 
6853  | 
using x[unfolded LIMSEQ_def, THEN spec[where x="e/2"]] by auto  | 
| 53282 | 6854  | 
then have N':"dist (z N) x < e / 2" by auto  | 
6855  | 
||
6856  | 
have *: "c * dist (z N) x \<le> dist (z N) x"  | 
|
| 52625 | 6857  | 
unfolding mult_le_cancel_right2  | 
| 33175 | 6858  | 
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
 | 
6859  | 
by (metis dist_eq_0_iff dist_nz order_less_asym less_le)  | 
| 52625 | 6860  | 
have "dist (f (z N)) (f x) \<le> c * dist (z N) x"  | 
6861  | 
using lipschitz[THEN bspec[where x="z N"], THEN bspec[where x=x]]  | 
|
6862  | 
using z_in_s[of N] `x\<in>s`  | 
|
6863  | 
using c  | 
|
6864  | 
by auto  | 
|
6865  | 
also have "\<dots> < e / 2"  | 
|
6866  | 
using N' and c using * by auto  | 
|
6867  | 
finally show False  | 
|
6868  | 
unfolding fzn  | 
|
| 33175 | 6869  | 
using N[THEN spec[where x="Suc N"]] and dist_triangle_half_r[of "z (Suc N)" "f x" e x]  | 
| 52625 | 6870  | 
unfolding e_def  | 
6871  | 
by auto  | 
|
| 33175 | 6872  | 
qed  | 
| 53282 | 6873  | 
then have "f x = x" unfolding e_def by auto  | 
| 33175 | 6874  | 
moreover  | 
| 52625 | 6875  | 
  {
 | 
6876  | 
fix y  | 
|
6877  | 
assume "f y = y" "y\<in>s"  | 
|
| 53282 | 6878  | 
then have "dist x y \<le> c * dist x y"  | 
| 52625 | 6879  | 
using lipschitz[THEN bspec[where x=x], THEN bspec[where x=y]]  | 
6880  | 
using `x\<in>s` and `f x = x`  | 
|
6881  | 
by auto  | 
|
| 53282 | 6882  | 
then have "dist x y = 0"  | 
| 52625 | 6883  | 
unfolding mult_le_cancel_right1  | 
6884  | 
using c and zero_le_dist[of x y]  | 
|
6885  | 
by auto  | 
|
| 53282 | 6886  | 
then have "y = x" by auto  | 
| 33175 | 6887  | 
}  | 
| 
34999
 
5312d2ffee3b
Changed 'bounded unique existential quantifiers' from a constant to syntax translation.
 
hoelzl 
parents: 
34964 
diff
changeset
 | 
6888  | 
ultimately show ?thesis using `x\<in>s` by blast+  | 
| 33175 | 6889  | 
qed  | 
6890  | 
||
| 53282 | 6891  | 
|
| 
44210
 
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
 
huffman 
parents: 
44207 
diff
changeset
 | 
6892  | 
subsection {* Edelstein fixed point theorem *}
 | 
| 33175 | 6893  | 
|
6894  | 
lemma edelstein_fix:  | 
|
| 
50970
 
3e5b67f85bf9
generalized theorem edelstein_fix to class metric_space
 
huffman 
parents: 
50955 
diff
changeset
 | 
6895  | 
fixes s :: "'a::metric_space set"  | 
| 52625 | 6896  | 
  assumes s: "compact s" "s \<noteq> {}"
 | 
6897  | 
and gs: "(g ` s) \<subseteq> s"  | 
|
6898  | 
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
 | 
6899  | 
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
 | 
6900  | 
proof -  | 
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6901  | 
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
 | 
6902  | 
  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
 | 
6903  | 
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
 | 
6904  | 
(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
 | 
6905  | 
|
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6906  | 
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
 | 
6907  | 
using dist by fastforce  | 
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6908  | 
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
 | 
6909  | 
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
 | 
6910  | 
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
 | 
6911  | 
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
 | 
6912  | 
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
 | 
6913  | 
(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
 | 
6914  | 
|
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6915  | 
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
 | 
6916  | 
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
 | 
6917  | 
|
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6918  | 
have "g a = a"  | 
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6919  | 
proof (rule ccontr)  | 
| 
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6920  | 
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
 | 
6921  | 
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
 | 
6922  | 
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
 | 
6923  | 
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
 | 
6924  | 
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
 | 
6925  | 
ultimately show False by auto  | 
| 33175 | 6926  | 
qed  | 
| 
51347
 
f8a00792fbc1
tuned proof of Edelstein fixed point theorem (use continuity of dist and attains_sup)
 
hoelzl 
parents: 
51346 
diff
changeset
 | 
6927  | 
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
 | 
6928  | 
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
 | 
6929  | 
ultimately show "\<exists>!x\<in>s. g x = x" using `a \<in> s` by blast  | 
| 33175 | 6930  | 
qed  | 
6931  | 
||
| 
44131
 
5fc334b94e00
declare tendsto_const [intro] (accidentally removed in 230a8665c919)
 
huffman 
parents: 
44129 
diff
changeset
 | 
6932  | 
declare tendsto_const [intro] (* FIXME: move *)  | 
| 
 
5fc334b94e00
declare tendsto_const [intro] (accidentally removed in 230a8665c919)
 
huffman 
parents: 
44129 
diff
changeset
 | 
6933  | 
|
| 
54775
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
6934  | 
no_notation  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
6935  | 
eucl_less (infix "<e" 50)  | 
| 
 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 
immler 
parents: 
54489 
diff
changeset
 | 
6936  | 
|
| 33175 | 6937  | 
end  |