| author | nipkow |
| Tue, 11 Sep 2018 17:53:08 +0200 | |
| changeset 68970 | b0dfe57dfa09 |
| parent 68617 | 75129a73aca3 |
| child 69000 | 7cb3ddd60fd6 |
| permissions | -rw-r--r-- |
| 63938 | 1 |
(* Author: L C Paulson, University of Cambridge |
| 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 |
||
| 67968 | 7 |
section \<open>Elementary topology in Euclidean space\<close> |
| 33175 | 8 |
|
9 |
theory Topology_Euclidean_Space |
|
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
10 |
imports |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66447
diff
changeset
|
11 |
"HOL-Library.Indicator_Function" |
|
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
66447
diff
changeset
|
12 |
"HOL-Library.Countable_Set" |
|
68188
2af1f142f855
move FuncSet back to HOL-Library (amending 493b818e8e10)
immler
parents:
68120
diff
changeset
|
13 |
"HOL-Library.FuncSet" |
| 50938 | 14 |
Linear_Algebra |
| 50087 | 15 |
Norm_Arith |
16 |
begin |
|
17 |
||
|
63593
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
18 |
(* FIXME: move elsewhere *) |
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
19 |
|
|
63967
2aa42596edc3
new material on paths, etc. Also rationalisation
paulson <lp15@cam.ac.uk>
parents:
63957
diff
changeset
|
20 |
lemma halfspace_Int_eq: |
|
2aa42596edc3
new material on paths, etc. Also rationalisation
paulson <lp15@cam.ac.uk>
parents:
63957
diff
changeset
|
21 |
"{x. a \<bullet> x \<le> b} \<inter> {x. b \<le> a \<bullet> x} = {x. a \<bullet> x = b}"
|
|
2aa42596edc3
new material on paths, etc. Also rationalisation
paulson <lp15@cam.ac.uk>
parents:
63957
diff
changeset
|
22 |
"{x. b \<le> a \<bullet> x} \<inter> {x. a \<bullet> x \<le> b} = {x. a \<bullet> x = b}"
|
|
2aa42596edc3
new material on paths, etc. Also rationalisation
paulson <lp15@cam.ac.uk>
parents:
63957
diff
changeset
|
23 |
by auto |
|
2aa42596edc3
new material on paths, etc. Also rationalisation
paulson <lp15@cam.ac.uk>
parents:
63957
diff
changeset
|
24 |
|
|
63593
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
25 |
definition (in monoid_add) support_on :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'b set"
|
| 64539 | 26 |
where "support_on s f = {x\<in>s. f x \<noteq> 0}"
|
|
63593
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
27 |
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
28 |
lemma in_support_on: "x \<in> support_on s f \<longleftrightarrow> x \<in> s \<and> f x \<noteq> 0" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
29 |
by (simp add: support_on_def) |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
30 |
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
31 |
lemma support_on_simps[simp]: |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
32 |
"support_on {} f = {}"
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
33 |
"support_on (insert x s) f = |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
34 |
(if f x = 0 then support_on s f else insert x (support_on s f))" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
35 |
"support_on (s \<union> t) f = support_on s f \<union> support_on t f" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
36 |
"support_on (s \<inter> t) f = support_on s f \<inter> support_on t f" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
37 |
"support_on (s - t) f = support_on s f - support_on t f" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
38 |
"support_on (f ` s) g = f ` (support_on s (g \<circ> f))" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
39 |
unfolding support_on_def by auto |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
40 |
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
41 |
lemma support_on_cong: |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
42 |
"(\<And>x. x \<in> s \<Longrightarrow> f x = 0 \<longleftrightarrow> g x = 0) \<Longrightarrow> support_on s f = support_on s g" |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
43 |
by (auto simp: support_on_def) |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
44 |
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
45 |
lemma support_on_if: "a \<noteq> 0 \<Longrightarrow> support_on A (\<lambda>x. if P x then a else 0) = {x\<in>A. P x}"
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
46 |
by (auto simp: support_on_def) |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
47 |
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
48 |
lemma support_on_if_subset: "support_on A (\<lambda>x. if P x then a else 0) \<subseteq> {x \<in> A. P x}"
|
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
49 |
by (auto simp: support_on_def) |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
50 |
|
| 68120 | 51 |
lemma finite_support[intro]: "finite S \<Longrightarrow> finite (support_on S f)" |
|
63593
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
52 |
unfolding support_on_def by auto |
|
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
53 |
|
| 64267 | 54 |
(* TODO: is supp_sum really needed? TODO: Generalize to Finite_Set.fold *) |
55 |
definition (in comm_monoid_add) supp_sum :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> 'a"
|
|
| 68120 | 56 |
where "supp_sum f S = (\<Sum>x\<in>support_on S f. f x)" |
| 64267 | 57 |
|
58 |
lemma supp_sum_empty[simp]: "supp_sum f {} = 0"
|
|
59 |
unfolding supp_sum_def by auto |
|
60 |
||
61 |
lemma supp_sum_insert[simp]: |
|
| 68120 | 62 |
"finite (support_on S f) \<Longrightarrow> |
63 |
supp_sum f (insert x S) = (if x \<in> S then supp_sum f S else f x + supp_sum f S)" |
|
| 64267 | 64 |
by (simp add: supp_sum_def in_support_on insert_absorb) |
65 |
||
66 |
lemma supp_sum_divide_distrib: "supp_sum f A / (r::'a::field) = supp_sum (\<lambda>n. f n / r) A" |
|
|
63593
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents:
63540
diff
changeset
|
67 |
by (cases "r = 0") |
| 64267 | 68 |
(auto simp: supp_sum_def sum_divide_distrib intro!: sum.cong support_on_cong) |
|
63305
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
69 |
|
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
70 |
(*END OF SUPPORT, ETC.*) |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
71 |
|
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61699
diff
changeset
|
72 |
lemma image_affinity_interval: |
|
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61699
diff
changeset
|
73 |
fixes c :: "'a::ordered_real_vector" |
| 68120 | 74 |
shows "((\<lambda>x. m *\<^sub>R x + c) ` {a..b}) =
|
75 |
(if {a..b}={} then {}
|
|
76 |
else if 0 \<le> m then {m *\<^sub>R a + c .. m *\<^sub>R b + c}
|
|
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61699
diff
changeset
|
77 |
else {m *\<^sub>R b + c .. m *\<^sub>R a + c})"
|
| 68120 | 78 |
(is "?lhs = ?rhs") |
79 |
proof (cases "m=0") |
|
80 |
case True |
|
81 |
then show ?thesis |
|
82 |
by force |
|
83 |
next |
|
84 |
case False |
|
85 |
show ?thesis |
|
86 |
proof |
|
87 |
show "?lhs \<subseteq> ?rhs" |
|
88 |
by (auto simp: scaleR_left_mono scaleR_left_mono_neg) |
|
89 |
show "?rhs \<subseteq> ?lhs" |
|
90 |
proof (clarsimp, intro conjI impI subsetI) |
|
91 |
show "\<lbrakk>0 \<le> m; a \<le> b; x \<in> {m *\<^sub>R a + c..m *\<^sub>R b + c}\<rbrakk>
|
|
92 |
\<Longrightarrow> x \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}" for x
|
|
93 |
apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI) |
|
94 |
using False apply (auto simp: le_diff_eq pos_le_divideRI) |
|
95 |
using diff_le_eq pos_le_divideR_eq by force |
|
96 |
show "\<lbrakk>\<not> 0 \<le> m; a \<le> b; x \<in> {m *\<^sub>R b + c..m *\<^sub>R a + c}\<rbrakk>
|
|
97 |
\<Longrightarrow> x \<in> (\<lambda>x. m *\<^sub>R x + c) ` {a..b}" for x
|
|
98 |
apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI) |
|
99 |
apply (auto simp: diff_le_eq neg_le_divideR_eq) |
|
100 |
using diff_eq_diff_less_eq linordered_field_class.sign_simps(11) minus_diff_eq not_less scaleR_le_cancel_left_neg by fastforce |
|
101 |
qed |
|
102 |
qed |
|
103 |
qed |
|
|
61738
c4f6031f1310
New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents:
61699
diff
changeset
|
104 |
|
| 53282 | 105 |
lemma countable_PiE: |
| 64910 | 106 |
"finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> countable (F i)) \<Longrightarrow> countable (Pi\<^sub>E I 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
|
107 |
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
|
108 |
|
| 64845 | 109 |
lemma open_sums: |
110 |
fixes T :: "('b::real_normed_vector) set"
|
|
111 |
assumes "open S \<or> open T" |
|
112 |
shows "open (\<Union>x\<in> S. \<Union>y \<in> T. {x + y})"
|
|
113 |
using assms |
|
114 |
proof |
|
115 |
assume S: "open S" |
|
116 |
show ?thesis |
|
117 |
proof (clarsimp simp: open_dist) |
|
118 |
fix x y |
|
119 |
assume "x \<in> S" "y \<in> T" |
|
120 |
with S obtain e where "e > 0" and e: "\<And>x'. dist x' x < e \<Longrightarrow> x' \<in> S" |
|
121 |
by (auto simp: open_dist) |
|
122 |
then have "\<And>z. dist z (x + y) < e \<Longrightarrow> \<exists>x\<in>S. \<exists>y\<in>T. z = x + y" |
|
123 |
by (metis \<open>y \<in> T\<close> diff_add_cancel dist_add_cancel2) |
|
124 |
then show "\<exists>e>0. \<forall>z. dist z (x + y) < e \<longrightarrow> (\<exists>x\<in>S. \<exists>y\<in>T. z = x + y)" |
|
125 |
using \<open>0 < e\<close> \<open>x \<in> S\<close> by blast |
|
126 |
qed |
|
127 |
next |
|
128 |
assume T: "open T" |
|
129 |
show ?thesis |
|
130 |
proof (clarsimp simp: open_dist) |
|
131 |
fix x y |
|
132 |
assume "x \<in> S" "y \<in> T" |
|
133 |
with T obtain e where "e > 0" and e: "\<And>x'. dist x' y < e \<Longrightarrow> x' \<in> T" |
|
134 |
by (auto simp: open_dist) |
|
135 |
then have "\<And>z. dist z (x + y) < e \<Longrightarrow> \<exists>x\<in>S. \<exists>y\<in>T. z = x + y" |
|
136 |
by (metis \<open>x \<in> S\<close> add_diff_cancel_left' add_diff_eq diff_diff_add dist_norm) |
|
137 |
then show "\<exists>e>0. \<forall>z. dist z (x + y) < e \<longrightarrow> (\<exists>x\<in>S. \<exists>y\<in>T. z = x + y)" |
|
138 |
using \<open>0 < e\<close> \<open>y \<in> T\<close> by blast |
|
139 |
qed |
|
140 |
qed |
|
141 |
||
| 53255 | 142 |
|
| 60420 | 143 |
subsection \<open>Topological Basis\<close> |
| 50087 | 144 |
|
145 |
context topological_space |
|
146 |
begin |
|
147 |
||
| 67962 | 148 |
definition%important "topological_basis B \<longleftrightarrow> |
| 53291 | 149 |
(\<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
|
150 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
151 |
lemma topological_basis: |
| 53291 | 152 |
"topological_basis B \<longleftrightarrow> (\<forall>x. open x \<longleftrightarrow> (\<exists>B'. B' \<subseteq> B \<and> \<Union>B' = x))" |
| 50998 | 153 |
unfolding topological_basis_def |
154 |
apply safe |
|
155 |
apply fastforce |
|
156 |
apply fastforce |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
157 |
apply (erule_tac x=x in allE, simp) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
158 |
apply (rule_tac x="{x}" in exI, auto)
|
| 50998 | 159 |
done |
160 |
||
| 50087 | 161 |
lemma topological_basis_iff: |
162 |
assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'" |
|
163 |
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'))" |
|
164 |
(is "_ \<longleftrightarrow> ?rhs") |
|
165 |
proof safe |
|
166 |
fix O' and x::'a |
|
167 |
assume H: "topological_basis B" "open O'" "x \<in> O'" |
|
| 53282 | 168 |
then have "(\<exists>B'\<subseteq>B. \<Union>B' = O')" by (simp add: topological_basis_def) |
| 50087 | 169 |
then obtain B' where "B' \<subseteq> B" "O' = \<Union>B'" by auto |
| 53282 | 170 |
then show "\<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'" using H by auto |
| 50087 | 171 |
next |
172 |
assume H: ?rhs |
|
| 53282 | 173 |
show "topological_basis B" |
174 |
using assms unfolding topological_basis_def |
|
| 50087 | 175 |
proof safe |
| 53640 | 176 |
fix O' :: "'a set" |
| 53282 | 177 |
assume "open O'" |
| 50087 | 178 |
with H obtain f where "\<forall>x\<in>O'. f x \<in> B \<and> x \<in> f x \<and> f x \<subseteq> O'" |
179 |
by (force intro: bchoice simp: Bex_def) |
|
| 53282 | 180 |
then show "\<exists>B'\<subseteq>B. \<Union>B' = O'" |
| 50087 | 181 |
by (auto intro: exI[where x="{f x |x. x \<in> O'}"])
|
182 |
qed |
|
183 |
qed |
|
184 |
||
185 |
lemma topological_basisI: |
|
186 |
assumes "\<And>B'. B' \<in> B \<Longrightarrow> open B'" |
|
| 53282 | 187 |
and "\<And>O' x. open O' \<Longrightarrow> x \<in> O' \<Longrightarrow> \<exists>B'\<in>B. x \<in> B' \<and> B' \<subseteq> O'" |
| 50087 | 188 |
shows "topological_basis B" |
189 |
using assms by (subst topological_basis_iff) auto |
|
190 |
||
191 |
lemma topological_basisE: |
|
192 |
fixes O' |
|
193 |
assumes "topological_basis B" |
|
| 53282 | 194 |
and "open O'" |
195 |
and "x \<in> O'" |
|
| 50087 | 196 |
obtains B' where "B' \<in> B" "x \<in> B'" "B' \<subseteq> O'" |
197 |
proof atomize_elim |
|
| 53282 | 198 |
from assms have "\<And>B'. B'\<in>B \<Longrightarrow> open B'" |
199 |
by (simp add: topological_basis_def) |
|
| 50087 | 200 |
with topological_basis_iff assms |
| 53282 | 201 |
show "\<exists>B'. B' \<in> B \<and> x \<in> B' \<and> B' \<subseteq> O'" |
202 |
using assms by (simp add: Bex_def) |
|
| 50087 | 203 |
qed |
204 |
||
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
205 |
lemma topological_basis_open: |
|
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
206 |
assumes "topological_basis B" |
| 53282 | 207 |
and "X \<in> B" |
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
208 |
shows "open X" |
| 53282 | 209 |
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
|
210 |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
211 |
lemma topological_basis_imp_subbasis: |
| 53255 | 212 |
assumes B: "topological_basis B" |
213 |
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
|
214 |
proof (intro ext iffI) |
| 53255 | 215 |
fix S :: "'a set" |
216 |
assume "open S" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
217 |
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
|
218 |
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
|
219 |
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
|
220 |
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
|
221 |
next |
| 53255 | 222 |
fix S :: "'a set" |
223 |
assume "generate_topology B S" |
|
224 |
then show "open S" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
225 |
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
|
226 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
227 |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
228 |
lemma basis_dense: |
| 53640 | 229 |
fixes B :: "'a set set" |
230 |
and f :: "'a set \<Rightarrow> 'a" |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
231 |
assumes "topological_basis B" |
| 53255 | 232 |
and choosefrom_basis: "\<And>B'. B' \<noteq> {} \<Longrightarrow> f B' \<in> B'"
|
| 55522 | 233 |
shows "\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>B' \<in> B. f B' \<in> X)"
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
234 |
proof (intro allI impI) |
| 53640 | 235 |
fix X :: "'a set" |
236 |
assume "open X" and "X \<noteq> {}"
|
|
| 60420 | 237 |
from topological_basisE[OF \<open>topological_basis B\<close> \<open>open X\<close> choosefrom_basis[OF \<open>X \<noteq> {}\<close>]]
|
| 55522 | 238 |
obtain B' where "B' \<in> B" "f X \<in> B'" "B' \<subseteq> X" . |
| 53255 | 239 |
then show "\<exists>B'\<in>B. f B' \<in> X" |
240 |
by (auto intro!: choosefrom_basis) |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
241 |
qed |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
242 |
|
| 50087 | 243 |
end |
244 |
||
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
245 |
lemma topological_basis_prod: |
| 53255 | 246 |
assumes A: "topological_basis A" |
247 |
and B: "topological_basis B" |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
248 |
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
|
249 |
unfolding topological_basis_def |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
250 |
proof (safe, simp_all del: ex_simps add: subset_image_iff ex_simps(1)[symmetric]) |
| 53255 | 251 |
fix S :: "('a \<times> 'b) set"
|
252 |
assume "open S" |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
253 |
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
|
254 |
proof (safe intro!: exI[of _ "{x\<in>A \<times> B. fst x \<times> snd x \<subseteq> S}"])
|
| 53255 | 255 |
fix x y |
256 |
assume "(x, y) \<in> S" |
|
| 60420 | 257 |
from open_prod_elim[OF \<open>open S\<close> this] |
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
258 |
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
|
259 |
by (metis mem_Sigma_iff) |
| 55522 | 260 |
moreover |
261 |
from A a obtain A0 where "A0 \<in> A" "x \<in> A0" "A0 \<subseteq> a" |
|
262 |
by (rule topological_basisE) |
|
263 |
moreover |
|
264 |
from B b obtain B0 where "B0 \<in> B" "y \<in> B0" "B0 \<subseteq> b" |
|
265 |
by (rule topological_basisE) |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
266 |
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
|
267 |
by (intro UN_I[of "(A0, B0)"]) auto |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
268 |
qed auto |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
269 |
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
|
270 |
|
| 53255 | 271 |
|
| 60420 | 272 |
subsection \<open>Countable Basis\<close> |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
273 |
|
| 67962 | 274 |
locale%important countable_basis = |
| 53640 | 275 |
fixes B :: "'a::topological_space set set" |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
276 |
assumes is_basis: "topological_basis B" |
| 53282 | 277 |
and countable_basis: "countable B" |
| 33175 | 278 |
begin |
279 |
||
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
280 |
lemma open_countable_basis_ex: |
| 50087 | 281 |
assumes "open X" |
| 61952 | 282 |
shows "\<exists>B' \<subseteq> B. X = \<Union>B'" |
| 53255 | 283 |
using assms countable_basis is_basis |
284 |
unfolding topological_basis_def by blast |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
285 |
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
286 |
lemma open_countable_basisE: |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
287 |
assumes "open X" |
| 61952 | 288 |
obtains B' where "B' \<subseteq> B" "X = \<Union>B'" |
| 53255 | 289 |
using assms open_countable_basis_ex |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
290 |
by atomize_elim simp |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
291 |
|
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
292 |
lemma countable_dense_exists: |
| 53291 | 293 |
"\<exists>D::'a set. countable D \<and> (\<forall>X. open X \<longrightarrow> X \<noteq> {} \<longrightarrow> (\<exists>d \<in> D. d \<in> X))"
|
| 50087 | 294 |
proof - |
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
295 |
let ?f = "(\<lambda>B'. SOME x. x \<in> B')" |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
296 |
have "countable (?f ` B)" using countable_basis by simp |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
297 |
with basis_dense[OF is_basis, of ?f] show ?thesis |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
298 |
by (intro exI[where x="?f ` B"]) (metis (mono_tags) all_not_in_conv imageI someI) |
| 50087 | 299 |
qed |
300 |
||
301 |
lemma countable_dense_setE: |
|
|
50245
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
302 |
obtains D :: "'a set" |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
303 |
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
|
304 |
using countable_dense_exists by blast |
|
dea9363887a6
based countable topological basis on Countable_Set
immler
parents:
50105
diff
changeset
|
305 |
|
| 50087 | 306 |
end |
307 |
||
| 50883 | 308 |
lemma (in first_countable_topology) first_countable_basisE: |
| 68120 | 309 |
fixes x :: 'a |
310 |
obtains \<A> where "countable \<A>" "\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A" "\<And>A. A \<in> \<A> \<Longrightarrow> open A" |
|
311 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>A\<in>\<A>. A \<subseteq> S)" |
|
312 |
proof - |
|
313 |
obtain \<A> where \<A>: "(\<forall>i::nat. x \<in> \<A> i \<and> open (\<A> i))" "(\<forall>S. open S \<and> x \<in> S \<longrightarrow> (\<exists>i. \<A> i \<subseteq> S))" |
|
314 |
using first_countable_basis[of x] by metis |
|
315 |
show thesis |
|
316 |
proof |
|
317 |
show "countable (range \<A>)" |
|
318 |
by simp |
|
319 |
qed (use \<A> in auto) |
|
320 |
qed |
|
| 50883 | 321 |
|
| 51105 | 322 |
lemma (in first_countable_topology) first_countable_basis_Int_stableE: |
| 68120 | 323 |
obtains \<A> where "countable \<A>" "\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A" "\<And>A. A \<in> \<A> \<Longrightarrow> open A" |
324 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> (\<exists>A\<in>\<A>. A \<subseteq> S)" |
|
325 |
"\<And>A B. A \<in> \<A> \<Longrightarrow> B \<in> \<A> \<Longrightarrow> A \<inter> B \<in> \<A>" |
|
| 51105 | 326 |
proof atomize_elim |
| 68120 | 327 |
obtain \<B> where \<B>: |
328 |
"countable \<B>" |
|
329 |
"\<And>B. B \<in> \<B> \<Longrightarrow> x \<in> B" |
|
330 |
"\<And>B. B \<in> \<B> \<Longrightarrow> open B" |
|
331 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>B\<in>\<B>. B \<subseteq> S" |
|
| 55522 | 332 |
by (rule first_countable_basisE) blast |
| 68120 | 333 |
define \<A> where [abs_def]: |
334 |
"\<A> = (\<lambda>N. \<Inter>((\<lambda>n. from_nat_into \<B> n) ` N)) ` (Collect finite::nat set set)" |
|
335 |
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> |
|
336 |
(\<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>)" |
|
337 |
proof (safe intro!: exI[where x=\<A>]) |
|
338 |
show "countable \<A>" |
|
339 |
unfolding \<A>_def by (intro countable_image countable_Collect_finite) |
|
340 |
fix A |
|
341 |
assume "A \<in> \<A>" |
|
342 |
then show "x \<in> A" "open A" |
|
343 |
using \<B>(4)[OF open_UNIV] by (auto simp: \<A>_def intro: \<B> from_nat_into) |
|
| 51105 | 344 |
next |
| 68120 | 345 |
let ?int = "\<lambda>N. \<Inter>(from_nat_into \<B> ` N)" |
346 |
fix A B |
|
347 |
assume "A \<in> \<A>" "B \<in> \<A>" |
|
348 |
then obtain N M where "A = ?int N" "B = ?int M" "finite (N \<union> M)" |
|
349 |
by (auto simp: \<A>_def) |
|
350 |
then show "A \<inter> B \<in> \<A>" |
|
351 |
by (auto simp: \<A>_def intro!: image_eqI[where x="N \<union> M"]) |
|
| 51105 | 352 |
next |
| 53255 | 353 |
fix S |
354 |
assume "open S" "x \<in> S" |
|
| 68120 | 355 |
then obtain a where a: "a\<in>\<B>" "a \<subseteq> S" using \<B> by blast |
356 |
then show "\<exists>a\<in>\<A>. a \<subseteq> S" using a \<B> |
|
357 |
by (intro bexI[where x=a]) (auto simp: \<A>_def intro: image_eqI[where x="{to_nat_on \<B> a}"])
|
|
| 51105 | 358 |
qed |
359 |
qed |
|
360 |
||
| 51473 | 361 |
lemma (in topological_space) first_countableI: |
| 68120 | 362 |
assumes "countable \<A>" |
363 |
and 1: "\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A" "\<And>A. A \<in> \<A> \<Longrightarrow> open A" |
|
364 |
and 2: "\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> \<exists>A\<in>\<A>. A \<subseteq> S" |
|
365 |
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))" |
|
366 |
proof (safe intro!: exI[of _ "from_nat_into \<A>"]) |
|
| 53255 | 367 |
fix i |
| 68120 | 368 |
have "\<A> \<noteq> {}" using 2[of UNIV] by auto
|
369 |
show "x \<in> from_nat_into \<A> i" "open (from_nat_into \<A> i)" |
|
370 |
using range_from_nat_into_subset[OF \<open>\<A> \<noteq> {}\<close>] 1 by auto
|
|
| 53255 | 371 |
next |
372 |
fix S |
|
373 |
assume "open S" "x\<in>S" from 2[OF this] |
|
| 68120 | 374 |
show "\<exists>i. from_nat_into \<A> i \<subseteq> S" |
375 |
using subset_range_from_nat_into[OF \<open>countable \<A>\<close>] by auto |
|
| 51473 | 376 |
qed |
| 51350 | 377 |
|
| 50883 | 378 |
instance prod :: (first_countable_topology, first_countable_topology) first_countable_topology |
379 |
proof |
|
380 |
fix x :: "'a \<times> 'b" |
|
| 68120 | 381 |
obtain \<A> where \<A>: |
382 |
"countable \<A>" |
|
383 |
"\<And>a. a \<in> \<A> \<Longrightarrow> fst x \<in> a" |
|
384 |
"\<And>a. a \<in> \<A> \<Longrightarrow> open a" |
|
385 |
"\<And>S. open S \<Longrightarrow> fst x \<in> S \<Longrightarrow> \<exists>a\<in>\<A>. a \<subseteq> S" |
|
| 55522 | 386 |
by (rule first_countable_basisE[of "fst x"]) blast |
387 |
obtain B where B: |
|
388 |
"countable B" |
|
389 |
"\<And>a. a \<in> B \<Longrightarrow> snd x \<in> a" |
|
390 |
"\<And>a. a \<in> B \<Longrightarrow> open a" |
|
391 |
"\<And>S. open S \<Longrightarrow> snd x \<in> S \<Longrightarrow> \<exists>a\<in>B. a \<subseteq> S" |
|
392 |
by (rule first_countable_basisE[of "snd x"]) blast |
|
| 68120 | 393 |
show "\<exists>\<A>::nat \<Rightarrow> ('a \<times> 'b) set.
|
394 |
(\<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))" |
|
395 |
proof (rule first_countableI[of "(\<lambda>(a, b). a \<times> b) ` (\<A> \<times> B)"], safe) |
|
| 53255 | 396 |
fix a b |
| 68120 | 397 |
assume x: "a \<in> \<A>" "b \<in> B" |
398 |
show "x \<in> a \<times> b" |
|
399 |
by (simp add: \<A>(2) B(2) mem_Times_iff x) |
|
400 |
show "open (a \<times> b)" |
|
401 |
by (simp add: \<A>(3) B(3) open_Times x) |
|
| 50883 | 402 |
next |
| 53255 | 403 |
fix S |
404 |
assume "open S" "x \<in> S" |
|
| 55522 | 405 |
then obtain a' b' where a'b': "open a'" "open b'" "x \<in> a' \<times> b'" "a' \<times> b' \<subseteq> S" |
406 |
by (rule open_prod_elim) |
|
407 |
moreover |
|
| 68120 | 408 |
from a'b' \<A>(4)[of a'] B(4)[of b'] |
409 |
obtain a b where "a \<in> \<A>" "a \<subseteq> a'" "b \<in> B" "b \<subseteq> b'" |
|
| 55522 | 410 |
by auto |
411 |
ultimately |
|
| 68120 | 412 |
show "\<exists>a\<in>(\<lambda>(a, b). a \<times> b) ` (\<A> \<times> B). a \<subseteq> S" |
| 50883 | 413 |
by (auto intro!: bexI[of _ "a \<times> b"] bexI[of _ a] bexI[of _ b]) |
| 68120 | 414 |
qed (simp add: \<A> B) |
| 50883 | 415 |
qed |
416 |
||
|
50881
ae630bab13da
renamed countable_basis_space to second_countable_topology
hoelzl
parents:
50526
diff
changeset
|
417 |
class second_countable_topology = topological_space + |
| 53282 | 418 |
assumes ex_countable_subbasis: |
419 |
"\<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
|
420 |
begin |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
421 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
422 |
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
|
423 |
proof - |
| 53255 | 424 |
from ex_countable_subbasis obtain B where B: "countable B" "open = generate_topology B" |
425 |
by blast |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
426 |
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
|
427 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
428 |
show ?thesis |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
429 |
proof (intro exI conjI) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
430 |
show "countable ?B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
431 |
by (intro countable_image countable_Collect_finite_subset B) |
| 53255 | 432 |
{
|
433 |
fix S |
|
434 |
assume "open S" |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
435 |
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
|
436 |
unfolding B |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
437 |
proof induct |
| 53255 | 438 |
case UNIV |
439 |
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
|
440 |
next |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
441 |
case (Int a b) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
442 |
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
|
443 |
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
|
444 |
by blast |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
445 |
show ?case |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
446 |
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
|
447 |
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
|
448 |
next |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
449 |
case (UN K) |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
450 |
then have "\<forall>k\<in>K. \<exists>B'\<subseteq>{b. finite b \<and> b \<subseteq> B}. UNION B' Inter = k" by auto
|
| 55522 | 451 |
then obtain k where |
452 |
"\<forall>ka\<in>K. k ka \<subseteq> {b. finite b \<and> b \<subseteq> B} \<and> UNION (k ka) Inter = ka"
|
|
453 |
unfolding bchoice_iff .. |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
454 |
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
|
455 |
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
|
456 |
next |
| 53255 | 457 |
case (Basis S) |
458 |
then show ?case |
|
|
51343
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
459 |
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
|
460 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
461 |
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
|
462 |
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
|
463 |
then show "topological_basis ?B" |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
464 |
unfolding topological_space_class.topological_basis_def |
| 53282 | 465 |
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
|
466 |
(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
|
467 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
468 |
qed |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
469 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
470 |
end |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
471 |
|
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
472 |
sublocale second_countable_topology < |
|
b61b32f62c78
use generate_topology for second countable topologies, does not require intersection stable basis
hoelzl
parents:
51342
diff
changeset
|
473 |
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
|
474 |
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
|
475 |
by unfold_locales safe |
|
50094
84ddcf5364b4
allow arbitrary enumerations of basis in locale for generation of borel sets
immler
parents:
50087
diff
changeset
|
476 |
|
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
477 |
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
|
478 |
proof |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
479 |
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
|
480 |
using ex_countable_basis by auto |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
481 |
moreover |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
482 |
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
|
483 |
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
|
484 |
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
|
485 |
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
|
486 |
topological_basis_imp_subbasis) |
|
50882
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
487 |
qed |
|
a382bf90867e
move prod instantiation of second_countable_topology to its definition
hoelzl
parents:
50881
diff
changeset
|
488 |
|
| 50883 | 489 |
instance second_countable_topology \<subseteq> first_countable_topology |
490 |
proof |
|
491 |
fix x :: 'a |
|
| 63040 | 492 |
define B :: "'a set set" where "B = (SOME B. countable B \<and> topological_basis B)" |
| 50883 | 493 |
then have B: "countable B" "topological_basis B" |
494 |
using countable_basis is_basis |
|
495 |
by (auto simp: countable_basis is_basis) |
|
| 53282 | 496 |
then show "\<exists>A::nat \<Rightarrow> 'a set. |
497 |
(\<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 | 498 |
by (intro first_countableI[of "{b\<in>B. x \<in> b}"])
|
499 |
(fastforce simp: topological_space_class.topological_basis_def)+ |
|
| 50883 | 500 |
qed |
501 |
||
|
64320
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents:
64287
diff
changeset
|
502 |
instance nat :: second_countable_topology |
|
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents:
64287
diff
changeset
|
503 |
proof |
|
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents:
64287
diff
changeset
|
504 |
show "\<exists>B::nat set set. countable B \<and> open = generate_topology B" |
|
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents:
64287
diff
changeset
|
505 |
by (intro exI[of _ "range lessThan \<union> range greaterThan"]) (auto simp: open_nat_def) |
|
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents:
64287
diff
changeset
|
506 |
qed |
| 53255 | 507 |
|
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
508 |
lemma countable_separating_set_linorder1: |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
509 |
shows "\<exists>B::('a::{linorder_topology, second_countable_topology} set). countable B \<and> (\<forall>x y. x < y \<longrightarrow> (\<exists>b \<in> B. x < b \<and> b \<le> y))"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
510 |
proof - |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
511 |
obtain A::"'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
512 |
define B1 where "B1 = {(LEAST x. x \<in> U)| U. U \<in> A}"
|
| 64911 | 513 |
then have "countable B1" using \<open>countable A\<close> by (simp add: Setcompr_eq_image) |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
514 |
define B2 where "B2 = {(SOME x. x \<in> U)| U. U \<in> A}"
|
| 64911 | 515 |
then have "countable B2" using \<open>countable A\<close> by (simp add: Setcompr_eq_image) |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
516 |
have "\<exists>b \<in> B1 \<union> B2. x < b \<and> b \<le> y" if "x < y" for x y |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
517 |
proof (cases) |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
518 |
assume "\<exists>z. x < z \<and> z < y" |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
519 |
then obtain z where z: "x < z \<and> z < y" by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
520 |
define U where "U = {x<..<y}"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
521 |
then have "open U" by simp |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
522 |
moreover have "z \<in> U" using z U_def by simp |
| 64911 | 523 |
ultimately obtain V where "V \<in> A" "z \<in> V" "V \<subseteq> U" using topological_basisE[OF \<open>topological_basis A\<close>] by auto |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
524 |
define w where "w = (SOME x. x \<in> V)" |
| 64911 | 525 |
then have "w \<in> V" using \<open>z \<in> V\<close> by (metis someI2) |
526 |
then have "x < w \<and> w \<le> y" using \<open>w \<in> V\<close> \<open>V \<subseteq> U\<close> U_def by fastforce |
|
527 |
moreover have "w \<in> B1 \<union> B2" using w_def B2_def \<open>V \<in> A\<close> by auto |
|
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
528 |
ultimately show ?thesis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
529 |
next |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
530 |
assume "\<not>(\<exists>z. x < z \<and> z < y)" |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
531 |
then have *: "\<And>z. z > x \<Longrightarrow> z \<ge> y" by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
532 |
define U where "U = {x<..}"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
533 |
then have "open U" by simp |
| 64911 | 534 |
moreover have "y \<in> U" using \<open>x < y\<close> U_def by simp |
535 |
ultimately obtain "V" where "V \<in> A" "y \<in> V" "V \<subseteq> U" using topological_basisE[OF \<open>topological_basis A\<close>] by auto |
|
536 |
have "U = {y..}" unfolding U_def using * \<open>x < y\<close> by auto
|
|
537 |
then have "V \<subseteq> {y..}" using \<open>V \<subseteq> U\<close> by simp
|
|
538 |
then have "(LEAST w. w \<in> V) = y" using \<open>y \<in> V\<close> by (meson Least_equality atLeast_iff subsetCE) |
|
539 |
then have "y \<in> B1 \<union> B2" using \<open>V \<in> A\<close> B1_def by auto |
|
540 |
moreover have "x < y \<and> y \<le> y" using \<open>x < y\<close> by simp |
|
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
541 |
ultimately show ?thesis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
542 |
qed |
| 64911 | 543 |
moreover have "countable (B1 \<union> B2)" using \<open>countable B1\<close> \<open>countable B2\<close> by simp |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
544 |
ultimately show ?thesis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
545 |
qed |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
546 |
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
547 |
lemma countable_separating_set_linorder2: |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
548 |
shows "\<exists>B::('a::{linorder_topology, second_countable_topology} set). countable B \<and> (\<forall>x y. x < y \<longrightarrow> (\<exists>b \<in> B. x \<le> b \<and> b < y))"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
549 |
proof - |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
550 |
obtain A::"'a set set" where "countable A" "topological_basis A" using ex_countable_basis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
551 |
define B1 where "B1 = {(GREATEST x. x \<in> U) | U. U \<in> A}"
|
| 64911 | 552 |
then have "countable B1" using \<open>countable A\<close> by (simp add: Setcompr_eq_image) |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
553 |
define B2 where "B2 = {(SOME x. x \<in> U)| U. U \<in> A}"
|
| 64911 | 554 |
then have "countable B2" using \<open>countable A\<close> by (simp add: Setcompr_eq_image) |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
555 |
have "\<exists>b \<in> B1 \<union> B2. x \<le> b \<and> b < y" if "x < y" for x y |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
556 |
proof (cases) |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
557 |
assume "\<exists>z. x < z \<and> z < y" |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
558 |
then obtain z where z: "x < z \<and> z < y" by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
559 |
define U where "U = {x<..<y}"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
560 |
then have "open U" by simp |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
561 |
moreover have "z \<in> U" using z U_def by simp |
| 64911 | 562 |
ultimately obtain "V" where "V \<in> A" "z \<in> V" "V \<subseteq> U" using topological_basisE[OF \<open>topological_basis A\<close>] by auto |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
563 |
define w where "w = (SOME x. x \<in> V)" |
| 64911 | 564 |
then have "w \<in> V" using \<open>z \<in> V\<close> by (metis someI2) |
565 |
then have "x \<le> w \<and> w < y" using \<open>w \<in> V\<close> \<open>V \<subseteq> U\<close> U_def by fastforce |
|
566 |
moreover have "w \<in> B1 \<union> B2" using w_def B2_def \<open>V \<in> A\<close> by auto |
|
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
567 |
ultimately show ?thesis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
568 |
next |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
569 |
assume "\<not>(\<exists>z. x < z \<and> z < y)" |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
570 |
then have *: "\<And>z. z < y \<Longrightarrow> z \<le> x" using leI by blast |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
571 |
define U where "U = {..<y}"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
572 |
then have "open U" by simp |
| 64911 | 573 |
moreover have "x \<in> U" using \<open>x < y\<close> U_def by simp |
574 |
ultimately obtain "V" where "V \<in> A" "x \<in> V" "V \<subseteq> U" using topological_basisE[OF \<open>topological_basis A\<close>] by auto |
|
575 |
have "U = {..x}" unfolding U_def using * \<open>x < y\<close> by auto
|
|
576 |
then have "V \<subseteq> {..x}" using \<open>V \<subseteq> U\<close> by simp
|
|
577 |
then have "(GREATEST x. x \<in> V) = x" using \<open>x \<in> V\<close> by (meson Greatest_equality atMost_iff subsetCE) |
|
578 |
then have "x \<in> B1 \<union> B2" using \<open>V \<in> A\<close> B1_def by auto |
|
579 |
moreover have "x \<le> x \<and> x < y" using \<open>x < y\<close> by simp |
|
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
580 |
ultimately show ?thesis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
581 |
qed |
| 64911 | 582 |
moreover have "countable (B1 \<union> B2)" using \<open>countable B1\<close> \<open>countable B2\<close> by simp |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
583 |
ultimately show ?thesis by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
584 |
qed |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
585 |
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
586 |
lemma countable_separating_set_dense_linorder: |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
587 |
shows "\<exists>B::('a::{linorder_topology, dense_linorder, second_countable_topology} set). countable B \<and> (\<forall>x y. x < y \<longrightarrow> (\<exists>b \<in> B. x < b \<and> b < y))"
|
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
588 |
proof - |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
589 |
obtain B::"'a set" where B: "countable B" "\<And>x y. x < y \<Longrightarrow> (\<exists>b \<in> B. x < b \<and> b \<le> y)" |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
590 |
using countable_separating_set_linorder1 by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
591 |
have "\<exists>b \<in> B. x < b \<and> b < y" if "x < y" for x y |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
592 |
proof - |
| 64911 | 593 |
obtain z where "x < z" "z < y" using \<open>x < y\<close> dense by blast |
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
594 |
then obtain b where "b \<in> B" "x < b \<and> b \<le> z" using B(2) by auto |
| 64911 | 595 |
then have "x < b \<and> b < y" using \<open>z < y\<close> by auto |
596 |
then show ?thesis using \<open>b \<in> B\<close> by auto |
|
|
64284
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
597 |
qed |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
598 |
then show ?thesis using B(1) by auto |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
599 |
qed |
|
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents:
64267
diff
changeset
|
600 |
|
| 67962 | 601 |
subsection%important \<open>Polish spaces\<close> |
| 60420 | 602 |
|
603 |
text \<open>Textbooks define Polish spaces as completely metrizable. |
|
604 |
We assume the topology to be complete for a given metric.\<close> |
|
| 50087 | 605 |
|
| 68617 | 606 |
class polish_space = complete_space + second_countable_topology |
| 50087 | 607 |
|
| 60420 | 608 |
subsection \<open>General notion of a topology as a value\<close> |
| 33175 | 609 |
|
| 67962 | 610 |
definition%important "istopology L \<longleftrightarrow> |
| 60585 | 611 |
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))"
|
| 53255 | 612 |
|
| 67962 | 613 |
typedef%important 'a topology = "{L::('a set) \<Rightarrow> bool. istopology L}"
|
| 33175 | 614 |
morphisms "openin" "topology" |
615 |
unfolding istopology_def by blast |
|
616 |
||
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
617 |
lemma istopology_openin[intro]: "istopology(openin U)" |
| 33175 | 618 |
using openin[of U] by blast |
619 |
||
620 |
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
|
621 |
using topology_inverse[unfolded mem_Collect_eq] . |
| 33175 | 622 |
|
623 |
lemma topology_inverse_iff: "istopology U \<longleftrightarrow> openin (topology U) = U" |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
624 |
using topology_inverse[of U] istopology_openin[of "topology U"] by auto |
| 33175 | 625 |
|
626 |
lemma topology_eq: "T1 = T2 \<longleftrightarrow> (\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S)" |
|
| 53255 | 627 |
proof |
628 |
assume "T1 = T2" |
|
629 |
then show "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" by simp |
|
630 |
next |
|
631 |
assume H: "\<forall>S. openin T1 S \<longleftrightarrow> openin T2 S" |
|
632 |
then have "openin T1 = openin T2" by (simp add: fun_eq_iff) |
|
633 |
then have "topology (openin T1) = topology (openin T2)" by simp |
|
634 |
then show "T1 = T2" unfolding openin_inverse . |
|
| 33175 | 635 |
qed |
636 |
||
| 60420 | 637 |
text\<open>Infer the "universe" from union of all sets in the topology.\<close> |
| 33175 | 638 |
|
| 53640 | 639 |
definition "topspace T = \<Union>{S. openin T S}"
|
| 33175 | 640 |
|
| 60420 | 641 |
subsubsection \<open>Main properties of open sets\<close> |
| 33175 | 642 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
643 |
proposition openin_clauses: |
| 33175 | 644 |
fixes U :: "'a topology" |
| 53282 | 645 |
shows |
646 |
"openin U {}"
|
|
647 |
"\<And>S T. openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S\<inter>T)" |
|
648 |
"\<And>K. (\<forall>S \<in> K. openin U S) \<Longrightarrow> openin U (\<Union>K)" |
|
649 |
using openin[of U] unfolding istopology_def mem_Collect_eq by fast+ |
|
| 33175 | 650 |
|
651 |
lemma openin_subset[intro]: "openin U S \<Longrightarrow> S \<subseteq> topspace U" |
|
652 |
unfolding topspace_def by blast |
|
| 53255 | 653 |
|
654 |
lemma openin_empty[simp]: "openin U {}"
|
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
655 |
by (rule openin_clauses) |
| 33175 | 656 |
|
657 |
lemma openin_Int[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<inter> T)" |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
658 |
by (rule openin_clauses) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
659 |
|
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
660 |
lemma openin_Union[intro]: "(\<And>S. S \<in> K \<Longrightarrow> openin U S) \<Longrightarrow> openin U (\<Union>K)" |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
661 |
using openin_clauses by blast |
| 33175 | 662 |
|
663 |
lemma openin_Un[intro]: "openin U S \<Longrightarrow> openin U T \<Longrightarrow> openin U (S \<union> T)" |
|
664 |
using openin_Union[of "{S,T}" U] by auto
|
|
665 |
||
| 53255 | 666 |
lemma openin_topspace[intro, simp]: "openin U (topspace U)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
667 |
by (force simp: openin_Union topspace_def) |
| 33175 | 668 |
|
| 49711 | 669 |
lemma openin_subopen: "openin U S \<longleftrightarrow> (\<forall>x \<in> S. \<exists>T. openin U T \<and> x \<in> T \<and> T \<subseteq> S)" |
670 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
| 36584 | 671 |
proof |
| 49711 | 672 |
assume ?lhs |
673 |
then show ?rhs by auto |
|
| 36584 | 674 |
next |
675 |
assume H: ?rhs |
|
676 |
let ?t = "\<Union>{T. openin U T \<and> T \<subseteq> S}"
|
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
677 |
have "openin U ?t" by (force simp: openin_Union) |
| 36584 | 678 |
also have "?t = S" using H by auto |
679 |
finally show "openin U S" . |
|
| 33175 | 680 |
qed |
681 |
||
| 64845 | 682 |
lemma openin_INT [intro]: |
683 |
assumes "finite I" |
|
684 |
"\<And>i. i \<in> I \<Longrightarrow> openin T (U i)" |
|
685 |
shows "openin T ((\<Inter>i \<in> I. U i) \<inter> topspace T)" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
686 |
using assms by (induct, auto simp: inf_sup_aci(2) openin_Int) |
| 64845 | 687 |
|
688 |
lemma openin_INT2 [intro]: |
|
689 |
assumes "finite I" "I \<noteq> {}"
|
|
690 |
"\<And>i. i \<in> I \<Longrightarrow> openin T (U i)" |
|
691 |
shows "openin T (\<Inter>i \<in> I. U i)" |
|
692 |
proof - |
|
693 |
have "(\<Inter>i \<in> I. U i) \<subseteq> topspace T" |
|
| 64911 | 694 |
using \<open>I \<noteq> {}\<close> openin_subset[OF assms(3)] by auto
|
| 64845 | 695 |
then show ?thesis |
696 |
using openin_INT[of _ _ U, OF assms(1) assms(3)] by (simp add: inf.absorb2 inf_commute) |
|
697 |
qed |
|
698 |
||
|
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
699 |
lemma openin_Inter [intro]: |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
700 |
assumes "finite \<F>" "\<F> \<noteq> {}" "\<And>X. X \<in> \<F> \<Longrightarrow> openin T X" shows "openin T (\<Inter>\<F>)"
|
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
701 |
by (metis (full_types) assms openin_INT2 image_ident) |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
702 |
|
| 49711 | 703 |
|
| 60420 | 704 |
subsubsection \<open>Closed sets\<close> |
| 33175 | 705 |
|
| 67962 | 706 |
definition%important "closedin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> openin U (topspace U - S)" |
| 33175 | 707 |
|
| 53255 | 708 |
lemma closedin_subset: "closedin U S \<Longrightarrow> S \<subseteq> topspace U" |
709 |
by (metis closedin_def) |
|
710 |
||
711 |
lemma closedin_empty[simp]: "closedin U {}"
|
|
712 |
by (simp add: closedin_def) |
|
713 |
||
714 |
lemma closedin_topspace[intro, simp]: "closedin U (topspace U)" |
|
715 |
by (simp add: closedin_def) |
|
716 |
||
| 33175 | 717 |
lemma closedin_Un[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<union> T)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
718 |
by (auto simp: Diff_Un closedin_def) |
| 33175 | 719 |
|
| 60585 | 720 |
lemma Diff_Inter[intro]: "A - \<Inter>S = \<Union>{A - s|s. s\<in>S}"
|
| 53255 | 721 |
by auto |
722 |
||
| 63955 | 723 |
lemma closedin_Union: |
724 |
assumes "finite S" "\<And>T. T \<in> S \<Longrightarrow> closedin U T" |
|
725 |
shows "closedin U (\<Union>S)" |
|
726 |
using assms by induction auto |
|
727 |
||
| 53255 | 728 |
lemma closedin_Inter[intro]: |
729 |
assumes Ke: "K \<noteq> {}"
|
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
730 |
and Kc: "\<And>S. S \<in>K \<Longrightarrow> closedin U S" |
| 60585 | 731 |
shows "closedin U (\<Inter>K)" |
| 53255 | 732 |
using Ke Kc unfolding closedin_def Diff_Inter by auto |
| 33175 | 733 |
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
734 |
lemma closedin_INT[intro]: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
735 |
assumes "A \<noteq> {}" "\<And>x. x \<in> A \<Longrightarrow> closedin U (B x)"
|
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
736 |
shows "closedin U (\<Inter>x\<in>A. B x)" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
737 |
apply (rule closedin_Inter) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
738 |
using assms |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
739 |
apply auto |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
740 |
done |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
741 |
|
| 33175 | 742 |
lemma closedin_Int[intro]: "closedin U S \<Longrightarrow> closedin U T \<Longrightarrow> closedin U (S \<inter> T)" |
743 |
using closedin_Inter[of "{S,T}" U] by auto
|
|
744 |
||
745 |
lemma openin_closedin_eq: "openin U S \<longleftrightarrow> S \<subseteq> topspace U \<and> closedin U (topspace U - S)" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
746 |
apply (auto simp: closedin_def Diff_Diff_Int inf_absorb2) |
| 33175 | 747 |
apply (metis openin_subset subset_eq) |
748 |
done |
|
749 |
||
| 53255 | 750 |
lemma openin_closedin: "S \<subseteq> topspace U \<Longrightarrow> (openin U S \<longleftrightarrow> closedin U (topspace U - S))" |
| 33175 | 751 |
by (simp add: openin_closedin_eq) |
752 |
||
| 53255 | 753 |
lemma openin_diff[intro]: |
754 |
assumes oS: "openin U S" |
|
755 |
and cT: "closedin U T" |
|
756 |
shows "openin U (S - T)" |
|
757 |
proof - |
|
| 33175 | 758 |
have "S - T = S \<inter> (topspace U - T)" using openin_subset[of U S] oS cT |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
759 |
by (auto simp: topspace_def openin_subset) |
| 53282 | 760 |
then show ?thesis using oS cT |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
761 |
by (auto simp: closedin_def) |
| 33175 | 762 |
qed |
763 |
||
| 53255 | 764 |
lemma closedin_diff[intro]: |
765 |
assumes oS: "closedin U S" |
|
766 |
and cT: "openin U T" |
|
767 |
shows "closedin U (S - T)" |
|
768 |
proof - |
|
769 |
have "S - T = S \<inter> (topspace U - T)" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
770 |
using closedin_subset[of U S] oS cT by (auto simp: topspace_def) |
| 53255 | 771 |
then show ?thesis |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
772 |
using oS cT by (auto simp: openin_closedin_eq) |
| 53255 | 773 |
qed |
774 |
||
| 33175 | 775 |
|
| 60420 | 776 |
subsubsection \<open>Subspace topology\<close> |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
777 |
|
| 67962 | 778 |
definition%important "subtopology U V = topology (\<lambda>T. \<exists>S. T = S \<inter> V \<and> openin U S)" |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
779 |
|
|
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
780 |
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
|
781 |
(is "istopology ?L") |
| 53255 | 782 |
proof - |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
783 |
have "?L {}" by blast
|
| 53255 | 784 |
{
|
785 |
fix A B |
|
786 |
assume A: "?L A" and B: "?L B" |
|
787 |
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" |
|
788 |
by blast |
|
789 |
have "A \<inter> B = (Sa \<inter> Sb) \<inter> V" "openin U (Sa \<inter> Sb)" |
|
790 |
using Sa Sb by blast+ |
|
791 |
then have "?L (A \<inter> B)" by blast |
|
792 |
} |
|
| 33175 | 793 |
moreover |
| 53255 | 794 |
{
|
| 53282 | 795 |
fix K |
796 |
assume K: "K \<subseteq> Collect ?L" |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
797 |
have th0: "Collect ?L = (\<lambda>S. S \<inter> V) ` Collect (openin U)" |
| 55775 | 798 |
by blast |
| 33175 | 799 |
from K[unfolded th0 subset_image_iff] |
| 53255 | 800 |
obtain Sk where Sk: "Sk \<subseteq> Collect (openin U)" "K = (\<lambda>S. S \<inter> V) ` Sk" |
801 |
by blast |
|
802 |
have "\<Union>K = (\<Union>Sk) \<inter> V" |
|
803 |
using Sk by auto |
|
| 60585 | 804 |
moreover have "openin U (\<Union>Sk)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
805 |
using Sk by (auto simp: subset_eq) |
| 53255 | 806 |
ultimately have "?L (\<Union>K)" by blast |
807 |
} |
|
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
808 |
ultimately show ?thesis |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
809 |
unfolding subset_eq mem_Collect_eq istopology_def by auto |
| 33175 | 810 |
qed |
811 |
||
| 53255 | 812 |
lemma openin_subtopology: "openin (subtopology U V) S \<longleftrightarrow> (\<exists>T. openin U T \<and> S = T \<inter> V)" |
| 33175 | 813 |
unfolding subtopology_def topology_inverse'[OF istopology_subtopology] |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
814 |
by auto |
| 33175 | 815 |
|
| 53255 | 816 |
lemma topspace_subtopology: "topspace (subtopology U V) = topspace U \<inter> V" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
817 |
by (auto simp: topspace_def openin_subtopology) |
| 33175 | 818 |
|
| 53255 | 819 |
lemma closedin_subtopology: "closedin (subtopology U V) S \<longleftrightarrow> (\<exists>T. closedin U T \<and> S = T \<inter> V)" |
| 33175 | 820 |
unfolding closedin_def topspace_subtopology |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
821 |
by (auto simp: openin_subtopology) |
| 33175 | 822 |
|
823 |
lemma openin_subtopology_refl: "openin (subtopology U V) V \<longleftrightarrow> V \<subseteq> topspace U" |
|
824 |
unfolding openin_subtopology |
|
| 55775 | 825 |
by auto (metis IntD1 in_mono openin_subset) |
| 49711 | 826 |
|
827 |
lemma subtopology_superset: |
|
828 |
assumes UV: "topspace U \<subseteq> V" |
|
| 33175 | 829 |
shows "subtopology U V = U" |
| 53255 | 830 |
proof - |
831 |
{
|
|
832 |
fix S |
|
833 |
{
|
|
834 |
fix T |
|
835 |
assume T: "openin U T" "S = T \<inter> V" |
|
836 |
from T openin_subset[OF T(1)] UV have eq: "S = T" |
|
837 |
by blast |
|
838 |
have "openin U S" |
|
839 |
unfolding eq using T by blast |
|
840 |
} |
|
| 33175 | 841 |
moreover |
| 53255 | 842 |
{
|
843 |
assume S: "openin U S" |
|
844 |
then have "\<exists>T. openin U T \<and> S = T \<inter> V" |
|
845 |
using openin_subset[OF S] UV by auto |
|
846 |
} |
|
847 |
ultimately have "(\<exists>T. openin U T \<and> S = T \<inter> V) \<longleftrightarrow> openin U S" |
|
848 |
by blast |
|
849 |
} |
|
850 |
then show ?thesis |
|
851 |
unfolding topology_eq openin_subtopology by blast |
|
| 33175 | 852 |
qed |
853 |
||
854 |
lemma subtopology_topspace[simp]: "subtopology U (topspace U) = U" |
|
855 |
by (simp add: subtopology_superset) |
|
856 |
||
857 |
lemma subtopology_UNIV[simp]: "subtopology U UNIV = U" |
|
858 |
by (simp add: subtopology_superset) |
|
859 |
||
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
860 |
lemma openin_subtopology_empty: |
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
861 |
"openin (subtopology U {}) S \<longleftrightarrow> S = {}"
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
862 |
by (metis Int_empty_right openin_empty openin_subtopology) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
863 |
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
864 |
lemma closedin_subtopology_empty: |
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
865 |
"closedin (subtopology U {}) S \<longleftrightarrow> S = {}"
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
866 |
by (metis Int_empty_right closedin_empty closedin_subtopology) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
867 |
|
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
868 |
lemma closedin_subtopology_refl [simp]: |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
869 |
"closedin (subtopology U X) X \<longleftrightarrow> X \<subseteq> topspace U" |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
870 |
by (metis closedin_def closedin_topspace inf.absorb_iff2 le_inf_iff topspace_subtopology) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
871 |
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
872 |
lemma openin_imp_subset: |
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
873 |
"openin (subtopology U S) T \<Longrightarrow> T \<subseteq> S" |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
874 |
by (metis Int_iff openin_subtopology subsetI) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
875 |
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
876 |
lemma closedin_imp_subset: |
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
877 |
"closedin (subtopology U S) T \<Longrightarrow> T \<subseteq> S" |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
878 |
by (simp add: closedin_def topspace_subtopology) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
879 |
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
880 |
lemma openin_subtopology_Un: |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
881 |
"\<lbrakk>openin (subtopology X T) S; openin (subtopology X U) S\<rbrakk> |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
882 |
\<Longrightarrow> openin (subtopology X (T \<union> U)) S" |
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
883 |
by (simp add: openin_subtopology) blast |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
884 |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
885 |
lemma closedin_subtopology_Un: |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
886 |
"\<lbrakk>closedin (subtopology X T) S; closedin (subtopology X U) S\<rbrakk> |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
887 |
\<Longrightarrow> closedin (subtopology X (T \<union> U)) S" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
888 |
by (simp add: closedin_subtopology) blast |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
889 |
|
| 53255 | 890 |
|
| 60420 | 891 |
subsubsection \<open>The standard Euclidean topology\<close> |
| 33175 | 892 |
|
| 67962 | 893 |
definition%important euclidean :: "'a::topological_space topology" |
| 53255 | 894 |
where "euclidean = topology open" |
| 33175 | 895 |
|
896 |
lemma open_openin: "open S \<longleftrightarrow> openin euclidean S" |
|
897 |
unfolding euclidean_def |
|
898 |
apply (rule cong[where x=S and y=S]) |
|
899 |
apply (rule topology_inverse[symmetric]) |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
900 |
apply (auto simp: istopology_def) |
|
44170
510ac30f44c0
make Multivariate_Analysis work with separate set type
huffman
parents:
44167
diff
changeset
|
901 |
done |
| 33175 | 902 |
|
| 64122 | 903 |
declare open_openin [symmetric, simp] |
904 |
||
|
63492
a662e8139804
More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents:
63469
diff
changeset
|
905 |
lemma topspace_euclidean [simp]: "topspace euclidean = UNIV" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
906 |
by (force simp: topspace_def) |
| 33175 | 907 |
|
908 |
lemma topspace_euclidean_subtopology[simp]: "topspace (subtopology euclidean S) = S" |
|
| 64122 | 909 |
by (simp add: topspace_subtopology) |
| 33175 | 910 |
|
911 |
lemma closed_closedin: "closed S \<longleftrightarrow> closedin euclidean S" |
|
| 64122 | 912 |
by (simp add: closed_def closedin_def Compl_eq_Diff_UNIV) |
| 33175 | 913 |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
914 |
declare closed_closedin [symmetric, simp] |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
915 |
|
| 33175 | 916 |
lemma open_subopen: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>T. open T \<and> x \<in> T \<and> T \<subseteq> S)" |
| 64122 | 917 |
using openI by auto |
| 33175 | 918 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
919 |
lemma openin_subtopology_self [simp]: "openin (subtopology euclidean S) S" |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
920 |
by (metis openin_topspace topspace_euclidean_subtopology) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
921 |
|
| 60420 | 922 |
text \<open>Basic "localization" results are handy for connectedness.\<close> |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
923 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
924 |
lemma openin_open: "openin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. open T \<and> (S = U \<inter> T))" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
925 |
by (auto simp: openin_subtopology) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
926 |
|
|
63305
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
927 |
lemma openin_Int_open: |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
928 |
"\<lbrakk>openin (subtopology euclidean U) S; open T\<rbrakk> |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
929 |
\<Longrightarrow> openin (subtopology euclidean U) (S \<inter> T)" |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
930 |
by (metis open_Int Int_assoc openin_open) |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
931 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
932 |
lemma openin_open_Int[intro]: "open S \<Longrightarrow> openin (subtopology euclidean U) (U \<inter> S)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
933 |
by (auto simp: openin_open) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
934 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
935 |
lemma open_openin_trans[trans]: |
| 53255 | 936 |
"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
|
937 |
by (metis Int_absorb1 openin_open_Int) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
938 |
|
| 53255 | 939 |
lemma open_subset: "S \<subseteq> T \<Longrightarrow> open S \<Longrightarrow> openin (subtopology euclidean T) S" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
940 |
by (auto simp: openin_open) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
941 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
942 |
lemma closedin_closed: "closedin (subtopology euclidean U) S \<longleftrightarrow> (\<exists>T. closed T \<and> S = U \<inter> T)" |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
943 |
by (simp add: closedin_subtopology Int_ac) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
944 |
|
| 53291 | 945 |
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
|
946 |
by (metis closedin_closed) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
947 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
948 |
lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
949 |
by (auto simp: closedin_closed) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
950 |
|
|
64791
05a2b3b20664
facts about ANRs, ENRs, covering spaces
paulson <lp15@cam.ac.uk>
parents:
64788
diff
changeset
|
951 |
lemma closedin_closed_subset: |
|
05a2b3b20664
facts about ANRs, ENRs, covering spaces
paulson <lp15@cam.ac.uk>
parents:
64788
diff
changeset
|
952 |
"\<lbrakk>closedin (subtopology euclidean U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk> |
|
05a2b3b20664
facts about ANRs, ENRs, covering spaces
paulson <lp15@cam.ac.uk>
parents:
64788
diff
changeset
|
953 |
\<Longrightarrow> closedin (subtopology euclidean T) S" |
|
05a2b3b20664
facts about ANRs, ENRs, covering spaces
paulson <lp15@cam.ac.uk>
parents:
64788
diff
changeset
|
954 |
by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE) |
|
05a2b3b20664
facts about ANRs, ENRs, covering spaces
paulson <lp15@cam.ac.uk>
parents:
64788
diff
changeset
|
955 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
956 |
lemma finite_imp_closedin: |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
957 |
fixes S :: "'a::t1_space set" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
958 |
shows "\<lbrakk>finite S; S \<subseteq> T\<rbrakk> \<Longrightarrow> closedin (subtopology euclidean T) S" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
959 |
by (simp add: finite_imp_closed closed_subset) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
960 |
|
|
63305
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
961 |
lemma closedin_singleton [simp]: |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
962 |
fixes a :: "'a::t1_space" |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
963 |
shows "closedin (subtopology euclidean U) {a} \<longleftrightarrow> a \<in> U"
|
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
964 |
using closedin_subset by (force intro: closed_subset) |
|
3b6975875633
Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
63301
diff
changeset
|
965 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
966 |
lemma openin_euclidean_subtopology_iff: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
967 |
fixes S U :: "'a::metric_space set" |
| 53255 | 968 |
shows "openin (subtopology euclidean U) S \<longleftrightarrow> |
969 |
S \<subseteq> U \<and> (\<forall>x\<in>S. \<exists>e>0. \<forall>x'\<in>U. dist x' x < e \<longrightarrow> x'\<in> S)" |
|
970 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
971 |
proof |
| 53255 | 972 |
assume ?lhs |
| 53282 | 973 |
then show ?rhs |
974 |
unfolding openin_open open_dist by blast |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
975 |
next |
| 63040 | 976 |
define T where "T = {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}"
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
977 |
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
|
978 |
unfolding T_def |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
979 |
apply clarsimp |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
980 |
apply (rule_tac x="d - dist x a" in exI) |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
981 |
apply (clarsimp simp add: less_diff_eq) |
| 55775 | 982 |
by (metis dist_commute dist_triangle_lt) |
| 53282 | 983 |
assume ?rhs then have 2: "S = U \<inter> T" |
|
60141
833adf7db7d8
New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents:
60040
diff
changeset
|
984 |
unfolding T_def |
| 55775 | 985 |
by auto (metis dist_self) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
986 |
from 1 2 show ?lhs |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
987 |
unfolding openin_open open_dist by fast |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
988 |
qed |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
989 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
990 |
lemma connected_openin: |
| 68120 | 991 |
"connected S \<longleftrightarrow> |
992 |
~(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and> |
|
993 |
openin (subtopology euclidean S) E2 \<and> |
|
994 |
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
|
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
995 |
apply (simp add: connected_def openin_open disjoint_iff_not_equal, safe) |
| 63988 | 996 |
apply (simp_all, blast+) (* SLOW *) |
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
997 |
done |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
998 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
999 |
lemma connected_openin_eq: |
| 68120 | 1000 |
"connected S \<longleftrightarrow> |
1001 |
~(\<exists>E1 E2. openin (subtopology euclidean S) E1 \<and> |
|
1002 |
openin (subtopology euclidean S) E2 \<and> |
|
1003 |
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
|
|
1004 |
E1 \<noteq> {} \<and> E2 \<noteq> {})"
|
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1005 |
apply (simp add: connected_openin, safe, blast) |
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
1006 |
by (metis Int_lower1 Un_subset_iff openin_open subset_antisym) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
1007 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1008 |
lemma connected_closedin: |
| 68120 | 1009 |
"connected S \<longleftrightarrow> |
1010 |
(\<nexists>E1 E2. |
|
1011 |
closedin (subtopology euclidean S) E1 \<and> |
|
1012 |
closedin (subtopology euclidean S) E2 \<and> |
|
1013 |
S \<subseteq> E1 \<union> E2 \<and> E1 \<inter> E2 = {} \<and> E1 \<noteq> {} \<and> E2 \<noteq> {})"
|
|
1014 |
(is "?lhs = ?rhs") |
|
1015 |
proof |
|
1016 |
assume ?lhs |
|
1017 |
then show ?rhs |
|
1018 |
by (auto simp add: connected_closed closedin_closed) |
|
1019 |
next |
|
1020 |
assume R: ?rhs |
|
1021 |
then show ?lhs |
|
1022 |
proof (clarsimp simp add: connected_closed closedin_closed) |
|
1023 |
fix A B |
|
1024 |
assume s_sub: "S \<subseteq> A \<union> B" "B \<inter> S \<noteq> {}"
|
|
1025 |
and disj: "A \<inter> B \<inter> S = {}"
|
|
1026 |
and cl: "closed A" "closed B" |
|
1027 |
have "S \<inter> (A \<union> B) = S" |
|
1028 |
using s_sub(1) by auto |
|
1029 |
have "S - A = B \<inter> S" |
|
1030 |
using Diff_subset_conv Un_Diff_Int disj s_sub(1) by auto |
|
1031 |
then have "S \<inter> A = {}"
|
|
1032 |
by (metis Diff_Diff_Int Diff_disjoint Un_Diff_Int R cl closedin_closed_Int inf_commute order_refl s_sub(2)) |
|
1033 |
then show "A \<inter> S = {}"
|
|
1034 |
by blast |
|
1035 |
qed |
|
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
1036 |
qed |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
1037 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1038 |
lemma connected_closedin_eq: |
| 68120 | 1039 |
"connected S \<longleftrightarrow> |
1040 |
~(\<exists>E1 E2. |
|
1041 |
closedin (subtopology euclidean S) E1 \<and> |
|
1042 |
closedin (subtopology euclidean S) E2 \<and> |
|
1043 |
E1 \<union> E2 = S \<and> E1 \<inter> E2 = {} \<and>
|
|
1044 |
E1 \<noteq> {} \<and> E2 \<noteq> {})"
|
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1045 |
apply (simp add: connected_closedin, safe, blast) |
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
1046 |
by (metis Int_lower1 Un_subset_iff closedin_closed subset_antisym) |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1047 |
|
| 60420 | 1048 |
text \<open>These "transitivity" results are handy too\<close> |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1049 |
|
| 53255 | 1050 |
lemma openin_trans[trans]: |
1051 |
"openin (subtopology euclidean T) S \<Longrightarrow> openin (subtopology euclidean U) T \<Longrightarrow> |
|
1052 |
openin (subtopology euclidean U) S" |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1053 |
unfolding open_openin openin_open by blast |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1054 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1055 |
lemma openin_open_trans: "openin (subtopology euclidean T) S \<Longrightarrow> open T \<Longrightarrow> open S" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1056 |
by (auto simp: openin_open intro: openin_trans) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1057 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1058 |
lemma closedin_trans[trans]: |
| 53255 | 1059 |
"closedin (subtopology euclidean T) S \<Longrightarrow> closedin (subtopology euclidean U) T \<Longrightarrow> |
1060 |
closedin (subtopology euclidean U) S" |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
1061 |
by (auto simp: closedin_closed closed_Inter Int_assoc) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1062 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1063 |
lemma closedin_closed_trans: "closedin (subtopology euclidean T) S \<Longrightarrow> closed T \<Longrightarrow> closed S" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1064 |
by (auto simp: closedin_closed intro: closedin_trans) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1065 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1066 |
lemma openin_subtopology_Int_subset: |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1067 |
"\<lbrakk>openin (subtopology euclidean u) (u \<inter> S); v \<subseteq> u\<rbrakk> \<Longrightarrow> openin (subtopology euclidean v) (v \<inter> S)" |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1068 |
by (auto simp: openin_subtopology) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1069 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1070 |
lemma openin_open_eq: "open s \<Longrightarrow> (openin (subtopology euclidean s) t \<longleftrightarrow> open t \<and> t \<subseteq> s)" |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1071 |
using open_subset openin_open_trans openin_subset by fastforce |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1072 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
1073 |
|
| 60420 | 1074 |
subsection \<open>Open and closed balls\<close> |
| 33175 | 1075 |
|
| 67962 | 1076 |
definition%important ball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" |
| 53255 | 1077 |
where "ball x e = {y. dist x y < e}"
|
1078 |
||
| 67962 | 1079 |
definition%important cball :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" |
| 53255 | 1080 |
where "cball x e = {y. dist x y \<le> e}"
|
| 33175 | 1081 |
|
| 67962 | 1082 |
definition%important sphere :: "'a::metric_space \<Rightarrow> real \<Rightarrow> 'a set" |
|
61762
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61738
diff
changeset
|
1083 |
where "sphere x e = {y. dist x y = e}"
|
|
d50b993b4fb9
Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents:
61738
diff
changeset
|
1084 |
|
|
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
|
1085 |
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
|
1086 |
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
|
1087 |
|
|
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
|
1088 |
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
|
1089 |
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
|
1090 |
|
| 61848 | 1091 |
lemma mem_sphere [simp]: "y \<in> sphere x e \<longleftrightarrow> dist x y = e" |
1092 |
by (simp add: sphere_def) |
|
1093 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1094 |
lemma ball_trivial [simp]: "ball x 0 = {}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1095 |
by (simp add: ball_def) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1096 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1097 |
lemma cball_trivial [simp]: "cball x 0 = {x}"
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1098 |
by (simp add: cball_def) |
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1099 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
1100 |
lemma sphere_trivial [simp]: "sphere x 0 = {x}"
|
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
1101 |
by (simp add: sphere_def) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
1102 |
|
| 64539 | 1103 |
lemma mem_ball_0 [simp]: "x \<in> ball 0 e \<longleftrightarrow> norm x < e" |
1104 |
for x :: "'a::real_normed_vector" |
|
| 33175 | 1105 |
by (simp add: dist_norm) |
1106 |
||
| 64539 | 1107 |
lemma mem_cball_0 [simp]: "x \<in> cball 0 e \<longleftrightarrow> norm x \<le> e" |
1108 |
for x :: "'a::real_normed_vector" |
|
| 33175 | 1109 |
by (simp add: dist_norm) |
1110 |
||
| 64539 | 1111 |
lemma disjoint_ballI: "dist x y \<ge> r+s \<Longrightarrow> ball x r \<inter> ball y s = {}"
|
| 64287 | 1112 |
using dist_triangle_less_add not_le by fastforce |
1113 |
||
| 64539 | 1114 |
lemma disjoint_cballI: "dist x y > r + s \<Longrightarrow> cball x r \<inter> cball y s = {}"
|
| 64287 | 1115 |
by (metis add_mono disjoint_iff_not_equal dist_triangle2 dual_order.trans leD mem_cball) |
1116 |
||
| 64539 | 1117 |
lemma mem_sphere_0 [simp]: "x \<in> sphere 0 e \<longleftrightarrow> norm x = e" |
1118 |
for x :: "'a::real_normed_vector" |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1119 |
by (simp add: dist_norm) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1120 |
|
| 64539 | 1121 |
lemma sphere_empty [simp]: "r < 0 \<Longrightarrow> sphere a r = {}"
|
1122 |
for a :: "'a::metric_space" |
|
1123 |
by auto |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
1124 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1125 |
lemma centre_in_ball [simp]: "x \<in> ball x e \<longleftrightarrow> 0 < e" |
|
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
|
1126 |
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
|
1127 |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
1128 |
lemma centre_in_cball [simp]: "x \<in> cball x e \<longleftrightarrow> 0 \<le> e" |
|
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
|
1129 |
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
|
1130 |
|
| 64539 | 1131 |
lemma ball_subset_cball [simp, intro]: "ball x e \<subseteq> cball x e" |
| 53255 | 1132 |
by (simp add: subset_eq) |
1133 |
||
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1134 |
lemma mem_ball_imp_mem_cball: "x \<in> ball y e \<Longrightarrow> x \<in> cball y e" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1135 |
by (auto simp: mem_ball mem_cball) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1136 |
|
|
61907
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
paulson <lp15@cam.ac.uk>
parents:
61880
diff
changeset
|
1137 |
lemma sphere_cball [simp,intro]: "sphere z r \<subseteq> cball z r" |
|
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
paulson <lp15@cam.ac.uk>
parents:
61880
diff
changeset
|
1138 |
by force |
|
f0c894ab18c9
Liouville theorem, Fundamental Theorem of Algebra, etc.
paulson <lp15@cam.ac.uk>
parents:
61880
diff
changeset
|
1139 |
|
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1140 |
lemma cball_diff_sphere: "cball a r - sphere a r = ball a r" |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1141 |
by auto |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1142 |
|
| 53282 | 1143 |
lemma subset_ball[intro]: "d \<le> e \<Longrightarrow> ball x d \<subseteq> ball x e" |
| 53255 | 1144 |
by (simp add: subset_eq) |
1145 |
||
| 53282 | 1146 |
lemma subset_cball[intro]: "d \<le> e \<Longrightarrow> cball x d \<subseteq> cball x e" |
| 53255 | 1147 |
by (simp add: subset_eq) |
1148 |
||
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1149 |
lemma mem_ball_leI: "x \<in> ball y e \<Longrightarrow> e \<le> f \<Longrightarrow> x \<in> ball y f" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1150 |
by (auto simp: mem_ball mem_cball) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1151 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1152 |
lemma mem_cball_leI: "x \<in> cball y e \<Longrightarrow> e \<le> f \<Longrightarrow> x \<in> cball y f" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1153 |
by (auto simp: mem_ball mem_cball) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1154 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1155 |
lemma cball_trans: "y \<in> cball z b \<Longrightarrow> x \<in> cball y a \<Longrightarrow> x \<in> cball z (b + a)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1156 |
unfolding mem_cball |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1157 |
proof - |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1158 |
have "dist z x \<le> dist z y + dist y x" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1159 |
by (rule dist_triangle) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1160 |
also assume "dist z y \<le> b" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1161 |
also assume "dist y x \<le> a" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1162 |
finally show "dist z x \<le> b + a" by arith |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1163 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1164 |
|
| 33175 | 1165 |
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
|
1166 |
by (simp add: set_eq_iff) arith |
| 33175 | 1167 |
|
1168 |
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
|
1169 |
by (simp add: set_eq_iff) |
| 33175 | 1170 |
|
|
64758
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1171 |
lemma cball_max_Un: "cball a (max r s) = cball a r \<union> cball a s" |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1172 |
by (simp add: set_eq_iff) arith |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1173 |
|
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1174 |
lemma cball_min_Int: "cball a (min r s) = cball a r \<inter> cball a s" |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1175 |
by (simp add: set_eq_iff) |
|
3b33d2fc5fc0
A few new lemmas and needed adaptations
paulson <lp15@cam.ac.uk>
parents:
64539
diff
changeset
|
1176 |
|
|
64788
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
1177 |
lemma cball_diff_eq_sphere: "cball a r - ball a r = sphere a r" |
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
1178 |
by (auto simp: cball_def ball_def dist_commute) |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
1179 |
|
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1180 |
lemma image_add_ball [simp]: |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1181 |
fixes a :: "'a::real_normed_vector" |
| 67399 | 1182 |
shows "(+) b ` ball a r = ball (a+b) r" |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1183 |
apply (intro equalityI subsetI) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1184 |
apply (force simp: dist_norm) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1185 |
apply (rule_tac x="x-b" in image_eqI) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1186 |
apply (auto simp: dist_norm algebra_simps) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1187 |
done |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1188 |
|
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1189 |
lemma image_add_cball [simp]: |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1190 |
fixes a :: "'a::real_normed_vector" |
| 67399 | 1191 |
shows "(+) b ` cball a r = cball (a+b) r" |
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1192 |
apply (intro equalityI subsetI) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1193 |
apply (force simp: dist_norm) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1194 |
apply (rule_tac x="x-b" in image_eqI) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1195 |
apply (auto simp: dist_norm algebra_simps) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1196 |
done |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
1197 |
|
| 54070 | 1198 |
lemma open_ball [intro, simp]: "open (ball x e)" |
1199 |
proof - |
|
1200 |
have "open (dist x -` {..<e})"
|
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
1201 |
by (intro open_vimage open_lessThan continuous_intros) |
| 54070 | 1202 |
also have "dist x -` {..<e} = ball x e"
|
1203 |
by auto |
|
1204 |
finally show ?thesis . |
|
1205 |
qed |
|
| 33175 | 1206 |
|
1207 |
lemma open_contains_ball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. ball x e \<subseteq> S)" |
|
| 63170 | 1208 |
by (simp add: open_dist subset_eq mem_ball Ball_def dist_commute) |
| 33175 | 1209 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
1210 |
lemma openI [intro?]: "(\<And>x. x\<in>S \<Longrightarrow> \<exists>e>0. ball x e \<subseteq> S) \<Longrightarrow> open S" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
1211 |
by (auto simp: open_contains_ball) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
1212 |
|
|
33714
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
1213 |
lemma openE[elim?]: |
| 53282 | 1214 |
assumes "open S" "x\<in>S" |
|
33714
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
1215 |
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
|
1216 |
using assms unfolding open_contains_ball by auto |
|
eb2574ac4173
Added new lemmas to Euclidean Space by Robert Himmelmann
hoelzl
parents:
33324
diff
changeset
|
1217 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
1218 |
lemma open_contains_ball_eq: "open S \<Longrightarrow> x\<in>S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)" |
| 33175 | 1219 |
by (metis open_contains_ball subset_eq centre_in_ball) |
1220 |
||
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1221 |
lemma openin_contains_ball: |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1222 |
"openin (subtopology euclidean t) s \<longleftrightarrow> |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1223 |
s \<subseteq> t \<and> (\<forall>x \<in> s. \<exists>e. 0 < e \<and> ball x e \<inter> t \<subseteq> s)" |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1224 |
(is "?lhs = ?rhs") |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1225 |
proof |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1226 |
assume ?lhs |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1227 |
then show ?rhs |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1228 |
apply (simp add: openin_open) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1229 |
apply (metis Int_commute Int_mono inf.cobounded2 open_contains_ball order_refl subsetCE) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1230 |
done |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1231 |
next |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1232 |
assume ?rhs |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1233 |
then show ?lhs |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1234 |
apply (simp add: openin_euclidean_subtopology_iff) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1235 |
by (metis (no_types) Int_iff dist_commute inf.absorb_iff2 mem_ball) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1236 |
qed |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1237 |
|
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1238 |
lemma openin_contains_cball: |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1239 |
"openin (subtopology euclidean t) s \<longleftrightarrow> |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1240 |
s \<subseteq> t \<and> |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1241 |
(\<forall>x \<in> s. \<exists>e. 0 < e \<and> cball x e \<inter> t \<subseteq> s)" |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1242 |
apply (simp add: openin_contains_ball) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1243 |
apply (rule iffI) |
|
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1244 |
apply (auto dest!: bspec) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1245 |
apply (rule_tac x="e/2" in exI, force+) |
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
1246 |
done |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
1247 |
|
| 33175 | 1248 |
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
|
1249 |
unfolding mem_ball set_eq_iff |
| 33175 | 1250 |
apply (simp add: not_less) |
| 52624 | 1251 |
apply (metis zero_le_dist order_trans dist_self) |
1252 |
done |
|
| 33175 | 1253 |
|
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
1254 |
lemma ball_empty: "e \<le> 0 \<Longrightarrow> ball x e = {}" by simp
|
| 33175 | 1255 |
|
|
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1256 |
lemma closed_cball [iff]: "closed (cball x e)" |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1257 |
proof - |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1258 |
have "closed (dist x -` {..e})"
|
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1259 |
by (intro closed_vimage closed_atMost continuous_intros) |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1260 |
also have "dist x -` {..e} = cball x e"
|
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1261 |
by auto |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1262 |
finally show ?thesis . |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1263 |
qed |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1264 |
|
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1265 |
lemma open_contains_cball: "open S \<longleftrightarrow> (\<forall>x\<in>S. \<exists>e>0. cball x e \<subseteq> S)" |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1266 |
proof - |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1267 |
{
|
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1268 |
fix x and e::real |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1269 |
assume "x\<in>S" "e>0" "ball x e \<subseteq> S" |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1270 |
then have "\<exists>d>0. cball x d \<subseteq> S" unfolding subset_eq by (rule_tac x="e/2" in exI, auto) |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1271 |
} |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1272 |
moreover |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1273 |
{
|
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1274 |
fix x and e::real |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1275 |
assume "x\<in>S" "e>0" "cball x e \<subseteq> S" |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1276 |
then have "\<exists>d>0. ball x d \<subseteq> S" |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1277 |
unfolding subset_eq |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1278 |
apply (rule_tac x="e/2" in exI, auto) |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1279 |
done |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1280 |
} |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1281 |
ultimately show ?thesis |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1282 |
unfolding open_contains_ball by auto |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1283 |
qed |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1284 |
|
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1285 |
lemma open_contains_cball_eq: "open S \<Longrightarrow> (\<forall>x. x \<in> S \<longleftrightarrow> (\<exists>e>0. cball x e \<subseteq> S))" |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1286 |
by (metis open_contains_cball subset_eq order_less_imp_le centre_in_cball) |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
1287 |
|
|
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
|
1288 |
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
|
1289 |
fixes x y :: "'a :: euclidean_space" |
| 67155 | 1290 |
shows "dist x y = L2_set (\<lambda>i. dist (x \<bullet> i) (y \<bullet> i)) Basis" |
1291 |
unfolding dist_norm norm_eq_sqrt_inner L2_set_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
|
1292 |
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
|
1293 |
|
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1294 |
lemma norm_nth_le: "norm (x \<bullet> i) \<le> norm x" if "i \<in> Basis" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1295 |
proof - |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1296 |
have "(x \<bullet> i)\<^sup>2 = (\<Sum>i\<in>{i}. (x \<bullet> i)\<^sup>2)"
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1297 |
by simp |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1298 |
also have "\<dots> \<le> (\<Sum>i\<in>Basis. (x \<bullet> i)\<^sup>2)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1299 |
by (intro sum_mono2) (auto simp: that) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1300 |
finally show ?thesis |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1301 |
unfolding norm_conv_dist euclidean_dist_l2[of x] L2_set_def |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1302 |
by (auto intro!: real_le_rsqrt) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1303 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1304 |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1305 |
lemma eventually_nhds_ball: "d > 0 \<Longrightarrow> eventually (\<lambda>x. x \<in> ball z d) (nhds z)" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1306 |
by (rule eventually_nhds_in_open) simp_all |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1307 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1308 |
lemma eventually_at_ball: "d > 0 \<Longrightarrow> eventually (\<lambda>t. t \<in> ball z d \<and> t \<in> A) (at z within A)" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1309 |
unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1310 |
|
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1311 |
lemma eventually_at_ball': "d > 0 \<Longrightarrow> eventually (\<lambda>t. t \<in> ball z d \<and> t \<noteq> z \<and> t \<in> A) (at z within A)" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1312 |
unfolding eventually_at by (intro exI[of _ d]) (simp_all add: dist_commute) |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
1313 |
|
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1314 |
lemma at_within_ball: "e > 0 \<Longrightarrow> dist x y < e \<Longrightarrow> at y within ball x e = at y" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1315 |
by (subst at_within_open) auto |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1316 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1317 |
lemma atLeastAtMost_eq_cball: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1318 |
fixes a b::real |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1319 |
shows "{a .. b} = cball ((a + b)/2) ((b - a)/2)"
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1320 |
by (auto simp: dist_real_def field_simps mem_cball) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1321 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1322 |
lemma greaterThanLessThan_eq_ball: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1323 |
fixes a b::real |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1324 |
shows "{a <..< b} = ball ((a + b)/2) ((b - a)/2)"
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1325 |
by (auto simp: dist_real_def field_simps mem_ball) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1326 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1327 |
|
| 60420 | 1328 |
subsection \<open>Boxes\<close> |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1329 |
|
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1330 |
abbreviation One :: "'a::euclidean_space" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1331 |
where "One \<equiv> \<Sum>Basis" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1332 |
|
|
63114
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1333 |
lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1334 |
proof - |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1335 |
have "dependent (Basis :: 'a set)" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1336 |
apply (simp add: dependent_finite) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1337 |
apply (rule_tac x="\<lambda>i. 1" in exI) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1338 |
using SOME_Basis apply (auto simp: assms) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1339 |
done |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1340 |
with independent_Basis show False by force |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1341 |
qed |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1342 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1343 |
corollary One_neq_0[iff]: "One \<noteq> 0" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1344 |
by (metis One_non_0) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1345 |
|
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1346 |
corollary Zero_neq_One[iff]: "0 \<noteq> One" |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1347 |
by (metis One_non_0) |
|
27afe7af7379
Lots of new material for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
63105
diff
changeset
|
1348 |
|
| 67962 | 1349 |
definition%important (in euclidean_space) eucl_less (infix "<e" 50) |
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
1350 |
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
|
1351 |
|
| 67962 | 1352 |
definition%important box_eucl_less: "box a b = {x. a <e x \<and> x <e b}"
|
1353 |
definition%important "cbox a b = {x. \<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i}"
|
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
1354 |
|
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
1355 |
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
|
1356 |
and in_box_eucl_less: "x \<in> box a b \<longleftrightarrow> a <e x \<and> x <e b" |
| 56188 | 1357 |
and mem_box: "x \<in> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i)" |
1358 |
"x \<in> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)" |
|
1359 |
by (auto simp: box_eucl_less eucl_less_def cbox_def) |
|
1360 |
||
|
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1361 |
lemma cbox_Pair_eq: "cbox (a, c) (b, d) = cbox a b \<times> cbox c d" |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1362 |
by (force simp: cbox_def Basis_prod_def) |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1363 |
|
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1364 |
lemma cbox_Pair_iff [iff]: "(x, y) \<in> cbox (a, c) (b, d) \<longleftrightarrow> x \<in> cbox a b \<and> y \<in> cbox c d" |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1365 |
by (force simp: cbox_Pair_eq) |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1366 |
|
|
65587
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
1367 |
lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\<lambda>(x,y). Complex x y) ` (cbox a b \<times> cbox c d)" |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
1368 |
apply (auto simp: cbox_def Basis_complex_def) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
1369 |
apply (rule_tac x = "(Re x, Im x)" in image_eqI) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
1370 |
using complex_eq by auto |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
1371 |
|
|
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1372 |
lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \<longleftrightarrow> cbox a b = {} \<or> cbox c d = {}"
|
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1373 |
by (force simp: cbox_Pair_eq) |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1374 |
|
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1375 |
lemma swap_cbox_Pair [simp]: "prod.swap ` cbox (c, a) (d, b) = cbox (a,c) (b,d)" |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1376 |
by auto |
|
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
1377 |
|
| 56188 | 1378 |
lemma mem_box_real[simp]: |
1379 |
"(x::real) \<in> box a b \<longleftrightarrow> a < x \<and> x < b" |
|
1380 |
"(x::real) \<in> cbox a b \<longleftrightarrow> a \<le> x \<and> x \<le> b" |
|
1381 |
by (auto simp: mem_box) |
|
1382 |
||
1383 |
lemma box_real[simp]: |
|
1384 |
fixes a b:: real |
|
1385 |
shows "box a b = {a <..< b}" "cbox a b = {a .. b}"
|
|
1386 |
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
|
1387 |
|
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1388 |
lemma box_Int_box: |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1389 |
fixes a :: "'a::euclidean_space" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1390 |
shows "box a b \<inter> box c d = |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1391 |
box (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1392 |
unfolding set_eq_iff and Int_iff and mem_box by auto |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1393 |
|
| 50087 | 1394 |
lemma rational_boxes: |
| 61076 | 1395 |
fixes x :: "'a::euclidean_space" |
| 53291 | 1396 |
assumes "e > 0" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1397 |
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 | 1398 |
proof - |
| 63040 | 1399 |
define e' where "e' = e / (2 * sqrt (real (DIM ('a))))"
|
| 53291 | 1400 |
then have e: "e' > 0" |
| 56541 | 1401 |
using assms by (auto simp: DIM_positive) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
1402 |
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 | 1403 |
proof |
| 53255 | 1404 |
fix i |
1405 |
from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e |
|
1406 |
show "?th i" by auto |
|
| 50087 | 1407 |
qed |
| 55522 | 1408 |
from choice[OF this] obtain a where |
1409 |
a: "\<forall>xa. a xa \<in> \<rat> \<and> a xa < x \<bullet> xa \<and> x \<bullet> xa - a xa < e'" .. |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
1410 |
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 | 1411 |
proof |
| 53255 | 1412 |
fix i |
1413 |
from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e |
|
1414 |
show "?th i" by auto |
|
| 50087 | 1415 |
qed |
| 55522 | 1416 |
from choice[OF this] obtain b where |
1417 |
b: "\<forall>xa. b xa \<in> \<rat> \<and> x \<bullet> xa < b xa \<and> b xa - x \<bullet> xa < e'" .. |
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
1418 |
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
|
1419 |
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
|
1420 |
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) |
| 53255 | 1421 |
fix y :: 'a |
1422 |
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
|
1423 |
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" |
| 67155 | 1424 |
unfolding L2_set_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
|
1425 |
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
|
| 64267 | 1426 |
proof (rule real_sqrt_less_mono, rule sum_strict_mono) |
| 53255 | 1427 |
fix i :: "'a" |
1428 |
assume i: "i \<in> Basis" |
|
1429 |
have "a i < y\<bullet>i \<and> y\<bullet>i < b i" |
|
1430 |
using * i by (auto simp: box_def) |
|
1431 |
moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" |
|
1432 |
using a by auto |
|
1433 |
moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" |
|
1434 |
using b by auto |
|
1435 |
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" |
|
1436 |
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
|
1437 |
then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
|
| 50087 | 1438 |
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
|
1439 |
then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
|
| 50087 | 1440 |
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
|
1441 |
then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
|
| 50087 | 1442 |
by (simp add: power_divide) |
1443 |
qed auto |
|
| 53255 | 1444 |
also have "\<dots> = e" |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1445 |
using \<open>0 < e\<close> by simp |
| 53255 | 1446 |
finally show "y \<in> ball x e" |
1447 |
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
|
1448 |
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
|
1449 |
qed |
| 51103 | 1450 |
|
|
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
|
1451 |
lemma open_UNION_box: |
| 61076 | 1452 |
fixes M :: "'a::euclidean_space set" |
| 53282 | 1453 |
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
|
1454 |
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
|
1455 |
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
|
1456 |
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
|
1457 |
shows "M = (\<Union>f\<in>I. box (a' f) (b' f))" |
| 52624 | 1458 |
proof - |
| 60462 | 1459 |
have "x \<in> (\<Union>f\<in>I. box (a' f) (b' f))" if "x \<in> M" for x |
1460 |
proof - |
|
| 52624 | 1461 |
obtain e where e: "e > 0" "ball x e \<subseteq> M" |
| 60420 | 1462 |
using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto |
| 53282 | 1463 |
moreover obtain a b where ab: |
1464 |
"x \<in> box a b" |
|
1465 |
"\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" |
|
1466 |
"\<forall>i\<in>Basis. b \<bullet> i \<in> \<rat>" |
|
1467 |
"box a b \<subseteq> ball x e" |
|
| 52624 | 1468 |
using rational_boxes[OF e(1)] by metis |
| 60462 | 1469 |
ultimately show ?thesis |
| 52624 | 1470 |
by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) |
1471 |
(auto simp: euclidean_representation I_def a'_def b'_def) |
|
| 60462 | 1472 |
qed |
| 52624 | 1473 |
then show ?thesis by (auto simp: I_def) |
1474 |
qed |
|
1475 |
||
|
66154
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1476 |
corollary open_countable_Union_open_box: |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1477 |
fixes S :: "'a :: euclidean_space set" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1478 |
assumes "open S" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1479 |
obtains \<D> where "countable \<D>" "\<D> \<subseteq> Pow S" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = box a b" "\<Union>\<D> = S" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1480 |
proof - |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1481 |
let ?a = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1482 |
let ?b = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1483 |
let ?I = "{f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. box (?a f) (?b f) \<subseteq> S}"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1484 |
let ?\<D> = "(\<lambda>f. box (?a f) (?b f)) ` ?I" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1485 |
show ?thesis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1486 |
proof |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1487 |
have "countable ?I" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1488 |
by (simp add: countable_PiE countable_rat) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1489 |
then show "countable ?\<D>" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1490 |
by blast |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1491 |
show "\<Union>?\<D> = S" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1492 |
using open_UNION_box [OF assms] by metis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1493 |
qed auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1494 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1495 |
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1496 |
lemma rational_cboxes: |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1497 |
fixes x :: "'a::euclidean_space" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1498 |
assumes "e > 0" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1499 |
shows "\<exists>a b. (\<forall>i\<in>Basis. a \<bullet> i \<in> \<rat> \<and> b \<bullet> i \<in> \<rat>) \<and> x \<in> cbox a b \<and> cbox a b \<subseteq> ball x e" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1500 |
proof - |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1501 |
define e' where "e' = e / (2 * sqrt (real (DIM ('a))))"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1502 |
then have e: "e' > 0" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1503 |
using assms by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1504 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x \<bullet> i \<and> x \<bullet> i - y < e'" (is "\<forall>i. ?th i") |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1505 |
proof |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1506 |
fix i |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1507 |
from Rats_dense_in_real[of "x \<bullet> i - e'" "x \<bullet> i"] e |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1508 |
show "?th i" by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1509 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1510 |
from choice[OF this] obtain a where |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1511 |
a: "\<forall>u. a u \<in> \<rat> \<and> a u < x \<bullet> u \<and> x \<bullet> u - a u < e'" .. |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1512 |
have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x \<bullet> i < y \<and> y - x \<bullet> i < e'" (is "\<forall>i. ?th i") |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1513 |
proof |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1514 |
fix i |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1515 |
from Rats_dense_in_real[of "x \<bullet> i" "x \<bullet> i + e'"] e |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1516 |
show "?th i" by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1517 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1518 |
from choice[OF this] obtain b where |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1519 |
b: "\<forall>u. b u \<in> \<rat> \<and> x \<bullet> u < b u \<and> b u - x \<bullet> u < e'" .. |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1520 |
let ?a = "\<Sum>i\<in>Basis. a i *\<^sub>R i" and ?b = "\<Sum>i\<in>Basis. b i *\<^sub>R i" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1521 |
show ?thesis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1522 |
proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1523 |
fix y :: 'a |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1524 |
assume *: "y \<in> cbox ?a ?b" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1525 |
have "dist x y = sqrt (\<Sum>i\<in>Basis. (dist (x \<bullet> i) (y \<bullet> i))\<^sup>2)" |
| 67155 | 1526 |
unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) |
|
66154
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1527 |
also have "\<dots> < sqrt (\<Sum>(i::'a)\<in>Basis. e^2 / real (DIM('a)))"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1528 |
proof (rule real_sqrt_less_mono, rule sum_strict_mono) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1529 |
fix i :: "'a" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1530 |
assume i: "i \<in> Basis" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1531 |
have "a i \<le> y\<bullet>i \<and> y\<bullet>i \<le> b i" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1532 |
using * i by (auto simp: cbox_def) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1533 |
moreover have "a i < x\<bullet>i" "x\<bullet>i - a i < e'" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1534 |
using a by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1535 |
moreover have "x\<bullet>i < b i" "b i - x\<bullet>i < e'" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1536 |
using b by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1537 |
ultimately have "\<bar>x\<bullet>i - y\<bullet>i\<bar> < 2 * e'" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1538 |
by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1539 |
then have "dist (x \<bullet> i) (y \<bullet> i) < e/sqrt (real (DIM('a)))"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1540 |
unfolding e'_def by (auto simp: dist_real_def) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1541 |
then have "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1542 |
by (rule power_strict_mono) auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1543 |
then show "(dist (x \<bullet> i) (y \<bullet> i))\<^sup>2 < e\<^sup>2 / real DIM('a)"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1544 |
by (simp add: power_divide) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1545 |
qed auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1546 |
also have "\<dots> = e" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1547 |
using \<open>0 < e\<close> by simp |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1548 |
finally show "y \<in> ball x e" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1549 |
by (auto simp: ball_def) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1550 |
next |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1551 |
show "x \<in> cbox (\<Sum>i\<in>Basis. a i *\<^sub>R i) (\<Sum>i\<in>Basis. b i *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1552 |
using a b less_imp_le by (auto simp: cbox_def) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1553 |
qed (use a b cbox_def in auto) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1554 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1555 |
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1556 |
lemma open_UNION_cbox: |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1557 |
fixes M :: "'a::euclidean_space set" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1558 |
assumes "open M" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1559 |
defines "a' \<equiv> \<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1560 |
defines "b' \<equiv> \<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1561 |
defines "I \<equiv> {f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. cbox (a' f) (b' f) \<subseteq> M}"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1562 |
shows "M = (\<Union>f\<in>I. cbox (a' f) (b' f))" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1563 |
proof - |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1564 |
have "x \<in> (\<Union>f\<in>I. cbox (a' f) (b' f))" if "x \<in> M" for x |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1565 |
proof - |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1566 |
obtain e where e: "e > 0" "ball x e \<subseteq> M" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1567 |
using openE[OF \<open>open M\<close> \<open>x \<in> M\<close>] by auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1568 |
moreover obtain a b where ab: "x \<in> cbox a b" "\<forall>i \<in> Basis. a \<bullet> i \<in> \<rat>" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1569 |
"\<forall>i \<in> Basis. b \<bullet> i \<in> \<rat>" "cbox a b \<subseteq> ball x e" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1570 |
using rational_cboxes[OF e(1)] by metis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1571 |
ultimately show ?thesis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1572 |
by (intro UN_I[of "\<lambda>i\<in>Basis. (a \<bullet> i, b \<bullet> i)"]) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1573 |
(auto simp: euclidean_representation I_def a'_def b'_def) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1574 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1575 |
then show ?thesis by (auto simp: I_def) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1576 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1577 |
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1578 |
corollary open_countable_Union_open_cbox: |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1579 |
fixes S :: "'a :: euclidean_space set" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1580 |
assumes "open S" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1581 |
obtains \<D> where "countable \<D>" "\<D> \<subseteq> Pow S" "\<And>X. X \<in> \<D> \<Longrightarrow> \<exists>a b. X = cbox a b" "\<Union>\<D> = S" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1582 |
proof - |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1583 |
let ?a = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. fst (f i) *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1584 |
let ?b = "\<lambda>f. (\<Sum>(i::'a)\<in>Basis. snd (f i) *\<^sub>R i)" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1585 |
let ?I = "{f\<in>Basis \<rightarrow>\<^sub>E \<rat> \<times> \<rat>. cbox (?a f) (?b f) \<subseteq> S}"
|
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1586 |
let ?\<D> = "(\<lambda>f. cbox (?a f) (?b f)) ` ?I" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1587 |
show ?thesis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1588 |
proof |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1589 |
have "countable ?I" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1590 |
by (simp add: countable_PiE countable_rat) |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1591 |
then show "countable ?\<D>" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1592 |
by blast |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1593 |
show "\<Union>?\<D> = S" |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1594 |
using open_UNION_cbox [OF assms] by metis |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1595 |
qed auto |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1596 |
qed |
|
bc5e6461f759
Tidying up integration theory and some new theorems
paulson <lp15@cam.ac.uk>
parents:
66112
diff
changeset
|
1597 |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1598 |
lemma box_eq_empty: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1599 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1600 |
shows "(box a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i))" (is ?th1)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1601 |
and "(cbox a b = {} \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i < a\<bullet>i))" (is ?th2)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1602 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1603 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1604 |
fix i x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1605 |
assume i: "i\<in>Basis" and as:"b\<bullet>i \<le> a\<bullet>i" and x:"x\<in>box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1606 |
then have "a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1607 |
unfolding mem_box by (auto simp: box_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1608 |
then have "a\<bullet>i < b\<bullet>i" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1609 |
then have False using as by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1610 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1611 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1612 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1613 |
assume as: "\<forall>i\<in>Basis. \<not> (b\<bullet>i \<le> a\<bullet>i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1614 |
let ?x = "(1/2) *\<^sub>R (a + b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1615 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1616 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1617 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1618 |
have "a\<bullet>i < b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1619 |
using as[THEN bspec[where x=i]] i by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1620 |
then have "a\<bullet>i < ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i < b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1621 |
by (auto simp: inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1622 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1623 |
then have "box a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1624 |
using mem_box(1)[of "?x" a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1625 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1626 |
ultimately show ?th1 by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1627 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1628 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1629 |
fix i x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1630 |
assume i: "i \<in> Basis" and as:"b\<bullet>i < a\<bullet>i" and x:"x\<in>cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1631 |
then have "a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1632 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1633 |
then have "a\<bullet>i \<le> b\<bullet>i" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1634 |
then have False using as by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1635 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1636 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1637 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1638 |
assume as:"\<forall>i\<in>Basis. \<not> (b\<bullet>i < a\<bullet>i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1639 |
let ?x = "(1/2) *\<^sub>R (a + b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1640 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1641 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1642 |
assume i:"i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1643 |
have "a\<bullet>i \<le> b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1644 |
using as[THEN bspec[where x=i]] i by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1645 |
then have "a\<bullet>i \<le> ((1/2) *\<^sub>R (a+b)) \<bullet> i" "((1/2) *\<^sub>R (a+b)) \<bullet> i \<le> b\<bullet>i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1646 |
by (auto simp: inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1647 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1648 |
then have "cbox a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1649 |
using mem_box(2)[of "?x" a b] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1650 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1651 |
ultimately show ?th2 by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1652 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1653 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1654 |
lemma box_ne_empty: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1655 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1656 |
shows "cbox a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i)"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1657 |
and "box a b \<noteq> {} \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1658 |
unfolding box_eq_empty[of a b] by fastforce+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1659 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1660 |
lemma |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1661 |
fixes a :: "'a::euclidean_space" |
|
66112
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents:
66089
diff
changeset
|
1662 |
shows cbox_sing [simp]: "cbox a a = {a}"
|
|
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
paulson <lp15@cam.ac.uk>
parents:
66089
diff
changeset
|
1663 |
and box_sing [simp]: "box a a = {}"
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1664 |
unfolding set_eq_iff mem_box eq_iff [symmetric] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1665 |
by (auto intro!: euclidean_eqI[where 'a='a]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1666 |
(metis all_not_in_conv nonempty_Basis) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1667 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1668 |
lemma subset_box_imp: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1669 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1670 |
shows "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1671 |
and "(\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i) \<Longrightarrow> cbox c d \<subseteq> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1672 |
and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1673 |
and "(\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i) \<Longrightarrow> box c d \<subseteq> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1674 |
unfolding subset_eq[unfolded Ball_def] unfolding mem_box |
| 58757 | 1675 |
by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1676 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1677 |
lemma box_subset_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1678 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1679 |
shows "box a b \<subseteq> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1680 |
unfolding subset_eq [unfolded Ball_def] mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1681 |
by (fast intro: less_imp_le) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1682 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1683 |
lemma subset_box: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1684 |
fixes a :: "'a::euclidean_space" |
| 64539 | 1685 |
shows "cbox c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th1) |
1686 |
and "cbox c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i \<le> d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < c\<bullet>i \<and> d\<bullet>i < b\<bullet>i)" (is ?th2) |
|
1687 |
and "box c d \<subseteq> cbox a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th3) |
|
1688 |
and "box c d \<subseteq> box a b \<longleftrightarrow> (\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i) \<longrightarrow> (\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i)" (is ?th4) |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1689 |
proof - |
| 68120 | 1690 |
let ?lesscd = "\<forall>i\<in>Basis. c\<bullet>i < d\<bullet>i" |
1691 |
let ?lerhs = "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
1692 |
show ?th1 ?th2 |
|
1693 |
by (fastforce simp: mem_box)+ |
|
1694 |
have acdb: "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
1695 |
if i: "i \<in> Basis" and box: "box c d \<subseteq> cbox a b" and cd: "\<And>i. i \<in> Basis \<Longrightarrow> c\<bullet>i < d\<bullet>i" for i |
|
1696 |
proof - |
|
1697 |
have "box c d \<noteq> {}"
|
|
1698 |
using that |
|
1699 |
unfolding box_eq_empty by force |
|
1700 |
{ let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((min (a\<bullet>j) (d\<bullet>j))+c\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
|
|
1701 |
assume *: "a\<bullet>i > c\<bullet>i" |
|
1702 |
then have "c \<bullet> j < ?x \<bullet> j \<and> ?x \<bullet> j < d \<bullet> j" if "j \<in> Basis" for j |
|
1703 |
using cd that by (fastforce simp add: i *) |
|
1704 |
then have "?x \<in> box c d" |
|
1705 |
unfolding mem_box by auto |
|
1706 |
moreover have "?x \<notin> cbox a b" |
|
1707 |
using i cd * by (force simp: mem_box) |
|
1708 |
ultimately have False using box by auto |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1709 |
} |
| 68120 | 1710 |
then have "a\<bullet>i \<le> c\<bullet>i" by force |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1711 |
moreover |
| 68120 | 1712 |
{ let ?x = "(\<Sum>j\<in>Basis. (if j=i then ((max (b\<bullet>j) (c\<bullet>j))+d\<bullet>j)/2 else (c\<bullet>j+d\<bullet>j)/2) *\<^sub>R j)::'a"
|
1713 |
assume *: "b\<bullet>i < d\<bullet>i" |
|
1714 |
then have "d \<bullet> j > ?x \<bullet> j \<and> ?x \<bullet> j > c \<bullet> j" if "j \<in> Basis" for j |
|
1715 |
using cd that by (fastforce simp add: i *) |
|
1716 |
then have "?x \<in> box c d" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1717 |
unfolding mem_box by auto |
| 68120 | 1718 |
moreover have "?x \<notin> cbox a b" |
1719 |
using i cd * by (force simp: mem_box) |
|
1720 |
ultimately have False using box by auto |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1721 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1722 |
then have "b\<bullet>i \<ge> d\<bullet>i" by (rule ccontr) auto |
| 68120 | 1723 |
ultimately show ?thesis by auto |
1724 |
qed |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1725 |
show ?th3 |
| 68120 | 1726 |
using acdb by (fastforce simp add: mem_box) |
1727 |
have acdb': "a\<bullet>i \<le> c\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i" |
|
1728 |
if "i \<in> Basis" "box c d \<subseteq> box a b" "\<And>i. i \<in> Basis \<Longrightarrow> c\<bullet>i < d\<bullet>i" for i |
|
1729 |
using box_subset_cbox[of a b] that acdb by auto |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1730 |
show ?th4 |
| 68120 | 1731 |
using acdb' by (fastforce simp add: mem_box) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1732 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1733 |
|
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1734 |
lemma eq_cbox: "cbox a b = cbox c d \<longleftrightarrow> cbox a b = {} \<and> cbox c d = {} \<or> a = c \<and> b = d"
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1735 |
(is "?lhs = ?rhs") |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1736 |
proof |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1737 |
assume ?lhs |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1738 |
then have "cbox a b \<subseteq> cbox c d" "cbox c d \<subseteq> cbox a b" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1739 |
by auto |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1740 |
then show ?rhs |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1741 |
by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI) |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1742 |
next |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1743 |
assume ?rhs |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1744 |
then show ?lhs |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1745 |
by force |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1746 |
qed |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1747 |
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1748 |
lemma eq_cbox_box [simp]: "cbox a b = box c d \<longleftrightarrow> cbox a b = {} \<and> box c d = {}"
|
| 64539 | 1749 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1750 |
proof |
| 68120 | 1751 |
assume L: ?lhs |
1752 |
then have "cbox a b \<subseteq> box c d" "box c d \<subseteq> cbox a b" |
|
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1753 |
by auto |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1754 |
then show ?rhs |
|
63957
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
hoelzl
parents:
63955
diff
changeset
|
1755 |
apply (simp add: subset_box) |
| 68120 | 1756 |
using L box_ne_empty box_sing apply (fastforce simp add:) |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1757 |
done |
| 68120 | 1758 |
qed force |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1759 |
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1760 |
lemma eq_box_cbox [simp]: "box a b = cbox c d \<longleftrightarrow> box a b = {} \<and> cbox c d = {}"
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1761 |
by (metis eq_cbox_box) |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1762 |
|
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1763 |
lemma eq_box: "box a b = box c d \<longleftrightarrow> box a b = {} \<and> box c d = {} \<or> a = c \<and> b = d"
|
| 64539 | 1764 |
(is "?lhs \<longleftrightarrow> ?rhs") |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1765 |
proof |
| 68120 | 1766 |
assume L: ?lhs |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1767 |
then have "box a b \<subseteq> box c d" "box c d \<subseteq> box a b" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1768 |
by auto |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1769 |
then show ?rhs |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1770 |
apply (simp add: subset_box) |
| 68120 | 1771 |
using box_ne_empty(2) L |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1772 |
apply auto |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1773 |
apply (meson euclidean_eqI less_eq_real_def not_less)+ |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1774 |
done |
| 68120 | 1775 |
qed force |
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1776 |
|
|
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1777 |
lemma subset_box_complex: |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1778 |
"cbox a b \<subseteq> cbox c d \<longleftrightarrow> |
|
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1779 |
(Re a \<le> Re b \<and> Im a \<le> Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1780 |
"cbox a b \<subseteq> box c d \<longleftrightarrow> |
|
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1781 |
(Re a \<le> Re b \<and> Im a \<le> Im b) \<longrightarrow> Re a > Re c \<and> Im a > Im c \<and> Re b < Re d \<and> Im b < Im d" |
|
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1782 |
"box a b \<subseteq> cbox c d \<longleftrightarrow> |
|
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1783 |
(Re a < Re b \<and> Im a < Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1784 |
"box a b \<subseteq> box c d \<longleftrightarrow> |
|
66466
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1785 |
(Re a < Re b \<and> Im a < Im b) \<longrightarrow> Re a \<ge> Re c \<and> Im a \<ge> Im c \<and> Re b \<le> Re d \<and> Im b \<le> Im d" |
|
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1786 |
by (subst subset_box; force simp: Basis_complex_def)+ |
|
aec5d9c88d69
More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
66453
diff
changeset
|
1787 |
|
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
1788 |
lemma Int_interval: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1789 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1790 |
shows "cbox a b \<inter> cbox c d = |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1791 |
cbox (\<Sum>i\<in>Basis. max (a\<bullet>i) (c\<bullet>i) *\<^sub>R i) (\<Sum>i\<in>Basis. min (b\<bullet>i) (d\<bullet>i) *\<^sub>R i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1792 |
unfolding set_eq_iff and Int_iff and mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1793 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1794 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1795 |
lemma disjoint_interval: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1796 |
fixes a::"'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1797 |
shows "cbox a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i < c\<bullet>i \<or> d\<bullet>i < a\<bullet>i))" (is ?th1)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1798 |
and "cbox a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i < a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th2)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1799 |
and "box a b \<inter> cbox c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i < c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th3)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1800 |
and "box a b \<inter> box c d = {} \<longleftrightarrow> (\<exists>i\<in>Basis. (b\<bullet>i \<le> a\<bullet>i \<or> d\<bullet>i \<le> c\<bullet>i \<or> b\<bullet>i \<le> c\<bullet>i \<or> d\<bullet>i \<le> a\<bullet>i))" (is ?th4)
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1801 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1802 |
let ?z = "(\<Sum>i\<in>Basis. (((max (a\<bullet>i) (c\<bullet>i)) + (min (b\<bullet>i) (d\<bullet>i))) / 2) *\<^sub>R i)::'a" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1803 |
have **: "\<And>P Q. (\<And>i :: 'a. i \<in> Basis \<Longrightarrow> Q ?z i \<Longrightarrow> P i) \<Longrightarrow> |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1804 |
(\<And>i x :: 'a. i \<in> Basis \<Longrightarrow> P i \<Longrightarrow> Q x i) \<Longrightarrow> (\<forall>x. \<exists>i\<in>Basis. Q x i) \<longleftrightarrow> (\<exists>i\<in>Basis. P i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1805 |
by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1806 |
note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1807 |
show ?th1 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1808 |
show ?th2 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1809 |
show ?th3 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1810 |
show ?th4 unfolding * by (intro **) auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1811 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1812 |
|
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1813 |
lemma UN_box_eq_UNIV: "(\<Union>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1814 |
proof - |
| 61942 | 1815 |
have "\<bar>x \<bullet> b\<bar> < real_of_int (\<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil> + 1)" |
| 60462 | 1816 |
if [simp]: "b \<in> Basis" for x b :: 'a |
1817 |
proof - |
|
| 61942 | 1818 |
have "\<bar>x \<bullet> b\<bar> \<le> real_of_int \<lceil>\<bar>x \<bullet> b\<bar>\<rceil>" |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1819 |
by (rule le_of_int_ceiling) |
| 61942 | 1820 |
also have "\<dots> \<le> real_of_int \<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil>" |
|
59587
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
nipkow
parents:
58877
diff
changeset
|
1821 |
by (auto intro!: ceiling_mono) |
| 61942 | 1822 |
also have "\<dots> < real_of_int (\<lceil>Max ((\<lambda>b. \<bar>x \<bullet> b\<bar>)`Basis)\<rceil> + 1)" |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1823 |
by simp |
| 60462 | 1824 |
finally show ?thesis . |
1825 |
qed |
|
1826 |
then have "\<exists>n::nat. \<forall>b\<in>Basis. \<bar>x \<bullet> b\<bar> < real n" for x :: 'a |
|
|
59587
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
nipkow
parents:
58877
diff
changeset
|
1827 |
by (metis order.strict_trans reals_Archimedean2) |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1828 |
moreover have "\<And>x b::'a. \<And>n::nat. \<bar>x \<bullet> b\<bar> < real n \<longleftrightarrow> - real n < x \<bullet> b \<and> x \<bullet> b < real n" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1829 |
by auto |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1830 |
ultimately show ?thesis |
| 64267 | 1831 |
by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1832 |
qed |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
1833 |
|
| 67968 | 1834 |
subsection \<open>Intervals in general, including infinite and mixtures of open and closed\<close> |
| 67962 | 1835 |
|
1836 |
definition%important "is_interval (s::('a::euclidean_space) set) \<longleftrightarrow>
|
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1837 |
(\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i\<in>Basis. ((a\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> b\<bullet>i) \<or> (b\<bullet>i \<le> x\<bullet>i \<and> x\<bullet>i \<le> a\<bullet>i))) \<longrightarrow> x \<in> s)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1838 |
|
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1839 |
lemma is_interval_1: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1840 |
"is_interval (s::real set) \<longleftrightarrow> (\<forall>a\<in>s. \<forall>b\<in>s. \<forall> x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> s)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1841 |
unfolding is_interval_def by auto |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1842 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1843 |
lemma is_interval_inter: "is_interval X \<Longrightarrow> is_interval Y \<Longrightarrow> is_interval (X \<inter> Y)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1844 |
unfolding is_interval_def |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1845 |
by blast |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1846 |
|
|
66089
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1847 |
lemma is_interval_cbox [simp]: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1848 |
and is_interval_box [simp]: "is_interval (box a b)" (is ?th2) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1849 |
unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1850 |
by (meson order_trans le_less_trans less_le_trans less_trans)+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1851 |
|
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1852 |
lemma is_interval_empty [iff]: "is_interval {}"
|
|
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1853 |
unfolding is_interval_def by simp |
|
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1854 |
|
|
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1855 |
lemma is_interval_univ [iff]: "is_interval UNIV" |
|
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
1856 |
unfolding is_interval_def by simp |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1857 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1858 |
lemma mem_is_intervalI: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1859 |
assumes "is_interval s" |
| 64539 | 1860 |
and "a \<in> s" "b \<in> s" |
1861 |
and "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i \<or> b \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> a \<bullet> i" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1862 |
shows "x \<in> s" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1863 |
by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1864 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1865 |
lemma interval_subst: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1866 |
fixes S::"'a::euclidean_space set" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1867 |
assumes "is_interval S" |
| 64539 | 1868 |
and "x \<in> S" "y j \<in> S" |
1869 |
and "j \<in> Basis" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1870 |
shows "(\<Sum>i\<in>Basis. (if i = j then y i \<bullet> i else x \<bullet> i) *\<^sub>R i) \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1871 |
by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1872 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1873 |
lemma mem_box_componentwiseI: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1874 |
fixes S::"'a::euclidean_space set" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1875 |
assumes "is_interval S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1876 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i \<in> ((\<lambda>x. x \<bullet> i) ` S)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1877 |
shows "x \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1878 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1879 |
from assms have "\<forall>i \<in> Basis. \<exists>s \<in> S. x \<bullet> i = s \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1880 |
by auto |
| 64539 | 1881 |
with finite_Basis obtain s and bs::"'a list" |
1882 |
where s: "\<And>i. i \<in> Basis \<Longrightarrow> x \<bullet> i = s i \<bullet> i" "\<And>i. i \<in> Basis \<Longrightarrow> s i \<in> S" |
|
1883 |
and bs: "set bs = Basis" "distinct bs" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1884 |
by (metis finite_distinct_list) |
| 64539 | 1885 |
from nonempty_Basis s obtain j where j: "j \<in> Basis" "s j \<in> S" |
1886 |
by blast |
|
| 63040 | 1887 |
define y where |
1888 |
"y = rec_list (s j) (\<lambda>j _ Y. (\<Sum>i\<in>Basis. (if i = j then s i \<bullet> i else Y \<bullet> i) *\<^sub>R i))" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1889 |
have "x = (\<Sum>i\<in>Basis. (if i \<in> set bs then s i \<bullet> i else s j \<bullet> i) *\<^sub>R i)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
1890 |
using bs by (auto simp: s(1)[symmetric] euclidean_representation) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1891 |
also have [symmetric]: "y bs = \<dots>" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1892 |
using bs(2) bs(1)[THEN equalityD1] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1893 |
by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1894 |
also have "y bs \<in> S" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1895 |
using bs(1)[THEN equalityD1] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1896 |
apply (induct bs) |
| 64539 | 1897 |
apply (auto simp: y_def j) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1898 |
apply (rule interval_subst[OF assms(1)]) |
| 64539 | 1899 |
apply (auto simp: s) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1900 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1901 |
finally show ?thesis . |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1902 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
1903 |
|
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62948
diff
changeset
|
1904 |
lemma cbox01_nonempty [simp]: "cbox 0 One \<noteq> {}"
|
| 64267 | 1905 |
by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg) |
|
63007
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62948
diff
changeset
|
1906 |
|
|
aa894a49f77d
new theorems about convex hulls, etc.; also, renamed some theorems
paulson <lp15@cam.ac.uk>
parents:
62948
diff
changeset
|
1907 |
lemma box01_nonempty [simp]: "box 0 One \<noteq> {}"
|
|
66089
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1908 |
by (simp add: box_ne_empty inner_Basis inner_sum_left) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
1909 |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
1910 |
lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)"
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
1911 |
using nonempty_Basis box01_nonempty box_eq_empty(1) box_ne_empty(1) by blast |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
1912 |
|
|
66089
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1913 |
lemma interval_subset_is_interval: |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1914 |
assumes "is_interval S" |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1915 |
shows "cbox a b \<subseteq> S \<longleftrightarrow> cbox a b = {} \<or> a \<in> S \<and> b \<in> S" (is "?lhs = ?rhs")
|
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1916 |
proof |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1917 |
assume ?lhs |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1918 |
then show ?rhs using box_ne_empty(1) mem_box(2) by fastforce |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1919 |
next |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1920 |
assume ?rhs |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1921 |
have "cbox a b \<subseteq> S" if "a \<in> S" "b \<in> S" |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1922 |
using assms unfolding is_interval_def |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1923 |
apply (clarsimp simp add: mem_box) |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1924 |
using that by blast |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1925 |
with \<open>?rhs\<close> show ?lhs |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1926 |
by blast |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1927 |
qed |
|
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents:
65587
diff
changeset
|
1928 |
|
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1929 |
lemma is_real_interval_union: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1930 |
"is_interval (X \<union> Y)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1931 |
if X: "is_interval X" and Y: "is_interval Y" and I: "(X \<noteq> {} \<Longrightarrow> Y \<noteq> {} \<Longrightarrow> X \<inter> Y \<noteq> {})"
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1932 |
for X Y::"real set" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1933 |
proof - |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1934 |
consider "X \<noteq> {}" "Y \<noteq> {}" | "X = {}" | "Y = {}" by blast
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1935 |
then show ?thesis |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1936 |
proof cases |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1937 |
case 1 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1938 |
then obtain r where "r \<in> X \<or> X \<inter> Y = {}" "r \<in> Y \<or> X \<inter> Y = {}"
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1939 |
by blast |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1940 |
then show ?thesis |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1941 |
using I 1 X Y unfolding is_interval_1 |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1942 |
by (metis (full_types) Un_iff le_cases) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1943 |
qed (use that in auto) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1944 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1945 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1946 |
lemma is_interval_translationI: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1947 |
assumes "is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1948 |
shows "is_interval ((+) x ` X)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1949 |
unfolding is_interval_def |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1950 |
proof safe |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1951 |
fix b d e |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1952 |
assume "b \<in> X" "d \<in> X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1953 |
"\<forall>i\<in>Basis. (x + b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (x + d) \<bullet> i \<or> |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1954 |
(x + d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (x + b) \<bullet> i" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1955 |
hence "e - x \<in> X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1956 |
by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "e - x"]) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1957 |
(auto simp: algebra_simps) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1958 |
thus "e \<in> (+) x ` X" by force |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1959 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1960 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1961 |
lemma is_interval_uminusI: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1962 |
assumes "is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1963 |
shows "is_interval (uminus ` X)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1964 |
unfolding is_interval_def |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1965 |
proof safe |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1966 |
fix b d e |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1967 |
assume "b \<in> X" "d \<in> X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1968 |
"\<forall>i\<in>Basis. (- b) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- d) \<bullet> i \<or> |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1969 |
(- d) \<bullet> i \<le> e \<bullet> i \<and> e \<bullet> i \<le> (- b) \<bullet> i" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1970 |
hence "- e \<in> X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1971 |
by (intro mem_is_intervalI[OF assms \<open>b \<in> X\<close> \<open>d \<in> X\<close>, of "- e"]) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1972 |
(auto simp: algebra_simps) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1973 |
thus "e \<in> uminus ` X" by force |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1974 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1975 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1976 |
lemma is_interval_uminus[simp]: "is_interval (uminus ` x) = is_interval x" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1977 |
using is_interval_uminusI[of x] is_interval_uminusI[of "uminus ` x"] |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1978 |
by (auto simp: image_image) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1979 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1980 |
lemma is_interval_neg_translationI: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1981 |
assumes "is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1982 |
shows "is_interval ((-) x ` X)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1983 |
proof - |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1984 |
have "(-) x ` X = (+) x ` uminus ` X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1985 |
by (force simp: algebra_simps) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1986 |
also have "is_interval \<dots>" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1987 |
by (metis is_interval_uminusI is_interval_translationI assms) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1988 |
finally show ?thesis . |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1989 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1990 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1991 |
lemma is_interval_translation[simp]: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1992 |
"is_interval ((+) x ` X) = is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1993 |
using is_interval_neg_translationI[of "(+) x ` X" x] |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1994 |
by (auto intro!: is_interval_translationI simp: image_image) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1995 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1996 |
lemma is_interval_minus_translation[simp]: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1997 |
shows "is_interval ((-) x ` X) = is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1998 |
proof - |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
1999 |
have "(-) x ` X = (+) x ` uminus ` X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2000 |
by (force simp: algebra_simps) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2001 |
also have "is_interval \<dots> = is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2002 |
by simp |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2003 |
finally show ?thesis . |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2004 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2005 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2006 |
lemma is_interval_minus_translation'[simp]: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2007 |
shows "is_interval ((\<lambda>x. x - c) ` X) = is_interval X" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2008 |
using is_interval_translation[of "-c" X] |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2009 |
by (metis image_cong uminus_add_conv_diff) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2010 |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2011 |
|
| 64539 | 2012 |
subsection \<open>Limit points\<close> |
| 33175 | 2013 |
|
| 67962 | 2014 |
definition%important (in topological_space) islimpt:: "'a \<Rightarrow> 'a set \<Rightarrow> bool" (infixr "islimpt" 60) |
| 53255 | 2015 |
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 | 2016 |
|
2017 |
lemma islimptI: |
|
2018 |
assumes "\<And>T. x \<in> T \<Longrightarrow> open T \<Longrightarrow> \<exists>y\<in>S. y \<in> T \<and> y \<noteq> x" |
|
2019 |
shows "x islimpt S" |
|
2020 |
using assms unfolding islimpt_def by auto |
|
2021 |
||
2022 |
lemma islimptE: |
|
2023 |
assumes "x islimpt S" and "x \<in> T" and "open T" |
|
2024 |
obtains y where "y \<in> S" and "y \<in> T" and "y \<noteq> x" |
|
2025 |
using assms unfolding islimpt_def by auto |
|
2026 |
||
| 44584 | 2027 |
lemma islimpt_iff_eventually: "x islimpt S \<longleftrightarrow> \<not> eventually (\<lambda>y. y \<notin> S) (at x)" |
2028 |
unfolding islimpt_def eventually_at_topological by auto |
|
2029 |
||
| 53255 | 2030 |
lemma islimpt_subset: "x islimpt S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> x islimpt T" |
| 44584 | 2031 |
unfolding islimpt_def by fast |
| 33175 | 2032 |
|
2033 |
lemma islimpt_approachable: |
|
2034 |
fixes x :: "'a::metric_space" |
|
2035 |
shows "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e)" |
|
| 44584 | 2036 |
unfolding islimpt_iff_eventually eventually_at by fast |
| 33175 | 2037 |
|
| 64539 | 2038 |
lemma islimpt_approachable_le: "x islimpt S \<longleftrightarrow> (\<forall>e>0. \<exists>x'\<in> S. x' \<noteq> x \<and> dist x' x \<le> e)" |
2039 |
for x :: "'a::metric_space" |
|
| 33175 | 2040 |
unfolding islimpt_approachable |
| 44584 | 2041 |
using approachable_lt_le [where f="\<lambda>y. dist y x" and P="\<lambda>y. y \<notin> S \<or> y = x", |
2042 |
THEN arg_cong [where f=Not]] |
|
2043 |
by (simp add: Bex_def conj_commute conj_left_commute) |
|
| 33175 | 2044 |
|
| 44571 | 2045 |
lemma islimpt_UNIV_iff: "x islimpt UNIV \<longleftrightarrow> \<not> open {x}"
|
2046 |
unfolding islimpt_def by (safe, fast, case_tac "T = {x}", fast, fast)
|
|
2047 |
||
| 51351 | 2048 |
lemma islimpt_punctured: "x islimpt S = x islimpt (S-{x})"
|
2049 |
unfolding islimpt_def by blast |
|
2050 |
||
| 60420 | 2051 |
text \<open>A perfect space has no isolated points.\<close> |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2052 |
|
| 64539 | 2053 |
lemma islimpt_UNIV [simp, intro]: "x islimpt UNIV" |
2054 |
for x :: "'a::perfect_space" |
|
| 44571 | 2055 |
unfolding islimpt_UNIV_iff by (rule not_open_singleton) |
| 33175 | 2056 |
|
| 64539 | 2057 |
lemma perfect_choose_dist: "0 < r \<Longrightarrow> \<exists>a. a \<noteq> x \<and> dist a x < r" |
2058 |
for x :: "'a::{perfect_space,metric_space}"
|
|
2059 |
using islimpt_UNIV [of x] by (simp add: islimpt_approachable) |
|
| 33175 | 2060 |
|
2061 |
lemma closed_limpt: "closed S \<longleftrightarrow> (\<forall>x. x islimpt S \<longrightarrow> x \<in> S)" |
|
2062 |
unfolding closed_def |
|
2063 |
apply (subst open_subopen) |
|
| 34105 | 2064 |
apply (simp add: islimpt_def subset_eq) |
| 52624 | 2065 |
apply (metis ComplE ComplI) |
2066 |
done |
|
| 33175 | 2067 |
|
2068 |
lemma islimpt_EMPTY[simp]: "\<not> x islimpt {}"
|
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2069 |
by (auto simp: islimpt_def) |
| 33175 | 2070 |
|
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2071 |
lemma finite_ball_include: |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2072 |
fixes a :: "'a::metric_space" |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2073 |
assumes "finite S" |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2074 |
shows "\<exists>e>0. S \<subseteq> ball a e" |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2075 |
using assms |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2076 |
proof induction |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2077 |
case (insert x S) |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2078 |
then obtain e0 where "e0>0" and e0:"S \<subseteq> ball a e0" by auto |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2079 |
define e where "e = max e0 (2 * dist a x)" |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2080 |
have "e>0" unfolding e_def using \<open>e0>0\<close> by auto |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2081 |
moreover have "insert x S \<subseteq> ball a e" |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2082 |
using e0 \<open>e>0\<close> unfolding e_def by auto |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2083 |
ultimately show ?case by auto |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2084 |
qed (auto intro: zero_less_one) |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2085 |
|
| 33175 | 2086 |
lemma finite_set_avoid: |
2087 |
fixes a :: "'a::metric_space" |
|
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2088 |
assumes "finite S" |
| 64539 | 2089 |
shows "\<exists>d>0. \<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x" |
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2090 |
using assms |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2091 |
proof induction |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2092 |
case (insert x S) |
|
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2093 |
then obtain d where "d > 0" and d: "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> d \<le> dist a x" |
| 53255 | 2094 |
by blast |
2095 |
show ?case |
|
2096 |
proof (cases "x = a") |
|
2097 |
case True |
|
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2098 |
with \<open>d > 0 \<close>d show ?thesis by auto |
| 53255 | 2099 |
next |
2100 |
case False |
|
| 33175 | 2101 |
let ?d = "min d (dist a x)" |
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2102 |
from False \<open>d > 0\<close> have dp: "?d > 0" |
| 64539 | 2103 |
by auto |
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2104 |
from d have d': "\<forall>x\<in>S. x \<noteq> a \<longrightarrow> ?d \<le> dist a x" |
| 53255 | 2105 |
by auto |
2106 |
with dp False show ?thesis |
|
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2107 |
by (metis insert_iff le_less min_less_iff_conj not_less) |
| 53255 | 2108 |
qed |
|
68532
f8b98d31ad45
Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents:
68188
diff
changeset
|
2109 |
qed (auto intro: zero_less_one) |
| 33175 | 2110 |
|
2111 |
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
|
2112 |
by (simp add: islimpt_iff_eventually eventually_conj_iff) |
| 33175 | 2113 |
|
2114 |
lemma discrete_imp_closed: |
|
2115 |
fixes S :: "'a::metric_space set" |
|
| 53255 | 2116 |
assumes e: "0 < e" |
2117 |
and d: "\<forall>x \<in> S. \<forall>y \<in> S. dist y x < e \<longrightarrow> y = x" |
|
| 33175 | 2118 |
shows "closed S" |
| 53255 | 2119 |
proof - |
| 68120 | 2120 |
have False if C: "\<And>e. e>0 \<Longrightarrow> \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" for x |
| 64539 | 2121 |
proof - |
| 33175 | 2122 |
from e have e2: "e/2 > 0" by arith |
| 53282 | 2123 |
from C[rule_format, OF e2] obtain y where y: "y \<in> S" "y \<noteq> x" "dist y x < e/2" |
| 53255 | 2124 |
by blast |
| 33175 | 2125 |
let ?m = "min (e/2) (dist x y) " |
| 53255 | 2126 |
from e2 y(2) have mp: "?m > 0" |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
2127 |
by simp |
| 68120 | 2128 |
from C[OF mp] obtain z where z: "z \<in> S" "z \<noteq> x" "dist z x < ?m" |
| 53255 | 2129 |
by blast |
| 64539 | 2130 |
from z y have "dist z y < e" |
2131 |
by (intro dist_triangle_lt [where z=x]) simp |
|
2132 |
from d[rule_format, OF y(1) z(1) this] y z show ?thesis |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2133 |
by (auto simp: dist_commute) |
| 64539 | 2134 |
qed |
| 53255 | 2135 |
then show ?thesis |
2136 |
by (metis islimpt_approachable closed_limpt [where 'a='a]) |
|
| 33175 | 2137 |
qed |
2138 |
||
| 64539 | 2139 |
lemma closed_of_nat_image: "closed (of_nat ` A :: 'a::real_normed_algebra_1 set)" |
|
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2140 |
by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_nat) |
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2141 |
|
| 64539 | 2142 |
lemma closed_of_int_image: "closed (of_int ` A :: 'a::real_normed_algebra_1 set)" |
|
61524
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2143 |
by (rule discrete_imp_closed[of 1]) (auto simp: dist_of_int) |
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2144 |
|
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2145 |
lemma closed_Nats [simp]: "closed (\<nat> :: 'a :: real_normed_algebra_1 set)" |
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2146 |
unfolding Nats_def by (rule closed_of_nat_image) |
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2147 |
|
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2148 |
lemma closed_Ints [simp]: "closed (\<int> :: 'a :: real_normed_algebra_1 set)" |
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2149 |
unfolding Ints_def by (rule closed_of_int_image) |
|
f2e51e704a96
added many small lemmas about setsum/setprod/powr/...
eberlm
parents:
61518
diff
changeset
|
2150 |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2151 |
lemma closed_subset_Ints: |
|
66286
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2152 |
fixes A :: "'a :: real_normed_algebra_1 set" |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2153 |
assumes "A \<subseteq> \<int>" |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2154 |
shows "closed A" |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2155 |
proof (intro discrete_imp_closed[OF zero_less_one] ballI impI, goal_cases) |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2156 |
case (1 x y) |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2157 |
with assms have "x \<in> \<int>" and "y \<in> \<int>" by auto |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2158 |
with \<open>dist y x < 1\<close> show "y = x" |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2159 |
by (auto elim!: Ints_cases simp: dist_of_int) |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2160 |
qed |
|
1c977b13414f
poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents:
66164
diff
changeset
|
2161 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2162 |
|
| 60420 | 2163 |
subsection \<open>Interior of a Set\<close> |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2164 |
|
| 67962 | 2165 |
definition%important "interior S = \<Union>{T. open T \<and> T \<subseteq> S}"
|
| 44519 | 2166 |
|
2167 |
lemma interiorI [intro?]: |
|
2168 |
assumes "open T" and "x \<in> T" and "T \<subseteq> S" |
|
2169 |
shows "x \<in> interior S" |
|
2170 |
using assms unfolding interior_def by fast |
|
2171 |
||
2172 |
lemma interiorE [elim?]: |
|
2173 |
assumes "x \<in> interior S" |
|
2174 |
obtains T where "open T" and "x \<in> T" and "T \<subseteq> S" |
|
2175 |
using assms unfolding interior_def by fast |
|
2176 |
||
2177 |
lemma open_interior [simp, intro]: "open (interior S)" |
|
2178 |
by (simp add: interior_def open_Union) |
|
2179 |
||
2180 |
lemma interior_subset: "interior S \<subseteq> S" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2181 |
by (auto simp: interior_def) |
| 44519 | 2182 |
|
2183 |
lemma interior_maximal: "T \<subseteq> S \<Longrightarrow> open T \<Longrightarrow> T \<subseteq> interior S" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2184 |
by (auto simp: interior_def) |
| 44519 | 2185 |
|
2186 |
lemma interior_open: "open S \<Longrightarrow> interior S = S" |
|
2187 |
by (intro equalityI interior_subset interior_maximal subset_refl) |
|
| 33175 | 2188 |
|
2189 |
lemma interior_eq: "interior S = S \<longleftrightarrow> open S" |
|
| 44519 | 2190 |
by (metis open_interior interior_open) |
2191 |
||
2192 |
lemma open_subset_interior: "open S \<Longrightarrow> S \<subseteq> interior T \<longleftrightarrow> S \<subseteq> T" |
|
| 33175 | 2193 |
by (metis interior_maximal interior_subset subset_trans) |
2194 |
||
| 44519 | 2195 |
lemma interior_empty [simp]: "interior {} = {}"
|
2196 |
using open_empty by (rule interior_open) |
|
2197 |
||
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
2198 |
lemma interior_UNIV [simp]: "interior UNIV = UNIV" |
|
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
2199 |
using open_UNIV by (rule interior_open) |
|
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
2200 |
|
| 44519 | 2201 |
lemma interior_interior [simp]: "interior (interior S) = interior S" |
2202 |
using open_interior by (rule interior_open) |
|
2203 |
||
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
2204 |
lemma interior_mono: "S \<subseteq> T \<Longrightarrow> interior S \<subseteq> interior T" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2205 |
by (auto simp: interior_def) |
| 44519 | 2206 |
|
2207 |
lemma interior_unique: |
|
2208 |
assumes "T \<subseteq> S" and "open T" |
|
2209 |
assumes "\<And>T'. T' \<subseteq> S \<Longrightarrow> open T' \<Longrightarrow> T' \<subseteq> T" |
|
2210 |
shows "interior S = T" |
|
2211 |
by (intro equalityI assms interior_subset open_interior interior_maximal) |
|
2212 |
||
| 64539 | 2213 |
lemma interior_singleton [simp]: "interior {a} = {}"
|
2214 |
for a :: "'a::perfect_space" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2215 |
apply (rule interior_unique, simp_all) |
| 64539 | 2216 |
using not_open_singleton subset_singletonD |
2217 |
apply fastforce |
|
2218 |
done |
|
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
2219 |
|
|
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
2220 |
lemma interior_Int [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
|
2221 |
by (intro equalityI Int_mono Int_greatest interior_mono Int_lower1 |
| 44519 | 2222 |
Int_lower2 interior_maximal interior_subset open_Int open_interior) |
2223 |
||
2224 |
lemma mem_interior: "x \<in> interior S \<longleftrightarrow> (\<exists>e>0. ball x e \<subseteq> S)" |
|
2225 |
using open_contains_ball_eq [where S="interior S"] |
|
2226 |
by (simp add: open_subset_interior) |
|
| 33175 | 2227 |
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2228 |
lemma eventually_nhds_in_nhd: "x \<in> interior s \<Longrightarrow> eventually (\<lambda>y. y \<in> s) (nhds x)" |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2229 |
using interior_subset[of s] by (subst eventually_nhds) blast |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2230 |
|
| 33175 | 2231 |
lemma interior_limit_point [intro]: |
2232 |
fixes x :: "'a::perfect_space" |
|
| 53255 | 2233 |
assumes x: "x \<in> interior S" |
2234 |
shows "x islimpt S" |
|
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2235 |
using x islimpt_UNIV [of x] |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2236 |
unfolding interior_def islimpt_def |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2237 |
apply (clarsimp, rename_tac T T') |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2238 |
apply (drule_tac x="T \<inter> T'" in spec) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2239 |
apply (auto simp: open_Int) |
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2240 |
done |
| 33175 | 2241 |
|
2242 |
lemma interior_closed_Un_empty_interior: |
|
| 53255 | 2243 |
assumes cS: "closed S" |
2244 |
and iT: "interior T = {}"
|
|
| 44519 | 2245 |
shows "interior (S \<union> T) = interior S" |
| 33175 | 2246 |
proof |
| 44519 | 2247 |
show "interior S \<subseteq> interior (S \<union> T)" |
| 53255 | 2248 |
by (rule interior_mono) (rule Un_upper1) |
| 33175 | 2249 |
show "interior (S \<union> T) \<subseteq> interior S" |
2250 |
proof |
|
| 53255 | 2251 |
fix x |
2252 |
assume "x \<in> interior (S \<union> T)" |
|
| 44519 | 2253 |
then obtain R where "open R" "x \<in> R" "R \<subseteq> S \<union> T" .. |
| 33175 | 2254 |
show "x \<in> interior S" |
2255 |
proof (rule ccontr) |
|
2256 |
assume "x \<notin> interior S" |
|
| 60420 | 2257 |
with \<open>x \<in> R\<close> \<open>open R\<close> obtain y where "y \<in> R - S" |
| 44519 | 2258 |
unfolding interior_def by fast |
| 60420 | 2259 |
from \<open>open R\<close> \<open>closed S\<close> have "open (R - S)" |
| 53282 | 2260 |
by (rule open_Diff) |
| 60420 | 2261 |
from \<open>R \<subseteq> S \<union> T\<close> have "R - S \<subseteq> T" |
| 53282 | 2262 |
by fast |
| 60420 | 2263 |
from \<open>y \<in> R - S\<close> \<open>open (R - S)\<close> \<open>R - S \<subseteq> T\<close> \<open>interior T = {}\<close> show False
|
| 53282 | 2264 |
unfolding interior_def by fast |
| 33175 | 2265 |
qed |
2266 |
qed |
|
2267 |
qed |
|
2268 |
||
| 44365 | 2269 |
lemma interior_Times: "interior (A \<times> B) = interior A \<times> interior B" |
2270 |
proof (rule interior_unique) |
|
2271 |
show "interior A \<times> interior B \<subseteq> A \<times> B" |
|
2272 |
by (intro Sigma_mono interior_subset) |
|
2273 |
show "open (interior A \<times> interior B)" |
|
2274 |
by (intro open_Times open_interior) |
|
| 53255 | 2275 |
fix T |
2276 |
assume "T \<subseteq> A \<times> B" and "open T" |
|
2277 |
then show "T \<subseteq> interior A \<times> interior B" |
|
| 53282 | 2278 |
proof safe |
| 53255 | 2279 |
fix x y |
2280 |
assume "(x, y) \<in> T" |
|
| 44519 | 2281 |
then obtain C D where "open C" "open D" "C \<times> D \<subseteq> T" "x \<in> C" "y \<in> D" |
| 60420 | 2282 |
using \<open>open T\<close> unfolding open_prod_def by fast |
| 53255 | 2283 |
then have "open C" "open D" "C \<subseteq> A" "D \<subseteq> B" "x \<in> C" "y \<in> D" |
| 60420 | 2284 |
using \<open>T \<subseteq> A \<times> B\<close> by auto |
| 53255 | 2285 |
then show "x \<in> interior A" and "y \<in> interior B" |
| 44519 | 2286 |
by (auto intro: interiorI) |
2287 |
qed |
|
| 44365 | 2288 |
qed |
2289 |
||
| 61245 | 2290 |
lemma interior_Ici: |
| 64539 | 2291 |
fixes x :: "'a :: {dense_linorder,linorder_topology}"
|
| 61245 | 2292 |
assumes "b < x" |
| 64539 | 2293 |
shows "interior {x ..} = {x <..}"
|
| 61245 | 2294 |
proof (rule interior_unique) |
| 64539 | 2295 |
fix T |
2296 |
assume "T \<subseteq> {x ..}" "open T"
|
|
| 61245 | 2297 |
moreover have "x \<notin> T" |
2298 |
proof |
|
2299 |
assume "x \<in> T" |
|
2300 |
obtain y where "y < x" "{y <.. x} \<subseteq> T"
|
|
2301 |
using open_left[OF \<open>open T\<close> \<open>x \<in> T\<close> \<open>b < x\<close>] by auto |
|
2302 |
with dense[OF \<open>y < x\<close>] obtain z where "z \<in> T" "z < x" |
|
2303 |
by (auto simp: subset_eq Ball_def) |
|
2304 |
with \<open>T \<subseteq> {x ..}\<close> show False by auto
|
|
2305 |
qed |
|
2306 |
ultimately show "T \<subseteq> {x <..}"
|
|
2307 |
by (auto simp: subset_eq less_le) |
|
2308 |
qed auto |
|
2309 |
||
2310 |
lemma interior_Iic: |
|
| 64539 | 2311 |
fixes x :: "'a ::{dense_linorder,linorder_topology}"
|
| 61245 | 2312 |
assumes "x < b" |
2313 |
shows "interior {.. x} = {..< x}"
|
|
2314 |
proof (rule interior_unique) |
|
| 64539 | 2315 |
fix T |
2316 |
assume "T \<subseteq> {.. x}" "open T"
|
|
| 61245 | 2317 |
moreover have "x \<notin> T" |
2318 |
proof |
|
2319 |
assume "x \<in> T" |
|
2320 |
obtain y where "x < y" "{x ..< y} \<subseteq> T"
|
|
2321 |
using open_right[OF \<open>open T\<close> \<open>x \<in> T\<close> \<open>x < b\<close>] by auto |
|
2322 |
with dense[OF \<open>x < y\<close>] obtain z where "z \<in> T" "x < z" |
|
2323 |
by (auto simp: subset_eq Ball_def less_le) |
|
2324 |
with \<open>T \<subseteq> {.. x}\<close> show False by auto
|
|
2325 |
qed |
|
2326 |
ultimately show "T \<subseteq> {..< x}"
|
|
2327 |
by (auto simp: subset_eq less_le) |
|
2328 |
qed auto |
|
| 33175 | 2329 |
|
| 64539 | 2330 |
|
| 60420 | 2331 |
subsection \<open>Closure of a Set\<close> |
| 33175 | 2332 |
|
| 67962 | 2333 |
definition%important "closure S = S \<union> {x | x. x islimpt S}"
|
| 33175 | 2334 |
|
| 44518 | 2335 |
lemma interior_closure: "interior S = - (closure (- S))" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2336 |
by (auto simp: interior_def closure_def islimpt_def) |
| 44518 | 2337 |
|
| 34105 | 2338 |
lemma closure_interior: "closure S = - interior (- S)" |
| 64539 | 2339 |
by (simp add: interior_closure) |
| 33175 | 2340 |
|
2341 |
lemma closed_closure[simp, intro]: "closed (closure S)" |
|
| 64539 | 2342 |
by (simp add: closure_interior closed_Compl) |
| 44518 | 2343 |
|
2344 |
lemma closure_subset: "S \<subseteq> closure S" |
|
| 64539 | 2345 |
by (simp add: closure_def) |
| 33175 | 2346 |
|
2347 |
lemma closure_hull: "closure S = closed hull S" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2348 |
by (auto simp: hull_def closure_interior interior_def) |
| 33175 | 2349 |
|
2350 |
lemma closure_eq: "closure S = S \<longleftrightarrow> closed S" |
|
| 44519 | 2351 |
unfolding closure_hull using closed_Inter by (rule hull_eq) |
2352 |
||
2353 |
lemma closure_closed [simp]: "closed S \<Longrightarrow> closure S = S" |
|
| 64539 | 2354 |
by (simp only: closure_eq) |
| 44519 | 2355 |
|
2356 |
lemma closure_closure [simp]: "closure (closure S) = closure S" |
|
| 44518 | 2357 |
unfolding closure_hull by (rule hull_hull) |
| 33175 | 2358 |
|
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
2359 |
lemma closure_mono: "S \<subseteq> T \<Longrightarrow> closure S \<subseteq> closure T" |
| 44518 | 2360 |
unfolding closure_hull by (rule hull_mono) |
| 33175 | 2361 |
|
| 44519 | 2362 |
lemma closure_minimal: "S \<subseteq> T \<Longrightarrow> closed T \<Longrightarrow> closure S \<subseteq> T" |
| 44518 | 2363 |
unfolding closure_hull by (rule hull_minimal) |
| 33175 | 2364 |
|
| 44519 | 2365 |
lemma closure_unique: |
| 53255 | 2366 |
assumes "S \<subseteq> T" |
2367 |
and "closed T" |
|
2368 |
and "\<And>T'. S \<subseteq> T' \<Longrightarrow> closed T' \<Longrightarrow> T \<subseteq> T'" |
|
| 44519 | 2369 |
shows "closure S = T" |
2370 |
using assms unfolding closure_hull by (rule hull_unique) |
|
2371 |
||
2372 |
lemma closure_empty [simp]: "closure {} = {}"
|
|
| 44518 | 2373 |
using closed_empty by (rule closure_closed) |
| 33175 | 2374 |
|
|
44522
2f7e9d890efe
rename subset_{interior,closure} to {interior,closure}_mono;
huffman
parents:
44519
diff
changeset
|
2375 |
lemma closure_UNIV [simp]: "closure UNIV = UNIV" |
| 44518 | 2376 |
using closed_UNIV by (rule closure_closed) |
2377 |
||
| 64122 | 2378 |
lemma closure_Un [simp]: "closure (S \<union> T) = closure S \<union> closure T" |
| 64539 | 2379 |
by (simp add: closure_interior) |
| 33175 | 2380 |
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
2381 |
lemma closure_eq_empty [iff]: "closure S = {} \<longleftrightarrow> S = {}"
|
| 64539 | 2382 |
using closure_empty closure_subset[of S] by blast |
| 33175 | 2383 |
|
2384 |
lemma closure_subset_eq: "closure S \<subseteq> S \<longleftrightarrow> closed S" |
|
| 64539 | 2385 |
using closure_eq[of S] closure_subset[of S] by simp |
2386 |
||
2387 |
lemma open_Int_closure_eq_empty: "open S \<Longrightarrow> (S \<inter> closure T) = {} \<longleftrightarrow> S \<inter> T = {}"
|
|
| 34105 | 2388 |
using open_subset_interior[of S "- T"] |
2389 |
using interior_subset[of "- T"] |
|
| 64539 | 2390 |
by (auto simp: closure_interior) |
2391 |
||
2392 |
lemma open_Int_closure_subset: "open S \<Longrightarrow> S \<inter> closure T \<subseteq> closure (S \<inter> T)" |
|
| 33175 | 2393 |
proof |
2394 |
fix x |
|
| 64539 | 2395 |
assume *: "open S" "x \<in> S \<inter> closure T" |
2396 |
have "x islimpt (S \<inter> T)" if **: "x islimpt T" |
|
2397 |
proof (rule islimptI) |
|
2398 |
fix A |
|
2399 |
assume "x \<in> A" "open A" |
|
2400 |
with * have "x \<in> A \<inter> S" "open (A \<inter> S)" |
|
2401 |
by (simp_all add: open_Int) |
|
2402 |
with ** obtain y where "y \<in> T" "y \<in> A \<inter> S" "y \<noteq> x" |
|
2403 |
by (rule islimptE) |
|
2404 |
then have "y \<in> S \<inter> T" "y \<in> A \<and> y \<noteq> x" |
|
2405 |
by simp_all |
|
2406 |
then show "\<exists>y\<in>(S \<inter> T). y \<in> A \<and> y \<noteq> x" .. |
|
2407 |
qed |
|
2408 |
with * show "x \<in> closure (S \<inter> T)" |
|
2409 |
unfolding closure_def by blast |
|
| 33175 | 2410 |
qed |
2411 |
||
| 44519 | 2412 |
lemma closure_complement: "closure (- S) = - interior S" |
| 64539 | 2413 |
by (simp add: closure_interior) |
| 33175 | 2414 |
|
| 44519 | 2415 |
lemma interior_complement: "interior (- S) = - closure S" |
| 64539 | 2416 |
by (simp add: closure_interior) |
| 64910 | 2417 |
|
| 64845 | 2418 |
lemma interior_diff: "interior(S - T) = interior S - closure T" |
2419 |
by (simp add: Diff_eq interior_complement) |
|
| 33175 | 2420 |
|
| 44365 | 2421 |
lemma closure_Times: "closure (A \<times> B) = closure A \<times> closure B" |
| 44519 | 2422 |
proof (rule closure_unique) |
| 44365 | 2423 |
show "A \<times> B \<subseteq> closure A \<times> closure B" |
2424 |
by (intro Sigma_mono closure_subset) |
|
2425 |
show "closed (closure A \<times> closure B)" |
|
2426 |
by (intro closed_Times closed_closure) |
|
| 53282 | 2427 |
fix T |
2428 |
assume "A \<times> B \<subseteq> T" and "closed T" |
|
2429 |
then show "closure A \<times> closure B \<subseteq> T" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2430 |
apply (simp add: closed_def open_prod_def, clarify) |
| 44365 | 2431 |
apply (rule ccontr) |
2432 |
apply (drule_tac x="(a, b)" in bspec, simp, clarify, rename_tac C D) |
|
2433 |
apply (simp add: closure_interior interior_def) |
|
2434 |
apply (drule_tac x=C in spec) |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2435 |
apply (drule_tac x=D in spec, auto) |
| 44365 | 2436 |
done |
2437 |
qed |
|
2438 |
||
|
66641
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2439 |
lemma closure_openin_Int_closure: |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2440 |
assumes ope: "openin (subtopology euclidean U) S" and "T \<subseteq> U" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2441 |
shows "closure(S \<inter> closure T) = closure(S \<inter> T)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2442 |
proof |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2443 |
obtain V where "open V" and S: "S = U \<inter> V" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2444 |
using ope using openin_open by metis |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2445 |
show "closure (S \<inter> closure T) \<subseteq> closure (S \<inter> T)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2446 |
proof (clarsimp simp: S) |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2447 |
fix x |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2448 |
assume "x \<in> closure (U \<inter> V \<inter> closure T)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2449 |
then have "V \<inter> closure T \<subseteq> A \<Longrightarrow> x \<in> closure A" for A |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2450 |
by (metis closure_mono subsetD inf.coboundedI2 inf_assoc) |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2451 |
then have "x \<in> closure (T \<inter> V)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2452 |
by (metis \<open>open V\<close> closure_closure inf_commute open_Int_closure_subset) |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2453 |
then show "x \<in> closure (U \<inter> V \<inter> T)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2454 |
by (metis \<open>T \<subseteq> U\<close> inf.absorb_iff2 inf_assoc inf_commute) |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2455 |
qed |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2456 |
next |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2457 |
show "closure (S \<inter> T) \<subseteq> closure (S \<inter> closure T)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2458 |
by (meson Int_mono closure_mono closure_subset order_refl) |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2459 |
qed |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
2460 |
|
| 67613 | 2461 |
lemma islimpt_in_closure: "(x islimpt S) = (x\<in>closure(S-{x}))"
|
| 51351 | 2462 |
unfolding closure_def using islimpt_punctured by blast |
2463 |
||
| 63301 | 2464 |
lemma connected_imp_connected_closure: "connected S \<Longrightarrow> connected (closure S)" |
| 64539 | 2465 |
by (rule connectedI) (meson closure_subset open_Int open_Int_closure_eq_empty subset_trans connectedD) |
2466 |
||
2467 |
lemma limpt_of_limpts: "x islimpt {y. y islimpt S} \<Longrightarrow> x islimpt S"
|
|
2468 |
for x :: "'a::metric_space" |
|
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2469 |
apply (clarsimp simp add: islimpt_approachable) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2470 |
apply (drule_tac x="e/2" in spec) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2471 |
apply (auto simp: simp del: less_divide_eq_numeral1) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2472 |
apply (drule_tac x="dist x' x" in spec) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2473 |
apply (auto simp: zero_less_dist_iff simp del: less_divide_eq_numeral1) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2474 |
apply (erule rev_bexI) |
| 64539 | 2475 |
apply (metis dist_commute dist_triangle_half_r less_trans less_irrefl) |
2476 |
done |
|
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2477 |
|
| 63301 | 2478 |
lemma closed_limpts: "closed {x::'a::metric_space. x islimpt S}"
|
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2479 |
using closed_limpt limpt_of_limpts by blast |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2480 |
|
| 64539 | 2481 |
lemma limpt_of_closure: "x islimpt closure S \<longleftrightarrow> x islimpt S" |
2482 |
for x :: "'a::metric_space" |
|
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2483 |
by (auto simp: closure_def islimpt_Un dest: limpt_of_limpts) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2484 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
2485 |
lemma closedin_limpt: |
| 64539 | 2486 |
"closedin (subtopology euclidean T) S \<longleftrightarrow> S \<subseteq> T \<and> (\<forall>x. x islimpt S \<and> x \<in> T \<longrightarrow> x \<in> S)" |
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2487 |
apply (simp add: closedin_closed, safe) |
| 64539 | 2488 |
apply (simp add: closed_limpt islimpt_subset) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2489 |
apply (rule_tac x="closure S" in exI, simp) |
|
61306
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2490 |
apply (force simp: closure_def) |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2491 |
done |
|
9dd394c866fc
New theorems about connected sets. And pairwise moved to Set.thy.
paulson <lp15@cam.ac.uk>
parents:
61284
diff
changeset
|
2492 |
|
| 64539 | 2493 |
lemma closedin_closed_eq: "closed S \<Longrightarrow> closedin (subtopology euclidean S) T \<longleftrightarrow> closed T \<and> T \<subseteq> S" |
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
2494 |
by (meson closedin_limpt closed_subset closedin_closed_trans) |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
2495 |
|
|
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
2496 |
lemma connected_closed_set: |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
2497 |
"closed S |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
2498 |
\<Longrightarrow> connected S \<longleftrightarrow> (\<nexists>A B. closed A \<and> closed B \<and> A \<noteq> {} \<and> B \<noteq> {} \<and> A \<union> B = S \<and> A \<inter> B = {})"
|
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
2499 |
unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
2500 |
|
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2501 |
text \<open>If a connnected set is written as the union of two nonempty closed sets, then these sets |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2502 |
have to intersect.\<close> |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2503 |
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2504 |
lemma connected_as_closed_union: |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2505 |
assumes "connected C" "C = A \<union> B" "closed A" "closed B" "A \<noteq> {}" "B \<noteq> {}"
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2506 |
shows "A \<inter> B \<noteq> {}"
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2507 |
by (metis assms closed_Un connected_closed_set) |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
2508 |
|
| 63301 | 2509 |
lemma closedin_subset_trans: |
| 64539 | 2510 |
"closedin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow> |
2511 |
closedin (subtopology euclidean T) S" |
|
2512 |
by (meson closedin_limpt subset_iff) |
|
| 63301 | 2513 |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
2514 |
lemma openin_subset_trans: |
| 64539 | 2515 |
"openin (subtopology euclidean U) S \<Longrightarrow> S \<subseteq> T \<Longrightarrow> T \<subseteq> U \<Longrightarrow> |
2516 |
openin (subtopology euclidean T) S" |
|
|
63469
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
2517 |
by (auto simp: openin_open) |
|
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
paulson <lp15@cam.ac.uk>
parents:
63332
diff
changeset
|
2518 |
|
| 64122 | 2519 |
lemma openin_Times: |
| 64539 | 2520 |
"openin (subtopology euclidean S) S' \<Longrightarrow> openin (subtopology euclidean T) T' \<Longrightarrow> |
2521 |
openin (subtopology euclidean (S \<times> T)) (S' \<times> T')" |
|
| 64122 | 2522 |
unfolding openin_open using open_Times by blast |
2523 |
||
2524 |
lemma Times_in_interior_subtopology: |
|
| 64539 | 2525 |
fixes U :: "('a::metric_space \<times> 'b::metric_space) set"
|
| 64122 | 2526 |
assumes "(x, y) \<in> U" "openin (subtopology euclidean (S \<times> T)) U" |
2527 |
obtains V W where "openin (subtopology euclidean S) V" "x \<in> V" |
|
2528 |
"openin (subtopology euclidean T) W" "y \<in> W" "(V \<times> W) \<subseteq> U" |
|
2529 |
proof - |
|
2530 |
from assms obtain e where "e > 0" and "U \<subseteq> S \<times> T" |
|
| 64539 | 2531 |
and e: "\<And>x' y'. \<lbrakk>x'\<in>S; y'\<in>T; dist (x', y') (x, y) < e\<rbrakk> \<Longrightarrow> (x', y') \<in> U" |
| 64122 | 2532 |
by (force simp: openin_euclidean_subtopology_iff) |
2533 |
with assms have "x \<in> S" "y \<in> T" |
|
2534 |
by auto |
|
2535 |
show ?thesis |
|
2536 |
proof |
|
2537 |
show "openin (subtopology euclidean S) (ball x (e/2) \<inter> S)" |
|
2538 |
by (simp add: Int_commute openin_open_Int) |
|
2539 |
show "x \<in> ball x (e / 2) \<inter> S" |
|
2540 |
by (simp add: \<open>0 < e\<close> \<open>x \<in> S\<close>) |
|
2541 |
show "openin (subtopology euclidean T) (ball y (e/2) \<inter> T)" |
|
2542 |
by (simp add: Int_commute openin_open_Int) |
|
2543 |
show "y \<in> ball y (e / 2) \<inter> T" |
|
2544 |
by (simp add: \<open>0 < e\<close> \<open>y \<in> T\<close>) |
|
2545 |
show "(ball x (e / 2) \<inter> S) \<times> (ball y (e / 2) \<inter> T) \<subseteq> U" |
|
2546 |
by clarify (simp add: e dist_Pair_Pair \<open>0 < e\<close> dist_commute sqrt_sum_squares_half_less) |
|
2547 |
qed |
|
2548 |
qed |
|
2549 |
||
2550 |
lemma openin_Times_eq: |
|
2551 |
fixes S :: "'a::metric_space set" and T :: "'b::metric_space set" |
|
2552 |
shows |
|
| 64539 | 2553 |
"openin (subtopology euclidean (S \<times> T)) (S' \<times> T') \<longleftrightarrow> |
2554 |
S' = {} \<or> T' = {} \<or> openin (subtopology euclidean S) S' \<and> openin (subtopology euclidean T) T'"
|
|
| 64122 | 2555 |
(is "?lhs = ?rhs") |
2556 |
proof (cases "S' = {} \<or> T' = {}")
|
|
2557 |
case True |
|
2558 |
then show ?thesis by auto |
|
2559 |
next |
|
2560 |
case False |
|
2561 |
then obtain x y where "x \<in> S'" "y \<in> T'" |
|
2562 |
by blast |
|
2563 |
show ?thesis |
|
2564 |
proof |
|
| 64539 | 2565 |
assume ?lhs |
| 64122 | 2566 |
have "openin (subtopology euclidean S) S'" |
2567 |
apply (subst openin_subopen, clarify) |
|
| 64539 | 2568 |
apply (rule Times_in_interior_subtopology [OF _ \<open>?lhs\<close>]) |
2569 |
using \<open>y \<in> T'\<close> |
|
2570 |
apply auto |
|
2571 |
done |
|
| 64122 | 2572 |
moreover have "openin (subtopology euclidean T) T'" |
2573 |
apply (subst openin_subopen, clarify) |
|
| 64539 | 2574 |
apply (rule Times_in_interior_subtopology [OF _ \<open>?lhs\<close>]) |
2575 |
using \<open>x \<in> S'\<close> |
|
2576 |
apply auto |
|
2577 |
done |
|
| 64122 | 2578 |
ultimately show ?rhs |
2579 |
by simp |
|
2580 |
next |
|
2581 |
assume ?rhs |
|
2582 |
with False show ?lhs |
|
2583 |
by (simp add: openin_Times) |
|
2584 |
qed |
|
2585 |
qed |
|
2586 |
||
| 63301 | 2587 |
lemma closedin_Times: |
| 64539 | 2588 |
"closedin (subtopology euclidean S) S' \<Longrightarrow> closedin (subtopology euclidean T) T' \<Longrightarrow> |
2589 |
closedin (subtopology euclidean (S \<times> T)) (S' \<times> T')" |
|
2590 |
unfolding closedin_closed using closed_Times by blast |
|
| 63301 | 2591 |
|
| 62083 | 2592 |
lemma bdd_below_closure: |
2593 |
fixes A :: "real set" |
|
2594 |
assumes "bdd_below A" |
|
2595 |
shows "bdd_below (closure A)" |
|
2596 |
proof - |
|
| 64539 | 2597 |
from assms obtain m where "\<And>x. x \<in> A \<Longrightarrow> m \<le> x" |
2598 |
by (auto simp: bdd_below_def) |
|
2599 |
then have "A \<subseteq> {m..}" by auto
|
|
2600 |
then have "closure A \<subseteq> {m..}"
|
|
2601 |
using closed_real_atLeast by (rule closure_minimal) |
|
2602 |
then show ?thesis |
|
2603 |
by (auto simp: bdd_below_def) |
|
2604 |
qed |
|
2605 |
||
2606 |
||
2607 |
subsection \<open>Frontier (also known as boundary)\<close> |
|
| 33175 | 2608 |
|
| 67962 | 2609 |
definition%important "frontier S = closure S - interior S" |
| 33175 | 2610 |
|
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
2611 |
lemma frontier_closed [iff]: "closed (frontier S)" |
| 33175 | 2612 |
by (simp add: frontier_def closed_Diff) |
2613 |
||
| 64539 | 2614 |
lemma frontier_closures: "frontier S = closure S \<inter> closure (- S)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2615 |
by (auto simp: frontier_def interior_closure) |
| 33175 | 2616 |
|
|
66939
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2617 |
lemma frontier_Int: "frontier(S \<inter> T) = closure(S \<inter> T) \<inter> (frontier S \<union> frontier T)" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2618 |
proof - |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2619 |
have "closure (S \<inter> T) \<subseteq> closure S" "closure (S \<inter> T) \<subseteq> closure T" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2620 |
by (simp_all add: closure_mono) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2621 |
then show ?thesis |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2622 |
by (auto simp: frontier_closures) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2623 |
qed |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2624 |
|
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2625 |
lemma frontier_Int_subset: "frontier(S \<inter> T) \<subseteq> frontier S \<union> frontier T" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2626 |
by (auto simp: frontier_Int) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2627 |
|
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2628 |
lemma frontier_Int_closed: |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2629 |
assumes "closed S" "closed T" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2630 |
shows "frontier(S \<inter> T) = (frontier S \<inter> T) \<union> (S \<inter> frontier T)" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2631 |
proof - |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2632 |
have "closure (S \<inter> T) = T \<inter> S" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2633 |
using assms by (simp add: Int_commute closed_Int) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2634 |
moreover have "T \<inter> (closure S \<inter> closure (- S)) = frontier S \<inter> T" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2635 |
by (simp add: Int_commute frontier_closures) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2636 |
ultimately show ?thesis |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2637 |
by (simp add: Int_Un_distrib Int_assoc Int_left_commute assms frontier_closures) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2638 |
qed |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2639 |
|
| 33175 | 2640 |
lemma frontier_straddle: |
2641 |
fixes a :: "'a::metric_space" |
|
| 44909 | 2642 |
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))" |
2643 |
unfolding frontier_def closure_interior |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2644 |
by (auto simp: mem_interior subset_eq ball_def) |
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
2645 |
|
| 33175 | 2646 |
lemma frontier_subset_closed: "closed S \<Longrightarrow> frontier S \<subseteq> S" |
2647 |
by (metis frontier_def closure_closed Diff_subset) |
|
2648 |
||
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
2649 |
lemma frontier_empty [simp]: "frontier {} = {}"
|
|
36362
06475a1547cb
fix lots of looping simp calls and other warnings
huffman
parents:
36360
diff
changeset
|
2650 |
by (simp add: frontier_def) |
| 33175 | 2651 |
|
2652 |
lemma frontier_subset_eq: "frontier S \<subseteq> S \<longleftrightarrow> closed S" |
|
| 58757 | 2653 |
proof - |
| 53255 | 2654 |
{
|
2655 |
assume "frontier S \<subseteq> S" |
|
2656 |
then have "closure S \<subseteq> S" |
|
2657 |
using interior_subset unfolding frontier_def by auto |
|
2658 |
then have "closed S" |
|
2659 |
using closure_subset_eq by auto |
|
| 33175 | 2660 |
} |
| 53255 | 2661 |
then show ?thesis using frontier_subset_closed[of S] .. |
| 33175 | 2662 |
qed |
2663 |
||
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
2664 |
lemma frontier_complement [simp]: "frontier (- S) = frontier S" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2665 |
by (auto simp: frontier_def closure_complement interior_complement) |
| 33175 | 2666 |
|
|
66939
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2667 |
lemma frontier_Un_subset: "frontier(S \<union> T) \<subseteq> frontier S \<union> frontier T" |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2668 |
by (metis compl_sup frontier_Int_subset frontier_complement) |
|
04678058308f
New results in topology, mostly from HOL Light's moretop.ml
paulson <lp15@cam.ac.uk>
parents:
66884
diff
changeset
|
2669 |
|
| 33175 | 2670 |
lemma frontier_disjoint_eq: "frontier S \<inter> S = {} \<longleftrightarrow> open S"
|
| 34105 | 2671 |
using frontier_complement frontier_subset_eq[of "- S"] |
2672 |
unfolding open_closed by auto |
|
| 33175 | 2673 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2674 |
lemma frontier_UNIV [simp]: "frontier UNIV = {}"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2675 |
using frontier_complement frontier_empty by fastforce |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2676 |
|
|
64788
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2677 |
lemma frontier_interiors: "frontier s = - interior(s) - interior(-s)" |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2678 |
by (simp add: Int_commute frontier_def interior_closure) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2679 |
|
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2680 |
lemma frontier_interior_subset: "frontier(interior S) \<subseteq> frontier S" |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2681 |
by (simp add: Diff_mono frontier_interiors interior_mono interior_subset) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2682 |
|
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2683 |
lemma connected_Int_frontier: |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2684 |
"\<lbrakk>connected s; s \<inter> t \<noteq> {}; s - t \<noteq> {}\<rbrakk> \<Longrightarrow> (s \<inter> frontier t \<noteq> {})"
|
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2685 |
apply (simp add: frontier_interiors connected_openin, safe) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2686 |
apply (drule_tac x="s \<inter> interior t" in spec, safe) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2687 |
apply (drule_tac [2] x="s \<inter> interior (-t)" in spec) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2688 |
apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2689 |
done |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2690 |
|
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2691 |
lemma closure_Un_frontier: "closure S = S \<union> frontier S" |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2692 |
proof - |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2693 |
have "S \<union> interior S = S" |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2694 |
using interior_subset by auto |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2695 |
then show ?thesis |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2696 |
using closure_subset by (auto simp: frontier_def) |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2697 |
qed |
|
19f3d4af7a7d
New material about path connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
64773
diff
changeset
|
2698 |
|
| 58757 | 2699 |
|
| 67962 | 2700 |
subsection%unimportant \<open>Filters and the ``eventually true'' quantifier\<close> |
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44076
diff
changeset
|
2701 |
|
| 64539 | 2702 |
definition indirection :: "'a::real_normed_vector \<Rightarrow> 'a \<Rightarrow> 'a filter" (infixr "indirection" 70) |
| 52624 | 2703 |
where "a indirection v = at a within {b. \<exists>c\<ge>0. b - a = scaleR c v}"
|
| 33175 | 2704 |
|
| 60420 | 2705 |
text \<open>Identify Trivial limits, where we can't approach arbitrarily closely.\<close> |
| 33175 | 2706 |
|
| 52624 | 2707 |
lemma trivial_limit_within: "trivial_limit (at a within S) \<longleftrightarrow> \<not> a islimpt S" |
| 33175 | 2708 |
proof |
2709 |
assume "trivial_limit (at a within S)" |
|
| 53255 | 2710 |
then show "\<not> a islimpt S" |
| 33175 | 2711 |
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
|
2712 |
unfolding eventually_at_topological |
| 33175 | 2713 |
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
|
2714 |
apply (clarsimp simp add: set_eq_iff) |
| 33175 | 2715 |
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
|
2716 |
apply (clarsimp, drule_tac x=y in bspec, simp_all) |
| 33175 | 2717 |
done |
2718 |
next |
|
2719 |
assume "\<not> a islimpt S" |
|
| 53255 | 2720 |
then show "trivial_limit (at a within S)" |
| 55775 | 2721 |
unfolding trivial_limit_def eventually_at_topological islimpt_def |
2722 |
by metis |
|
| 33175 | 2723 |
qed |
2724 |
||
2725 |
lemma trivial_limit_at_iff: "trivial_limit (at a) \<longleftrightarrow> \<not> a islimpt UNIV" |
|
| 45031 | 2726 |
using trivial_limit_within [of a UNIV] by simp |
| 33175 | 2727 |
|
| 64539 | 2728 |
lemma trivial_limit_at: "\<not> trivial_limit (at a)" |
2729 |
for a :: "'a::perfect_space" |
|
| 44571 | 2730 |
by (rule at_neq_bot) |
| 33175 | 2731 |
|
2732 |
lemma trivial_limit_at_infinity: |
|
|
44081
730f7cced3a6
rename type 'a net to 'a filter, following standard mathematical terminology
huffman
parents:
44076
diff
changeset
|
2733 |
"\<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
|
2734 |
unfolding trivial_limit_def eventually_at_infinity |
|
246493d61204
define nets directly as filters, instead of as filter bases
huffman
parents:
36336
diff
changeset
|
2735 |
apply clarsimp |
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2736 |
apply (subgoal_tac "\<exists>x::'a. x \<noteq> 0", clarify) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2737 |
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
|
2738 |
apply (cut_tac islimpt_UNIV [of "0::'a", unfolded islimpt_def]) |
|
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
2739 |
apply (drule_tac x=UNIV in spec, simp) |
| 33175 | 2740 |
done |
2741 |
||
| 53640 | 2742 |
lemma not_trivial_limit_within: "\<not> trivial_limit (at x within S) = (x \<in> closure (S - {x}))"
|
| 64539 | 2743 |
using islimpt_in_closure by (metis trivial_limit_within) |
2744 |
||
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2745 |
lemma not_in_closure_trivial_limitI: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2746 |
"x \<notin> closure s \<Longrightarrow> trivial_limit (at x within s)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2747 |
using not_trivial_limit_within[of x s] |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2748 |
by safe (metis Diff_empty Diff_insert0 closure_subset contra_subsetD) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2749 |
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2750 |
lemma filterlim_at_within_closure_implies_filterlim: "filterlim f l (at x within s)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2751 |
if "x \<in> closure s \<Longrightarrow> filterlim f l (at x within s)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2752 |
by (metis bot.extremum filterlim_filtercomap filterlim_mono not_in_closure_trivial_limitI that) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
2753 |
|
| 64539 | 2754 |
lemma at_within_eq_bot_iff: "at c within A = bot \<longleftrightarrow> c \<notin> closure (A - {c})"
|
|
61531
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2755 |
using not_trivial_limit_within[of c A] by blast |
|
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents:
61524
diff
changeset
|
2756 |
|
| 60420 | 2757 |
text \<open>Some property holds "sufficiently close" to the limit point.\<close> |
| 33175 | 2758 |
|
2759 |
lemma trivial_limit_eventually: "trivial_limit net \<Longrightarrow> eventually P net" |
|
| 45031 | 2760 |
by simp |
| 33175 | 2761 |
|
2762 |
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
|
2763 |
by (simp add: filter_eq_iff) |
| 33175 | 2764 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2765 |
|
| 60420 | 2766 |
subsection \<open>Limits\<close> |
| 33175 | 2767 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
2768 |
proposition Lim: "(f \<longlongrightarrow> l) net \<longleftrightarrow> trivial_limit net \<or> (\<forall>e>0. eventually (\<lambda>x. dist (f x) l < e) net)" |
| 64539 | 2769 |
by (auto simp: tendsto_iff trivial_limit_eq) |
2770 |
||
2771 |
text \<open>Show that they yield usual definitions in the various cases.\<close> |
|
| 33175 | 2772 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
2773 |
proposition Lim_within_le: "(f \<longlongrightarrow> l)(at a within S) \<longleftrightarrow> |
| 53640 | 2774 |
(\<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)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2775 |
by (auto simp: tendsto_iff eventually_at_le) |
| 33175 | 2776 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
2777 |
proposition Lim_within: "(f \<longlongrightarrow> l) (at a within S) \<longleftrightarrow> |
| 53640 | 2778 |
(\<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)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2779 |
by (auto simp: tendsto_iff eventually_at) |
| 33175 | 2780 |
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2781 |
corollary Lim_withinI [intro?]: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2782 |
assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>d>0. \<forall>x \<in> S. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l \<le> e" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2783 |
shows "(f \<longlongrightarrow> l) (at a within S)" |
| 64539 | 2784 |
apply (simp add: Lim_within, clarify) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2785 |
apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) |
| 64539 | 2786 |
done |
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2787 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
2788 |
proposition Lim_at: "(f \<longlongrightarrow> l) (at a) \<longleftrightarrow> |
| 53640 | 2789 |
(\<forall>e >0. \<exists>d>0. \<forall>x. 0 < dist x a \<and> dist x a < d \<longrightarrow> dist (f x) l < e)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2790 |
by (auto simp: tendsto_iff eventually_at) |
| 33175 | 2791 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
2792 |
proposition Lim_at_infinity: "(f \<longlongrightarrow> l) at_infinity \<longleftrightarrow> (\<forall>e>0. \<exists>b. \<forall>x. norm x \<ge> b \<longrightarrow> dist (f x) l < e)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2793 |
by (auto simp: tendsto_iff eventually_at_infinity) |
| 33175 | 2794 |
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2795 |
corollary Lim_at_infinityI [intro?]: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2796 |
assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>B. \<forall>x. norm x \<ge> B \<longrightarrow> dist (f x) l \<le> e" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2797 |
shows "(f \<longlongrightarrow> l) at_infinity" |
| 64539 | 2798 |
apply (simp add: Lim_at_infinity, clarify) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
2799 |
apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) |
| 64539 | 2800 |
done |
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
2801 |
|
| 61973 | 2802 |
lemma Lim_eventually: "eventually (\<lambda>x. f x = l) net \<Longrightarrow> (f \<longlongrightarrow> l) net" |
| 64539 | 2803 |
by (rule topological_tendstoI) (auto elim: eventually_mono) |
| 33175 | 2804 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2805 |
lemma Lim_transform_within_set: |
| 63301 | 2806 |
fixes a :: "'a::metric_space" and l :: "'b::metric_space" |
2807 |
shows "\<lbrakk>(f \<longlongrightarrow> l) (at a within S); eventually (\<lambda>x. x \<in> S \<longleftrightarrow> x \<in> T) (at a)\<rbrakk> |
|
2808 |
\<Longrightarrow> (f \<longlongrightarrow> l) (at a within T)" |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2809 |
apply (clarsimp simp: eventually_at Lim_within) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2810 |
apply (drule_tac x=e in spec, clarify) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2811 |
apply (rename_tac k) |
| 63301 | 2812 |
apply (rule_tac x="min d k" in exI, simp) |
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2813 |
done |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2814 |
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2815 |
lemma Lim_transform_within_set_eq: |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2816 |
fixes a l :: "'a::real_normed_vector" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2817 |
shows "eventually (\<lambda>x. x \<in> s \<longleftrightarrow> x \<in> t) (at a) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2818 |
\<Longrightarrow> ((f \<longlongrightarrow> l) (at a within s) \<longleftrightarrow> (f \<longlongrightarrow> l) (at a within t))" |
| 64539 | 2819 |
by (force intro: Lim_transform_within_set elim: eventually_mono) |
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
2820 |
|
| 63301 | 2821 |
lemma Lim_transform_within_openin: |
2822 |
fixes a :: "'a::metric_space" |
|
2823 |
assumes f: "(f \<longlongrightarrow> l) (at a within T)" |
|
| 64539 | 2824 |
and "openin (subtopology euclidean T) S" "a \<in> S" |
2825 |
and eq: "\<And>x. \<lbrakk>x \<in> S; x \<noteq> a\<rbrakk> \<Longrightarrow> f x = g x" |
|
| 63301 | 2826 |
shows "(g \<longlongrightarrow> l) (at a within T)" |
2827 |
proof - |
|
2828 |
obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "ball a \<epsilon> \<inter> T \<subseteq> S" |
|
2829 |
using assms by (force simp: openin_contains_ball) |
|
2830 |
then have "a \<in> ball a \<epsilon>" |
|
| 64539 | 2831 |
by simp |
| 63301 | 2832 |
show ?thesis |
| 64539 | 2833 |
by (rule Lim_transform_within [OF f \<open>0 < \<epsilon>\<close> eq]) (use \<epsilon> in \<open>auto simp: dist_commute subset_iff\<close>) |
| 63301 | 2834 |
qed |
2835 |
||
2836 |
lemma continuous_transform_within_openin: |
|
2837 |
fixes a :: "'a::metric_space" |
|
2838 |
assumes "continuous (at a within T) f" |
|
| 64539 | 2839 |
and "openin (subtopology euclidean T) S" "a \<in> S" |
2840 |
and eq: "\<And>x. x \<in> S \<Longrightarrow> f x = g x" |
|
| 63301 | 2841 |
shows "continuous (at a within T) g" |
| 64539 | 2842 |
using assms by (simp add: Lim_transform_within_openin continuous_within) |
2843 |
||
2844 |
text \<open>The expected monotonicity property.\<close> |
|
| 33175 | 2845 |
|
| 53255 | 2846 |
lemma Lim_Un: |
| 61973 | 2847 |
assumes "(f \<longlongrightarrow> l) (at x within S)" "(f \<longlongrightarrow> l) (at x within T)" |
2848 |
shows "(f \<longlongrightarrow> l) (at x within (S \<union> T))" |
|
| 53860 | 2849 |
using assms unfolding at_within_union by (rule filterlim_sup) |
| 33175 | 2850 |
|
2851 |
lemma Lim_Un_univ: |
|
| 61973 | 2852 |
"(f \<longlongrightarrow> l) (at x within S) \<Longrightarrow> (f \<longlongrightarrow> l) (at x within T) \<Longrightarrow> |
2853 |
S \<union> T = UNIV \<Longrightarrow> (f \<longlongrightarrow> 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
|
2854 |
by (metis Lim_Un) |
| 33175 | 2855 |
|
| 64539 | 2856 |
text \<open>Interrelations between restricted and unrestricted limits.\<close> |
2857 |
||
2858 |
lemma Lim_at_imp_Lim_at_within: "(f \<longlongrightarrow> l) (at x) \<Longrightarrow> (f \<longlongrightarrow> l) (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
|
2859 |
by (metis order_refl filterlim_mono subset_UNIV at_le) |
| 33175 | 2860 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2861 |
lemma eventually_within_interior: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2862 |
assumes "x \<in> interior S" |
| 53255 | 2863 |
shows "eventually P (at x within S) \<longleftrightarrow> eventually P (at x)" |
2864 |
(is "?lhs = ?rhs") |
|
2865 |
proof |
|
| 44519 | 2866 |
from assms obtain T where T: "open T" "x \<in> T" "T \<subseteq> S" .. |
| 53255 | 2867 |
{
|
| 64539 | 2868 |
assume ?lhs |
| 53640 | 2869 |
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" |
| 64539 | 2870 |
by (auto simp: eventually_at_topological) |
| 53640 | 2871 |
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
|
2872 |
by auto |
| 64539 | 2873 |
then show ?rhs |
2874 |
by (auto simp: eventually_at_topological) |
|
| 53255 | 2875 |
next |
| 64539 | 2876 |
assume ?rhs |
2877 |
then show ?lhs |
|
| 61810 | 2878 |
by (auto elim: eventually_mono simp: eventually_at_filter) |
| 52624 | 2879 |
} |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2880 |
qed |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2881 |
|
| 64539 | 2882 |
lemma at_within_interior: "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
|
2883 |
unfolding filter_eq_iff by (intro allI eventually_within_interior) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
2884 |
|
| 43338 | 2885 |
lemma Lim_within_LIMSEQ: |
| 53862 | 2886 |
fixes a :: "'a::first_countable_topology" |
| 61969 | 2887 |
assumes "\<forall>S. (\<forall>n. S n \<noteq> a \<and> S n \<in> T) \<and> S \<longlonglongrightarrow> a \<longrightarrow> (\<lambda>n. X (S n)) \<longlonglongrightarrow> L" |
| 61973 | 2888 |
shows "(X \<longlongrightarrow> L) (at a within T)" |
| 44584 | 2889 |
using assms unfolding tendsto_def [where l=L] |
2890 |
by (simp add: sequentially_imp_eventually_within) |
|
| 43338 | 2891 |
|
2892 |
lemma Lim_right_bound: |
|
| 51773 | 2893 |
fixes f :: "'a :: {linorder_topology, conditionally_complete_linorder, no_top} \<Rightarrow>
|
2894 |
'b::{linorder_topology, conditionally_complete_linorder}"
|
|
| 43338 | 2895 |
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 | 2896 |
and bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a" |
| 61973 | 2897 |
shows "(f \<longlongrightarrow> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
|
| 53640 | 2898 |
proof (cases "{x<..} \<inter> I = {}")
|
2899 |
case True |
|
| 53859 | 2900 |
then show ?thesis by simp |
| 43338 | 2901 |
next |
| 53640 | 2902 |
case False |
| 43338 | 2903 |
show ?thesis |
|
51518
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
hoelzl
parents:
51481
diff
changeset
|
2904 |
proof (rule order_tendstoI) |
| 53282 | 2905 |
fix a |
2906 |
assume a: "a < Inf (f ` ({x<..} \<inter> I))"
|
|
| 53255 | 2907 |
{
|
2908 |
fix y |
|
2909 |
assume "y \<in> {x<..} \<inter> I"
|
|
| 53640 | 2910 |
with False bnd have "Inf (f ` ({x<..} \<inter> I)) \<le> f y"
|
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
2911 |
by (auto intro!: cInf_lower bdd_belowI2) |
| 53255 | 2912 |
with a have "a < f y" |
2913 |
by (blast intro: less_le_trans) |
|
2914 |
} |
|
|
51518
6a56b7088a6a
separate SupInf into Conditional_Complete_Lattice, move instantiation of real to RealDef
hoelzl
parents:
51481
diff
changeset
|
2915 |
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
|
2916 |
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
|
2917 |
next |
| 53255 | 2918 |
fix a |
2919 |
assume "Inf (f ` ({x<..} \<inter> I)) < a"
|
|
| 53640 | 2920 |
from cInf_lessD[OF _ this] False obtain y where y: "x < y" "y \<in> I" "f y < a" |
| 53255 | 2921 |
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
|
2922 |
then have "eventually (\<lambda>x. x \<in> I \<longrightarrow> f x < a) (at_right x)" |
| 60420 | 2923 |
unfolding eventually_at_right[OF \<open>x < y\<close>] by (metis less_imp_le le_less_trans mono) |
|
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
|
2924 |
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
|
2925 |
unfolding eventually_at_filter by eventually_elim simp |
| 43338 | 2926 |
qed |
2927 |
qed |
|
2928 |
||
| 64539 | 2929 |
text \<open>Another limit point characterization.\<close> |
| 33175 | 2930 |
|
| 63301 | 2931 |
lemma limpt_sequential_inj: |
2932 |
fixes x :: "'a::metric_space" |
|
2933 |
shows "x islimpt S \<longleftrightarrow> |
|
2934 |
(\<exists>f. (\<forall>n::nat. f n \<in> S - {x}) \<and> inj f \<and> (f \<longlongrightarrow> x) sequentially)"
|
|
2935 |
(is "?lhs = ?rhs") |
|
2936 |
proof |
|
2937 |
assume ?lhs |
|
2938 |
then have "\<forall>e>0. \<exists>x'\<in>S. x' \<noteq> x \<and> dist x' x < e" |
|
2939 |
by (force simp: islimpt_approachable) |
|
2940 |
then obtain y where y: "\<And>e. e>0 \<Longrightarrow> y e \<in> S \<and> y e \<noteq> x \<and> dist (y e) x < e" |
|
2941 |
by metis |
|
2942 |
define f where "f \<equiv> rec_nat (y 1) (\<lambda>n fn. y (min (inverse(2 ^ (Suc n))) (dist fn x)))" |
|
2943 |
have [simp]: "f 0 = y 1" |
|
2944 |
"f(Suc n) = y (min (inverse(2 ^ (Suc n))) (dist (f n) x))" for n |
|
2945 |
by (simp_all add: f_def) |
|
2946 |
have f: "f n \<in> S \<and> (f n \<noteq> x) \<and> dist (f n) x < inverse(2 ^ n)" for n |
|
2947 |
proof (induction n) |
|
2948 |
case 0 show ?case |
|
2949 |
by (simp add: y) |
|
2950 |
next |
|
2951 |
case (Suc n) then show ?case |
|
2952 |
apply (auto simp: y) |
|
2953 |
by (metis half_gt_zero_iff inverse_positive_iff_positive less_divide_eq_numeral1(1) min_less_iff_conj y zero_less_dist_iff zero_less_numeral zero_less_power) |
|
2954 |
qed |
|
2955 |
show ?rhs |
|
2956 |
proof (rule_tac x=f in exI, intro conjI allI) |
|
2957 |
show "\<And>n. f n \<in> S - {x}"
|
|
2958 |
using f by blast |
|
2959 |
have "dist (f n) x < dist (f m) x" if "m < n" for m n |
|
2960 |
using that |
|
2961 |
proof (induction n) |
|
2962 |
case 0 then show ?case by simp |
|
2963 |
next |
|
2964 |
case (Suc n) |
|
2965 |
then consider "m < n" | "m = n" using less_Suc_eq by blast |
|
2966 |
then show ?case |
|
2967 |
proof cases |
|
2968 |
assume "m < n" |
|
2969 |
have "dist (f(Suc n)) x = dist (y (min (inverse(2 ^ (Suc n))) (dist (f n) x))) x" |
|
2970 |
by simp |
|
| 68120 | 2971 |
also have "\<dots> < dist (f n) x" |
| 63301 | 2972 |
by (metis dist_pos_lt f min.strict_order_iff min_less_iff_conj y) |
| 68120 | 2973 |
also have "\<dots> < dist (f m) x" |
| 63301 | 2974 |
using Suc.IH \<open>m < n\<close> by blast |
2975 |
finally show ?thesis . |
|
2976 |
next |
|
2977 |
assume "m = n" then show ?case |
|
2978 |
by simp (metis dist_pos_lt f half_gt_zero_iff inverse_positive_iff_positive min_less_iff_conj y zero_less_numeral zero_less_power) |
|
2979 |
qed |
|
2980 |
qed |
|
2981 |
then show "inj f" |
|
2982 |
by (metis less_irrefl linorder_injI) |
|
2983 |
show "f \<longlonglongrightarrow> x" |
|
2984 |
apply (rule tendstoI) |
|
2985 |
apply (rule_tac c="nat (ceiling(1/e))" in eventually_sequentiallyI) |
|
2986 |
apply (rule less_trans [OF f [THEN conjunct2, THEN conjunct2]]) |
|
2987 |
apply (simp add: field_simps) |
|
2988 |
by (meson le_less_trans mult_less_cancel_left not_le of_nat_less_two_power) |
|
2989 |
qed |
|
2990 |
next |
|
2991 |
assume ?rhs |
|
2992 |
then show ?lhs |
|
2993 |
by (fastforce simp add: islimpt_approachable lim_sequentially) |
|
2994 |
qed |
|
2995 |
||
2996 |
(*could prove directly from islimpt_sequential_inj, but only for metric spaces*) |
|
| 33175 | 2997 |
lemma islimpt_sequential: |
| 50883 | 2998 |
fixes x :: "'a::first_countable_topology" |
| 61973 | 2999 |
shows "x islimpt S \<longleftrightarrow> (\<exists>f. (\<forall>n::nat. f n \<in> S - {x}) \<and> (f \<longlongrightarrow> x) sequentially)"
|
| 33175 | 3000 |
(is "?lhs = ?rhs") |
3001 |
proof |
|
3002 |
assume ?lhs |
|
| 55522 | 3003 |
from countable_basis_at_decseq[of x] obtain A where A: |
3004 |
"\<And>i. open (A i)" |
|
3005 |
"\<And>i. x \<in> A i" |
|
3006 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
|
3007 |
by blast |
|
| 63040 | 3008 |
define f where "f n = (SOME y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y)" for n |
| 53255 | 3009 |
{
|
3010 |
fix n |
|
| 60420 | 3011 |
from \<open>?lhs\<close> have "\<exists>y. y \<in> S \<and> y \<in> A n \<and> x \<noteq> y" |
| 50883 | 3012 |
unfolding islimpt_def using A(1,2)[of n] by auto |
3013 |
then have "f n \<in> S \<and> f n \<in> A n \<and> x \<noteq> f n" |
|
3014 |
unfolding f_def by (rule someI_ex) |
|
| 53255 | 3015 |
then have "f n \<in> S" "f n \<in> A n" "x \<noteq> f n" by auto |
3016 |
} |
|
| 50883 | 3017 |
then have "\<forall>n. f n \<in> S - {x}" by auto
|
| 61969 | 3018 |
moreover have "(\<lambda>n. f n) \<longlonglongrightarrow> x" |
| 50883 | 3019 |
proof (rule topological_tendstoI) |
| 53255 | 3020 |
fix S |
3021 |
assume "open S" "x \<in> S" |
|
| 60420 | 3022 |
from A(3)[OF this] \<open>\<And>n. f n \<in> A n\<close> |
| 53255 | 3023 |
show "eventually (\<lambda>x. f x \<in> S) sequentially" |
| 61810 | 3024 |
by (auto elim!: eventually_mono) |
| 44584 | 3025 |
qed |
3026 |
ultimately show ?rhs by fast |
|
| 33175 | 3027 |
next |
3028 |
assume ?rhs |
|
| 61969 | 3029 |
then obtain f :: "nat \<Rightarrow> 'a" where f: "\<And>n. f n \<in> S - {x}" and lim: "f \<longlonglongrightarrow> x"
|
| 53255 | 3030 |
by auto |
| 50883 | 3031 |
show ?lhs |
3032 |
unfolding islimpt_def |
|
3033 |
proof safe |
|
| 53255 | 3034 |
fix T |
3035 |
assume "open T" "x \<in> T" |
|
| 50883 | 3036 |
from lim[THEN topological_tendstoD, OF this] f |
3037 |
show "\<exists>y\<in>S. y \<in> T \<and> y \<noteq> x" |
|
3038 |
unfolding eventually_sequentially by auto |
|
3039 |
qed |
|
| 33175 | 3040 |
qed |
3041 |
||
3042 |
lemma Lim_null: |
|
3043 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 61973 | 3044 |
shows "(f \<longlongrightarrow> l) net \<longleftrightarrow> ((\<lambda>x. f(x) - l) \<longlongrightarrow> 0) net" |
| 33175 | 3045 |
by (simp add: Lim dist_norm) |
3046 |
||
3047 |
lemma Lim_null_comparison: |
|
3048 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 61973 | 3049 |
assumes "eventually (\<lambda>x. norm (f x) \<le> g x) net" "(g \<longlongrightarrow> 0) net" |
3050 |
shows "(f \<longlongrightarrow> 0) net" |
|
| 53282 | 3051 |
using assms(2) |
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
3052 |
proof (rule metric_tendsto_imp_tendsto) |
|
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
3053 |
show "eventually (\<lambda>x. dist (f x) 0 \<le> dist (g x) 0) net" |
| 61810 | 3054 |
using assms(1) by (rule eventually_mono) (simp add: dist_norm) |
| 33175 | 3055 |
qed |
3056 |
||
3057 |
lemma Lim_transform_bound: |
|
3058 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 53255 | 3059 |
and g :: "'a \<Rightarrow> 'c::real_normed_vector" |
| 53640 | 3060 |
assumes "eventually (\<lambda>n. norm (f n) \<le> norm (g n)) net" |
| 61973 | 3061 |
and "(g \<longlongrightarrow> 0) net" |
3062 |
shows "(f \<longlongrightarrow> 0) net" |
|
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
3063 |
using assms(1) tendsto_norm_zero [OF assms(2)] |
|
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
3064 |
by (rule Lim_null_comparison) |
| 33175 | 3065 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3066 |
lemma lim_null_mult_right_bounded: |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3067 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3068 |
assumes f: "(f \<longlongrightarrow> 0) F" and g: "eventually (\<lambda>x. norm(g x) \<le> B) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3069 |
shows "((\<lambda>z. f z * g z) \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3070 |
proof - |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3071 |
have *: "((\<lambda>x. norm (f x) * B) \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3072 |
by (simp add: f tendsto_mult_left_zero tendsto_norm_zero) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3073 |
have "((\<lambda>x. norm (f x) * norm (g x)) \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3074 |
apply (rule Lim_null_comparison [OF _ *]) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3075 |
apply (simp add: eventually_mono [OF g] mult_left_mono) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3076 |
done |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3077 |
then show ?thesis |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3078 |
by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3079 |
qed |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3080 |
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3081 |
lemma lim_null_mult_left_bounded: |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3082 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3083 |
assumes g: "eventually (\<lambda>x. norm(g x) \<le> B) F" and f: "(f \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3084 |
shows "((\<lambda>z. g z * f z) \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3085 |
proof - |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3086 |
have *: "((\<lambda>x. B * norm (f x)) \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3087 |
by (simp add: f tendsto_mult_right_zero tendsto_norm_zero) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3088 |
have "((\<lambda>x. norm (g x) * norm (f x)) \<longlongrightarrow> 0) F" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3089 |
apply (rule Lim_null_comparison [OF _ *]) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3090 |
apply (simp add: eventually_mono [OF g] mult_right_mono) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3091 |
done |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3092 |
then show ?thesis |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3093 |
by (subst tendsto_norm_zero_iff [symmetric]) (simp add: norm_mult) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3094 |
qed |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3095 |
|
| 63128 | 3096 |
lemma lim_null_scaleR_bounded: |
3097 |
assumes f: "(f \<longlongrightarrow> 0) net" and gB: "eventually (\<lambda>a. f a = 0 \<or> norm(g a) \<le> B) net" |
|
3098 |
shows "((\<lambda>n. f n *\<^sub>R g n) \<longlongrightarrow> 0) net" |
|
3099 |
proof |
|
3100 |
fix \<epsilon>::real |
|
3101 |
assume "0 < \<epsilon>" |
|
3102 |
then have B: "0 < \<epsilon> / (abs B + 1)" by simp |
|
3103 |
have *: "\<bar>f x\<bar> * norm (g x) < \<epsilon>" if f: "\<bar>f x\<bar> * (\<bar>B\<bar> + 1) < \<epsilon>" and g: "norm (g x) \<le> B" for x |
|
3104 |
proof - |
|
3105 |
have "\<bar>f x\<bar> * norm (g x) \<le> \<bar>f x\<bar> * B" |
|
3106 |
by (simp add: mult_left_mono g) |
|
| 68120 | 3107 |
also have "\<dots> \<le> \<bar>f x\<bar> * (\<bar>B\<bar> + 1)" |
| 63128 | 3108 |
by (simp add: mult_left_mono) |
| 68120 | 3109 |
also have "\<dots> < \<epsilon>" |
| 63128 | 3110 |
by (rule f) |
3111 |
finally show ?thesis . |
|
3112 |
qed |
|
3113 |
show "\<forall>\<^sub>F x in net. dist (f x *\<^sub>R g x) 0 < \<epsilon>" |
|
3114 |
apply (rule eventually_mono [OF eventually_conj [OF tendstoD [OF f B] gB] ]) |
|
3115 |
apply (auto simp: \<open>0 < \<epsilon>\<close> divide_simps * split: if_split_asm) |
|
3116 |
done |
|
3117 |
qed |
|
3118 |
||
| 60420 | 3119 |
text\<open>Deducing things about the limit from the elements.\<close> |
| 33175 | 3120 |
|
3121 |
lemma Lim_in_closed_set: |
|
| 53255 | 3122 |
assumes "closed S" |
3123 |
and "eventually (\<lambda>x. f(x) \<in> S) net" |
|
| 61973 | 3124 |
and "\<not> trivial_limit net" "(f \<longlongrightarrow> l) net" |
| 33175 | 3125 |
shows "l \<in> S" |
3126 |
proof (rule ccontr) |
|
3127 |
assume "l \<notin> S" |
|
| 60420 | 3128 |
with \<open>closed S\<close> have "open (- S)" "l \<in> - S" |
| 33175 | 3129 |
by (simp_all add: open_Compl) |
3130 |
with assms(4) have "eventually (\<lambda>x. f x \<in> - S) net" |
|
3131 |
by (rule topological_tendstoD) |
|
3132 |
with assms(2) have "eventually (\<lambda>x. False) net" |
|
3133 |
by (rule eventually_elim2) simp |
|
3134 |
with assms(3) show "False" |
|
3135 |
by (simp add: eventually_False) |
|
3136 |
qed |
|
3137 |
||
| 60420 | 3138 |
text\<open>Need to prove closed(cball(x,e)) before deducing this as a corollary.\<close> |
| 33175 | 3139 |
|
3140 |
lemma Lim_dist_ubound: |
|
| 53255 | 3141 |
assumes "\<not>(trivial_limit net)" |
| 61973 | 3142 |
and "(f \<longlongrightarrow> l) net" |
| 53640 | 3143 |
and "eventually (\<lambda>x. dist a (f x) \<le> e) net" |
3144 |
shows "dist a l \<le> e" |
|
| 56290 | 3145 |
using assms by (fast intro: tendsto_le tendsto_intros) |
| 33175 | 3146 |
|
3147 |
lemma Lim_norm_ubound: |
|
3148 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 61973 | 3149 |
assumes "\<not>(trivial_limit net)" "(f \<longlongrightarrow> l) net" "eventually (\<lambda>x. norm(f x) \<le> e) net" |
| 53255 | 3150 |
shows "norm(l) \<le> e" |
| 56290 | 3151 |
using assms by (fast intro: tendsto_le tendsto_intros) |
| 33175 | 3152 |
|
3153 |
lemma Lim_norm_lbound: |
|
3154 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
| 53640 | 3155 |
assumes "\<not> trivial_limit net" |
| 61973 | 3156 |
and "(f \<longlongrightarrow> l) net" |
| 53640 | 3157 |
and "eventually (\<lambda>x. e \<le> norm (f x)) net" |
| 33175 | 3158 |
shows "e \<le> norm l" |
| 56290 | 3159 |
using assms by (fast intro: tendsto_le tendsto_intros) |
| 33175 | 3160 |
|
| 60420 | 3161 |
text\<open>Limit under bilinear function\<close> |
| 33175 | 3162 |
|
3163 |
lemma Lim_bilinear: |
|
| 61973 | 3164 |
assumes "(f \<longlongrightarrow> l) net" |
3165 |
and "(g \<longlongrightarrow> m) net" |
|
| 53282 | 3166 |
and "bounded_bilinear h" |
| 61973 | 3167 |
shows "((\<lambda>x. h (f x) (g x)) \<longlongrightarrow> (h l m)) net" |
3168 |
using \<open>bounded_bilinear h\<close> \<open>(f \<longlongrightarrow> l) net\<close> \<open>(g \<longlongrightarrow> m) net\<close> |
|
| 52624 | 3169 |
by (rule bounded_bilinear.tendsto) |
| 33175 | 3170 |
|
| 60420 | 3171 |
text\<open>These are special for limits out of the same vector space.\<close> |
| 33175 | 3172 |
|
| 61973 | 3173 |
lemma Lim_within_id: "(id \<longlongrightarrow> 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
|
3174 |
unfolding id_def by (rule tendsto_ident_at) |
| 33175 | 3175 |
|
| 61973 | 3176 |
lemma Lim_at_id: "(id \<longlongrightarrow> a) (at a)" |
| 45031 | 3177 |
unfolding id_def by (rule tendsto_ident_at) |
| 33175 | 3178 |
|
3179 |
lemma Lim_at_zero: |
|
3180 |
fixes a :: "'a::real_normed_vector" |
|
| 53291 | 3181 |
and l :: "'b::topological_space" |
| 61973 | 3182 |
shows "(f \<longlongrightarrow> l) (at a) \<longleftrightarrow> ((\<lambda>x. f(a + x)) \<longlongrightarrow> l) (at 0)" |
|
44252
10362a07eb7c
Topology_Euclidean_Space.thy: simplify some proofs
huffman
parents:
44250
diff
changeset
|
3183 |
using LIM_offset_zero LIM_offset_zero_cancel .. |
| 33175 | 3184 |
|
| 60420 | 3185 |
text\<open>It's also sometimes useful to extract the limit point from the filter.\<close> |
| 33175 | 3186 |
|
| 52624 | 3187 |
abbreviation netlimit :: "'a::t2_space filter \<Rightarrow> 'a" |
3188 |
where "netlimit F \<equiv> Lim F (\<lambda>x. x)" |
|
| 33175 | 3189 |
|
| 53282 | 3190 |
lemma netlimit_within: "\<not> trivial_limit (at a within S) \<Longrightarrow> netlimit (at a within S) = a" |
| 51365 | 3191 |
by (rule tendsto_Lim) (auto intro: tendsto_intros) |
| 33175 | 3192 |
|
|
67979
53323937ee25
new material about vec, real^1, etc.
paulson <lp15@cam.ac.uk>
parents:
67968
diff
changeset
|
3193 |
lemma netlimit_at [simp]: |
|
44072
5b970711fb39
class perfect_space inherits from topological_space;
huffman
parents:
43338
diff
changeset
|
3194 |
fixes a :: "'a::{perfect_space,t2_space}"
|
| 33175 | 3195 |
shows "netlimit (at a) = a" |
| 45031 | 3196 |
using netlimit_within [of a UNIV] by simp |
| 33175 | 3197 |
|
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3198 |
lemma lim_within_interior: |
| 61973 | 3199 |
"x \<in> interior S \<Longrightarrow> (f \<longlongrightarrow> l) (at x within S) \<longleftrightarrow> (f \<longlongrightarrow> 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
|
3200 |
by (metis at_within_interior) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3201 |
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3202 |
lemma netlimit_within_interior: |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3203 |
fixes x :: "'a::{t2_space,perfect_space}"
|
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3204 |
assumes "x \<in> interior S" |
|
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3205 |
shows "netlimit (at x within S) = x" |
| 52624 | 3206 |
using assms by (metis at_within_interior netlimit_at) |
|
44210
eba74571833b
Topology_Euclidean_Space.thy: organize section headings
huffman
parents:
44207
diff
changeset
|
3207 |
|
|
61824
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3208 |
lemma netlimit_at_vector: |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3209 |
fixes a :: "'a::real_normed_vector" |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3210 |
shows "netlimit (at a) = a" |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3211 |
proof (cases "\<exists>x. x \<noteq> a") |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3212 |
case True then obtain x where x: "x \<noteq> a" .. |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3213 |
have "\<not> trivial_limit (at a)" |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3214 |
unfolding trivial_limit_def eventually_at dist_norm |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3215 |
apply clarsimp |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3216 |
apply (rule_tac x="a + scaleR (d / 2) (sgn (x - a))" in exI) |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3217 |
apply (simp add: norm_sgn sgn_zero_iff x) |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3218 |
done |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3219 |
then show ?thesis |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3220 |
by (rule netlimit_within [of a UNIV]) |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3221 |
qed simp |
|
dcbe9f756ae0
not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents:
61810
diff
changeset
|
3222 |
|
| 33175 | 3223 |
|
| 60420 | 3224 |
text\<open>Useful lemmas on closure and set of possible sequential limits.\<close> |
| 33175 | 3225 |
|
3226 |
lemma closure_sequential: |
|
| 50883 | 3227 |
fixes l :: "'a::first_countable_topology" |
| 61973 | 3228 |
shows "l \<in> closure S \<longleftrightarrow> (\<exists>x. (\<forall>n. x n \<in> S) \<and> (x \<longlongrightarrow> l) sequentially)" |
| 53291 | 3229 |
(is "?lhs = ?rhs") |
| 33175 | 3230 |
proof |
| 53282 | 3231 |
assume "?lhs" |
3232 |
moreover |
|
3233 |
{
|
|
3234 |
assume "l \<in> S" |
|
3235 |
then have "?rhs" using tendsto_const[of l sequentially] by auto |
|
| 52624 | 3236 |
} |
3237 |
moreover |
|
| 53282 | 3238 |
{
|
3239 |
assume "l islimpt S" |
|
3240 |
then have "?rhs" unfolding islimpt_sequential by auto |
|
| 52624 | 3241 |
} |
3242 |
ultimately show "?rhs" |
|
3243 |
unfolding closure_def by auto |
|
| 33175 | 3244 |
next |
3245 |
assume "?rhs" |
|
| 53282 | 3246 |
then show "?lhs" unfolding closure_def islimpt_sequential by auto |
| 33175 | 3247 |
qed |
3248 |
||
3249 |
lemma closed_sequential_limits: |
|
| 50883 | 3250 |
fixes S :: "'a::first_countable_topology set" |
| 61973 | 3251 |
shows "closed S \<longleftrightarrow> (\<forall>x l. (\<forall>n. x n \<in> S) \<and> (x \<longlongrightarrow> l) sequentially \<longrightarrow> l \<in> S)" |
| 55775 | 3252 |
by (metis closure_sequential closure_subset_eq subset_iff) |
| 33175 | 3253 |
|
3254 |
lemma closure_approachable: |
|
3255 |
fixes S :: "'a::metric_space set" |
|
3256 |
shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e)" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3257 |
apply (auto simp: closure_def islimpt_approachable) |
| 52624 | 3258 |
apply (metis dist_self) |
3259 |
done |
|
| 33175 | 3260 |
|
|
66641
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3261 |
lemma closure_approachable_le: |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3262 |
fixes S :: "'a::metric_space set" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3263 |
shows "x \<in> closure S \<longleftrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x \<le> e)" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3264 |
unfolding closure_approachable |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3265 |
using dense by force |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3266 |
|
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
3267 |
lemma closure_approachableD: |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
3268 |
assumes "x \<in> closure S" "e>0" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
3269 |
shows "\<exists>y\<in>S. dist x y < e" |
| 68120 | 3270 |
using assms unfolding closure_approachable by (auto simp: dist_commute) |
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
3271 |
|
| 33175 | 3272 |
lemma closed_approachable: |
3273 |
fixes S :: "'a::metric_space set" |
|
| 53291 | 3274 |
shows "closed S \<Longrightarrow> (\<forall>e>0. \<exists>y\<in>S. dist y x < e) \<longleftrightarrow> x \<in> S" |
| 33175 | 3275 |
by (metis closure_closed closure_approachable) |
3276 |
||
| 51351 | 3277 |
lemma closure_contains_Inf: |
3278 |
fixes S :: "real set" |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
3279 |
assumes "S \<noteq> {}" "bdd_below S"
|
| 51351 | 3280 |
shows "Inf S \<in> closure S" |
| 52624 | 3281 |
proof - |
| 51351 | 3282 |
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
|
3283 |
using cInf_lower[of _ S] assms by metis |
| 52624 | 3284 |
{
|
| 53282 | 3285 |
fix e :: real |
3286 |
assume "e > 0" |
|
| 52624 | 3287 |
then have "Inf S < Inf S + e" by simp |
3288 |
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
|
3289 |
by (subst (asm) cInf_less_iff) auto |
| 52624 | 3290 |
with * have "\<exists>x\<in>S. dist x (Inf S) < e" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3291 |
by (intro bexI[of _ x]) (auto simp: dist_real_def) |
| 52624 | 3292 |
} |
3293 |
then show ?thesis unfolding closure_approachable by auto |
|
| 51351 | 3294 |
qed |
3295 |
||
|
66641
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3296 |
lemma closure_Int_ballI: |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3297 |
fixes S :: "'a :: metric_space set" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3298 |
assumes "\<And>U. \<lbrakk>openin (subtopology euclidean S) U; U \<noteq> {}\<rbrakk> \<Longrightarrow> T \<inter> U \<noteq> {}"
|
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3299 |
shows "S \<subseteq> closure T" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3300 |
proof (clarsimp simp: closure_approachable dist_commute) |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3301 |
fix x and e::real |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3302 |
assume "x \<in> S" "0 < e" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3303 |
with assms [of "S \<inter> ball x e"] show "\<exists>y\<in>T. dist x y < e" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3304 |
by force |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3305 |
qed |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
3306 |
|
| 51351 | 3307 |
lemma closed_contains_Inf: |
3308 |
fixes S :: "real set" |
|
|
54258
adfc759263ab
use bdd_above and bdd_below for conditionally complete lattices
hoelzl
parents:
54230
diff
changeset
|
3309 |
shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
|
| 63092 | 3310 |
by (metis closure_contains_Inf closure_closed) |
| 51351 | 3311 |
|
| 62083 | 3312 |
lemma closed_subset_contains_Inf: |
3313 |
fixes A C :: "real set" |
|
3314 |
shows "closed C \<Longrightarrow> A \<subseteq> C \<Longrightarrow> A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> Inf A \<in> C"
|
|
3315 |
by (metis closure_contains_Inf closure_minimal subset_eq) |
|
3316 |
||
3317 |
lemma atLeastAtMost_subset_contains_Inf: |
|
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
3318 |
fixes A :: "real set" and a b :: real |
| 62083 | 3319 |
shows "A \<noteq> {} \<Longrightarrow> a \<le> b \<Longrightarrow> A \<subseteq> {a..b} \<Longrightarrow> Inf A \<in> {a..b}"
|
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
3320 |
by (rule closed_subset_contains_Inf) |
| 62083 | 3321 |
(auto intro: closed_real_atLeastAtMost intro!: bdd_belowI[of A a]) |
3322 |
||
| 51351 | 3323 |
lemma not_trivial_limit_within_ball: |
| 53640 | 3324 |
"\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
|
| 60462 | 3325 |
(is "?lhs \<longleftrightarrow> ?rhs") |
3326 |
proof |
|
3327 |
show ?rhs if ?lhs |
|
3328 |
proof - |
|
| 53282 | 3329 |
{
|
3330 |
fix e :: real |
|
3331 |
assume "e > 0" |
|
| 53640 | 3332 |
then obtain y where "y \<in> S - {x}" and "dist y x < e"
|
| 60420 | 3333 |
using \<open>?lhs\<close> not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
|
| 51351 | 3334 |
by auto |
| 53640 | 3335 |
then have "y \<in> S \<inter> ball x e - {x}"
|
| 51351 | 3336 |
unfolding ball_def by (simp add: dist_commute) |
| 53640 | 3337 |
then have "S \<inter> ball x e - {x} \<noteq> {}" by blast
|
| 52624 | 3338 |
} |
| 60462 | 3339 |
then show ?thesis by auto |
3340 |
qed |
|
3341 |
show ?lhs if ?rhs |
|
3342 |
proof - |
|
| 53282 | 3343 |
{
|
3344 |
fix e :: real |
|
3345 |
assume "e > 0" |
|
| 53640 | 3346 |
then obtain y where "y \<in> S \<inter> ball x e - {x}"
|
| 60420 | 3347 |
using \<open>?rhs\<close> by blast |
| 53640 | 3348 |
then have "y \<in> S - {x}" and "dist y x < e"
|
3349 |
unfolding ball_def by (simp_all add: dist_commute) |
|
3350 |
then have "\<exists>y \<in> S - {x}. dist y x < e"
|
|
| 53282 | 3351 |
by auto |
| 51351 | 3352 |
} |
| 60462 | 3353 |
then show ?thesis |
| 53282 | 3354 |
using not_trivial_limit_within[of x S] closure_approachable[of x "S - {x}"]
|
3355 |
by auto |
|
| 60462 | 3356 |
qed |
| 51351 | 3357 |
qed |
3358 |
||
|
67685
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3359 |
lemma tendsto_If_within_closures: |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3360 |
assumes f: "x \<in> s \<union> (closure s \<inter> closure t) \<Longrightarrow> |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3361 |
(f \<longlongrightarrow> l x) (at x within s \<union> (closure s \<inter> closure t))" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3362 |
assumes g: "x \<in> t \<union> (closure s \<inter> closure t) \<Longrightarrow> |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3363 |
(g \<longlongrightarrow> l x) (at x within t \<union> (closure s \<inter> closure t))" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3364 |
assumes "x \<in> s \<union> t" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3365 |
shows "((\<lambda>x. if x \<in> s then f x else g x) \<longlongrightarrow> l x) (at x within s \<union> t)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3366 |
proof - |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3367 |
have *: "(s \<union> t) \<inter> {x. x \<in> s} = s" "(s \<union> t) \<inter> {x. x \<notin> s} = t - s"
|
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3368 |
by auto |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3369 |
have "(f \<longlongrightarrow> l x) (at x within s)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3370 |
by (rule filterlim_at_within_closure_implies_filterlim) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3371 |
(use \<open>x \<in> _\<close> in \<open>auto simp: inf_commute closure_def intro: tendsto_within_subset[OF f]\<close>) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3372 |
moreover |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3373 |
have "(g \<longlongrightarrow> l x) (at x within t - s)" |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3374 |
by (rule filterlim_at_within_closure_implies_filterlim) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3375 |
(use \<open>x \<in> _\<close> in |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3376 |
\<open>auto intro!: tendsto_within_subset[OF g] simp: closure_def intro: islimpt_subset\<close>) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3377 |
ultimately show ?thesis |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3378 |
by (intro filterlim_at_within_If) (simp_all only: *) |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3379 |
qed |
|
bdff8bf0a75b
moved theorems from AFP/Affine_Arithmetic and AFP/Ordinary_Differential_Equations
immler
parents:
67613
diff
changeset
|
3380 |
|
| 52624 | 3381 |
|
| 60420 | 3382 |
subsection \<open>Boundedness\<close> |
| 33175 | 3383 |
|
3384 |
(* FIXME: This has to be unified with BSEQ!! *) |
|
| 67962 | 3385 |
definition%important (in metric_space) bounded :: "'a set \<Rightarrow> bool" |
| 52624 | 3386 |
where "bounded S \<longleftrightarrow> (\<exists>x e. \<forall>y\<in>S. dist x y \<le> e)" |
| 33175 | 3387 |
|
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3388 |
lemma bounded_subset_cball: "bounded S \<longleftrightarrow> (\<exists>e x. S \<subseteq> cball x e \<and> 0 \<le> e)" |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3389 |
unfolding bounded_def subset_eq by auto (meson order_trans zero_le_dist) |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3390 |
|
| 33175 | 3391 |
lemma bounded_any_center: "bounded S \<longleftrightarrow> (\<exists>e. \<forall>y\<in>S. dist a y \<le> e)" |
| 52624 | 3392 |
unfolding bounded_def |
|
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57448
diff
changeset
|
3393 |
by auto (metis add.commute add_le_cancel_right dist_commute dist_triangle_le) |
| 33175 | 3394 |
|
3395 |
lemma bounded_iff: "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. norm x \<le> a)" |
|
| 52624 | 3396 |
unfolding bounded_any_center [where a=0] |
3397 |
by (simp add: dist_norm) |
|
| 33175 | 3398 |
|
|
61552
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
3399 |
lemma bdd_above_norm: "bdd_above (norm ` X) \<longleftrightarrow> bounded X" |
|
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
3400 |
by (simp add: bounded_iff bdd_above_def) |
|
980dd46a03fb
Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents:
61531
diff
changeset
|
3401 |
|
|
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3402 |
lemma bounded_norm_comp: "bounded ((\<lambda>x. norm (f x)) ` S) = bounded (f ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3403 |
by (simp add: bounded_iff) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3404 |
|
|
65036
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents:
64911
diff
changeset
|
3405 |
lemma boundedI: |
|
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents:
64911
diff
changeset
|
3406 |
assumes "\<And>x. x \<in> S \<Longrightarrow> norm x \<le> B" |
|
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents:
64911
diff
changeset
|
3407 |
shows "bounded S" |
|
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents:
64911
diff
changeset
|
3408 |
using assms bounded_iff by blast |
| 50104 | 3409 |
|
| 50948 | 3410 |
lemma bounded_empty [simp]: "bounded {}"
|
3411 |
by (simp add: bounded_def) |
|
3412 |
||
| 53291 | 3413 |
lemma bounded_subset: "bounded T \<Longrightarrow> S \<subseteq> T \<Longrightarrow> bounded S" |
| 33175 | 3414 |
by (metis bounded_def subset_eq) |
3415 |
||
| 53291 | 3416 |
lemma bounded_interior[intro]: "bounded S \<Longrightarrow> bounded(interior S)" |
| 33175 | 3417 |
by (metis bounded_subset interior_subset) |
3418 |
||
| 52624 | 3419 |
lemma bounded_closure[intro]: |
3420 |
assumes "bounded S" |
|
3421 |
shows "bounded (closure S)" |
|
3422 |
proof - |
|
3423 |
from assms obtain x and a where a: "\<forall>y\<in>S. dist x y \<le> a" |
|
3424 |
unfolding bounded_def by auto |
|
3425 |
{
|
|
3426 |
fix y |
|
3427 |
assume "y \<in> closure S" |
|
| 61973 | 3428 |
then obtain f where f: "\<forall>n. f n \<in> S" "(f \<longlongrightarrow> y) sequentially" |
| 33175 | 3429 |
unfolding closure_sequential by auto |
3430 |
have "\<forall>n. f n \<in> S \<longrightarrow> dist x (f n) \<le> a" using a by simp |
|
| 53282 | 3431 |
then have "eventually (\<lambda>n. dist x (f n) \<le> a) sequentially" |
| 61810 | 3432 |
by (simp add: f(1)) |
| 33175 | 3433 |
have "dist x y \<le> a" |
3434 |
apply (rule Lim_dist_ubound [of sequentially f]) |
|
3435 |
apply (rule trivial_limit_sequentially) |
|
3436 |
apply (rule f(2)) |
|
3437 |
apply fact |
|
3438 |
done |
|
3439 |
} |
|
| 53282 | 3440 |
then show ?thesis |
3441 |
unfolding bounded_def by auto |
|
| 33175 | 3442 |
qed |
3443 |
||
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
3444 |
lemma bounded_closure_image: "bounded (f ` closure S) \<Longrightarrow> bounded (f ` S)" |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
3445 |
by (simp add: bounded_subset closure_subset image_mono) |
|
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
3446 |
|
| 33175 | 3447 |
lemma bounded_cball[simp,intro]: "bounded (cball x e)" |
3448 |
apply (simp add: bounded_def) |
|
3449 |
apply (rule_tac x=x in exI) |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3450 |
apply (rule_tac x=e in exI, auto) |
| 33175 | 3451 |
done |
3452 |
||
| 53640 | 3453 |
lemma bounded_ball[simp,intro]: "bounded (ball x e)" |
| 33175 | 3454 |
by (metis ball_subset_cball bounded_cball bounded_subset) |
3455 |
||
3456 |
lemma bounded_Un[simp]: "bounded (S \<union> T) \<longleftrightarrow> bounded S \<and> bounded T" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3457 |
by (auto simp: bounded_def) (metis Un_iff bounded_any_center le_max_iff_disj) |
| 33175 | 3458 |
|
| 53640 | 3459 |
lemma bounded_Union[intro]: "finite F \<Longrightarrow> \<forall>S\<in>F. bounded S \<Longrightarrow> bounded (\<Union>F)" |
| 52624 | 3460 |
by (induct rule: finite_induct[of F]) auto |
| 33175 | 3461 |
|
| 50955 | 3462 |
lemma bounded_UN [intro]: "finite A \<Longrightarrow> \<forall>x\<in>A. bounded (B x) \<Longrightarrow> bounded (\<Union>x\<in>A. B x)" |
| 52624 | 3463 |
by (induct set: finite) auto |
| 50955 | 3464 |
|
| 50948 | 3465 |
lemma bounded_insert [simp]: "bounded (insert x S) \<longleftrightarrow> bounded S" |
3466 |
proof - |
|
| 53640 | 3467 |
have "\<forall>y\<in>{x}. dist x y \<le> 0"
|
3468 |
by simp |
|
3469 |
then have "bounded {x}"
|
|
3470 |
unfolding bounded_def by fast |
|
3471 |
then show ?thesis |
|
3472 |
by (metis insert_is_Un bounded_Un) |
|
| 50948 | 3473 |
qed |
3474 |
||
|
65587
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3475 |
lemma bounded_subset_ballI: "S \<subseteq> ball x r \<Longrightarrow> bounded S" |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3476 |
by (meson bounded_ball bounded_subset) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3477 |
|
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3478 |
lemma bounded_subset_ballD: |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3479 |
assumes "bounded S" shows "\<exists>r. 0 < r \<and> S \<subseteq> ball x r" |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3480 |
proof - |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3481 |
obtain e::real and y where "S \<subseteq> cball y e" "0 \<le> e" |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3482 |
using assms by (auto simp: bounded_subset_cball) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3483 |
then show ?thesis |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3484 |
apply (rule_tac x="dist x y + e + 1" in exI) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3485 |
apply (simp add: add.commute add_pos_nonneg) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3486 |
apply (erule subset_trans) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3487 |
apply (clarsimp simp add: cball_def) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3488 |
by (metis add_le_cancel_right add_strict_increasing dist_commute dist_triangle_le zero_less_one) |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3489 |
qed |
|
16a8991ab398
New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents:
65585
diff
changeset
|
3490 |
|
| 50948 | 3491 |
lemma finite_imp_bounded [intro]: "finite S \<Longrightarrow> bounded S" |
| 52624 | 3492 |
by (induct set: finite) simp_all |
| 50948 | 3493 |
|
| 53640 | 3494 |
lemma bounded_pos: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x \<le> b)" |
| 33175 | 3495 |
apply (simp add: bounded_iff) |
| 61945 | 3496 |
apply (subgoal_tac "\<And>x (y::real). 0 < 1 + \<bar>y\<bar> \<and> (x \<le> y \<longrightarrow> x \<le> 1 + \<bar>y\<bar>)") |
| 52624 | 3497 |
apply metis |
3498 |
apply arith |
|
3499 |
done |
|
| 33175 | 3500 |
|
| 60762 | 3501 |
lemma bounded_pos_less: "bounded S \<longleftrightarrow> (\<exists>b>0. \<forall>x\<in> S. norm x < b)" |
3502 |
apply (simp add: bounded_pos) |
|
3503 |
apply (safe; rule_tac x="b+1" in exI; force) |
|
3504 |
done |
|
3505 |
||
| 53640 | 3506 |
lemma Bseq_eq_bounded: |
3507 |
fixes f :: "nat \<Rightarrow> 'a::real_normed_vector" |
|
3508 |
shows "Bseq f \<longleftrightarrow> bounded (range f)" |
|
| 50972 | 3509 |
unfolding Bseq_def bounded_pos by auto |
3510 |
||
| 33175 | 3511 |
lemma bounded_Int[intro]: "bounded S \<or> bounded T \<Longrightarrow> bounded (S \<inter> T)" |
3512 |
by (metis Int_lower1 Int_lower2 bounded_subset) |
|
3513 |
||
| 53291 | 3514 |
lemma bounded_diff[intro]: "bounded S \<Longrightarrow> bounded (S - T)" |
| 52624 | 3515 |
by (metis Diff_subset bounded_subset) |
| 33175 | 3516 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
3517 |
lemma not_bounded_UNIV[simp]: |
| 33175 | 3518 |
"\<not> bounded (UNIV :: 'a::{real_normed_vector, perfect_space} set)"
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3519 |
proof (auto simp: bounded_pos not_le) |
| 33175 | 3520 |
obtain x :: 'a where "x \<noteq> 0" |
3521 |
using perfect_choose_dist [OF zero_less_one] by fast |
|
| 53640 | 3522 |
fix b :: real |
3523 |
assume b: "b >0" |
|
3524 |
have b1: "b +1 \<ge> 0" |
|
3525 |
using b by simp |
|
| 60420 | 3526 |
with \<open>x \<noteq> 0\<close> have "b < norm (scaleR (b + 1) (sgn x))" |
| 33175 | 3527 |
by (simp add: norm_sgn) |
3528 |
then show "\<exists>x::'a. b < norm x" .. |
|
3529 |
qed |
|
3530 |
||
|
61426
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3531 |
corollary cobounded_imp_unbounded: |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3532 |
fixes S :: "'a::{real_normed_vector, perfect_space} set"
|
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3533 |
shows "bounded (- S) \<Longrightarrow> ~ (bounded S)" |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3534 |
using bounded_Un [of S "-S"] by (simp add: sup_compl_top) |
|
d53db136e8fd
new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents:
61424
diff
changeset
|
3535 |
|
|
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3536 |
lemma bounded_dist_comp: |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3537 |
assumes "bounded (f ` S)" "bounded (g ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3538 |
shows "bounded ((\<lambda>x. dist (f x) (g x)) ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3539 |
proof - |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3540 |
from assms obtain M1 M2 where *: "dist (f x) undefined \<le> M1" "dist undefined (g x) \<le> M2" if "x \<in> S" for x |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3541 |
by (auto simp: bounded_any_center[of _ undefined] dist_commute) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3542 |
have "dist (f x) (g x) \<le> M1 + M2" if "x \<in> S" for x |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3543 |
using *[OF that] |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3544 |
by (rule order_trans[OF dist_triangle add_mono]) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3545 |
then show ?thesis |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3546 |
by (auto intro!: boundedI) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3547 |
qed |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3548 |
|
| 33175 | 3549 |
lemma bounded_linear_image: |
| 53291 | 3550 |
assumes "bounded S" |
3551 |
and "bounded_linear f" |
|
3552 |
shows "bounded (f ` S)" |
|
| 52624 | 3553 |
proof - |
| 67984 | 3554 |
from assms(1) obtain b where "b > 0" and b: "\<forall>x\<in>S. norm x \<le> b" |
| 52624 | 3555 |
unfolding bounded_pos by auto |
| 53640 | 3556 |
from assms(2) obtain B where B: "B > 0" "\<forall>x. norm (f x) \<le> B * norm x" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3557 |
using bounded_linear.pos_bounded by (auto simp: ac_simps) |
| 67984 | 3558 |
show ?thesis |
| 53282 | 3559 |
unfolding bounded_pos |
| 67984 | 3560 |
proof (intro exI, safe) |
3561 |
show "norm (f x) \<le> B * b" if "x \<in> S" for x |
|
3562 |
by (meson B b less_imp_le mult_left_mono order_trans that) |
|
3563 |
qed (use \<open>b > 0\<close> \<open>B > 0\<close> in auto) |
|
| 33175 | 3564 |
qed |
3565 |
||
3566 |
lemma bounded_scaling: |
|
3567 |
fixes S :: "'a::real_normed_vector set" |
|
3568 |
shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. c *\<^sub>R x) ` S)" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3569 |
apply (rule bounded_linear_image, assumption) |
|
44282
f0de18b62d63
remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents:
44252
diff
changeset
|
3570 |
apply (rule bounded_linear_scaleR_right) |
| 33175 | 3571 |
done |
3572 |
||
|
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3573 |
lemma bounded_scaleR_comp: |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3574 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3575 |
assumes "bounded (f ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3576 |
shows "bounded ((\<lambda>x. r *\<^sub>R f x) ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3577 |
using bounded_scaling[of "f ` S" r] assms |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3578 |
by (auto simp: image_image) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3579 |
|
| 33175 | 3580 |
lemma bounded_translation: |
3581 |
fixes S :: "'a::real_normed_vector set" |
|
| 52624 | 3582 |
assumes "bounded S" |
3583 |
shows "bounded ((\<lambda>x. a + x) ` S)" |
|
| 53282 | 3584 |
proof - |
| 53640 | 3585 |
from assms obtain b where b: "b > 0" "\<forall>x\<in>S. norm x \<le> b" |
| 52624 | 3586 |
unfolding bounded_pos by auto |
3587 |
{
|
|
3588 |
fix x |
|
| 53640 | 3589 |
assume "x \<in> S" |
| 53282 | 3590 |
then have "norm (a + x) \<le> b + norm a" |
| 52624 | 3591 |
using norm_triangle_ineq[of a x] b by auto |
| 33175 | 3592 |
} |
| 53282 | 3593 |
then show ?thesis |
| 52624 | 3594 |
unfolding bounded_pos |
3595 |
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
|
3596 |
by (auto intro!: exI[of _ "b + norm a"]) |
| 33175 | 3597 |
qed |
3598 |
||
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
3599 |
lemma bounded_translation_minus: |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
3600 |
fixes S :: "'a::real_normed_vector set" |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
3601 |
shows "bounded S \<Longrightarrow> bounded ((\<lambda>x. x - a) ` S)" |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
3602 |
using bounded_translation [of S "-a"] by simp |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
3603 |
|
|
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
3604 |
lemma bounded_uminus [simp]: |
| 62466 | 3605 |
fixes X :: "'a::real_normed_vector set" |
|
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
3606 |
shows "bounded (uminus ` X) \<longleftrightarrow> bounded X" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3607 |
by (auto simp: bounded_def dist_norm; rule_tac x="-x" in exI; force simp: add.commute norm_minus_commute) |
|
60615
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents:
60585
diff
changeset
|
3608 |
|
|
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3609 |
lemma uminus_bounded_comp [simp]: |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3610 |
fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3611 |
shows "bounded ((\<lambda>x. - f x) ` S) \<longleftrightarrow> bounded (f ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3612 |
using bounded_uminus[of "f ` S"] |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3613 |
by (auto simp: image_image) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3614 |
|
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3615 |
lemma bounded_plus_comp: |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3616 |
fixes f g::"'a \<Rightarrow> 'b::real_normed_vector" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3617 |
assumes "bounded (f ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3618 |
assumes "bounded (g ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3619 |
shows "bounded ((\<lambda>x. f x + g x) ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3620 |
proof - |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3621 |
{
|
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3622 |
fix B C |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3623 |
assume "\<And>x. x\<in>S \<Longrightarrow> norm (f x) \<le> B" "\<And>x. x\<in>S \<Longrightarrow> norm (g x) \<le> C" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3624 |
then have "\<And>x. x \<in> S \<Longrightarrow> norm (f x + g x) \<le> B + C" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3625 |
by (auto intro!: norm_triangle_le add_mono) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3626 |
} then show ?thesis |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3627 |
using assms by (fastforce simp: bounded_iff) |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3628 |
qed |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3629 |
|
|
67683
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3630 |
lemma bounded_plus: |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3631 |
fixes S ::"'a::real_normed_vector set" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3632 |
assumes "bounded S" "bounded T" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3633 |
shows "bounded ((\<lambda>(x,y). x + y) ` (S \<times> T))" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3634 |
using bounded_plus_comp [of fst "S \<times> T" snd] assms |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3635 |
by (auto simp: split_def split: if_split_asm) |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3636 |
|
|
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3637 |
lemma bounded_minus_comp: |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3638 |
"bounded (f ` S) \<Longrightarrow> bounded (g ` S) \<Longrightarrow> bounded ((\<lambda>x. f x - g x) ` S)" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3639 |
for f g::"'a \<Rightarrow> 'b::real_normed_vector" |
|
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3640 |
using bounded_plus_comp[of "f" S "\<lambda>x. - g x"] |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
3641 |
by auto |
|
65204
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
immler
parents:
65038
diff
changeset
|
3642 |
|
|
67683
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3643 |
lemma bounded_minus: |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3644 |
fixes S ::"'a::real_normed_vector set" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3645 |
assumes "bounded S" "bounded T" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3646 |
shows "bounded ((\<lambda>(x,y). x - y) ` (S \<times> T))" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3647 |
using bounded_minus_comp [of fst "S \<times> T" snd] assms |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3648 |
by (auto simp: split_def split: if_split_asm) |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
3649 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
3650 |
|
| 60420 | 3651 |
subsection \<open>Compactness\<close> |
3652 |
||
3653 |
subsubsection \<open>Bolzano-Weierstrass property\<close> |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3654 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
3655 |
proposition heine_borel_imp_bolzano_weierstrass: |
| 53640 | 3656 |
assumes "compact s" |
3657 |
and "infinite t" |
|
3658 |
and "t \<subseteq> s" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3659 |
shows "\<exists>x \<in> s. x islimpt t" |
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
3660 |
proof (rule ccontr) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3661 |
assume "\<not> (\<exists>x \<in> s. x islimpt t)" |
| 53640 | 3662 |
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 | 3663 |
unfolding islimpt_def |
| 53282 | 3664 |
using bchoice[of s "\<lambda> x T. x \<in> T \<and> open T \<and> (\<forall>y\<in>t. y \<in> T \<longrightarrow> y = x)"] |
3665 |
by auto |
|
| 53640 | 3666 |
obtain g where g: "g \<subseteq> {t. \<exists>x. x \<in> s \<and> t = f x}" "finite g" "s \<subseteq> \<Union>g"
|
| 52624 | 3667 |
using assms(1)[unfolded compact_eq_heine_borel, THEN spec[where x="{t. \<exists>x. x\<in>s \<and> t = f x}"]]
|
3668 |
using f by auto |
|
| 53640 | 3669 |
from g(1,3) have g':"\<forall>x\<in>g. \<exists>xa \<in> s. x = f xa" |
3670 |
by auto |
|
| 52624 | 3671 |
{
|
3672 |
fix x y |
|
| 53640 | 3673 |
assume "x \<in> t" "y \<in> t" "f x = f y" |
| 53282 | 3674 |
then have "x \<in> f x" "y \<in> f x \<longrightarrow> y = x" |
| 60420 | 3675 |
using f[THEN bspec[where x=x]] and \<open>t \<subseteq> s\<close> by auto |
| 53282 | 3676 |
then have "x = y" |
| 60420 | 3677 |
using \<open>f x = f y\<close> and f[THEN bspec[where x=y]] and \<open>y \<in> t\<close> and \<open>t \<subseteq> s\<close> |
| 53640 | 3678 |
by auto |
| 52624 | 3679 |
} |
| 53282 | 3680 |
then have "inj_on f t" |
| 52624 | 3681 |
unfolding inj_on_def by simp |
| 53282 | 3682 |
then have "infinite (f ` t)" |
| 52624 | 3683 |
using assms(2) using finite_imageD by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3684 |
moreover |
| 52624 | 3685 |
{
|
3686 |
fix x |
|
| 53640 | 3687 |
assume "x \<in> t" "f x \<notin> g" |
| 60420 | 3688 |
from g(3) assms(3) \<open>x \<in> t\<close> obtain h where "h \<in> g" and "x \<in> h" |
| 53640 | 3689 |
by auto |
3690 |
then obtain y where "y \<in> s" "h = f y" |
|
| 52624 | 3691 |
using g'[THEN bspec[where x=h]] by auto |
| 53282 | 3692 |
then have "y = x" |
| 60420 | 3693 |
using f[THEN bspec[where x=y]] and \<open>x\<in>t\<close> and \<open>x\<in>h\<close>[unfolded \<open>h = f y\<close>] |
| 53640 | 3694 |
by auto |
| 53282 | 3695 |
then have False |
| 60420 | 3696 |
using \<open>f x \<notin> g\<close> \<open>h \<in> g\<close> unfolding \<open>h = f y\<close> |
| 53640 | 3697 |
by auto |
| 52624 | 3698 |
} |
| 53282 | 3699 |
then have "f ` t \<subseteq> g" by auto |
| 52624 | 3700 |
ultimately show False |
3701 |
using g(2) using finite_subset by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3702 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3703 |
|
|
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
|
3704 |
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
|
3705 |
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
|
3706 |
assumes l: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> range f)" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
3707 |
shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3708 |
proof - |
| 55522 | 3709 |
from countable_basis_at_decseq[of l] |
3710 |
obtain A where A: |
|
3711 |
"\<And>i. open (A i)" |
|
3712 |
"\<And>i. l \<in> A i" |
|
3713 |
"\<And>S. open S \<Longrightarrow> l \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
|
3714 |
by blast |
|
| 63040 | 3715 |
define s where "s n i = (SOME j. i < j \<and> f j \<in> A (Suc n))" for n i |
| 52624 | 3716 |
{
|
3717 |
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
|
3718 |
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
|
3719 |
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
|
3720 |
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
|
3721 |
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
|
3722 |
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
|
3723 |
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
|
3724 |
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
|
3725 |
by (auto simp: not_le) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3726 |
then have "i < s n i" "f (s n i) \<in> A (Suc n)" |
| 52624 | 3727 |
unfolding s_def by (auto intro: someI2_ex) |
3728 |
} |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3729 |
note s = this |
| 63040 | 3730 |
define r where "r = rec_nat (s 0 0) s" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
3731 |
have "strict_mono r" |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
3732 |
by (auto simp: r_def s strict_mono_Suc_iff) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3733 |
moreover |
| 61969 | 3734 |
have "(\<lambda>n. f (r n)) \<longlonglongrightarrow> l" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3735 |
proof (rule topological_tendstoI) |
| 53282 | 3736 |
fix S |
3737 |
assume "open S" "l \<in> S" |
|
| 53640 | 3738 |
with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
3739 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3740 |
moreover |
| 52624 | 3741 |
{
|
3742 |
fix i |
|
| 53282 | 3743 |
assume "Suc 0 \<le> i" |
3744 |
then have "f (r i) \<in> A i" |
|
| 52624 | 3745 |
by (cases i) (simp_all add: r_def s) |
3746 |
} |
|
3747 |
then have "eventually (\<lambda>i. f (r i) \<in> A i) sequentially" |
|
3748 |
by (auto simp: eventually_sequentially) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3749 |
ultimately show "eventually (\<lambda>i. f (r i) \<in> S) sequentially" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3750 |
by eventually_elim auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3751 |
qed |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
3752 |
ultimately show "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3753 |
by (auto simp: convergent_def comp_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3754 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3755 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3756 |
lemma sequence_infinite_lemma: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3757 |
fixes f :: "nat \<Rightarrow> 'a::t1_space" |
| 53282 | 3758 |
assumes "\<forall>n. f n \<noteq> l" |
| 61973 | 3759 |
and "(f \<longlongrightarrow> l) sequentially" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3760 |
shows "infinite (range f)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3761 |
proof |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3762 |
assume "finite (range f)" |
| 53640 | 3763 |
then have "closed (range f)" |
3764 |
by (rule finite_imp_closed) |
|
3765 |
then have "open (- range f)" |
|
3766 |
by (rule open_Compl) |
|
3767 |
from assms(1) have "l \<in> - range f" |
|
3768 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3769 |
from assms(2) have "eventually (\<lambda>n. f n \<in> - range f) sequentially" |
| 60420 | 3770 |
using \<open>open (- range f)\<close> \<open>l \<in> - range f\<close> |
| 53640 | 3771 |
by (rule topological_tendstoD) |
3772 |
then show False |
|
3773 |
unfolding eventually_sequentially |
|
3774 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3775 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3776 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3777 |
lemma closure_insert: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3778 |
fixes x :: "'a::t1_space" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3779 |
shows "closure (insert x s) = insert x (closure s)" |
| 52624 | 3780 |
apply (rule closure_unique) |
3781 |
apply (rule insert_mono [OF closure_subset]) |
|
3782 |
apply (rule closed_insert [OF closed_closure]) |
|
3783 |
apply (simp add: closure_minimal) |
|
3784 |
done |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3785 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3786 |
lemma islimpt_insert: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3787 |
fixes x :: "'a::t1_space" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3788 |
shows "x islimpt (insert a s) \<longleftrightarrow> x islimpt s" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3789 |
proof |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3790 |
assume *: "x islimpt (insert a s)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3791 |
show "x islimpt s" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3792 |
proof (rule islimptI) |
| 53282 | 3793 |
fix t |
3794 |
assume t: "x \<in> t" "open t" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3795 |
show "\<exists>y\<in>s. y \<in> t \<and> y \<noteq> x" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3796 |
proof (cases "x = a") |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3797 |
case True |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3798 |
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
|
3799 |
using * t by (rule islimptE) |
| 60420 | 3800 |
with \<open>x = a\<close> show ?thesis by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3801 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3802 |
case False |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3803 |
with t have t': "x \<in> t - {a}" "open (t - {a})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3804 |
by (simp_all add: open_Diff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3805 |
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
|
3806 |
using * t' by (rule islimptE) |
| 53282 | 3807 |
then show ?thesis by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3808 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3809 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3810 |
next |
| 53282 | 3811 |
assume "x islimpt s" |
3812 |
then show "x islimpt (insert a s)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3813 |
by (rule islimpt_subset) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3814 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3815 |
|
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3816 |
lemma islimpt_finite: |
|
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3817 |
fixes x :: "'a::t1_space" |
|
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3818 |
shows "finite s \<Longrightarrow> \<not> x islimpt s" |
| 52624 | 3819 |
by (induct set: finite) (simp_all add: islimpt_insert) |
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3820 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3821 |
lemma islimpt_Un_finite: |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3822 |
fixes x :: "'a::t1_space" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3823 |
shows "finite s \<Longrightarrow> x islimpt (s \<union> t) \<longleftrightarrow> x islimpt t" |
| 52624 | 3824 |
by (simp add: islimpt_Un islimpt_finite) |
|
50897
078590669527
generalize lemma islimpt_finite to class t1_space
huffman
parents:
50884
diff
changeset
|
3825 |
|
|
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
|
3826 |
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
|
3827 |
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
|
3828 |
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
|
3829 |
proof (safe intro!: islimptI) |
| 53282 | 3830 |
fix U |
3831 |
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
|
3832 |
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
|
3833 |
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
|
3834 |
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
|
3835 |
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
|
3836 |
next |
| 53282 | 3837 |
fix T |
3838 |
assume *: "\<forall>U. l\<in>U \<longrightarrow> open U \<longrightarrow> infinite (U \<inter> S)" "l \<in> T" "open T" |
|
3839 |
then have "infinite (T \<inter> S - {l})"
|
|
3840 |
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
|
3841 |
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
|
3842 |
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
|
3843 |
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
|
3844 |
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
|
3845 |
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
|
3846 |
|
| 63938 | 3847 |
corollary infinite_openin: |
3848 |
fixes S :: "'a :: t1_space set" |
|
3849 |
shows "\<lbrakk>openin (subtopology euclidean U) S; x \<in> S; x islimpt U\<rbrakk> \<Longrightarrow> infinite S" |
|
3850 |
by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute) |
|
3851 |
||
|
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
|
3852 |
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
|
3853 |
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
|
3854 |
assumes l: "l islimpt (range f)" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
3855 |
shows "\<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
|
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
|
3856 |
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
|
3857 |
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
|
3858 |
|
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3859 |
lemma islimpt_eq_infinite_ball: "x islimpt S \<longleftrightarrow> (\<forall>e>0. infinite(S \<inter> ball x e))" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3860 |
apply (simp add: islimpt_eq_acc_point, safe) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3861 |
apply (metis Int_commute open_ball centre_in_ball) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3862 |
by (metis open_contains_ball Int_mono finite_subset inf_commute subset_refl) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3863 |
|
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3864 |
lemma islimpt_eq_infinite_cball: "x islimpt S \<longleftrightarrow> (\<forall>e>0. infinite(S \<inter> cball x e))" |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3865 |
apply (simp add: islimpt_eq_infinite_ball, safe) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3866 |
apply (meson Int_mono ball_subset_cball finite_subset order_refl) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3867 |
by (metis open_ball centre_in_ball finite_Int inf.absorb_iff2 inf_assoc open_contains_cball_eq) |
|
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
3868 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3869 |
lemma sequence_unique_limpt: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3870 |
fixes f :: "nat \<Rightarrow> 'a::t2_space" |
| 61973 | 3871 |
assumes "(f \<longlongrightarrow> l) sequentially" |
| 53282 | 3872 |
and "l' islimpt (range f)" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3873 |
shows "l' = l" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3874 |
proof (rule ccontr) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3875 |
assume "l' \<noteq> l" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3876 |
obtain s t where "open s" "open t" "l' \<in> s" "l \<in> t" "s \<inter> t = {}"
|
| 60420 | 3877 |
using hausdorff [OF \<open>l' \<noteq> l\<close>] by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3878 |
have "eventually (\<lambda>n. f n \<in> t) sequentially" |
| 60420 | 3879 |
using assms(1) \<open>open t\<close> \<open>l \<in> t\<close> by (rule topological_tendstoD) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3880 |
then obtain N where "\<forall>n\<ge>N. f n \<in> t" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3881 |
unfolding eventually_sequentially by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3882 |
|
| 53282 | 3883 |
have "UNIV = {..<N} \<union> {N..}"
|
3884 |
by auto |
|
3885 |
then have "l' islimpt (f ` ({..<N} \<union> {N..}))"
|
|
3886 |
using assms(2) by simp |
|
3887 |
then have "l' islimpt (f ` {..<N} \<union> f ` {N..})"
|
|
3888 |
by (simp add: image_Un) |
|
3889 |
then have "l' islimpt (f ` {N..})"
|
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3890 |
by (simp add: islimpt_Un_finite) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3891 |
then obtain y where "y \<in> f ` {N..}" "y \<in> s" "y \<noteq> l'"
|
| 60420 | 3892 |
using \<open>l' \<in> s\<close> \<open>open s\<close> by (rule islimptE) |
| 53282 | 3893 |
then obtain n where "N \<le> n" "f n \<in> s" "f n \<noteq> l'" |
3894 |
by auto |
|
| 60420 | 3895 |
with \<open>\<forall>n\<ge>N. f n \<in> t\<close> have "f n \<in> s \<inter> t" |
| 53282 | 3896 |
by simp |
| 60420 | 3897 |
with \<open>s \<inter> t = {}\<close> show False
|
| 53282 | 3898 |
by simp |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3899 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3900 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3901 |
lemma bolzano_weierstrass_imp_closed: |
| 53640 | 3902 |
fixes s :: "'a::{first_countable_topology,t2_space} set"
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3903 |
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
|
3904 |
shows "closed s" |
| 52624 | 3905 |
proof - |
3906 |
{
|
|
3907 |
fix x l |
|
| 61973 | 3908 |
assume as: "\<forall>n::nat. x n \<in> s" "(x \<longlongrightarrow> l) sequentially" |
| 53282 | 3909 |
then have "l \<in> s" |
| 52624 | 3910 |
proof (cases "\<forall>n. x n \<noteq> l") |
3911 |
case False |
|
| 53282 | 3912 |
then show "l\<in>s" using as(1) by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3913 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3914 |
case True note cas = this |
| 52624 | 3915 |
with as(2) have "infinite (range x)" |
3916 |
using sequence_infinite_lemma[of x l] by auto |
|
3917 |
then obtain l' where "l'\<in>s" "l' islimpt (range x)" |
|
3918 |
using assms[THEN spec[where x="range x"]] as(1) by auto |
|
| 53282 | 3919 |
then show "l\<in>s" using sequence_unique_limpt[of x l l'] |
| 52624 | 3920 |
using as cas by auto |
3921 |
qed |
|
3922 |
} |
|
| 53282 | 3923 |
then show ?thesis |
3924 |
unfolding closed_sequential_limits by fast |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3925 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3926 |
|
| 50944 | 3927 |
lemma compact_imp_bounded: |
| 52624 | 3928 |
assumes "compact U" |
3929 |
shows "bounded U" |
|
| 50944 | 3930 |
proof - |
| 52624 | 3931 |
have "compact U" "\<forall>x\<in>U. open (ball x 1)" "U \<subseteq> (\<Union>x\<in>U. ball x 1)" |
3932 |
using assms by auto |
|
| 50944 | 3933 |
then obtain D where D: "D \<subseteq> U" "finite D" "U \<subseteq> (\<Union>x\<in>D. ball x 1)" |
|
65585
a043de9ad41e
Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents:
65204
diff
changeset
|
3934 |
by (metis compactE_image) |
| 60420 | 3935 |
from \<open>finite D\<close> have "bounded (\<Union>x\<in>D. ball x 1)" |
| 50955 | 3936 |
by (simp add: bounded_UN) |
| 60420 | 3937 |
then show "bounded U" using \<open>U \<subseteq> (\<Union>x\<in>D. ball x 1)\<close> |
| 50955 | 3938 |
by (rule bounded_subset) |
| 50944 | 3939 |
qed |
3940 |
||
| 60420 | 3941 |
text\<open>In particular, some common special cases.\<close> |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3942 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3943 |
lemma compact_Un [intro]: |
| 53291 | 3944 |
assumes "compact s" |
3945 |
and "compact t" |
|
| 53282 | 3946 |
shows " compact (s \<union> t)" |
| 50898 | 3947 |
proof (rule compactI) |
| 52624 | 3948 |
fix f |
3949 |
assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f" |
|
| 60420 | 3950 |
from * \<open>compact s\<close> obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'" |
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
55927
diff
changeset
|
3951 |
unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) |
| 52624 | 3952 |
moreover |
| 60420 | 3953 |
from * \<open>compact t\<close> obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'" |
|
56073
29e308b56d23
enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents:
55927
diff
changeset
|
3954 |
unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3955 |
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
|
3956 |
by (auto intro!: exI[of _ "s' \<union> t'"]) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3957 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3958 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3959 |
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
|
3960 |
by (induct set: finite) auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3961 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3962 |
lemma compact_UN [intro]: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3963 |
"finite A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> compact (B x)) \<Longrightarrow> compact (\<Union>x\<in>A. B x)" |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
3964 |
by (rule compact_Union) auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3965 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3966 |
lemma closed_Int_compact [intro]: |
| 53282 | 3967 |
assumes "closed s" |
3968 |
and "compact t" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3969 |
shows "compact (s \<inter> t)" |
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3970 |
using compact_Int_closed [of t s] assms |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3971 |
by (simp add: Int_commute) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3972 |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3973 |
lemma compact_Int [intro]: |
| 50898 | 3974 |
fixes s t :: "'a :: t2_space set" |
| 53282 | 3975 |
assumes "compact s" |
3976 |
and "compact t" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3977 |
shows "compact (s \<inter> t)" |
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3978 |
using assms by (intro compact_Int_closed compact_imp_closed) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3979 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3980 |
lemma compact_sing [simp]: "compact {a}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3981 |
unfolding compact_eq_heine_borel by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3982 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3983 |
lemma compact_insert [simp]: |
| 53282 | 3984 |
assumes "compact s" |
3985 |
shows "compact (insert x s)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3986 |
proof - |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3987 |
have "compact ({x} \<union> s)"
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
3988 |
using compact_sing assms by (rule compact_Un) |
| 53282 | 3989 |
then show ?thesis by simp |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3990 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3991 |
|
| 52624 | 3992 |
lemma finite_imp_compact: "finite s \<Longrightarrow> compact s" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3993 |
by (induct set: finite) simp_all |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3994 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3995 |
lemma open_delete: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3996 |
fixes s :: "'a::t1_space set" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3997 |
shows "open s \<Longrightarrow> open (s - {x})"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3998 |
by (simp add: open_Diff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
3999 |
|
|
62948
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4000 |
lemma openin_delete: |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4001 |
fixes a :: "'a :: t1_space" |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4002 |
shows "openin (subtopology euclidean u) s |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4003 |
\<Longrightarrow> openin (subtopology euclidean u) (s - {a})"
|
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4004 |
by (metis Int_Diff open_delete openin_open) |
|
7700f467892b
lots of new theorems for multivariate analysis
paulson <lp15@cam.ac.uk>
parents:
62843
diff
changeset
|
4005 |
|
| 60420 | 4006 |
text\<open>Compactness expressed with filters\<close> |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4007 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4008 |
lemma closure_iff_nhds_not_empty: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4009 |
"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
|
4010 |
proof safe |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4011 |
assume x: "x \<in> closure X" |
| 53282 | 4012 |
fix S A |
4013 |
assume "open S" "x \<in> S" "X \<inter> A = {}" "S \<subseteq> A"
|
|
4014 |
then have "x \<notin> closure (-S)" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4015 |
by (auto simp: closure_complement subset_eq[symmetric] intro: interiorI) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4016 |
with x have "x \<in> closure X - closure (-S)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4017 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4018 |
also have "\<dots> \<subseteq> closure (X \<inter> S)" |
| 63128 | 4019 |
using \<open>open S\<close> open_Int_closure_subset[of S X] by (simp add: closed_Compl ac_simps) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4020 |
finally have "X \<inter> S \<noteq> {}" by auto
|
| 60420 | 4021 |
then show False using \<open>X \<inter> A = {}\<close> \<open>S \<subseteq> A\<close> by auto
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4022 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4023 |
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
|
4024 |
from this[THEN spec, of "- X", THEN spec, of "- closure X"] |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4025 |
show "x \<in> closure X" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4026 |
by (simp add: closure_subset open_Compl) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4027 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4028 |
|
|
66641
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
4029 |
corollary closure_Int_ball_not_empty: |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
4030 |
assumes "S \<subseteq> closure T" "x \<in> S" "r > 0" |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
4031 |
shows "T \<inter> ball x r \<noteq> {}"
|
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
4032 |
using assms centre_in_ball closure_iff_nhds_not_empty by blast |
|
ff2e0115fea4
Simplicial complexes and triangulations; Baire Category Theorem
paulson <lp15@cam.ac.uk>
parents:
66466
diff
changeset
|
4033 |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4034 |
lemma compact_filter: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4035 |
"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
|
4036 |
proof (intro allI iffI impI compact_fip[THEN iffD2] notI) |
| 53282 | 4037 |
fix F |
4038 |
assume "compact U" |
|
4039 |
assume F: "F \<noteq> bot" "eventually (\<lambda>x. x \<in> U) F" |
|
4040 |
then have "U \<noteq> {}"
|
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4041 |
by (auto simp: eventually_False) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4042 |
|
| 63040 | 4043 |
define Z where "Z = closure ` {A. eventually (\<lambda>x. x \<in> A) F}"
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4044 |
then have "\<forall>z\<in>Z. closed z" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4045 |
by auto |
| 53282 | 4046 |
moreover |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4047 |
have ev_Z: "\<And>z. z \<in> Z \<Longrightarrow> eventually (\<lambda>x. x \<in> z) F" |
| 61810 | 4048 |
unfolding Z_def by (auto elim: eventually_mono intro: set_mp[OF closure_subset]) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4049 |
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
|
4050 |
proof (intro allI impI) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4051 |
fix B assume "finite B" "B \<subseteq> Z" |
| 60420 | 4052 |
with \<open>finite B\<close> ev_Z F(2) have "eventually (\<lambda>x. x \<in> U \<inter> (\<Inter>B)) F" |
|
60040
1fa1023b13b9
move MOST and INFM in Infinite_Set to Filter; change them to abbreviations over the cofinite filter
hoelzl
parents:
60017
diff
changeset
|
4053 |
by (auto simp: eventually_ball_finite_distrib eventually_conj_iff) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4054 |
with F show "U \<inter> \<Inter>B \<noteq> {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4055 |
by (intro notI) (simp add: eventually_False) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4056 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4057 |
ultimately have "U \<inter> \<Inter>Z \<noteq> {}"
|
| 60420 | 4058 |
using \<open>compact U\<close> unfolding compact_fip by blast |
| 53282 | 4059 |
then obtain x where "x \<in> U" and x: "\<And>z. z \<in> Z \<Longrightarrow> x \<in> z" |
4060 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4061 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4062 |
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
|
4063 |
unfolding eventually_inf eventually_nhds |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4064 |
proof safe |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4065 |
fix P Q R S |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4066 |
assume "eventually R F" "open S" "x \<in> S" |
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
4067 |
with open_Int_closure_eq_empty[of S "{x. R x}"] x[of "closure {x. R x}"]
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4068 |
have "S \<inter> {x. R x} \<noteq> {}" by (auto simp: Z_def)
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4069 |
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
|
4070 |
ultimately show False by (auto simp: set_eq_iff) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4071 |
qed |
| 60420 | 4072 |
with \<open>x \<in> U\<close> show "\<exists>x\<in>U. inf (nhds x) F \<noteq> bot" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4073 |
by (metis eventually_bot) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4074 |
next |
| 53282 | 4075 |
fix A |
4076 |
assume A: "\<forall>a\<in>A. closed a" "\<forall>B\<subseteq>A. finite B \<longrightarrow> U \<inter> \<Inter>B \<noteq> {}" "U \<inter> \<Inter>A = {}"
|
|
| 63040 | 4077 |
define F where "F = (INF a:insert U A. principal a)" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4078 |
have "F \<noteq> bot" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4079 |
unfolding F_def |
| 57276 | 4080 |
proof (rule INF_filter_not_bot) |
| 63540 | 4081 |
fix X |
4082 |
assume X: "X \<subseteq> insert U A" "finite X" |
|
4083 |
with A(2)[THEN spec, of "X - {U}"] have "U \<inter> \<Inter>(X - {U}) \<noteq> {}"
|
|
| 53282 | 4084 |
by auto |
| 63540 | 4085 |
with X show "(INF a:X. principal a) \<noteq> bot" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4086 |
by (auto simp: INF_principal_finite principal_eq_bot_iff) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4087 |
qed |
| 57276 | 4088 |
moreover |
4089 |
have "F \<le> principal U" |
|
4090 |
unfolding F_def by auto |
|
4091 |
then have "eventually (\<lambda>x. x \<in> U) F" |
|
4092 |
by (auto simp: le_filter_def eventually_principal) |
|
| 53282 | 4093 |
moreover |
4094 |
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
|
4095 |
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
|
4096 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4097 |
|
| 57276 | 4098 |
{ fix V assume "V \<in> A"
|
4099 |
then have "F \<le> principal V" |
|
4100 |
unfolding F_def by (intro INF_lower2[of V]) auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4101 |
then have V: "eventually (\<lambda>x. x \<in> V) F" |
| 57276 | 4102 |
by (auto simp: le_filter_def eventually_principal) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4103 |
have "x \<in> closure V" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4104 |
unfolding closure_iff_nhds_not_empty |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4105 |
proof (intro impI allI) |
| 53282 | 4106 |
fix S A |
4107 |
assume "open S" "x \<in> S" "S \<subseteq> A" |
|
4108 |
then have "eventually (\<lambda>x. x \<in> A) (nhds x)" |
|
4109 |
by (auto simp: eventually_nhds) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4110 |
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
|
4111 |
by (auto simp: eventually_inf) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4112 |
with x show "V \<inter> A \<noteq> {}"
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4113 |
by (auto simp del: Int_iff simp add: trivial_limit_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4114 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4115 |
then have "x \<in> V" |
| 60420 | 4116 |
using \<open>V \<in> A\<close> A(1) by simp |
| 53282 | 4117 |
} |
| 60420 | 4118 |
with \<open>x\<in>U\<close> have "x \<in> U \<inter> \<Inter>A" by auto |
4119 |
with \<open>U \<inter> \<Inter>A = {}\<close> show False by auto
|
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4120 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4121 |
|
| 67962 | 4122 |
definition%important "countably_compact U \<longleftrightarrow> |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4123 |
(\<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
|
4124 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4125 |
lemma countably_compactE: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4126 |
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
|
4127 |
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
|
4128 |
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
|
4129 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4130 |
lemma countably_compactI: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4131 |
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
|
4132 |
shows "countably_compact s" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4133 |
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
|
4134 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4135 |
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
|
4136 |
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
|
4137 |
|
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4138 |
lemma countably_compact_imp_compact: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4139 |
assumes "countably_compact U" |
| 53282 | 4140 |
and ccover: "countable B" "\<forall>b\<in>B. open b" |
4141 |
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
|
4142 |
shows "compact U" |
| 60420 | 4143 |
using \<open>countably_compact U\<close> |
| 53282 | 4144 |
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
|
4145 |
proof safe |
| 53282 | 4146 |
fix A |
4147 |
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
|
4148 |
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)" |
| 63040 | 4149 |
moreover define C where "C = {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
|
4150 |
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
|
4151 |
unfolding C_def using ccover by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4152 |
moreover |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4153 |
have "\<Union>A \<inter> U \<subseteq> \<Union>C" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4154 |
proof safe |
| 53282 | 4155 |
fix x a |
4156 |
assume "x \<in> U" "x \<in> a" "a \<in> A" |
|
4157 |
with basis[of a x] A obtain b where "b \<in> B" "x \<in> b" "b \<inter> U \<subseteq> a" |
|
4158 |
by blast |
|
| 60420 | 4159 |
with \<open>a \<in> A\<close> show "x \<in> \<Union>C" |
| 53282 | 4160 |
unfolding C_def by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4161 |
qed |
| 60420 | 4162 |
then have "U \<subseteq> \<Union>C" using \<open>U \<subseteq> \<Union>A\<close> by auto |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
4163 |
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
|
4164 |
using * by metis |
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
4165 |
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
|
4166 |
by (auto simp: C_def) |
| 55522 | 4167 |
then obtain f where "\<forall>t\<in>T. f t \<in> A \<and> t \<inter> U \<subseteq> f t" |
4168 |
unfolding bchoice_iff Bex_def .. |
|
|
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53291
diff
changeset
|
4169 |
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
|
4170 |
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
|
4171 |
qed |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4172 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4173 |
proposition countably_compact_imp_compact_second_countable: |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4174 |
"countably_compact U \<Longrightarrow> compact (U :: 'a :: second_countable_topology set)" |
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4175 |
proof (rule countably_compact_imp_compact) |
| 53282 | 4176 |
fix T and x :: 'a |
4177 |
assume "open T" "x \<in> T" |
|
| 55522 | 4178 |
from topological_basisE[OF is_basis this] obtain b where |
4179 |
"b \<in> (SOME B. countable B \<and> topological_basis B)" "x \<in> b" "b \<subseteq> T" . |
|
| 53282 | 4180 |
then show "\<exists>b\<in>SOME B. countable B \<and> topological_basis B. x \<in> b \<and> b \<inter> U \<subseteq> T" |
| 55522 | 4181 |
by blast |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4182 |
qed (insert countable_basis topological_basis_open[OF is_basis], auto) |
| 36437 | 4183 |
|
|
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
|
4184 |
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
|
4185 |
"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
|
4186 |
using countably_compact_imp_compact_second_countable compact_imp_countably_compact by blast |
| 53282 | 4187 |
|
| 60420 | 4188 |
subsubsection\<open>Sequential compactness\<close> |
| 33175 | 4189 |
|
| 67962 | 4190 |
definition%important seq_compact :: "'a::topological_space set \<Rightarrow> bool" |
| 53282 | 4191 |
where "seq_compact S \<longleftrightarrow> |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4192 |
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially))" |
| 33175 | 4193 |
|
| 54070 | 4194 |
lemma seq_compactI: |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4195 |
assumes "\<And>f. \<forall>n. f n \<in> S \<Longrightarrow> \<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 54070 | 4196 |
shows "seq_compact S" |
4197 |
unfolding seq_compact_def using assms by fast |
|
4198 |
||
4199 |
lemma seq_compactE: |
|
4200 |
assumes "seq_compact S" "\<forall>n. f n \<in> S" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4201 |
obtains l r where "l \<in> S" "strict_mono (r :: nat \<Rightarrow> nat)" "((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 54070 | 4202 |
using assms unfolding seq_compact_def by fast |
4203 |
||
4204 |
lemma closed_sequentially: (* TODO: move upwards *) |
|
| 61969 | 4205 |
assumes "closed s" and "\<forall>n. f n \<in> s" and "f \<longlonglongrightarrow> l" |
| 54070 | 4206 |
shows "l \<in> s" |
4207 |
proof (rule ccontr) |
|
4208 |
assume "l \<notin> s" |
|
| 61969 | 4209 |
with \<open>closed s\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "eventually (\<lambda>n. f n \<in> - s) sequentially" |
| 54070 | 4210 |
by (fast intro: topological_tendstoD) |
| 60420 | 4211 |
with \<open>\<forall>n. f n \<in> s\<close> show "False" |
| 54070 | 4212 |
by simp |
4213 |
qed |
|
4214 |
||
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
4215 |
lemma seq_compact_Int_closed: |
| 54070 | 4216 |
assumes "seq_compact s" and "closed t" |
4217 |
shows "seq_compact (s \<inter> t)" |
|
4218 |
proof (rule seq_compactI) |
|
4219 |
fix f assume "\<forall>n::nat. f n \<in> s \<inter> t" |
|
4220 |
hence "\<forall>n. f n \<in> s" and "\<forall>n. f n \<in> t" |
|
4221 |
by simp_all |
|
| 60420 | 4222 |
from \<open>seq_compact s\<close> and \<open>\<forall>n. f n \<in> s\<close> |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4223 |
obtain l r where "l \<in> s" and r: "strict_mono r" and l: "(f \<circ> r) \<longlonglongrightarrow> l" |
| 54070 | 4224 |
by (rule seq_compactE) |
| 60420 | 4225 |
from \<open>\<forall>n. f n \<in> t\<close> have "\<forall>n. (f \<circ> r) n \<in> t" |
| 54070 | 4226 |
by simp |
| 60420 | 4227 |
from \<open>closed t\<close> and this and l have "l \<in> t" |
| 54070 | 4228 |
by (rule closed_sequentially) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4229 |
with \<open>l \<in> s\<close> and r and l show "\<exists>l\<in>s \<inter> t. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
| 54070 | 4230 |
by fast |
4231 |
qed |
|
4232 |
||
4233 |
lemma seq_compact_closed_subset: |
|
4234 |
assumes "closed s" and "s \<subseteq> t" and "seq_compact t" |
|
4235 |
shows "seq_compact s" |
|
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
4236 |
using assms seq_compact_Int_closed [of t s] by (simp add: Int_absorb1) |
| 54070 | 4237 |
|
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4238 |
lemma seq_compact_imp_countably_compact: |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4239 |
fixes U :: "'a :: first_countable_topology set" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4240 |
assumes "seq_compact U" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4241 |
shows "countably_compact U" |
|
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4242 |
proof (safe intro!: countably_compactI) |
| 52624 | 4243 |
fix A |
4244 |
assume A: "\<forall>a\<in>A. open a" "U \<subseteq> \<Union>A" "countable A" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4245 |
have subseq: "\<And>X. range X \<subseteq> U \<Longrightarrow> \<exists>r x. x \<in> U \<and> strict_mono (r :: nat \<Rightarrow> nat) \<and> (X \<circ> r) \<longlonglongrightarrow> x" |
| 60420 | 4246 |
using \<open>seq_compact U\<close> by (fastforce simp: seq_compact_def subset_eq) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4247 |
show "\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4248 |
proof cases |
| 52624 | 4249 |
assume "finite A" |
4250 |
with A show ?thesis by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4251 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4252 |
assume "infinite A" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4253 |
then have "A \<noteq> {}" by auto
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4254 |
show ?thesis |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4255 |
proof (rule ccontr) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4256 |
assume "\<not> (\<exists>T\<subseteq>A. finite T \<and> U \<subseteq> \<Union>T)" |
| 53282 | 4257 |
then have "\<forall>T. \<exists>x. T \<subseteq> A \<and> finite T \<longrightarrow> (x \<in> U - \<Union>T)" |
4258 |
by auto |
|
4259 |
then obtain X' where T: "\<And>T. T \<subseteq> A \<Longrightarrow> finite T \<Longrightarrow> X' T \<in> U - \<Union>T" |
|
4260 |
by metis |
|
| 63040 | 4261 |
define X where "X n = X' (from_nat_into A ` {.. n})" for n
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4262 |
have X: "\<And>n. X n \<in> U - (\<Union>i\<le>n. from_nat_into A i)" |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
4263 |
using \<open>A \<noteq> {}\<close> unfolding X_def by (intro T) (auto intro: from_nat_into)
|
| 53282 | 4264 |
then have "range X \<subseteq> U" |
4265 |
by auto |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4266 |
with subseq[of X] obtain r x where "x \<in> U" and r: "strict_mono r" "(X \<circ> r) \<longlonglongrightarrow> x" |
| 53282 | 4267 |
by auto |
| 60420 | 4268 |
from \<open>x\<in>U\<close> \<open>U \<subseteq> \<Union>A\<close> from_nat_into_surj[OF \<open>countable A\<close>] |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4269 |
obtain n where "x \<in> from_nat_into A n" by auto |
| 60420 | 4270 |
with r(2) A(1) from_nat_into[OF \<open>A \<noteq> {}\<close>, of n]
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4271 |
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
|
4272 |
unfolding tendsto_def by (auto simp: comp_def) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4273 |
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
|
4274 |
by (auto simp: eventually_sequentially) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4275 |
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
|
4276 |
by auto |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4277 |
moreover from \<open>strict_mono r\<close>[THEN seq_suble, of "max n N"] have "\<exists>i. n \<le> r i \<and> N \<le> i" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4278 |
by (auto intro!: exI[of _ "max n N"]) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4279 |
ultimately show False |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4280 |
by auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4281 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4282 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4283 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4284 |
|
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4285 |
lemma compact_imp_seq_compact: |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4286 |
fixes U :: "'a :: first_countable_topology set" |
| 53282 | 4287 |
assumes "compact U" |
4288 |
shows "seq_compact U" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4289 |
unfolding seq_compact_def |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4290 |
proof safe |
| 52624 | 4291 |
fix X :: "nat \<Rightarrow> 'a" |
4292 |
assume "\<forall>n. X n \<in> U" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4293 |
then have "eventually (\<lambda>x. x \<in> U) (filtermap X sequentially)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4294 |
by (auto simp: eventually_filtermap) |
| 52624 | 4295 |
moreover |
4296 |
have "filtermap X sequentially \<noteq> bot" |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4297 |
by (simp add: trivial_limit_def eventually_filtermap) |
| 52624 | 4298 |
ultimately |
4299 |
obtain x where "x \<in> U" and x: "inf (nhds x) (filtermap X sequentially) \<noteq> bot" (is "?F \<noteq> _") |
|
| 60420 | 4300 |
using \<open>compact U\<close> by (auto simp: compact_filter) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4301 |
|
| 55522 | 4302 |
from countable_basis_at_decseq[of x] |
4303 |
obtain A where A: |
|
4304 |
"\<And>i. open (A i)" |
|
4305 |
"\<And>i. x \<in> A i" |
|
4306 |
"\<And>S. open S \<Longrightarrow> x \<in> S \<Longrightarrow> eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
|
4307 |
by blast |
|
| 63040 | 4308 |
define s where "s n i = (SOME j. i < j \<and> X j \<in> A (Suc n))" for n i |
| 52624 | 4309 |
{
|
4310 |
fix n i |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4311 |
have "\<exists>a. i < a \<and> X a \<in> A (Suc n)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4312 |
proof (rule ccontr) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4313 |
assume "\<not> (\<exists>a>i. X a \<in> A (Suc n))" |
| 53282 | 4314 |
then have "\<And>a. Suc i \<le> a \<Longrightarrow> X a \<notin> A (Suc n)" |
4315 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4316 |
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
|
4317 |
by (auto simp: eventually_filtermap eventually_sequentially) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4318 |
moreover have "eventually (\<lambda>x. x \<in> A (Suc n)) (nhds x)" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4319 |
using A(1,2)[of "Suc n"] by (auto simp: eventually_nhds) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4320 |
ultimately have "eventually (\<lambda>x. False) ?F" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4321 |
by (auto simp: eventually_inf) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4322 |
with x show False |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4323 |
by (simp add: eventually_False) |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4324 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4325 |
then have "i < s n i" "X (s n i) \<in> A (Suc n)" |
| 52624 | 4326 |
unfolding s_def by (auto intro: someI2_ex) |
4327 |
} |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4328 |
note s = this |
| 63040 | 4329 |
define r where "r = rec_nat (s 0 0) s" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4330 |
have "strict_mono r" |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4331 |
by (auto simp: r_def s strict_mono_Suc_iff) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4332 |
moreover |
| 61969 | 4333 |
have "(\<lambda>n. X (r n)) \<longlonglongrightarrow> x" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4334 |
proof (rule topological_tendstoI) |
| 52624 | 4335 |
fix S |
4336 |
assume "open S" "x \<in> S" |
|
| 53282 | 4337 |
with A(3) have "eventually (\<lambda>i. A i \<subseteq> S) sequentially" |
4338 |
by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4339 |
moreover |
| 52624 | 4340 |
{
|
4341 |
fix i |
|
4342 |
assume "Suc 0 \<le> i" |
|
4343 |
then have "X (r i) \<in> A i" |
|
4344 |
by (cases i) (simp_all add: r_def s) |
|
4345 |
} |
|
4346 |
then have "eventually (\<lambda>i. X (r i) \<in> A i) sequentially" |
|
4347 |
by (auto simp: eventually_sequentially) |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4348 |
ultimately show "eventually (\<lambda>i. X (r i) \<in> S) sequentially" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4349 |
by eventually_elim auto |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4350 |
qed |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4351 |
ultimately show "\<exists>x \<in> U. \<exists>r. strict_mono r \<and> (X \<circ> r) \<longlonglongrightarrow> x" |
| 60420 | 4352 |
using \<open>x \<in> U\<close> by (auto simp: convergent_def comp_def) |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4353 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4354 |
|
|
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
|
4355 |
lemma countably_compact_imp_acc_point: |
| 53291 | 4356 |
assumes "countably_compact s" |
4357 |
and "countable t" |
|
4358 |
and "infinite t" |
|
4359 |
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
|
4360 |
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
|
4361 |
proof (rule ccontr) |
| 63040 | 4362 |
define C where "C = (\<lambda>F. interior (F \<union> (- t))) ` {F. finite F \<and> F \<subseteq> t }"
|
| 60420 | 4363 |
note \<open>countably_compact s\<close> |
| 53282 | 4364 |
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
|
4365 |
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
|
4366 |
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
|
4367 |
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
|
4368 |
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
|
4369 |
have "s \<subseteq> \<Union>C" |
| 60420 | 4370 |
using \<open>t \<subseteq> s\<close> |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
4371 |
unfolding C_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
|
4372 |
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
|
4373 |
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
|
4374 |
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
|
4375 |
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
|
4376 |
moreover |
| 60420 | 4377 |
from \<open>countable t\<close> have "countable C" |
|
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
|
4378 |
unfolding C_def by (auto intro: countable_Collect_finite_subset) |
| 55522 | 4379 |
ultimately |
4380 |
obtain D where "D \<subseteq> C" "finite D" "s \<subseteq> \<Union>D" |
|
4381 |
by (rule countably_compactE) |
|
| 53282 | 4382 |
then obtain E where E: "E \<subseteq> {F. finite F \<and> F \<subseteq> t }" "finite E"
|
4383 |
and s: "s \<subseteq> (\<Union>F\<in>E. interior (F \<union> (- t)))" |
|
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
4384 |
by (metis (lifting) finite_subset_image C_def) |
| 60420 | 4385 |
from s \<open>t \<subseteq> s\<close> have "t \<subseteq> \<Union>E" |
|
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
|
4386 |
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
|
4387 |
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
|
4388 |
using E by auto |
| 60420 | 4389 |
ultimately show False using \<open>infinite t\<close> |
| 53282 | 4390 |
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
|
4391 |
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
|
4392 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
4393 |
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
|
4394 |
fixes s :: "'a::first_countable_topology set" |
| 53291 | 4395 |
assumes "\<forall>t. infinite t \<and> countable t \<and> t \<subseteq> s \<longrightarrow> |
4396 |
(\<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
|
4397 |
shows "seq_compact s" |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4398 |
proof - |
| 52624 | 4399 |
{
|
4400 |
fix f :: "nat \<Rightarrow> 'a" |
|
4401 |
assume f: "\<forall>n. f n \<in> s" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4402 |
have "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4403 |
proof (cases "finite (range f)") |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4404 |
case True |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4405 |
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
|
4406 |
using pigeonhole_infinite[OF _ True] by auto |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4407 |
then obtain r :: "nat \<Rightarrow> nat" where "strict_mono r" and fr: "\<forall>n. f (r n) = f l" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4408 |
using infinite_enumerate by blast |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4409 |
then have "strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> f l" |
|
58729
e8ecc79aee43
add tendsto_const and tendsto_ident_at as simp and intro rules
hoelzl
parents:
58184
diff
changeset
|
4410 |
by (simp add: fr o_def) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4411 |
with f show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
|
50941
3690724028b1
add countable compacteness; replace finite_range_imp_infinite_repeats by pigeonhole_infinite
hoelzl
parents:
50940
diff
changeset
|
4412 |
by auto |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4413 |
next |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4414 |
case False |
| 53282 | 4415 |
with f assms have "\<exists>x\<in>s. \<forall>U. x\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" |
4416 |
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
|
4417 |
then obtain l where "l \<in> s" "\<forall>U. l\<in>U \<and> open U \<longrightarrow> infinite (U \<inter> range f)" .. |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4418 |
from this(2) have "\<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
|
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
|
4419 |
using acc_point_range_imp_convergent_subsequence[of l f] by auto |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4420 |
with \<open>l \<in> s\<close> show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" .. |
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4421 |
qed |
|
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4422 |
} |
| 53282 | 4423 |
then show ?thesis |
4424 |
unfolding seq_compact_def by auto |
|
|
50884
2b21b4e2d7cb
differentiate (cover) compactness and sequential compactness
hoelzl
parents:
50883
diff
changeset
|
4425 |
qed |
| 44075 | 4426 |
|
|
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
|
4427 |
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
|
4428 |
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
|
4429 |
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
|
4430 |
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
|
4431 |
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
|
4432 |
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
|
4433 |
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
|
4434 |
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
|
4435 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
4436 |
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
|
4437 |
fixes s :: "'a :: first_countable_topology set" |
| 53291 | 4438 |
shows "seq_compact s \<longleftrightarrow> |
4439 |
(\<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
|
4440 |
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
|
4441 |
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
|
4442 |
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
|
4443 |
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
|
4444 |
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
|
4445 |
|
|
88a00a1c7c2c
use accumulation point characterization (avoids t1_space restriction for equivalence of countable and sequential compactness); remove heine_borel_lemma
hoelzl
parents:
50942
diff
changeset
|
4446 |
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
|
4447 |
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
|
4448 |
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
|
4449 |
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
|
4450 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4451 |
proposition bolzano_weierstrass_imp_seq_compact: |
|
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
|
4452 |
fixes s :: "'a::{t1_space, first_countable_topology} set"
|
| 64539 | 4453 |
shows "\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> seq_compact s" |
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4454 |
by (rule countable_acc_point_imp_seq_compact) (metis islimpt_eq_acc_point) |
|
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
|
4455 |
|
| 64539 | 4456 |
|
| 60420 | 4457 |
subsubsection\<open>Totally bounded\<close> |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4458 |
|
| 64539 | 4459 |
lemma cauchy_def: "Cauchy s \<longleftrightarrow> (\<forall>e>0. \<exists>N. \<forall>m n. m \<ge> N \<and> n \<ge> N \<longrightarrow> dist (s m) (s n) < e)" |
| 52624 | 4460 |
unfolding Cauchy_def by metis |
4461 |
||
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4462 |
proposition seq_compact_imp_totally_bounded: |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4463 |
assumes "seq_compact s" |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4464 |
shows "\<forall>e>0. \<exists>k. finite k \<and> k \<subseteq> s \<and> s \<subseteq> (\<Union>x\<in>k. ball x e)" |
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4465 |
proof - |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4466 |
{ fix e::real assume "e > 0" assume *: "\<And>k. finite k \<Longrightarrow> k \<subseteq> s \<Longrightarrow> \<not> s \<subseteq> (\<Union>x\<in>k. ball x e)"
|
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4467 |
let ?Q = "\<lambda>x n r. r \<in> s \<and> (\<forall>m < (n::nat). \<not> (dist (x m) r < e))" |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4468 |
have "\<exists>x. \<forall>n::nat. ?Q x n (x n)" |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4469 |
proof (rule dependent_wellorder_choice) |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4470 |
fix n x assume "\<And>y. y < n \<Longrightarrow> ?Q x y (x y)" |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4471 |
then have "\<not> s \<subseteq> (\<Union>x\<in>x ` {0..<n}. ball x e)"
|
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4472 |
using *[of "x ` {0 ..< n}"] by (auto simp: subset_eq)
|
| 52624 | 4473 |
then obtain z where z:"z\<in>s" "z \<notin> (\<Union>x\<in>x ` {0..<n}. ball x e)"
|
4474 |
unfolding subset_eq by auto |
|
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4475 |
show "\<exists>r. ?Q x n r" |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4476 |
using z by auto |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4477 |
qed simp |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4478 |
then obtain x where "\<forall>n::nat. x n \<in> s" and x:"\<And>n m. m < n \<Longrightarrow> \<not> (dist (x m) (x n) < e)" |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4479 |
by blast |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4480 |
then obtain l r where "l \<in> s" and r:"strict_mono r" and "((x \<circ> r) \<longlongrightarrow> l) sequentially" |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4481 |
using assms by (metis seq_compact_def) |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4482 |
from this(3) have "Cauchy (x \<circ> r)" |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4483 |
using LIMSEQ_imp_Cauchy by auto |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4484 |
then obtain N::nat where "\<And>m n. N \<le> m \<Longrightarrow> N \<le> n \<Longrightarrow> dist ((x \<circ> r) m) ((x \<circ> r) n) < e" |
| 60420 | 4485 |
unfolding cauchy_def using \<open>e > 0\<close> by blast |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4486 |
then have False |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4487 |
using x[of "r N" "r (N+1)"] r by (auto simp: strict_mono_def) } |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4488 |
then show ?thesis |
|
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4489 |
by metis |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4490 |
qed |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4491 |
|
| 60420 | 4492 |
subsubsection\<open>Heine-Borel theorem\<close> |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4493 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4494 |
proposition seq_compact_imp_heine_borel: |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4495 |
fixes s :: "'a :: metric_space set" |
| 53282 | 4496 |
assumes "seq_compact s" |
4497 |
shows "compact s" |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4498 |
proof - |
| 60420 | 4499 |
from seq_compact_imp_totally_bounded[OF \<open>seq_compact s\<close>] |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4500 |
obtain f where f: "\<forall>e>0. finite (f e) \<and> f e \<subseteq> s \<and> s \<subseteq> (\<Union>x\<in>f e. ball x e)" |
| 55522 | 4501 |
unfolding choice_iff' .. |
| 63040 | 4502 |
define K where "K = (\<lambda>(x, r). ball x r) ` ((\<Union>e \<in> \<rat> \<inter> {0 <..}. f e) \<times> \<rat>)"
|
|
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
|
4503 |
have "countably_compact s" |
| 60420 | 4504 |
using \<open>seq_compact s\<close> by (rule seq_compact_imp_countably_compact) |
|
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
|
4505 |
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
|
4506 |
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
|
4507 |
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
|
4508 |
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
|
4509 |
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
|
4510 |
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
|
4511 |
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
|
4512 |
next |
| 53282 | 4513 |
fix T x |
4514 |
assume T: "open T" "x \<in> T" and x: "x \<in> s" |
|
4515 |
from openE[OF T] obtain e where "0 < e" "ball x e \<subseteq> T" |
|
4516 |
by auto |
|
4517 |
then have "0 < e / 2" "ball x (e / 2) \<subseteq> T" |
|
4518 |
by auto |
|
| 60420 | 4519 |
from Rats_dense_in_real[OF \<open>0 < e / 2\<close>] obtain r where "r \<in> \<rat>" "0 < r" "r < e / 2" |
| 53282 | 4520 |
by auto |
| 60420 | 4521 |
from f[rule_format, of r] \<open>0 < r\<close> \<open>x \<in> s\<close> obtain k where "k \<in> f r" "x \<in> ball k r" |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62131
diff
changeset
|
4522 |
by auto |
| 60420 | 4523 |
from \<open>r \<in> \<rat>\<close> \<open>0 < r\<close> \<open>k \<in> f r\<close> have "ball k r \<in> K" |
| 53282 | 4524 |
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
|
4525 |
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
|
4526 |
proof (rule bexI[rotated], safe) |
| 53282 | 4527 |
fix y |
4528 |
assume "y \<in> ball k r" |
|
| 60420 | 4529 |
with \<open>r < e / 2\<close> \<open>x \<in> ball k r\<close> have "dist x y < e" |
|
62397
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
4530 |
by (intro dist_triangle_half_r [of k _ e]) (auto simp: dist_commute) |
| 60420 | 4531 |
with \<open>ball x e \<subseteq> T\<close> show "y \<in> T" |
| 53282 | 4532 |
by auto |
4533 |
next |
|
4534 |
show "x \<in> ball k r" by fact |
|
4535 |
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
|
4536 |
qed |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4537 |
qed |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4538 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4539 |
proposition compact_eq_seq_compact_metric: |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4540 |
"compact (s :: 'a::metric_space set) \<longleftrightarrow> seq_compact s" |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4541 |
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
|
4542 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4543 |
proposition compact_def: \<comment> \<open>this is the definition of compactness in HOL Light\<close> |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4544 |
"compact (S :: 'a::metric_space set) \<longleftrightarrow> |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4545 |
(\<forall>f. (\<forall>n. f n \<in> S) \<longrightarrow> (\<exists>l\<in>S. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l))" |
|
50940
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4546 |
unfolding compact_eq_seq_compact_metric seq_compact_def by auto |
|
a7c273a83d27
group compactness-eq-seq-compactness lemmas together
hoelzl
parents:
50939
diff
changeset
|
4547 |
|
| 60420 | 4548 |
subsubsection \<open>Complete the chain of compactness variants\<close> |
| 50944 | 4549 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4550 |
proposition compact_eq_bolzano_weierstrass: |
| 50944 | 4551 |
fixes s :: "'a::metric_space set" |
| 53282 | 4552 |
shows "compact s \<longleftrightarrow> (\<forall>t. infinite t \<and> t \<subseteq> s --> (\<exists>x \<in> s. x islimpt t))" |
4553 |
(is "?lhs = ?rhs") |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4554 |
proof |
| 52624 | 4555 |
assume ?lhs |
| 53282 | 4556 |
then show ?rhs |
4557 |
using heine_borel_imp_bolzano_weierstrass[of s] by auto |
|
| 50944 | 4558 |
next |
| 52624 | 4559 |
assume ?rhs |
| 53282 | 4560 |
then show ?lhs |
| 50944 | 4561 |
unfolding compact_eq_seq_compact_metric by (rule bolzano_weierstrass_imp_seq_compact) |
4562 |
qed |
|
4563 |
||
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4564 |
proposition bolzano_weierstrass_imp_bounded: |
| 53282 | 4565 |
"\<forall>t. infinite t \<and> t \<subseteq> s \<longrightarrow> (\<exists>x \<in> s. x islimpt t) \<Longrightarrow> bounded s" |
| 50944 | 4566 |
using compact_imp_bounded unfolding compact_eq_bolzano_weierstrass . |
4567 |
||
|
64773
223b2ebdda79
Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents:
64758
diff
changeset
|
4568 |
|
| 60420 | 4569 |
subsection \<open>Metric spaces with the Heine-Borel property\<close> |
4570 |
||
4571 |
text \<open> |
|
| 33175 | 4572 |
A metric space (or topological vector space) is said to have the |
4573 |
Heine-Borel property if every closed and bounded subset is compact. |
|
| 60420 | 4574 |
\<close> |
| 33175 | 4575 |
|
| 68617 | 4576 |
class heine_borel = metric_space + |
| 33175 | 4577 |
assumes bounded_imp_convergent_subsequence: |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4578 |
"bounded (range f) \<Longrightarrow> \<exists>l r. strict_mono (r::nat\<Rightarrow>nat) \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 33175 | 4579 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4580 |
proposition bounded_closed_imp_seq_compact: |
| 33175 | 4581 |
fixes s::"'a::heine_borel set" |
| 53282 | 4582 |
assumes "bounded s" |
4583 |
and "closed s" |
|
4584 |
shows "seq_compact s" |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4585 |
proof (unfold seq_compact_def, clarify) |
| 53282 | 4586 |
fix f :: "nat \<Rightarrow> 'a" |
4587 |
assume f: "\<forall>n. f n \<in> s" |
|
| 60420 | 4588 |
with \<open>bounded s\<close> have "bounded (range f)" |
| 53282 | 4589 |
by (auto intro: bounded_subset) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4590 |
obtain l r where r: "strict_mono (r :: nat \<Rightarrow> nat)" and l: "((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 60420 | 4591 |
using bounded_imp_convergent_subsequence [OF \<open>bounded (range f)\<close>] by auto |
| 53282 | 4592 |
from f have fr: "\<forall>n. (f \<circ> r) n \<in> s" |
4593 |
by simp |
|
| 60420 | 4594 |
have "l \<in> s" using \<open>closed s\<close> fr l |
| 54070 | 4595 |
by (rule closed_sequentially) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4596 |
show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 60420 | 4597 |
using \<open>l \<in> s\<close> r l by blast |
| 33175 | 4598 |
qed |
4599 |
||
| 50944 | 4600 |
lemma compact_eq_bounded_closed: |
4601 |
fixes s :: "'a::heine_borel set" |
|
| 53291 | 4602 |
shows "compact s \<longleftrightarrow> bounded s \<and> closed s" |
4603 |
(is "?lhs = ?rhs") |
|
| 50944 | 4604 |
proof |
| 52624 | 4605 |
assume ?lhs |
| 53282 | 4606 |
then show ?rhs |
| 52624 | 4607 |
using compact_imp_closed compact_imp_bounded |
4608 |
by blast |
|
| 50944 | 4609 |
next |
| 52624 | 4610 |
assume ?rhs |
| 53282 | 4611 |
then show ?lhs |
| 52624 | 4612 |
using bounded_closed_imp_seq_compact[of s] |
4613 |
unfolding compact_eq_seq_compact_metric |
|
4614 |
by auto |
|
| 50944 | 4615 |
qed |
4616 |
||
|
66793
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4617 |
lemma compact_Inter: |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4618 |
fixes \<F> :: "'a :: heine_borel set set" |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4619 |
assumes com: "\<And>S. S \<in> \<F> \<Longrightarrow> compact S" and "\<F> \<noteq> {}"
|
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4620 |
shows "compact(\<Inter> \<F>)" |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4621 |
using assms |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4622 |
by (meson Inf_lower all_not_in_conv bounded_subset closed_Inter compact_eq_bounded_closed) |
|
deabce3ccf1f
new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents:
66643
diff
changeset
|
4623 |
|
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
4624 |
lemma compact_closure [simp]: |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
4625 |
fixes S :: "'a::heine_borel set" |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
4626 |
shows "compact(closure S) \<longleftrightarrow> bounded S" |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
4627 |
by (meson bounded_closure bounded_subset closed_closure closure_subset compact_eq_bounded_closed) |
|
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
4628 |
|
|
62381
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
4629 |
lemma not_compact_UNIV[simp]: |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
4630 |
fixes s :: "'a::{real_normed_vector,perfect_space,heine_borel} set"
|
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
4631 |
shows "~ compact (UNIV::'a set)" |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
4632 |
by (simp add: compact_eq_bounded_closed) |
|
a6479cb85944
New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents:
62343
diff
changeset
|
4633 |
|
|
67683
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4634 |
text\<open>Representing sets as the union of a chain of compact sets.\<close> |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4635 |
lemma closed_Union_compact_subsets: |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4636 |
fixes S :: "'a::{heine_borel,real_normed_vector} set"
|
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4637 |
assumes "closed S" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4638 |
obtains F where "\<And>n. compact(F n)" "\<And>n. F n \<subseteq> S" "\<And>n. F n \<subseteq> F(Suc n)" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4639 |
"(\<Union>n. F n) = S" "\<And>K. \<lbrakk>compact K; K \<subseteq> S\<rbrakk> \<Longrightarrow> \<exists>N. \<forall>n \<ge> N. K \<subseteq> F n" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4640 |
proof |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4641 |
show "compact (S \<inter> cball 0 (of_nat n))" for n |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4642 |
using assms compact_eq_bounded_closed by auto |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4643 |
next |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4644 |
show "(\<Union>n. S \<inter> cball 0 (real n)) = S" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4645 |
by (auto simp: real_arch_simple) |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4646 |
next |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4647 |
fix K :: "'a set" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4648 |
assume "compact K" "K \<subseteq> S" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4649 |
then obtain N where "K \<subseteq> cball 0 N" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4650 |
by (meson bounded_pos mem_cball_0 compact_imp_bounded subsetI) |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4651 |
then show "\<exists>N. \<forall>n\<ge>N. K \<subseteq> S \<inter> cball 0 (real n)" |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4652 |
by (metis of_nat_le_iff Int_subset_iff \<open>K \<subseteq> S\<close> real_arch_simple subset_cball subset_trans) |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4653 |
qed auto |
|
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
4654 |
|
| 67962 | 4655 |
instance%important real :: heine_borel |
4656 |
proof%unimportant |
|
| 50998 | 4657 |
fix f :: "nat \<Rightarrow> real" |
4658 |
assume f: "bounded (range f)" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4659 |
obtain r :: "nat \<Rightarrow> nat" where r: "strict_mono r" "monoseq (f \<circ> r)" |
| 50972 | 4660 |
unfolding comp_def by (metis seq_monosub) |
4661 |
then have "Bseq (f \<circ> r)" |
|
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
4662 |
unfolding Bseq_eq_bounded using f by (force intro: bounded_subset) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4663 |
with r show "\<exists>l r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
| 50972 | 4664 |
using Bseq_monoseq_convergent[of "f \<circ> r"] by (auto simp: convergent_def) |
| 33175 | 4665 |
qed |
4666 |
||
| 62127 | 4667 |
lemma compact_lemma_general: |
4668 |
fixes f :: "nat \<Rightarrow> 'a" |
|
4669 |
fixes proj::"'a \<Rightarrow> 'b \<Rightarrow> 'c::heine_borel" (infixl "proj" 60) |
|
4670 |
fixes unproj:: "('b \<Rightarrow> 'c) \<Rightarrow> 'a"
|
|
4671 |
assumes finite_basis: "finite basis" |
|
4672 |
assumes bounded_proj: "\<And>k. k \<in> basis \<Longrightarrow> bounded ((\<lambda>x. x proj k) ` range f)" |
|
4673 |
assumes proj_unproj: "\<And>e k. k \<in> basis \<Longrightarrow> (unproj e) proj k = e k" |
|
4674 |
assumes unproj_proj: "\<And>x. unproj (\<lambda>k. x proj k) = x" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4675 |
shows "\<forall>d\<subseteq>basis. \<exists>l::'a. \<exists> r::nat\<Rightarrow>nat. |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4676 |
strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) proj i) (l proj 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
|
4677 |
proof safe |
| 62127 | 4678 |
fix d :: "'b set" |
4679 |
assume d: "d \<subseteq> basis" |
|
4680 |
with finite_basis have "finite d" |
|
| 53282 | 4681 |
by (blast intro: finite_subset) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4682 |
from this d show "\<exists>l::'a. \<exists>r::nat\<Rightarrow>nat. strict_mono r \<and> |
| 62127 | 4683 |
(\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) proj i) (l proj i) < e) sequentially)" |
| 52624 | 4684 |
proof (induct d) |
4685 |
case empty |
|
| 53282 | 4686 |
then show ?case |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4687 |
unfolding strict_mono_def by auto |
| 52624 | 4688 |
next |
4689 |
case (insert k d) |
|
| 62127 | 4690 |
have k[intro]: "k \<in> basis" |
| 53282 | 4691 |
using insert by auto |
| 62127 | 4692 |
have s': "bounded ((\<lambda>x. x proj k) ` range f)" |
4693 |
using k |
|
4694 |
by (rule bounded_proj) |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4695 |
obtain l1::"'a" and r1 where r1: "strict_mono r1" |
| 62127 | 4696 |
and lr1: "\<forall>e > 0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj 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
|
4697 |
using insert(3) using insert(4) by auto |
| 62127 | 4698 |
have f': "\<forall>n. f (r1 n) proj k \<in> (\<lambda>x. x proj k) ` range f" |
| 53282 | 4699 |
by simp |
| 62127 | 4700 |
have "bounded (range (\<lambda>i. f (r1 i) proj k))" |
| 50998 | 4701 |
by (metis (lifting) bounded_subset f' image_subsetI s') |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4702 |
then obtain l2 r2 where r2:"strict_mono r2" and lr2:"((\<lambda>i. f (r1 (r2 i)) proj k) \<longlongrightarrow> l2) sequentially" |
| 62127 | 4703 |
using bounded_imp_convergent_subsequence[of "\<lambda>i. f (r1 i) proj k"] |
| 53282 | 4704 |
by (auto simp: o_def) |
| 63040 | 4705 |
define r where "r = r1 \<circ> r2" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4706 |
have r:"strict_mono r" |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4707 |
using r1 and r2 unfolding r_def o_def strict_mono_def by auto |
| 33175 | 4708 |
moreover |
| 63040 | 4709 |
define l where "l = unproj (\<lambda>i. if i = k then l2 else l1 proj i)" |
| 52624 | 4710 |
{
|
4711 |
fix e::real |
|
| 53282 | 4712 |
assume "e > 0" |
| 62127 | 4713 |
from lr1 \<open>e > 0\<close> have N1: "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 n) proj i) (l1 proj i) < e) sequentially" |
| 52624 | 4714 |
by blast |
| 62127 | 4715 |
from lr2 \<open>e > 0\<close> have N2:"eventually (\<lambda>n. dist (f (r1 (r2 n)) proj k) l2 < e) sequentially" |
| 52624 | 4716 |
by (rule tendstoD) |
| 62127 | 4717 |
from r2 N1 have N1': "eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r1 (r2 n)) proj i) (l1 proj i) < e) sequentially" |
| 33175 | 4718 |
by (rule eventually_subseq) |
| 62127 | 4719 |
have "eventually (\<lambda>n. \<forall>i\<in>(insert k d). dist (f (r n) proj i) (l proj i) < e) sequentially" |
| 53282 | 4720 |
using N1' N2 |
| 62127 | 4721 |
by eventually_elim (insert insert.prems, auto simp: l_def r_def o_def proj_unproj) |
| 33175 | 4722 |
} |
4723 |
ultimately show ?case by auto |
|
4724 |
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
|
4725 |
qed |
|
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
hoelzl
parents:
37452
diff
changeset
|
4726 |
|
| 62127 | 4727 |
lemma compact_lemma: |
4728 |
fixes f :: "nat \<Rightarrow> 'a::euclidean_space" |
|
4729 |
assumes "bounded (range f)" |
|
4730 |
shows "\<forall>d\<subseteq>Basis. \<exists>l::'a. \<exists> r. |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4731 |
strict_mono r \<and> (\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>d. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially)" |
| 62127 | 4732 |
by (rule compact_lemma_general[where unproj="\<lambda>e. \<Sum>i\<in>Basis. e i *\<^sub>R i"]) |
4733 |
(auto intro!: assms bounded_linear_inner_left bounded_linear_image |
|
4734 |
simp: euclidean_representation) |
|
4735 |
||
| 67962 | 4736 |
instance%important euclidean_space \<subseteq> heine_borel |
4737 |
proof%unimportant |
|
| 50998 | 4738 |
fix f :: "nat \<Rightarrow> 'a" |
4739 |
assume f: "bounded (range f)" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4740 |
then obtain l::'a and r where r: "strict_mono 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
|
4741 |
and l: "\<forall>e>0. eventually (\<lambda>n. \<forall>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i) < e) sequentially" |
| 50998 | 4742 |
using compact_lemma [OF f] by blast |
| 52624 | 4743 |
{
|
4744 |
fix e::real |
|
| 53282 | 4745 |
assume "e > 0" |
| 56541 | 4746 |
hence "e / real_of_nat DIM('a) > 0" by (simp add: DIM_positive)
|
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50324
diff
changeset
|
4747 |
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 | 4748 |
by simp |
4749 |
moreover |
|
| 52624 | 4750 |
{
|
4751 |
fix n |
|
4752 |
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
|
4753 |
have "dist (f (r n)) l \<le> (\<Sum>i\<in>Basis. dist (f (r n) \<bullet> i) (l \<bullet> i))" |
| 52624 | 4754 |
apply (subst euclidean_dist_l2) |
4755 |
using zero_le_dist |
|
| 67155 | 4756 |
apply (rule L2_set_le_sum) |
| 53282 | 4757 |
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
|
4758 |
also have "\<dots> < (\<Sum>i\<in>(Basis::'a set). e / (real_of_nat DIM('a)))"
|
| 64267 | 4759 |
apply (rule sum_strict_mono) |
| 52624 | 4760 |
using n |
| 53282 | 4761 |
apply auto |
4762 |
done |
|
4763 |
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
|
4764 |
by auto |
| 33175 | 4765 |
} |
4766 |
ultimately have "eventually (\<lambda>n. dist (f (r n)) l < e) sequentially" |
|
| 61810 | 4767 |
by (rule eventually_mono) |
| 33175 | 4768 |
} |
| 61973 | 4769 |
then have *: "((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 52624 | 4770 |
unfolding o_def tendsto_iff by simp |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4771 |
with r show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 52624 | 4772 |
by auto |
| 33175 | 4773 |
qed |
4774 |
||
4775 |
lemma bounded_fst: "bounded s \<Longrightarrow> bounded (fst ` s)" |
|
| 52624 | 4776 |
unfolding bounded_def |
| 55775 | 4777 |
by (metis (erased, hide_lams) dist_fst_le image_iff order_trans) |
| 33175 | 4778 |
|
4779 |
lemma bounded_snd: "bounded s \<Longrightarrow> bounded (snd ` s)" |
|
| 52624 | 4780 |
unfolding bounded_def |
| 55775 | 4781 |
by (metis (no_types, hide_lams) dist_snd_le image_iff order.trans) |
| 33175 | 4782 |
|
| 67962 | 4783 |
instance%important prod :: (heine_borel, heine_borel) heine_borel |
4784 |
proof%unimportant |
|
| 50998 | 4785 |
fix f :: "nat \<Rightarrow> 'a \<times> 'b" |
4786 |
assume f: "bounded (range f)" |
|
|
56154
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4787 |
then have "bounded (fst ` range f)" |
|
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4788 |
by (rule bounded_fst) |
|
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4789 |
then have s1: "bounded (range (fst \<circ> f))" |
|
f0a927235162
more complete set of lemmas wrt. image and composition
haftmann
parents:
56073
diff
changeset
|
4790 |
by (simp add: image_comp) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4791 |
obtain l1 r1 where r1: "strict_mono r1" and l1: "(\<lambda>n. fst (f (r1 n))) \<longlonglongrightarrow> l1" |
| 50998 | 4792 |
using bounded_imp_convergent_subsequence [OF s1] unfolding o_def by fast |
4793 |
from f have s2: "bounded (range (snd \<circ> f \<circ> r1))" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4794 |
by (auto simp: image_comp intro: bounded_snd bounded_subset) |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4795 |
obtain l2 r2 where r2: "strict_mono r2" and l2: "((\<lambda>n. snd (f (r1 (r2 n)))) \<longlongrightarrow> l2) sequentially" |
| 50998 | 4796 |
using bounded_imp_convergent_subsequence [OF s2] |
| 33175 | 4797 |
unfolding o_def by fast |
| 61973 | 4798 |
have l1': "((\<lambda>n. fst (f (r1 (r2 n)))) \<longlongrightarrow> l1) sequentially" |
| 50972 | 4799 |
using LIMSEQ_subseq_LIMSEQ [OF l1 r2] unfolding o_def . |
| 61973 | 4800 |
have l: "((f \<circ> (r1 \<circ> r2)) \<longlongrightarrow> (l1, l2)) sequentially" |
| 33175 | 4801 |
using tendsto_Pair [OF l1' l2] unfolding o_def by simp |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4802 |
have r: "strict_mono (r1 \<circ> r2)" |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4803 |
using r1 r2 unfolding strict_mono_def by simp |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4804 |
show "\<exists>l r. strict_mono r \<and> ((f \<circ> r) \<longlongrightarrow> l) sequentially" |
| 33175 | 4805 |
using l r by fast |
4806 |
qed |
|
| 64910 | 4807 |
|
| 60420 | 4808 |
subsubsection \<open>Completeness\<close> |
| 33175 | 4809 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4810 |
proposition (in metric_space) completeI: |
| 61969 | 4811 |
assumes "\<And>f. \<forall>n. f n \<in> s \<Longrightarrow> Cauchy f \<Longrightarrow> \<exists>l\<in>s. f \<longlonglongrightarrow> l" |
| 54070 | 4812 |
shows "complete s" |
4813 |
using assms unfolding complete_def by fast |
|
4814 |
||
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4815 |
proposition (in metric_space) completeE: |
| 54070 | 4816 |
assumes "complete s" and "\<forall>n. f n \<in> s" and "Cauchy f" |
| 61969 | 4817 |
obtains l where "l \<in> s" and "f \<longlonglongrightarrow> l" |
| 54070 | 4818 |
using assms unfolding complete_def by fast |
4819 |
||
| 62101 | 4820 |
(* TODO: generalize to uniform spaces *) |
| 52624 | 4821 |
lemma compact_imp_complete: |
| 62101 | 4822 |
fixes s :: "'a::metric_space set" |
| 52624 | 4823 |
assumes "compact s" |
4824 |
shows "complete s" |
|
4825 |
proof - |
|
4826 |
{
|
|
4827 |
fix f |
|
4828 |
assume as: "(\<forall>n::nat. f n \<in> s)" "Cauchy f" |
|
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4829 |
from as(1) obtain l r where lr: "l\<in>s" "strict_mono r" "(f \<circ> r) \<longlonglongrightarrow> l" |
| 50971 | 4830 |
using assms unfolding compact_def by blast |
4831 |
||
4832 |
note lr' = seq_suble [OF lr(2)] |
|
| 52624 | 4833 |
{
|
| 53282 | 4834 |
fix e :: real |
4835 |
assume "e > 0" |
|
| 52624 | 4836 |
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" |
4837 |
unfolding cauchy_def |
|
| 60420 | 4838 |
using \<open>e > 0\<close> |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4839 |
apply (erule_tac x="e/2" in allE, auto) |
| 52624 | 4840 |
done |
|
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
59815
diff
changeset
|
4841 |
from lr(3)[unfolded lim_sequentially, THEN spec[where x="e/2"]] |
| 53282 | 4842 |
obtain M where M:"\<forall>n\<ge>M. dist ((f \<circ> r) n) l < e/2" |
| 60420 | 4843 |
using \<open>e > 0\<close> by auto |
| 52624 | 4844 |
{
|
| 53282 | 4845 |
fix n :: nat |
4846 |
assume n: "n \<ge> max N M" |
|
4847 |
have "dist ((f \<circ> r) n) l < e/2" |
|
4848 |
using n M by auto |
|
4849 |
moreover have "r n \<ge> N" |
|
4850 |
using lr'[of n] n by auto |
|
4851 |
then have "dist (f n) ((f \<circ> r) n) < e / 2" |
|
4852 |
using N and n by auto |
|
| 52624 | 4853 |
ultimately have "dist (f n) l < e" |
| 53282 | 4854 |
using dist_triangle_half_r[of "f (r n)" "f n" e l] |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4855 |
by (auto simp: dist_commute) |
| 52624 | 4856 |
} |
| 53282 | 4857 |
then have "\<exists>N. \<forall>n\<ge>N. dist (f n) l < e" by blast |
| 52624 | 4858 |
} |
| 61973 | 4859 |
then have "\<exists>l\<in>s. (f \<longlongrightarrow> l) sequentially" using \<open>l\<in>s\<close> |
|
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
59815
diff
changeset
|
4860 |
unfolding lim_sequentially by auto |
| 52624 | 4861 |
} |
| 53282 | 4862 |
then show ?thesis unfolding complete_def by auto |
| 50971 | 4863 |
qed |
4864 |
||
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4865 |
proposition compact_eq_totally_bounded: |
|
58184
db1381d811ab
cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents:
57865
diff
changeset
|
4866 |
"compact s \<longleftrightarrow> complete s \<and> (\<forall>e>0. \<exists>k. finite k \<and> s \<subseteq> (\<Union>x\<in>k. ball x e))" |
| 50971 | 4867 |
(is "_ \<longleftrightarrow> ?rhs") |
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
4868 |
proof |
| 50971 | 4869 |
assume assms: "?rhs" |
4870 |
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)" |
|
4871 |
by (auto simp: choice_iff') |
|
4872 |
||
4873 |
show "compact s" |
|
4874 |
proof cases |
|
| 53282 | 4875 |
assume "s = {}"
|
4876 |
then show "compact s" by (simp add: compact_def) |
|
| 50971 | 4877 |
next |
4878 |
assume "s \<noteq> {}"
|
|
4879 |
show ?thesis |
|
4880 |
unfolding compact_def |
|
4881 |
proof safe |
|
| 53282 | 4882 |
fix f :: "nat \<Rightarrow> 'a" |
4883 |
assume f: "\<forall>n. f n \<in> s" |
|
4884 |
||
| 63040 | 4885 |
define e where "e n = 1 / (2 * Suc n)" for n |
| 50971 | 4886 |
then have [simp]: "\<And>n. 0 < e n" by auto |
| 63040 | 4887 |
define B where "B n U = (SOME b. infinite {n. f n \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U))" for n U
|
| 53282 | 4888 |
{
|
4889 |
fix n U |
|
4890 |
assume "infinite {n. f n \<in> U}"
|
|
| 50971 | 4891 |
then have "\<exists>b\<in>k (e n). infinite {i\<in>{n. f n \<in> U}. f i \<in> ball b (e n)}"
|
4892 |
using k f by (intro pigeonhole_infinite_rel) (auto simp: subset_eq) |
|
| 55522 | 4893 |
then obtain a where |
4894 |
"a \<in> k (e n)" |
|
4895 |
"infinite {i \<in> {n. f n \<in> U}. f i \<in> ball a (e n)}" ..
|
|
| 50971 | 4896 |
then have "\<exists>b. infinite {i. f i \<in> b} \<and> (\<exists>x. b \<subseteq> ball x (e n) \<inter> U)"
|
4897 |
by (intro exI[of _ "ball a (e n) \<inter> U"] exI[of _ a]) (auto simp: ac_simps) |
|
4898 |
from someI_ex[OF this] |
|
4899 |
have "infinite {i. f i \<in> B n U}" "\<exists>x. B n U \<subseteq> ball x (e n) \<inter> U"
|
|
| 53282 | 4900 |
unfolding B_def by auto |
4901 |
} |
|
| 50971 | 4902 |
note B = this |
4903 |
||
| 63040 | 4904 |
define F where "F = rec_nat (B 0 UNIV) B" |
| 53282 | 4905 |
{
|
4906 |
fix n |
|
4907 |
have "infinite {i. f i \<in> F n}"
|
|
4908 |
by (induct n) (auto simp: F_def B) |
|
4909 |
} |
|
| 50971 | 4910 |
then have F: "\<And>n. \<exists>x. F (Suc n) \<subseteq> ball x (e n) \<inter> F n" |
4911 |
using B by (simp add: F_def) |
|
4912 |
then have F_dec: "\<And>m n. m \<le> n \<Longrightarrow> F n \<subseteq> F m" |
|
4913 |
using decseq_SucI[of F] by (auto simp: decseq_def) |
|
4914 |
||
4915 |
obtain sel where sel: "\<And>k i. i < sel k i" "\<And>k i. f (sel k i) \<in> F k" |
|
4916 |
proof (atomize_elim, unfold all_conj_distrib[symmetric], intro choice allI) |
|
4917 |
fix k i |
|
4918 |
have "infinite ({n. f n \<in> F k} - {.. i})"
|
|
| 60420 | 4919 |
using \<open>infinite {n. f n \<in> F k}\<close> by auto
|
| 50971 | 4920 |
from infinite_imp_nonempty[OF this] |
4921 |
show "\<exists>x>i. f x \<in> F k" |
|
4922 |
by (simp add: set_eq_iff not_le conj_commute) |
|
4923 |
qed |
|
4924 |
||
| 63040 | 4925 |
define t where "t = rec_nat (sel 0 0) (\<lambda>n i. sel (Suc n) i)" |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4926 |
have "strict_mono t" |
|
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4927 |
unfolding strict_mono_Suc_iff by (simp add: t_def sel) |
| 50971 | 4928 |
moreover have "\<forall>i. (f \<circ> t) i \<in> s" |
4929 |
using f by auto |
|
4930 |
moreover |
|
| 53282 | 4931 |
{
|
4932 |
fix n |
|
4933 |
have "(f \<circ> t) n \<in> F n" |
|
4934 |
by (cases n) (simp_all add: t_def sel) |
|
4935 |
} |
|
| 50971 | 4936 |
note t = this |
4937 |
||
4938 |
have "Cauchy (f \<circ> t)" |
|
4939 |
proof (safe intro!: metric_CauchyI exI elim!: nat_approx_posE) |
|
| 53282 | 4940 |
fix r :: real and N n m |
4941 |
assume "1 / Suc N < r" "Suc N \<le> n" "Suc N \<le> m" |
|
| 50971 | 4942 |
then have "(f \<circ> t) n \<in> F (Suc N)" "(f \<circ> t) m \<in> F (Suc N)" "2 * e N < r" |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
4943 |
using F_dec t by (auto simp: e_def field_simps of_nat_Suc) |
| 50971 | 4944 |
with F[of N] obtain x where "dist x ((f \<circ> t) n) < e N" "dist x ((f \<circ> t) m) < e N" |
4945 |
by (auto simp: subset_eq) |
|
| 60420 | 4946 |
with dist_triangle[of "(f \<circ> t) m" "(f \<circ> t) n" x] \<open>2 * e N < r\<close> |
| 50971 | 4947 |
show "dist ((f \<circ> t) m) ((f \<circ> t) n) < r" |
4948 |
by (simp add: dist_commute) |
|
4949 |
qed |
|
4950 |
||
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
66408
diff
changeset
|
4951 |
ultimately show "\<exists>l\<in>s. \<exists>r. strict_mono r \<and> (f \<circ> r) \<longlonglongrightarrow> l" |
| 50971 | 4952 |
using assms unfolding complete_def by blast |
4953 |
qed |
|
4954 |
qed |
|
4955 |
qed (metis compact_imp_complete compact_imp_seq_compact seq_compact_imp_totally_bounded) |
|
| 33175 | 4956 |
|
| 53282 | 4957 |
lemma cauchy_imp_bounded: |
4958 |
assumes "Cauchy s" |
|
4959 |
shows "bounded (range s)" |
|
4960 |
proof - |
|
4961 |
from assms obtain N :: nat where "\<forall>m n. N \<le> m \<and> N \<le> n \<longrightarrow> dist (s m) (s n) < 1" |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4962 |
unfolding cauchy_def by force |
| 53282 | 4963 |
then have N:"\<forall>n. N \<le> n \<longrightarrow> dist (s N) (s n) < 1" by auto |
| 33175 | 4964 |
moreover |
| 52624 | 4965 |
have "bounded (s ` {0..N})"
|
4966 |
using finite_imp_bounded[of "s ` {1..N}"] by auto
|
|
| 33175 | 4967 |
then obtain a where a:"\<forall>x\<in>s ` {0..N}. dist (s N) x \<le> a"
|
4968 |
unfolding bounded_any_center [where a="s N"] by auto |
|
4969 |
ultimately show "?thesis" |
|
4970 |
unfolding bounded_any_center [where a="s N"] |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4971 |
apply (rule_tac x="max a 1" in exI, auto) |
| 52624 | 4972 |
apply (erule_tac x=y in allE) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
4973 |
apply (erule_tac x=y in ballE, auto) |
| 52624 | 4974 |
done |
| 33175 | 4975 |
qed |
4976 |
||
4977 |
instance heine_borel < complete_space |
|
4978 |
proof |
|
4979 |
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f" |
|
| 53282 | 4980 |
then have "bounded (range f)" |
| 34104 | 4981 |
by (rule cauchy_imp_bounded) |
| 53282 | 4982 |
then have "compact (closure (range f))" |
| 50971 | 4983 |
unfolding compact_eq_bounded_closed by auto |
| 53282 | 4984 |
then have "complete (closure (range f))" |
| 50971 | 4985 |
by (rule compact_imp_complete) |
| 33175 | 4986 |
moreover have "\<forall>n. f n \<in> closure (range f)" |
4987 |
using closure_subset [of "range f"] by auto |
|
| 61973 | 4988 |
ultimately have "\<exists>l\<in>closure (range f). (f \<longlongrightarrow> l) sequentially" |
| 60420 | 4989 |
using \<open>Cauchy f\<close> unfolding complete_def by auto |
| 33175 | 4990 |
then show "convergent f" |
|
36660
1cc4ab4b7ff7
make (X ----> L) an abbreviation for (X ---> L) sequentially
huffman
parents:
36659
diff
changeset
|
4991 |
unfolding convergent_def by auto |
| 33175 | 4992 |
qed |
4993 |
||
| 44632 | 4994 |
instance euclidean_space \<subseteq> banach .. |
4995 |
||
| 54070 | 4996 |
lemma complete_UNIV: "complete (UNIV :: ('a::complete_space) set)"
|
4997 |
proof (rule completeI) |
|
4998 |
fix f :: "nat \<Rightarrow> 'a" assume "Cauchy f" |
|
| 53282 | 4999 |
then have "convergent f" by (rule Cauchy_convergent) |
| 61969 | 5000 |
then show "\<exists>l\<in>UNIV. f \<longlonglongrightarrow> l" unfolding convergent_def by simp |
| 53282 | 5001 |
qed |
5002 |
||
5003 |
lemma complete_imp_closed: |
|
| 64287 | 5004 |
fixes S :: "'a::metric_space set" |
5005 |
assumes "complete S" |
|
5006 |
shows "closed S" |
|
| 54070 | 5007 |
proof (unfold closed_sequential_limits, clarify) |
| 64287 | 5008 |
fix f x assume "\<forall>n. f n \<in> S" and "f \<longlonglongrightarrow> x" |
| 61969 | 5009 |
from \<open>f \<longlonglongrightarrow> x\<close> have "Cauchy f" |
| 54070 | 5010 |
by (rule LIMSEQ_imp_Cauchy) |
| 64287 | 5011 |
with \<open>complete S\<close> and \<open>\<forall>n. f n \<in> S\<close> obtain l where "l \<in> S" and "f \<longlonglongrightarrow> l" |
| 54070 | 5012 |
by (rule completeE) |
| 61969 | 5013 |
from \<open>f \<longlonglongrightarrow> x\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "x = l" |
| 54070 | 5014 |
by (rule LIMSEQ_unique) |
| 64287 | 5015 |
with \<open>l \<in> S\<close> show "x \<in> S" |
| 54070 | 5016 |
by simp |
5017 |
qed |
|
5018 |
||
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
5019 |
lemma complete_Int_closed: |
| 64287 | 5020 |
fixes S :: "'a::metric_space set" |
5021 |
assumes "complete S" and "closed t" |
|
5022 |
shows "complete (S \<inter> t)" |
|
| 54070 | 5023 |
proof (rule completeI) |
| 64287 | 5024 |
fix f assume "\<forall>n. f n \<in> S \<inter> t" and "Cauchy f" |
5025 |
then have "\<forall>n. f n \<in> S" and "\<forall>n. f n \<in> t" |
|
| 54070 | 5026 |
by simp_all |
| 64287 | 5027 |
from \<open>complete S\<close> obtain l where "l \<in> S" and "f \<longlonglongrightarrow> l" |
5028 |
using \<open>\<forall>n. f n \<in> S\<close> and \<open>Cauchy f\<close> by (rule completeE) |
|
| 61969 | 5029 |
from \<open>closed t\<close> and \<open>\<forall>n. f n \<in> t\<close> and \<open>f \<longlonglongrightarrow> l\<close> have "l \<in> t" |
| 54070 | 5030 |
by (rule closed_sequentially) |
| 64287 | 5031 |
with \<open>l \<in> S\<close> and \<open>f \<longlonglongrightarrow> l\<close> show "\<exists>l\<in>S \<inter> t. f \<longlonglongrightarrow> l" |
| 54070 | 5032 |
by fast |
5033 |
qed |
|
5034 |
||
5035 |
lemma complete_closed_subset: |
|
| 64287 | 5036 |
fixes S :: "'a::metric_space set" |
5037 |
assumes "closed S" and "S \<subseteq> t" and "complete t" |
|
5038 |
shows "complete S" |
|
5039 |
using assms complete_Int_closed [of t S] by (simp add: Int_absorb1) |
|
| 33175 | 5040 |
|
5041 |
lemma complete_eq_closed: |
|
| 64287 | 5042 |
fixes S :: "('a::complete_space) set"
|
5043 |
shows "complete S \<longleftrightarrow> closed S" |
|
| 33175 | 5044 |
proof |
| 64287 | 5045 |
assume "closed S" then show "complete S" |
| 54070 | 5046 |
using subset_UNIV complete_UNIV by (rule complete_closed_subset) |
| 33175 | 5047 |
next |
| 64287 | 5048 |
assume "complete S" then show "closed S" |
| 54070 | 5049 |
by (rule complete_imp_closed) |
| 33175 | 5050 |
qed |
5051 |
||
| 64287 | 5052 |
lemma convergent_eq_Cauchy: |
5053 |
fixes S :: "nat \<Rightarrow> 'a::complete_space" |
|
5054 |
shows "(\<exists>l. (S \<longlongrightarrow> l) sequentially) \<longleftrightarrow> Cauchy S" |
|
| 44632 | 5055 |
unfolding Cauchy_convergent_iff convergent_def .. |
| 33175 | 5056 |
|
5057 |
lemma convergent_imp_bounded: |
|
| 64287 | 5058 |
fixes S :: "nat \<Rightarrow> 'a::metric_space" |
5059 |
shows "(S \<longlongrightarrow> l) sequentially \<Longrightarrow> bounded (range S)" |
|
|
50939
ae7cd20ed118
replace convergent_imp_cauchy by LIMSEQ_imp_Cauchy
hoelzl
parents:
50938
diff
changeset
|
5060 |
by (intro cauchy_imp_bounded LIMSEQ_imp_Cauchy) |
| 33175 | 5061 |
|
5062 |
lemma frontier_subset_compact: |
|
| 64287 | 5063 |
fixes S :: "'a::heine_borel set" |
5064 |
shows "compact S \<Longrightarrow> frontier S \<subseteq> S" |
|
| 33175 | 5065 |
using frontier_subset_closed compact_eq_bounded_closed |
5066 |
by blast |
|
|
63075
60a42a4166af
lemmas about dimension, hyperplanes, span, etc.
paulson <lp15@cam.ac.uk>
parents:
63040
diff
changeset
|
5067 |
|
| 60420 | 5068 |
subsection \<open>Continuity\<close> |
5069 |
||
5070 |
text\<open>Derive the epsilon-delta forms, which we often use as "definitions"\<close> |
|
| 33175 | 5071 |
|
|
68607
67bb59e49834
make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents:
68532
diff
changeset
|
5072 |
proposition continuous_within_eps_delta: |
| 33175 | 5073 |
"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)" |
|
67683
817944aeac3f
Lots of new material about matrices, etc.
paulson <lp15@cam.ac.uk>
parents:
67613
diff
changeset
|
5074 |
unfolding continuous_within and Lim_within by fastforce |
| 53282 | 5075 |
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5076 |
corollary continuous_at_eps_delta: |
| 53282 | 5077 |
"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 | 5078 |
using continuous_within_eps_delta [of x UNIV f] by simp |
| 33175 | 5079 |
|
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
5080 |
lemma continuous_at_right_real_increasing: |
|
57448
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5081 |
fixes f :: "real \<Rightarrow> real" |
|
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5082 |
assumes nondecF: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y" |
|
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5083 |
shows "continuous (at_right a) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f (a + d) - f a < e)" |
|
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5084 |
apply (simp add: greaterThan_def dist_real_def continuous_within Lim_within_le) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5085 |
apply (intro all_cong ex_cong, safe) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5086 |
apply (erule_tac x="a + d" in allE, simp) |
|
57448
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5087 |
apply (simp add: nondecF field_simps) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5088 |
apply (drule nondecF, simp) |
|
57448
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5089 |
done |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
5090 |
|
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
5091 |
lemma continuous_at_left_real_increasing: |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
5092 |
assumes nondecF: "\<And> x y. x \<le> y \<Longrightarrow> f x \<le> ((f y) :: real)" |
|
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
5093 |
shows "(continuous (at_left (a :: real)) f) = (\<forall>e > 0. \<exists>delta > 0. f a - f (a - delta) < e)" |
|
57448
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5094 |
apply (simp add: lessThan_def dist_real_def continuous_within Lim_within_le) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5095 |
apply (intro all_cong ex_cong, safe) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5096 |
apply (erule_tac x="a - d" in allE, simp) |
|
57448
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5097 |
apply (simp add: nondecF field_simps) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5098 |
apply (cut_tac x="a - d" and y=x in nondecF, simp_all) |
|
57448
159e45728ceb
more equalities of topological filters; strengthen dependent_nat_choice; tuned a couple of proofs
hoelzl
parents:
57447
diff
changeset
|
5099 |
done |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
57418
diff
changeset
|
5100 |
|
| 60420 | 5101 |
text\<open>Versions in terms of open balls.\<close> |
| 33175 | 5102 |
|
5103 |
lemma continuous_within_ball: |
|
| 53282 | 5104 |
"continuous (at x within s) f \<longleftrightarrow> |
5105 |
(\<forall>e > 0. \<exists>d > 0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e)" |
|
5106 |
(is "?lhs = ?rhs") |
|
| 33175 | 5107 |
proof |
5108 |
assume ?lhs |
|
| 53282 | 5109 |
{
|
5110 |
fix e :: real |
|
5111 |
assume "e > 0" |
|
| 33175 | 5112 |
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" |
| 60420 | 5113 |
using \<open>?lhs\<close>[unfolded continuous_within Lim_within] by auto |
| 53282 | 5114 |
{
|
5115 |
fix y |
|
5116 |
assume "y \<in> f ` (ball x d \<inter> s)" |
|
5117 |
then have "y \<in> ball (f x) e" |
|
5118 |
using d(2) |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5119 |
apply (auto simp: dist_commute) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5120 |
apply (erule_tac x=xa in ballE, auto) |
| 60420 | 5121 |
using \<open>e > 0\<close> |
| 53282 | 5122 |
apply auto |
5123 |
done |
|
| 33175 | 5124 |
} |
| 53282 | 5125 |
then have "\<exists>d>0. f ` (ball x d \<inter> s) \<subseteq> ball (f x) e" |
| 60420 | 5126 |
using \<open>d > 0\<close> |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5127 |
unfolding subset_eq ball_def by (auto simp: dist_commute) |
| 53282 | 5128 |
} |
5129 |
then show ?rhs by auto |
|
| 33175 | 5130 |
next |
| 53282 | 5131 |
assume ?rhs |
5132 |
then show ?lhs |
|
5133 |
unfolding continuous_within Lim_within ball_def subset_eq |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5134 |
apply (auto simp: dist_commute) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5135 |
apply (erule_tac x=e in allE, auto) |
| 53282 | 5136 |
done |
| 33175 | 5137 |
qed |
5138 |
||
5139 |
lemma continuous_at_ball: |
|
5140 |
"continuous (at x) f \<longleftrightarrow> (\<forall>e>0. \<exists>d>0. f ` (ball x d) \<subseteq> ball (f x) e)" (is "?lhs = ?rhs") |
|
5141 |
proof |
|
| 53282 | 5142 |
assume ?lhs |
5143 |
then show ?rhs |
|
5144 |
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball |
|
5145 |
apply auto |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5146 |
apply (erule_tac x=e in allE, auto) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5147 |
apply (rule_tac x=d in exI, auto) |
| 53282 | 5148 |
apply (erule_tac x=xa in allE) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5149 |
apply (auto simp: dist_commute) |
| 53282 | 5150 |
done |
| 33175 | 5151 |
next |
| 53282 | 5152 |
assume ?rhs |
5153 |
then show ?lhs |
|
5154 |
unfolding continuous_at Lim_at subset_eq Ball_def Bex_def image_iff mem_ball |
|
5155 |
apply auto |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5156 |
apply (erule_tac x=e in allE, auto) |
|
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5157 |
apply (rule_tac x=d in exI, auto) |
| 53282 | 5158 |
apply (erule_tac x="f xa" in allE) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5159 |
apply (auto simp: dist_commute) |
| 53282 | 5160 |
done |
| 33175 | 5161 |
qed |
5162 |
||
| 60420 | 5163 |
text\<open>Define setwise continuity in terms of limits within the set.\<close> |
| 33175 | 5164 |
|
| 36359 | 5165 |
lemma continuous_on_iff: |
5166 |
"continuous_on s f \<longleftrightarrow> |
|
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
5167 |
(\<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 | 5168 |
unfolding continuous_on_def Lim_within |
| 55775 | 5169 |
by (metis dist_pos_lt dist_self) |
| 53282 | 5170 |
|
|
62131
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5171 |
lemma continuous_within_E: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5172 |
assumes "continuous (at x within s) f" "e>0" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5173 |
obtains d where "d>0" "\<And>x'. \<lbrakk>x'\<in> s; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5174 |
using assms apply (simp add: continuous_within_eps_delta) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5175 |
apply (drule spec [of _ e], clarify) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5176 |
apply (rule_tac d="d/2" in that, auto) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5177 |
done |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5178 |
|
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5179 |
lemma continuous_onI [intro?]: |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5180 |
assumes "\<And>x e. \<lbrakk>e > 0; x \<in> s\<rbrakk> \<Longrightarrow> \<exists>d>0. \<forall>x'\<in>s. dist x' x < d \<longrightarrow> dist (f x') (f x) \<le> e" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5181 |
shows "continuous_on s f" |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5182 |
apply (simp add: continuous_on_iff, clarify) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5183 |
apply (rule ex_forward [OF assms [OF half_gt_zero]], auto) |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5184 |
done |
|
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents:
62127
diff
changeset
|
5185 |
|
| 60420 | 5186 |
text\<open>Some simple consequential lemmas.\<close> |
| 33175 | 5187 |
|
|
62397
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5188 |
lemma continuous_onE: |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
5189 |
assumes "continuous_on s f" "x\<in>s" "e>0" |
|
62397
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5190 |
obtains d where "d>0" "\<And>x'. \<lbrakk>x' \<in> s; dist x' x \<le> d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e" |
|
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5191 |
using assms |
|
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5192 |
apply (simp add: continuous_on_iff) |
|
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5193 |
apply (elim ballE allE) |
|
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5194 |
apply (auto intro: that [where d="d/2" for d]) |
|
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62381
diff
changeset
|
5195 |
done |
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
5196 |
|
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5197 |
lemma uniformly_continuous_onE: |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5198 |
assumes "uniformly_continuous_on s f" "0 < e" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5199 |
obtains d where "d>0" "\<And>x x'. \<lbrakk>x\<in>s; x'\<in>s; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e" |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5200 |
using assms |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5201 |
by (auto simp: uniformly_continuous_on_def) |
|
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5202 |
|
| 33175 | 5203 |
lemma continuous_at_imp_continuous_within: |
| 53282 | 5204 |
"continuous (at x) f \<Longrightarrow> continuous (at x within s) f" |
| 60762 | 5205 |
unfolding continuous_within continuous_at using Lim_at_imp_Lim_at_within by auto |
| 33175 | 5206 |
|
| 61973 | 5207 |
lemma Lim_trivial_limit: "trivial_limit net \<Longrightarrow> (f \<longlongrightarrow> 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
|
5208 |
by simp |
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
5209 |
|
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
5210 |
lemmas continuous_on = continuous_on_def \<comment> \<open>legacy theorem name\<close> |
| 33175 | 5211 |
|
5212 |
lemma continuous_within_subset: |
|
| 53282 | 5213 |
"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
|
5214 |
unfolding continuous_within by(metis tendsto_within_subset) |
| 33175 | 5215 |
|
5216 |
lemma continuous_on_interior: |
|
| 53282 | 5217 |
"continuous_on s f \<Longrightarrow> x \<in> interior s \<Longrightarrow> continuous (at x) f" |
| 55775 | 5218 |
by (metis continuous_on_eq_continuous_at continuous_on_subset interiorE) |
| 33175 | 5219 |
|
5220 |
lemma continuous_on_eq: |
|
| 61204 | 5221 |
"\<lbrakk>continuous_on s f; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<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
|
5222 |
unfolding continuous_on_def tendsto_def eventually_at_topological |
|
36440
89a70297564d
simplify definition of continuous_on; generalize some lemmas
huffman
parents:
36439
diff
changeset
|
5223 |
by simp |
| 33175 | 5224 |
|
| 60420 | 5225 |
text \<open>Characterization of various kinds of continuity in terms of sequences.\<close> |
| 33175 | 5226 |
|
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5227 |
lemma continuous_within_sequentiallyI: |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5228 |
fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5229 |
assumes "\<And>u::nat \<Rightarrow> 'a. u \<longlonglongrightarrow> a \<Longrightarrow> (\<forall>n. u n \<in> s) \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5230 |
shows "continuous (at a within s) f" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5231 |
using assms unfolding continuous_within tendsto_def[where l = "f a"] |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5232 |
by (auto intro!: sequentially_imp_eventually_within) |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5233 |
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5234 |
lemma continuous_within_tendsto_compose: |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5235 |
fixes f::"'a::t2_space \<Rightarrow> 'b::topological_space" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5236 |
assumes "continuous (at a within s) f" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5237 |
"eventually (\<lambda>n. x n \<in> s) F" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5238 |
"(x \<longlongrightarrow> a) F " |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5239 |
shows "((\<lambda>n. f (x n)) \<longlongrightarrow> f a) F" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5240 |
proof - |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5241 |
have *: "filterlim x (inf (nhds a) (principal s)) F" |
| 68120 | 5242 |
using assms(2) assms(3) unfolding at_within_def filterlim_inf by (auto simp: filterlim_principal eventually_mono) |
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5243 |
show ?thesis |
| 68120 | 5244 |
by (auto simp: assms(1) continuous_within[symmetric] tendsto_at_within_iff_tendsto_nhds[symmetric] intro!: filterlim_compose[OF _ *]) |
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5245 |
qed |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5246 |
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5247 |
lemma continuous_within_tendsto_compose': |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5248 |
fixes f::"'a::t2_space \<Rightarrow> 'b::topological_space" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5249 |
assumes "continuous (at a within s) f" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5250 |
"\<And>n. x n \<in> s" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5251 |
"(x \<longlongrightarrow> a) F " |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5252 |
shows "((\<lambda>n. f (x n)) \<longlongrightarrow> f a) F" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5253 |
by (auto intro!: continuous_within_tendsto_compose[OF assms(1)] simp add: assms) |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5254 |
|
| 33175 | 5255 |
lemma continuous_within_sequentially: |
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5256 |
fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
|
| 33175 | 5257 |
shows "continuous (at a within s) f \<longleftrightarrow> |
| 61973 | 5258 |
(\<forall>x. (\<forall>n::nat. x n \<in> s) \<and> (x \<longlongrightarrow> a) sequentially |
5259 |
\<longrightarrow> ((f \<circ> x) \<longlongrightarrow> f a) sequentially)" |
|
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5260 |
using continuous_within_tendsto_compose'[of a s f _ sequentially] |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5261 |
continuous_within_sequentiallyI[of a s f] |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5262 |
by (auto simp: o_def) |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5263 |
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5264 |
lemma continuous_at_sequentiallyI: |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5265 |
fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5266 |
assumes "\<And>u. u \<longlonglongrightarrow> a \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5267 |
shows "continuous (at a) f" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5268 |
using continuous_within_sequentiallyI[of a UNIV f] assms by auto |
| 33175 | 5269 |
|
5270 |
lemma continuous_at_sequentially: |
|
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
5271 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 53291 | 5272 |
shows "continuous (at a) f \<longleftrightarrow> |
| 61973 | 5273 |
(\<forall>x. (x \<longlongrightarrow> a) sequentially --> ((f \<circ> x) \<longlongrightarrow> f a) sequentially)" |
| 45031 | 5274 |
using continuous_within_sequentially[of a UNIV f] by simp |
| 33175 | 5275 |
|
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5276 |
lemma continuous_on_sequentiallyI: |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5277 |
fixes f :: "'a::{first_countable_topology, t2_space} \<Rightarrow> 'b::topological_space"
|
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5278 |
assumes "\<And>u a. (\<forall>n. u n \<in> s) \<Longrightarrow> a \<in> s \<Longrightarrow> u \<longlonglongrightarrow> a \<Longrightarrow> (\<lambda>n. f (u n)) \<longlonglongrightarrow> f a" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5279 |
shows "continuous_on s f" |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5280 |
using assms unfolding continuous_on_eq_continuous_within |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5281 |
using continuous_within_sequentiallyI[of _ s f] by auto |
|
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5282 |
|
| 33175 | 5283 |
lemma continuous_on_sequentially: |
|
44533
7abe4a59f75d
generalize and simplify proof of continuous_within_sequentially
huffman
parents:
44531
diff
changeset
|
5284 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
| 36359 | 5285 |
shows "continuous_on s f \<longleftrightarrow> |
| 61973 | 5286 |
(\<forall>x. \<forall>a \<in> s. (\<forall>n. x(n) \<in> s) \<and> (x \<longlongrightarrow> a) sequentially |
5287 |
--> ((f \<circ> x) \<longlongrightarrow> f a) sequentially)" |
|
|
67727
ce3e87a51488
moved Lipschitz continuity from AFP/Ordinary_Differential_Equations and AFP/Gromov_Hyperbolicity; moved lemmas from AFP/Gromov_Hyperbolicity/Library_Complements
immler
parents:
67686
diff
changeset
|
5288 |
(is "?lhs = ?rhs") |
| 33175 | 5289 |
proof |
| 53282 | 5290 |
assume ?rhs |
5291 |
then show ?lhs |
|
5292 |
using continuous_within_sequentially[of _ s f] |
|
5293 |
unfolding continuous_on_eq_continuous_within |
|
5294 |
by auto |
|
| 33175 | 5295 |
next |
| 53282 | 5296 |
assume ?lhs |
5297 |
then show ?rhs |
|
5298 |
unfolding continuous_on_eq_continuous_within |
|
5299 |
using continuous_within_sequentially[of _ s f] |
|
5300 |
by auto |
|
| 33175 | 5301 |
qed |
5302 |
||
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5303 |
lemma uniformly_continuous_on_sequentially: |
| 36441 | 5304 |
"uniformly_continuous_on s f \<longleftrightarrow> (\<forall>x y. (\<forall>n. x n \<in> s) \<and> (\<forall>n. y n \<in> s) \<and> |
| 62101 | 5305 |
(\<lambda>n. dist (x n) (y n)) \<longlonglongrightarrow> 0 \<longrightarrow> (\<lambda>n. dist (f(x n)) (f(y n))) \<longlonglongrightarrow> 0)" (is "?lhs = ?rhs") |
| 33175 | 5306 |
proof |
5307 |
assume ?lhs |
|
| 53282 | 5308 |
{
|
5309 |
fix x y |
|
5310 |
assume x: "\<forall>n. x n \<in> s" |
|
5311 |
and y: "\<forall>n. y n \<in> s" |
|
| 61973 | 5312 |
and xy: "((\<lambda>n. dist (x n) (y n)) \<longlongrightarrow> 0) sequentially" |
| 53282 | 5313 |
{
|
5314 |
fix e :: real |
|
5315 |
assume "e > 0" |
|
5316 |
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" |
|
| 60420 | 5317 |
using \<open>?lhs\<close>[unfolded uniformly_continuous_on_def, THEN spec[where x=e]] by auto |
| 53282 | 5318 |
obtain N where N: "\<forall>n\<ge>N. dist (x n) (y n) < d" |
| 60420 | 5319 |
using xy[unfolded lim_sequentially dist_norm] and \<open>d>0\<close> by auto |
| 53282 | 5320 |
{
|
5321 |
fix n |
|
5322 |
assume "n\<ge>N" |
|
5323 |
then have "dist (f (x n)) (f (y n)) < e" |
|
5324 |
using N[THEN spec[where x=n]] |
|
5325 |
using d[THEN bspec[where x="x n"], THEN bspec[where x="y n"]] |
|
5326 |
using x and y |
|
| 63170 | 5327 |
by (simp add: dist_commute) |
| 53282 | 5328 |
} |
5329 |
then have "\<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e" |
|
5330 |
by auto |
|
5331 |
} |
|
| 61973 | 5332 |
then have "((\<lambda>n. dist (f(x n)) (f(y n))) \<longlongrightarrow> 0) sequentially" |
|
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
59815
diff
changeset
|
5333 |
unfolding lim_sequentially and dist_real_def by auto |
| 53282 | 5334 |
} |
5335 |
then show ?rhs by auto |
|
| 33175 | 5336 |
next |
5337 |
assume ?rhs |
|
| 53282 | 5338 |
{
|
5339 |
assume "\<not> ?lhs" |
|
5340 |
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" |
|
5341 |
unfolding uniformly_continuous_on_def by auto |
|
5342 |
then obtain fa where fa: |
|
5343 |
"\<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" |
|
5344 |
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"] |
|
5345 |
unfolding Bex_def |
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5346 |
by (auto simp: dist_commute) |
| 63040 | 5347 |
define x where "x n = fst (fa (inverse (real n + 1)))" for n |
5348 |
define y where "y n = snd (fa (inverse (real n + 1)))" for n |
|
| 53282 | 5349 |
have xyn: "\<forall>n. x n \<in> s \<and> y n \<in> s" |
5350 |
and xy0: "\<forall>n. dist (x n) (y n) < inverse (real n + 1)" |
|
5351 |
and fxy:"\<forall>n. \<not> dist (f (x n)) (f (y n)) < e" |
|
5352 |
unfolding x_def and y_def using fa |
|
5353 |
by auto |
|
5354 |
{
|
|
5355 |
fix e :: real |
|
5356 |
assume "e > 0" |
|
5357 |
then obtain N :: nat where "N \<noteq> 0" and N: "0 < inverse (real N) \<and> inverse (real N) < e" |
|
|
62623
dbc62f86a1a9
rationalisation of theorem names esp about "real Archimedian" etc.
paulson <lp15@cam.ac.uk>
parents:
62620
diff
changeset
|
5358 |
unfolding real_arch_inverse[of e] by auto |
| 53282 | 5359 |
{
|
5360 |
fix n :: nat |
|
5361 |
assume "n \<ge> N" |
|
5362 |
then have "inverse (real n + 1) < inverse (real N)" |
|
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
5363 |
using of_nat_0_le_iff and \<open>N\<noteq>0\<close> by auto |
| 33175 | 5364 |
also have "\<dots> < e" using N by auto |
5365 |
finally have "inverse (real n + 1) < e" by auto |
|
| 53282 | 5366 |
then have "dist (x n) (y n) < e" |
5367 |
using xy0[THEN spec[where x=n]] by auto |
|
5368 |
} |
|
5369 |
then have "\<exists>N. \<forall>n\<ge>N. dist (x n) (y n) < e" by auto |
|
5370 |
} |
|
5371 |
then have "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. dist (f (x n)) (f (y n)) < e" |
|
| 60420 | 5372 |
using \<open>?rhs\<close>[THEN spec[where x=x], THEN spec[where x=y]] and xyn |
|
60017
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents:
59815
diff
changeset
|
5373 |
unfolding lim_sequentially dist_real_def by auto |
| 60420 | 5374 |
then have False using fxy and \<open>e>0\<close> by auto |
| 53282 | 5375 |
} |
5376 |
then show ?lhs |
|
5377 |
unfolding uniformly_continuous_on_def by blast |
|
| 33175 | 5378 |
qed |
5379 |
||
| 64287 | 5380 |
lemma continuous_closed_imp_Cauchy_continuous: |
5381 |
fixes S :: "('a::complete_space) set"
|
|
| 68120 | 5382 |
shows "\<lbrakk>continuous_on S f; closed S; Cauchy \<sigma>; \<And>n. (\<sigma> n) \<in> S\<rbrakk> \<Longrightarrow> Cauchy(f \<circ> \<sigma>)" |
| 64287 | 5383 |
apply (simp add: complete_eq_closed [symmetric] continuous_on_sequentially) |
5384 |
by (meson LIMSEQ_imp_Cauchy complete_def) |
|
|
61915
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
immler
parents:
61907
diff
changeset
|
5385 |
|
| 60420 | 5386 |
text\<open>The usual transformation theorems.\<close> |
| 33175 | 5387 |
|
5388 |
lemma continuous_transform_within: |
|
| 36667 | 5389 |
fixes f g :: "'a::metric_space \<Rightarrow> 'b::topological_space" |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
5390 |
assumes "continuous (at x within s) f" |
|
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
5391 |
and "0 < d" |
| 53282 | 5392 |
and "x \<in> s" |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
5393 |
and "\<And>x'. \<lbrakk>x' \<in> s; dist x' x < d\<rbrakk> \<Longrightarrow> f x' = g x'" |
| 33175 | 5394 |
shows "continuous (at x within s) g" |
|
62087
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
paulson
parents:
62083
diff
changeset
|
5395 |
using assms |
| 53282 | 5396 |
unfolding continuous_within |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5397 |
by (force intro: Lim_transform_within) |
| 62101 | 5398 |
|
| 53282 | 5399 |
|
| 67962 | 5400 |
subsubsection%unimportant \<open>Structural rules for pointwise continuity\<close> |
| 33175 | 5401 |
|
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
5402 |
lemma continuous_infnorm[continuous_intros]: |
| 53282 | 5403 |
"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
|
5404 |
unfolding continuous_def by (rule tendsto_infnorm) |
| 33175 | 5405 |
|
|
51361
21e5b6efb317
changed continuous_intros into a named theorems collection
hoelzl
parents:
51351
diff
changeset
|
5406 |
lemma continuous_inner[continuous_intros]: |
| 53282 | 5407 |
assumes "continuous F f" |
5408 |
and "continuous F g" |
|
|
44647
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
5409 |
shows "continuous F (\<lambda>x. inner (f x) (g x))" |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
5410 |
using assms unfolding continuous_def by (rule tendsto_inner) |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
5411 |
|
| 67962 | 5412 |
subsubsection%unimportant \<open>Structural rules for setwise continuity\<close> |
| 33175 | 5413 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5414 |
lemma continuous_on_infnorm[continuous_intros]: |
| 53282 | 5415 |
"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
|
5416 |
unfolding continuous_on by (fast intro: tendsto_infnorm) |
|
e4de7750cdeb
modernize lemmas about 'continuous' and 'continuous_on';
huffman
parents:
44632
diff
changeset
|
5417 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5418 |
lemma continuous_on_inner[continuous_intros]: |
|
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
5419 |
fixes g :: "'a::topological_space \<Rightarrow> 'b::real_inner" |
| 53282 | 5420 |
assumes "continuous_on s f" |
5421 |
and "continuous_on s g" |
|
|
44531
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
5422 |
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
|
5423 |
using bounded_bilinear_inner assms |
|
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
5424 |
by (rule bounded_bilinear.continuous_on) |
|
1d477a2b1572
replace some continuous_on lemmas with more general versions
huffman
parents:
44530
diff
changeset
|
5425 |
|
| 67962 | 5426 |
subsubsection%unimportant \<open>Structural rules for uniform continuity\<close> |
| 33175 | 5427 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5428 |
lemma uniformly_continuous_on_dist[continuous_intros]: |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5429 |
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
|
5430 |
assumes "uniformly_continuous_on s f" |
| 53282 | 5431 |
and "uniformly_continuous_on s g" |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5432 |
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
|
5433 |
proof - |
| 53282 | 5434 |
{
|
5435 |
fix a b c d :: 'b |
|
5436 |
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
|
5437 |
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
|
5438 |
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
|
5439 |
by arith |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5440 |
} note le = this |
| 53282 | 5441 |
{
|
5442 |
fix x y |
|
| 61969 | 5443 |
assume f: "(\<lambda>n. dist (f (x n)) (f (y n))) \<longlonglongrightarrow> 0" |
5444 |
assume g: "(\<lambda>n. dist (g (x n)) (g (y n))) \<longlonglongrightarrow> 0" |
|
5445 |
have "(\<lambda>n. \<bar>dist (f (x n)) (g (x n)) - dist (f (y n)) (g (y n))\<bar>) \<longlonglongrightarrow> 0" |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5446 |
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
|
5447 |
simp add: le) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5448 |
} |
| 53282 | 5449 |
then show ?thesis |
5450 |
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
|
5451 |
unfolding dist_real_def by simp |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5452 |
qed |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5453 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5454 |
lemma uniformly_continuous_on_norm[continuous_intros]: |
| 62101 | 5455 |
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
|
5456 |
assumes "uniformly_continuous_on s f" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5457 |
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
|
5458 |
unfolding norm_conv_dist using assms |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5459 |
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
|
5460 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5461 |
lemma (in bounded_linear) uniformly_continuous_on[continuous_intros]: |
| 62101 | 5462 |
fixes g :: "_::metric_space \<Rightarrow> _" |
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5463 |
assumes "uniformly_continuous_on s g" |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5464 |
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
|
5465 |
using assms unfolding uniformly_continuous_on_sequentially |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5466 |
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
|
5467 |
by (auto intro: tendsto_zero) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5468 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5469 |
lemma uniformly_continuous_on_cmul[continuous_intros]: |
| 36441 | 5470 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
| 33175 | 5471 |
assumes "uniformly_continuous_on s f" |
5472 |
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
|
5473 |
using bounded_linear_scaleR_right assms |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5474 |
by (rule bounded_linear.uniformly_continuous_on) |
| 33175 | 5475 |
|
5476 |
lemma dist_minus: |
|
5477 |
fixes x y :: "'a::real_normed_vector" |
|
5478 |
shows "dist (- x) (- y) = dist x y" |
|
5479 |
unfolding dist_norm minus_diff_minus norm_minus_cancel .. |
|
5480 |
||
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5481 |
lemma uniformly_continuous_on_minus[continuous_intros]: |
| 33175 | 5482 |
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
|
5483 |
shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. - f x)" |
| 33175 | 5484 |
unfolding uniformly_continuous_on_def dist_minus . |
5485 |
||
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5486 |
lemma uniformly_continuous_on_add[continuous_intros]: |
| 36441 | 5487 |
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
|
5488 |
assumes "uniformly_continuous_on s f" |
| 53282 | 5489 |
and "uniformly_continuous_on s g" |
| 33175 | 5490 |
shows "uniformly_continuous_on s (\<lambda>x. f x + g x)" |
| 53282 | 5491 |
using assms |
5492 |
unfolding uniformly_continuous_on_sequentially |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5493 |
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
|
5494 |
by (auto intro: tendsto_add_zero) |
|
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5495 |
|
|
56371
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
hoelzl
parents:
56290
diff
changeset
|
5496 |
lemma uniformly_continuous_on_diff[continuous_intros]: |
| 36441 | 5497 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::real_normed_vector" |
| 53282 | 5498 |
assumes "uniformly_continuous_on s f" |
5499 |
and "uniformly_continuous_on s g" |
|
|
44648
897f32a827f2
simplify some proofs about uniform continuity, and add some new ones;
huffman
parents:
44647
diff
changeset
|
5500 |
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
|
5501 |
using assms uniformly_continuous_on_add [of s f "- g"] |
|
b1d955791529
more simplification rules on unary and binary minus
haftmann
parents:
54070
diff
changeset
|
5502 |
by (simp add: fun_Compl_def uniformly_continuous_on_minus) |
| 33175 | 5503 |
|
| 62101 | 5504 |
text \<open>Continuity in terms of open preimages.\<close> |
| 33175 | 5505 |
|
5506 |
lemma continuous_at_open: |
|
| 53282 | 5507 |
"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)))" |
5508 |
unfolding continuous_within_topological [of x UNIV f] |
|
5509 |
unfolding imp_conjL |
|
5510 |
by (intro all_cong imp_cong ex_cong conj_cong refl) auto |
|
| 33175 | 5511 |
|
| 51351 | 5512 |
lemma continuous_imp_tendsto: |
| 53282 | 5513 |
assumes "continuous (at x0) f" |
| 61969 | 5514 |
and "x \<longlonglongrightarrow> x0" |
5515 |
shows "(f \<circ> x) \<longlonglongrightarrow> (f x0)" |
|
| 51351 | 5516 |
proof (rule topological_tendstoI) |
5517 |
fix S |
|
5518 |
assume "open S" "f x0 \<in> S" |
|
5519 |
then obtain T where T_def: "open T" "x0 \<in> T" "\<forall>x\<in>T. f x \<in> S" |
|
5520 |
using assms continuous_at_open by metis |
|
5521 |
then have "eventually (\<lambda>n. x n \<in> T) sequentially" |
|
5522 |
using assms T_def by (auto simp: tendsto_def) |
|
5523 |
then show "eventually (\<lambda>n. (f \<circ> x) n \<in> S) sequentially" |
|
| 61810 | 5524 |
using T_def by (auto elim!: eventually_mono) |
| 51351 | 5525 |
qed |
5526 |
||
| 33175 | 5527 |
lemma continuous_on_open: |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5528 |
"continuous_on S f \<longleftrightarrow> |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5529 |
(\<forall>T. openin (subtopology euclidean (f ` S)) T \<longrightarrow> |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5530 |
openin (subtopology euclidean S) (S \<inter> f -` T))" |
|
51481
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5531 |
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
|
5532 |
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) |
| 36441 | 5533 |
|
| 63301 | 5534 |
lemma continuous_on_open_gen: |
5535 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
5536 |
assumes "f ` S \<subseteq> T" |
|
5537 |
shows "continuous_on S f \<longleftrightarrow> |
|
5538 |
(\<forall>U. openin (subtopology euclidean T) U |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5539 |
\<longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` U))" |
| 63301 | 5540 |
(is "?lhs = ?rhs") |
5541 |
proof |
|
5542 |
assume ?lhs |
|
5543 |
then show ?rhs |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5544 |
apply (clarsimp simp: openin_euclidean_subtopology_iff continuous_on_iff) |
| 63301 | 5545 |
by (metis assms image_subset_iff) |
5546 |
next |
|
5547 |
have ope: "openin (subtopology euclidean T) (ball y e \<inter> T)" for y e |
|
5548 |
by (simp add: Int_commute openin_open_Int) |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5549 |
assume R [rule_format]: ?rhs |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5550 |
show ?lhs |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5551 |
proof (clarsimp simp add: continuous_on_iff) |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5552 |
fix x and e::real |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5553 |
assume "x \<in> S" and "0 < e" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5554 |
then have x: "x \<in> S \<inter> (f -` ball (f x) e \<inter> f -` T)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5555 |
using assms by auto |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5556 |
show "\<exists>d>0. \<forall>x'\<in>S. dist x' x < d \<longrightarrow> dist (f x') (f x) < e" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5557 |
using R [of "ball (f x) e \<inter> T"] x |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5558 |
by (fastforce simp add: ope openin_euclidean_subtopology_iff [of S] dist_commute) |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5559 |
qed |
| 63301 | 5560 |
qed |
5561 |
||
5562 |
lemma continuous_openin_preimage: |
|
5563 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
5564 |
shows |
|
5565 |
"\<lbrakk>continuous_on S f; f ` S \<subseteq> T; openin (subtopology euclidean T) U\<rbrakk> |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5566 |
\<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` U)" |
| 63301 | 5567 |
by (simp add: continuous_on_open_gen) |
5568 |
||
| 60420 | 5569 |
text \<open>Similarly in terms of closed sets.\<close> |
| 33175 | 5570 |
|
5571 |
lemma continuous_on_closed: |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5572 |
"continuous_on S f \<longleftrightarrow> |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5573 |
(\<forall>T. closedin (subtopology euclidean (f ` S)) T \<longrightarrow> |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5574 |
closedin (subtopology euclidean S) (S \<inter> f -` T))" |
|
51481
ef949192e5d6
move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents:
51480
diff
changeset
|
5575 |
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
|
5576 |
by (simp add: imp_ex imageI conj_commute eq_commute cong: conj_cong) |
| 33175 | 5577 |
|
| 63301 | 5578 |
lemma continuous_on_closed_gen: |
5579 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
5580 |
assumes "f ` S \<subseteq> T" |
|
5581 |
shows "continuous_on S f \<longleftrightarrow> |
|
5582 |
(\<forall>U. closedin (subtopology euclidean T) U |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5583 |
\<longrightarrow> closedin (subtopology euclidean S) (S \<inter> f -` U))" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5584 |
(is "?lhs = ?rhs") |
| 63301 | 5585 |
proof - |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5586 |
have *: "U \<subseteq> T \<Longrightarrow> S \<inter> f -` (T - U) = S - (S \<inter> f -` U)" for U |
| 63301 | 5587 |
using assms by blast |
5588 |
show ?thesis |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5589 |
proof |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5590 |
assume L: ?lhs |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5591 |
show ?rhs |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5592 |
proof clarify |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5593 |
fix U |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5594 |
assume "closedin (subtopology euclidean T) U" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5595 |
then show "closedin (subtopology euclidean S) (S \<inter> f -` U)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5596 |
using L unfolding continuous_on_open_gen [OF assms] |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5597 |
by (metis * closedin_def inf_le1 topspace_euclidean_subtopology) |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5598 |
qed |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5599 |
next |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5600 |
assume R [rule_format]: ?rhs |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5601 |
show ?lhs |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5602 |
unfolding continuous_on_open_gen [OF assms] |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5603 |
by (metis * R inf_le1 openin_closedin_eq topspace_euclidean_subtopology) |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5604 |
qed |
| 63301 | 5605 |
qed |
5606 |
||
5607 |
lemma continuous_closedin_preimage_gen: |
|
5608 |
fixes f :: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
5609 |
assumes "continuous_on S f" "f ` S \<subseteq> T" "closedin (subtopology euclidean T) U" |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5610 |
shows "closedin (subtopology euclidean S) (S \<inter> f -` U)" |
| 63301 | 5611 |
using assms continuous_on_closed_gen by blast |
5612 |
||
5613 |
lemma continuous_on_imp_closedin: |
|
5614 |
assumes "continuous_on S f" "closedin (subtopology euclidean (f ` S)) T" |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5615 |
shows "closedin (subtopology euclidean S) (S \<inter> f -` T)" |
| 63301 | 5616 |
using assms continuous_on_closed by blast |
5617 |
||
| 67968 | 5618 |
subsection%unimportant \<open>Half-global and completely global cases\<close> |
| 33175 | 5619 |
|
| 63301 | 5620 |
lemma continuous_openin_preimage_gen: |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5621 |
assumes "continuous_on S f" "open T" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5622 |
shows "openin (subtopology euclidean S) (S \<inter> f -` T)" |
| 53282 | 5623 |
proof - |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5624 |
have *: "(S \<inter> f -` T) = (S \<inter> f -` (T \<inter> f ` S))" |
| 53282 | 5625 |
by auto |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5626 |
have "openin (subtopology euclidean (f ` S)) (T \<inter> f ` S)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5627 |
using openin_open_Int[of T "f ` S", OF assms(2)] unfolding openin_open by auto |
| 53282 | 5628 |
then show ?thesis |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5629 |
using assms(1)[unfolded continuous_on_open, THEN spec[where x="T \<inter> f ` S"]] |
| 53282 | 5630 |
using * by auto |
| 33175 | 5631 |
qed |
5632 |
||
|
61518
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
paulson
parents:
61426
diff
changeset
|
5633 |
lemma continuous_closedin_preimage: |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5634 |
assumes "continuous_on S f" and "closed T" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5635 |
shows "closedin (subtopology euclidean S) (S \<inter> f -` T)" |
| 53282 | 5636 |
proof - |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5637 |
have *: "(S \<inter> f -` T) = (S \<inter> f -` (T \<inter> f ` S))" |
| 53282 | 5638 |
by auto |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5639 |
have "closedin (subtopology euclidean (f ` S)) (T \<inter> f ` S)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5640 |
using closedin_closed_Int[of T "f ` S", OF assms(2)] |
| 63170 | 5641 |
by (simp add: Int_commute) |
| 53282 | 5642 |
then show ?thesis |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5643 |
using assms(1)[unfolded continuous_on_closed, THEN spec[where x="T \<inter> f ` S"]] |
| 53282 | 5644 |
using * by auto |
| 33175 | 5645 |
qed |
5646 |
||
| 63955 | 5647 |
lemma continuous_openin_preimage_eq: |
5648 |
"continuous_on S f \<longleftrightarrow> |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5649 |
(\<forall>T. open T \<longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` T))" |
| 63955 | 5650 |
apply safe |
5651 |
apply (simp add: continuous_openin_preimage_gen) |
|
5652 |
apply (fastforce simp add: continuous_on_open openin_open) |
|
5653 |
done |
|
5654 |
||
5655 |
lemma continuous_closedin_preimage_eq: |
|
5656 |
"continuous_on S f \<longleftrightarrow> |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5657 |
(\<forall>T. closed T \<longrightarrow> closedin (subtopology euclidean S) (S \<inter> f -` T))" |
| 63955 | 5658 |
apply safe |
5659 |
apply (simp add: continuous_closedin_preimage) |
|
5660 |
apply (fastforce simp add: continuous_on_closed closedin_closed) |
|
5661 |
done |
|
5662 |
||
| 33175 | 5663 |
lemma continuous_open_preimage: |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5664 |
assumes contf: "continuous_on S f" and "open S" "open T" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5665 |
shows "open (S \<inter> f -` T)" |
| 33175 | 5666 |
proof- |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5667 |
obtain U where "open U" "(S \<inter> f -` T) = S \<inter> U" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5668 |
using continuous_openin_preimage_gen[OF contf \<open>open T\<close>] |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5669 |
unfolding openin_open by auto |
| 53282 | 5670 |
then show ?thesis |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5671 |
using open_Int[of S U, OF \<open>open S\<close>] by auto |
| 33175 | 5672 |
qed |
5673 |
||
5674 |
lemma continuous_closed_preimage: |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5675 |
assumes contf: "continuous_on S f" and "closed S" "closed T" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5676 |
shows "closed (S \<inter> f -` T)" |
| 33175 | 5677 |
proof- |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5678 |
obtain U where "closed U" "(S \<inter> f -` T) = S \<inter> U" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5679 |
using continuous_closedin_preimage[OF contf \<open>closed T\<close>] |
| 53282 | 5680 |
unfolding closedin_closed by auto |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5681 |
then show ?thesis using closed_Int[of S U, OF \<open>closed S\<close>] by auto |
| 33175 | 5682 |
qed |
5683 |
||
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5684 |
lemma continuous_open_vimage: "open S \<Longrightarrow> (\<And>x. continuous (at x) f) \<Longrightarrow> open (f -` S)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5685 |
by (metis continuous_on_eq_continuous_within open_vimage) |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5686 |
|
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5687 |
lemma continuous_closed_vimage: "closed S \<Longrightarrow> (\<And>x. continuous (at x) f) \<Longrightarrow> closed (f -` S)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5688 |
by (simp add: closed_vimage continuous_on_eq_continuous_within) |
| 33175 | 5689 |
|
| 36441 | 5690 |
lemma interior_image_subset: |
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
5691 |
assumes "inj f" "\<And>x. continuous (at x) f" |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5692 |
shows "interior (f ` S) \<subseteq> f ` (interior S)" |
| 44519 | 5693 |
proof |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5694 |
fix x assume "x \<in> interior (f ` S)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5695 |
then obtain T where as: "open T" "x \<in> T" "T \<subseteq> f ` S" .. |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5696 |
then have "x \<in> f ` S" by auto |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5697 |
then obtain y where y: "y \<in> S" "x = f y" by auto |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5698 |
have "open (f -` T)" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5699 |
using assms \<open>open T\<close> by (simp add: continuous_at_imp_continuous_on open_vimage) |
| 44519 | 5700 |
moreover have "y \<in> vimage f T" |
| 60420 | 5701 |
using \<open>x = f y\<close> \<open>x \<in> T\<close> by simp |
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5702 |
moreover have "vimage f T \<subseteq> S" |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5703 |
using \<open>T \<subseteq> image f S\<close> \<open>inj f\<close> unfolding inj_on_def subset_eq by auto |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5704 |
ultimately have "y \<in> interior S" .. |
|
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
5705 |
with \<open>x = f y\<close> show "x \<in> f ` interior S" .. |
| 60420 | 5706 |
qed |
5707 |
||
| 67968 | 5708 |
subsection%unimportant \<open>Topological properties of linear functions\<close> |
| 36437 | 5709 |
|
5710 |
lemma linear_lim_0: |
|
| 53282 | 5711 |
assumes "bounded_linear f" |
| 61973 | 5712 |
shows "(f \<longlongrightarrow> 0) (at (0))" |
| 53282 | 5713 |
proof - |
| 36437 | 5714 |
interpret f: bounded_linear f by fact |
| 61973 | 5715 |
have "(f \<longlongrightarrow> f 0) (at 0)" |
| 36437 | 5716 |
using tendsto_ident_at by (rule f.tendsto) |
| 53282 | 5717 |
then show ?thesis unfolding f.zero . |
| 36437 | 5718 |
qed |
5719 |
||
5720 |
lemma linear_continuous_at: |
|
| 53282 | 5721 |
assumes "bounded_linear f" |
5722 |
shows "continuous (at a) f" |
|
| 36437 | 5723 |
unfolding continuous_at using assms |
5724 |
apply (rule bounded_linear.tendsto) |
|
5725 |
apply (rule tendsto_ident_at) |
|
5726 |
done |
|
5727 |
||
5728 |
lemma linear_continuous_within: |
|
| 53291 | 5729 |
"bounded_linear f \<Longrightarrow> continuous (at x within s) f" |
| 36437 | 5730 |
using continuous_at_imp_continuous_within[of x f s] using linear_continuous_at[of f] by auto |
5731 |
||
5732 |
lemma linear_continuous_on: |
|
| 53291 | 5733 |
"bounded_linear f \<Longrightarrow> continuous_on s f" |
| 36437 | 5734 |
using continuous_at_imp_continuous_on[of s f] using linear_continuous_at[of f] by auto |
5735 |
||
| 67962 | 5736 |
subsection%unimportant \<open>Intervals\<close> |
|
60974
6a6f15d8fbc4
New material and fixes related to the forthcoming Stone-Weierstrass development
paulson <lp15@cam.ac.uk>
parents:
60762
diff
changeset
|
5737 |
|
| 60420 | 5738 |
text \<open>Openness of halfspaces.\<close> |
| 33175 | 5739 |
|
5740 |
lemma open_halfspace_lt: "open {x. inner a x < b}"
|
|
| 63332 | 5741 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
| 33175 | 5742 |
|
5743 |
lemma open_halfspace_gt: "open {x. inner a x > b}"
|
|
| 63332 | 5744 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
| 33175 | 5745 |
|
| 53282 | 5746 |
lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\<bullet>i < a}"
|
| 63332 | 5747 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
| 33175 | 5748 |
|
| 53282 | 5749 |
lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\<bullet>i > a}"
|
| 63332 | 5750 |
by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) |
| 33175 | 5751 |
|
| 60420 | 5752 |
text \<open>This gives a simple derivation of limit component bounds.\<close> |
| 33175 | 5753 |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5754 |
lemma open_box[intro]: "open (box a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5755 |
proof - |
| 67399 | 5756 |
have "open (\<Inter>i\<in>Basis. ((\<bullet>) i) -` {a \<bullet> i <..< b \<bullet> i})"
|
|
62533
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents:
62466
diff
changeset
|
5757 |
by (auto intro!: continuous_open_vimage continuous_inner continuous_ident continuous_const) |
| 67399 | 5758 |
also have "(\<Inter>i\<in>Basis. ((\<bullet>) i) -` {a \<bullet> i <..< b \<bullet> i}) = box a b"
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5759 |
by (auto simp: box_def inner_commute) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5760 |
finally show ?thesis . |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5761 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5762 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5763 |
instance euclidean_space \<subseteq> second_countable_topology |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5764 |
proof |
| 63040 | 5765 |
define a where "a f = (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5766 |
then have a: "\<And>f. (\<Sum>i\<in>Basis. fst (f i) *\<^sub>R i) = a f" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5767 |
by simp |
| 63040 | 5768 |
define b where "b f = (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i)" for f :: "'a \<Rightarrow> real \<times> real" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5769 |
then have b: "\<And>f. (\<Sum>i\<in>Basis. snd (f i) *\<^sub>R i) = b f" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5770 |
by simp |
| 63040 | 5771 |
define B where "B = (\<lambda>f. box (a f) (b f)) ` (Basis \<rightarrow>\<^sub>E (\<rat> \<times> \<rat>))" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5772 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5773 |
have "Ball B open" by (simp add: B_def open_box) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5774 |
moreover have "(\<forall>A. open A \<longrightarrow> (\<exists>B'\<subseteq>B. \<Union>B' = A))" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5775 |
proof safe |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5776 |
fix A::"'a set" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5777 |
assume "open A" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5778 |
show "\<exists>B'\<subseteq>B. \<Union>B' = A" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5779 |
apply (rule exI[of _ "{b\<in>B. b \<subseteq> A}"])
|
| 60420 | 5780 |
apply (subst (3) open_UNION_box[OF \<open>open A\<close>]) |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5781 |
apply (auto simp: a b B_def) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5782 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5783 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5784 |
ultimately |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5785 |
have "topological_basis B" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5786 |
unfolding topological_basis_def by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5787 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5788 |
have "countable B" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5789 |
unfolding B_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5790 |
by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5791 |
ultimately show "\<exists>B::'a set set. countable B \<and> open = generate_topology B" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5792 |
by (blast intro: topological_basis_imp_subbasis) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5793 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5794 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5795 |
instance euclidean_space \<subseteq> polish_space .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5796 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5797 |
lemma closed_cbox[intro]: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5798 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5799 |
shows "closed (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5800 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5801 |
have "closed (\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i})"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5802 |
by (intro closed_INT ballI continuous_closed_vimage allI |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5803 |
linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5804 |
also have "(\<Inter>i\<in>Basis. (\<lambda>x. x\<bullet>i) -` {a\<bullet>i .. b\<bullet>i}) = cbox a b"
|
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5805 |
by (auto simp: cbox_def) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5806 |
finally show "closed (cbox a b)" . |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5807 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5808 |
|
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
5809 |
lemma interior_cbox [simp]: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5810 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5811 |
shows "interior (cbox a b) = box a b" (is "?L = ?R") |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5812 |
proof(rule subset_antisym) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5813 |
show "?R \<subseteq> ?L" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5814 |
using box_subset_cbox open_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5815 |
by (rule interior_maximal) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5816 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5817 |
fix x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5818 |
assume "x \<in> interior (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5819 |
then obtain s where s: "open s" "x \<in> s" "s \<subseteq> cbox a b" .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5820 |
then obtain e where "e>0" and e:"\<forall>x'. dist x' x < e \<longrightarrow> x' \<in> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5821 |
unfolding open_dist and subset_eq by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5822 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5823 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5824 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5825 |
have "dist (x - (e / 2) *\<^sub>R i) x < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5826 |
and "dist (x + (e / 2) *\<^sub>R i) x < e" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5827 |
unfolding dist_norm |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5828 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5829 |
unfolding norm_minus_cancel |
| 60420 | 5830 |
using norm_Basis[OF i] \<open>e>0\<close> |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5831 |
apply auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5832 |
done |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5833 |
then have "a \<bullet> i \<le> (x - (e / 2) *\<^sub>R i) \<bullet> i" and "(x + (e / 2) *\<^sub>R i) \<bullet> i \<le> b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5834 |
using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5835 |
and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5836 |
unfolding mem_box |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5837 |
using i |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5838 |
by blast+ |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5839 |
then have "a \<bullet> i < x \<bullet> i" and "x \<bullet> i < b \<bullet> i" |
| 60420 | 5840 |
using \<open>e>0\<close> i |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5841 |
by (auto simp: inner_diff_left inner_Basis inner_add_left) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5842 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5843 |
then have "x \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5844 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5845 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5846 |
then show "?L \<subseteq> ?R" .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5847 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5848 |
|
|
63928
d81fb5b46a5c
new material about topological concepts, etc
paulson <lp15@cam.ac.uk>
parents:
63881
diff
changeset
|
5849 |
lemma bounded_cbox [simp]: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5850 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5851 |
shows "bounded (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5852 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5853 |
let ?b = "\<Sum>i\<in>Basis. \<bar>a\<bullet>i\<bar> + \<bar>b\<bullet>i\<bar>" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5854 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5855 |
fix x :: "'a" |
| 68120 | 5856 |
assume "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5857 |
then have "(\<Sum>i\<in>Basis. \<bar>x \<bullet> i\<bar>) \<le> ?b" |
| 68120 | 5858 |
by (force simp: intro!: sum_mono) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5859 |
then have "norm x \<le> ?b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5860 |
using norm_le_l1[of x] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5861 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5862 |
then show ?thesis |
| 68120 | 5863 |
unfolding cbox_def bounded_iff by force |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5864 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5865 |
|
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
5866 |
lemma bounded_box [simp]: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5867 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5868 |
shows "bounded (box a b)" |
| 68120 | 5869 |
using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5870 |
by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5871 |
|
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
5872 |
lemma not_interval_UNIV [simp]: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5873 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5874 |
shows "cbox a b \<noteq> UNIV" "box a b \<noteq> UNIV" |
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
5875 |
using bounded_box[of a b] bounded_cbox[of a b] by force+ |
|
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
5876 |
|
|
63945
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
5877 |
lemma not_interval_UNIV2 [simp]: |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
5878 |
fixes a :: "'a::euclidean_space" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
5879 |
shows "UNIV \<noteq> cbox a b" "UNIV \<noteq> box a b" |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
5880 |
using bounded_box[of a b] bounded_cbox[of a b] by force+ |
|
444eafb6e864
a few new theorems and a renaming
paulson <lp15@cam.ac.uk>
parents:
63938
diff
changeset
|
5881 |
|
|
62618
f7f2467ab854
Refactoring (moving theorems into better locations), plus a bit of new material
paulson <lp15@cam.ac.uk>
parents:
62533
diff
changeset
|
5882 |
lemma compact_cbox [simp]: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5883 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5884 |
shows "compact (cbox a b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5885 |
using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5886 |
by (auto simp: compact_eq_seq_compact_metric) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5887 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5888 |
lemma box_midpoint: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5889 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5890 |
assumes "box a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5891 |
shows "((1/2) *\<^sub>R (a + b)) \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5892 |
proof - |
| 68120 | 5893 |
have "a \<bullet> i < ((1 / 2) *\<^sub>R (a + b)) \<bullet> i \<and> ((1 / 2) *\<^sub>R (a + b)) \<bullet> i < b \<bullet> i" if "i \<in> Basis" for i |
5894 |
using assms that by (auto simp: inner_add_left box_ne_empty) |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5895 |
then show ?thesis unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5896 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5897 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5898 |
lemma open_cbox_convex: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5899 |
fixes x :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5900 |
assumes x: "x \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5901 |
and y: "y \<in> cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5902 |
and e: "0 < e" "e \<le> 1" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5903 |
shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<in> box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5904 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5905 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5906 |
fix i :: 'a |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5907 |
assume i: "i \<in> Basis" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5908 |
have "a \<bullet> i = e * (a \<bullet> i) + (1 - e) * (a \<bullet> i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5909 |
unfolding left_diff_distrib by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5910 |
also have "\<dots> < e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" |
| 68120 | 5911 |
proof (rule add_less_le_mono) |
5912 |
show "e * (a \<bullet> i) < e * (x \<bullet> i)" |
|
5913 |
using \<open>0 < e\<close> i mem_box(1) x by auto |
|
5914 |
show "(1 - e) * (a \<bullet> i) \<le> (1 - e) * (y \<bullet> i)" |
|
5915 |
by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y) |
|
5916 |
qed |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5917 |
finally have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5918 |
unfolding inner_simps by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5919 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5920 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5921 |
have "b \<bullet> i = e * (b\<bullet>i) + (1 - e) * (b\<bullet>i)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5922 |
unfolding left_diff_distrib by simp |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5923 |
also have "\<dots> > e * (x \<bullet> i) + (1 - e) * (y \<bullet> i)" |
|
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
5924 |
proof (rule add_less_le_mono) |
|
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
5925 |
show "e * (x \<bullet> i) < e * (b \<bullet> i)" |
| 68120 | 5926 |
using \<open>0 < e\<close> i mem_box(1) x by auto |
|
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
5927 |
show "(1 - e) * (y \<bullet> i) \<le> (1 - e) * (b \<bullet> i)" |
| 68120 | 5928 |
by (meson diff_ge_0_iff_ge \<open>e \<le> 1\<close> i mem_box(2) mult_left_mono y) |
|
66827
c94531b5007d
Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents:
66794
diff
changeset
|
5929 |
qed |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5930 |
finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5931 |
unfolding inner_simps by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5932 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5933 |
ultimately have "a \<bullet> i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i \<and> (e *\<^sub>R x + (1 - e) *\<^sub>R y) \<bullet> i < b \<bullet> i" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5934 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5935 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5936 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5937 |
unfolding mem_box by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5938 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5939 |
|
|
63881
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5940 |
lemma closure_cbox [simp]: "closure (cbox a b) = cbox a b" |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5941 |
by (simp add: closed_cbox) |
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5942 |
|
|
b746b19197bd
lots of new results about topology, affine dimension etc
paulson <lp15@cam.ac.uk>
parents:
63627
diff
changeset
|
5943 |
lemma closure_box [simp]: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5944 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5945 |
assumes "box a b \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5946 |
shows "closure (box a b) = cbox a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5947 |
proof - |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5948 |
have ab: "a <e b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5949 |
using assms by (simp add: eucl_less_def box_ne_empty) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5950 |
let ?c = "(1 / 2) *\<^sub>R (a + b)" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5951 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5952 |
fix x |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5953 |
assume as:"x \<in> cbox a b" |
| 63040 | 5954 |
define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5955 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5956 |
fix n |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5957 |
assume fn: "f n <e b \<longrightarrow> a <e f n \<longrightarrow> f n = x" and xc: "x \<noteq> ?c" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5958 |
have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \<le> 1" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5959 |
unfolding inverse_le_1_iff by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5960 |
have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5961 |
x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5962 |
by (auto simp: algebra_simps) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5963 |
then have "f n <e b" and "a <e f n" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5964 |
using open_cbox_convex[OF box_midpoint[OF assms] as *] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5965 |
unfolding f_def by (auto simp: box_def eucl_less_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5966 |
then have False |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5967 |
using fn unfolding f_def using xc by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5968 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5969 |
moreover |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5970 |
{
|
| 61973 | 5971 |
assume "\<not> (f \<longlongrightarrow> x) sequentially" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5972 |
{
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5973 |
fix e :: real |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5974 |
assume "e > 0" |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
5975 |
then obtain N :: nat where N: "inverse (real (N + 1)) < e" |
| 68120 | 5976 |
using reals_Archimedean by auto |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
5977 |
have "inverse (real n + 1) < e" if "N \<le> n" for n |
|
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
61552
diff
changeset
|
5978 |
by (auto intro!: that le_less_trans [OF _ N]) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5979 |
then have "\<exists>N::nat. \<forall>n\<ge>N. inverse (real n + 1) < e" by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5980 |
} |
| 61973 | 5981 |
then have "((\<lambda>n. inverse (real n + 1)) \<longlongrightarrow> 0) sequentially" |
|
66643
f7e38b8583a0
Correction of typos and a bit of streamlining
paulson <lp15@cam.ac.uk>
parents:
66641
diff
changeset
|
5982 |
unfolding lim_sequentially by(auto simp: dist_norm) |
| 61973 | 5983 |
then have "(f \<longlongrightarrow> x) sequentially" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5984 |
unfolding f_def |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5985 |
using tendsto_add[OF tendsto_const, of "\<lambda>n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5986 |
using tendsto_scaleR [OF _ tendsto_const, of "\<lambda>n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5987 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5988 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5989 |
ultimately have "x \<in> closure (box a b)" |
| 68120 | 5990 |
using as box_midpoint[OF assms] |
5991 |
unfolding closure_def islimpt_sequential |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5992 |
by (cases "x=?c") (auto simp: in_box_eucl_less) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5993 |
} |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5994 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5995 |
using closure_minimal[OF box_subset_cbox, of a b] by blast |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5996 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5997 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
5998 |
lemma bounded_subset_box_symmetric: |
| 68120 | 5999 |
fixes S :: "('a::euclidean_space) set"
|
6000 |
assumes "bounded S" |
|
6001 |
obtains a where "S \<subseteq> box (-a) a" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6002 |
proof - |
| 68120 | 6003 |
obtain b where "b>0" and b: "\<forall>x\<in>S. norm x \<le> b" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6004 |
using assms[unfolded bounded_pos] by auto |
| 63040 | 6005 |
define a :: 'a where "a = (\<Sum>i\<in>Basis. (b + 1) *\<^sub>R i)" |
| 68120 | 6006 |
have "(-a)\<bullet>i < x\<bullet>i" and "x\<bullet>i < a\<bullet>i" if "x \<in> S" and i: "i \<in> Basis" for x i |
6007 |
using b Basis_le_norm[OF i, of x] that by (auto simp: a_def) |
|
6008 |
then have "S \<subseteq> box (-a) a" |
|
6009 |
by (auto simp: simp add: box_def) |
|
6010 |
then show ?thesis .. |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6011 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6012 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6013 |
lemma bounded_subset_cbox_symmetric: |
| 68120 | 6014 |
fixes S :: "('a::euclidean_space) set"
|
6015 |
assumes "bounded S" |
|
6016 |
obtains a where "S \<subseteq> cbox (-a) a" |
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6017 |
proof - |
| 68120 | 6018 |
obtain a where "S \<subseteq> box (-a) a" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6019 |
using bounded_subset_box_symmetric[OF assms] by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6020 |
then show ?thesis |
| 68120 | 6021 |
by (meson box_subset_cbox dual_order.trans that) |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6022 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6023 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6024 |
lemma frontier_cbox: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6025 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6026 |
shows "frontier (cbox a b) = cbox a b - box a b" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6027 |
unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6028 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6029 |
lemma frontier_box: |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6030 |
fixes a b :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6031 |
shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6032 |
proof (cases "box a b = {}")
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6033 |
case True |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6034 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6035 |
using frontier_empty by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6036 |
next |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6037 |
case False |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6038 |
then show ?thesis |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6039 |
unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6040 |
by auto |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6041 |
qed |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6042 |
|
|
66884
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents:
66827
diff
changeset
|
6043 |
lemma Int_interval_mixed_eq_empty: |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6044 |
fixes a :: "'a::euclidean_space" |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6045 |
assumes "box c d \<noteq> {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6046 |
shows "box a b \<inter> cbox c d = {} \<longleftrightarrow> box a b \<inter> box c d = {}"
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6047 |
unfolding closure_box[OF assms, symmetric] |
|
62843
313d3b697c9a
Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents:
62623
diff
changeset
|
6048 |
unfolding open_Int_closure_eq_empty[OF open_box] .. |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6049 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6050 |
lemma eucl_less_eq_halfspaces: |
| 61076 | 6051 |
fixes a :: "'a::euclidean_space" |
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6052 |
shows "{x. x <e a} = (\<Inter>i\<in>Basis. {x. x \<bullet> i < a \<bullet> i})"
|
| 68120 | 6053 |
"{x. a <e x} = (\<Inter>i\<in>Basis. {x. a \<bullet> i < x \<bullet> i})"
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6054 |
by (auto simp: eucl_less_def) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6055 |
|
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6056 |
lemma open_Collect_eucl_less[simp, intro]: |
| 61076 | 6057 |
fixes a :: "'a::euclidean_space" |
| 68120 | 6058 |
shows "open {x. x <e a}" "open {x. a <e x}"
|
|
56189
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6059 |
by (auto simp: eucl_less_eq_halfspaces open_halfspace_component_lt open_halfspace_component_gt) |
|
c4daa97ac57a
removed dependencies on theory Ordered_Euclidean_Space
immler
parents:
56188
diff
changeset
|
6060 |
|
|
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
6061 |
no_notation |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
6062 |
eucl_less (infix "<e" 50) |
|
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54489
diff
changeset
|
6063 |
|
| 33175 | 6064 |
end |