author | paulson <lp15@cam.ac.uk> |
Tue, 02 May 2023 12:51:05 +0100 | |
changeset 77934 | 01c88cf514fc |
parent 71172 | 575b3a818de5 |
child 77943 | ffdad62bc235 |
permissions | -rw-r--r-- |
69874 | 1 |
theory Abstract_Limits |
2 |
imports |
|
3 |
Abstract_Topology |
|
4 |
begin |
|
5 |
||
6 |
subsection\<open>nhdsin and atin\<close> |
|
7 |
||
8 |
definition nhdsin :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a filter" |
|
9 |
where "nhdsin X a = |
|
70337
48609a6af1a0
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
70065
diff
changeset
|
10 |
(if a \<in> topspace X then (INF S\<in>{S. openin X S \<and> a \<in> S}. principal S) else bot)" |
69874 | 11 |
|
12 |
definition atin :: "'a topology \<Rightarrow> 'a \<Rightarrow> 'a filter" |
|
13 |
where "atin X a \<equiv> inf (nhdsin X a) (principal (topspace X - {a}))" |
|
14 |
||
15 |
||
16 |
lemma nhdsin_degenerate [simp]: "a \<notin> topspace X \<Longrightarrow> nhdsin X a = bot" |
|
17 |
and atin_degenerate [simp]: "a \<notin> topspace X \<Longrightarrow> atin X a = bot" |
|
18 |
by (simp_all add: nhdsin_def atin_def) |
|
19 |
||
20 |
lemma eventually_nhdsin: |
|
21 |
"eventually P (nhdsin X a) \<longleftrightarrow> a \<notin> topspace X \<or> (\<exists>S. openin X S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" |
|
22 |
proof (cases "a \<in> topspace X") |
|
23 |
case True |
|
70337
48609a6af1a0
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
70065
diff
changeset
|
24 |
hence "nhdsin X a = (INF S\<in>{S. openin X S \<and> a \<in> S}. principal S)" |
69874 | 25 |
by (simp add: nhdsin_def) |
26 |
also have "eventually P \<dots> \<longleftrightarrow> (\<exists>S. openin X S \<and> a \<in> S \<and> (\<forall>x\<in>S. P x))" |
|
27 |
using True by (subst eventually_INF_base) (auto simp: eventually_principal) |
|
28 |
finally show ?thesis using True by simp |
|
29 |
qed auto |
|
30 |
||
31 |
lemma eventually_atin: |
|
32 |
"eventually P (atin X a) \<longleftrightarrow> a \<notin> topspace X \<or> |
|
33 |
(\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x \<in> U - {a}. P x))" |
|
34 |
proof (cases "a \<in> topspace X") |
|
35 |
case True |
|
36 |
hence "eventually P (atin X a) \<longleftrightarrow> (\<exists>S. openin X S \<and> |
|
37 |
a \<in> S \<and> (\<forall>x\<in>S. x \<in> topspace X \<and> x \<noteq> a \<longrightarrow> P x))" |
|
38 |
by (simp add: atin_def eventually_inf_principal eventually_nhdsin) |
|
39 |
also have "\<dots> \<longleftrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> (\<forall>x \<in> U - {a}. P x))" |
|
40 |
using openin_subset by (intro ex_cong) auto |
|
41 |
finally show ?thesis by (simp add: True) |
|
42 |
qed auto |
|
43 |
||
44 |
||
45 |
subsection\<open>Limits in a topological space\<close> |
|
46 |
||
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
47 |
definition limitin :: "'a topology \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b filter \<Rightarrow> bool" where |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
48 |
"limitin X f l F \<equiv> l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> eventually (\<lambda>x. f x \<in> U) F)" |
69874 | 49 |
|
77934 | 50 |
lemma limitinD: "\<lbrakk>limitin X f l F; openin X U; l \<in> U\<rbrakk> \<Longrightarrow> eventually (\<lambda>x. f x \<in> U) F" |
51 |
by (simp add: limitin_def) |
|
52 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
53 |
lemma limitin_canonical_iff [simp]: "limitin euclidean f l F \<longleftrightarrow> (f \<longlongrightarrow> l) F" |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
54 |
by (auto simp: limitin_def tendsto_def) |
69874 | 55 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
56 |
lemma limitin_topspace: "limitin X f l F \<Longrightarrow> l \<in> topspace X" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
57 |
by (simp add: limitin_def) |
69874 | 58 |
|
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
59 |
lemma limitin_const_iff [simp]: "limitin X (\<lambda>a. l) l F \<longleftrightarrow> l \<in> topspace X" |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
60 |
by (simp add: limitin_def) |
69874 | 61 |
|
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
62 |
lemma limitin_const: "limitin euclidean (\<lambda>a. l) l F" |
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
63 |
by simp |
69874 | 64 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
65 |
lemma limitin_eventually: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
66 |
"\<lbrakk>l \<in> topspace X; eventually (\<lambda>x. f x = l) F\<rbrakk> \<Longrightarrow> limitin X f l F" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
67 |
by (auto simp: limitin_def eventually_mono) |
69874 | 68 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
69 |
lemma limitin_subsequence: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
70 |
"\<lbrakk>strict_mono r; limitin X f l sequentially\<rbrakk> \<Longrightarrow> limitin X (f \<circ> r) l sequentially" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
71 |
unfolding limitin_def using eventually_subseq by fastforce |
69874 | 72 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
73 |
lemma limitin_subtopology: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
74 |
"limitin (subtopology X S) f l F |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
75 |
\<longleftrightarrow> l \<in> S \<and> eventually (\<lambda>a. f a \<in> S) F \<and> limitin X f l F" (is "?lhs = ?rhs") |
69874 | 76 |
proof (cases "l \<in> S \<inter> topspace X") |
77 |
case True |
|
78 |
show ?thesis |
|
79 |
proof |
|
80 |
assume L: ?lhs |
|
81 |
with True |
|
82 |
have "\<forall>\<^sub>F b in F. f b \<in> topspace X \<inter> S" |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
83 |
by (metis (no_types) limitin_def openin_topspace topspace_subtopology) |
69874 | 84 |
with L show ?rhs |
71172 | 85 |
apply (clarsimp simp add: limitin_def eventually_mono openin_subtopology_alt) |
69874 | 86 |
apply (drule_tac x="S \<inter> U" in spec, force simp: elim: eventually_mono) |
87 |
done |
|
88 |
next |
|
89 |
assume ?rhs |
|
90 |
then show ?lhs |
|
91 |
using eventually_elim2 |
|
71172 | 92 |
by (fastforce simp add: limitin_def openin_subtopology_alt) |
69874 | 93 |
qed |
71172 | 94 |
qed (auto simp: limitin_def) |
69874 | 95 |
|
96 |
||
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
97 |
lemma limitin_canonical_iff_gen [simp]: |
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
98 |
assumes "open S" |
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
99 |
shows "limitin (top_of_set S) f l F \<longleftrightarrow> (f \<longlongrightarrow> l) F \<and> l \<in> S" |
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
100 |
using assms by (auto simp: limitin_subtopology tendsto_def) |
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
101 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
102 |
lemma limitin_sequentially: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
103 |
"limitin X S l sequentially \<longleftrightarrow> |
69874 | 104 |
l \<in> topspace X \<and> (\<forall>U. openin X U \<and> l \<in> U \<longrightarrow> (\<exists>N. \<forall>n. N \<le> n \<longrightarrow> S n \<in> U))" |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
105 |
by (simp add: limitin_def eventually_sequentially) |
69874 | 106 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
107 |
lemma limitin_sequentially_offset: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
108 |
"limitin X f l sequentially \<Longrightarrow> limitin X (\<lambda>i. f (i + k)) l sequentially" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
109 |
unfolding limitin_sequentially |
69874 | 110 |
by (metis add.commute le_add2 order_trans) |
111 |
||
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
112 |
lemma limitin_sequentially_offset_rev: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
113 |
assumes "limitin X (\<lambda>i. f (i + k)) l sequentially" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
114 |
shows "limitin X f l sequentially" |
69874 | 115 |
proof - |
116 |
have "\<exists>N. \<forall>n\<ge>N. f n \<in> U" if U: "openin X U" "l \<in> U" for U |
|
117 |
proof - |
|
118 |
obtain N where "\<And>n. n\<ge>N \<Longrightarrow> f (n + k) \<in> U" |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
119 |
using assms U unfolding limitin_sequentially by blast |
69874 | 120 |
then have "\<forall>n\<ge>N+k. f n \<in> U" |
121 |
by (metis add_leD2 le_add_diff_inverse ordered_cancel_comm_monoid_diff_class.le_diff_conv2 add.commute) |
|
122 |
then show ?thesis .. |
|
123 |
qed |
|
124 |
with assms show ?thesis |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
125 |
unfolding limitin_sequentially |
69874 | 126 |
by simp |
127 |
qed |
|
128 |
||
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
129 |
lemma limitin_atin: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
130 |
"limitin Y f y (atin X x) \<longleftrightarrow> |
69874 | 131 |
y \<in> topspace Y \<and> |
132 |
(x \<in> topspace X |
|
133 |
\<longrightarrow> (\<forall>V. openin Y V \<and> y \<in> V |
|
134 |
\<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> f ` (U - {x}) \<subseteq> V)))" |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
135 |
by (auto simp: limitin_def eventually_atin image_subset_iff) |
69874 | 136 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
137 |
lemma limitin_atin_self: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
138 |
"limitin Y f (f a) (atin X a) \<longleftrightarrow> |
69874 | 139 |
f a \<in> topspace Y \<and> |
140 |
(a \<in> topspace X |
|
141 |
\<longrightarrow> (\<forall>V. openin Y V \<and> f a \<in> V |
|
142 |
\<longrightarrow> (\<exists>U. openin X U \<and> a \<in> U \<and> f ` U \<subseteq> V)))" |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
143 |
unfolding limitin_atin by fastforce |
69874 | 144 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
145 |
lemma limitin_trivial: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
146 |
"\<lbrakk>trivial_limit F; y \<in> topspace X\<rbrakk> \<Longrightarrow> limitin X f y F" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
147 |
by (simp add: limitin_def) |
69874 | 148 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
149 |
lemma limitin_transform_eventually: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
150 |
"\<lbrakk>eventually (\<lambda>x. f x = g x) F; limitin X f l F\<rbrakk> \<Longrightarrow> limitin X g l F" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
151 |
unfolding limitin_def using eventually_elim2 by fastforce |
69874 | 152 |
|
153 |
lemma continuous_map_limit: |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
154 |
assumes "continuous_map X Y g" and f: "limitin X f l F" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
155 |
shows "limitin Y (g \<circ> f) (g l) F" |
69874 | 156 |
proof - |
157 |
have "g l \<in> topspace Y" |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
158 |
by (meson assms continuous_map_def limitin_topspace) |
69874 | 159 |
moreover |
160 |
have "\<And>U. \<lbrakk>\<forall>V. openin X V \<and> l \<in> V \<longrightarrow> (\<forall>\<^sub>F x in F. f x \<in> V); openin Y U; g l \<in> U\<rbrakk> |
|
161 |
\<Longrightarrow> \<forall>\<^sub>F x in F. g (f x) \<in> U" |
|
162 |
using assms eventually_mono |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
163 |
by (fastforce simp: limitin_def dest!: openin_continuous_map_preimage) |
69874 | 164 |
ultimately show ?thesis |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
165 |
using f by (fastforce simp add: limitin_def) |
69874 | 166 |
qed |
167 |
||
168 |
||
169 |
subsection\<open>Pointwise continuity in topological spaces\<close> |
|
170 |
||
171 |
definition topcontinuous_at where |
|
172 |
"topcontinuous_at X Y f x \<longleftrightarrow> |
|
173 |
x \<in> topspace X \<and> |
|
174 |
(\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and> |
|
175 |
(\<forall>V. openin Y V \<and> f x \<in> V |
|
176 |
\<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y \<in> U. f y \<in> V)))" |
|
177 |
||
178 |
lemma topcontinuous_at_atin: |
|
179 |
"topcontinuous_at X Y f x \<longleftrightarrow> |
|
180 |
x \<in> topspace X \<and> |
|
181 |
(\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and> |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
182 |
limitin Y f (f x) (atin X x)" |
69874 | 183 |
unfolding topcontinuous_at_def |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
184 |
by (fastforce simp add: limitin_atin)+ |
69874 | 185 |
|
186 |
lemma continuous_map_eq_topcontinuous_at: |
|
187 |
"continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. topcontinuous_at X Y f x)" |
|
188 |
(is "?lhs = ?rhs") |
|
189 |
proof |
|
190 |
assume ?lhs |
|
191 |
then show ?rhs |
|
192 |
by (auto simp: continuous_map_def topcontinuous_at_def) |
|
193 |
next |
|
194 |
assume R: ?rhs |
|
195 |
then show ?lhs |
|
196 |
apply (auto simp: continuous_map_def topcontinuous_at_def) |
|
197 |
apply (subst openin_subopen, safe) |
|
198 |
apply (drule bspec, assumption) |
|
199 |
using openin_subset[of X] apply (auto simp: subset_iff dest!: spec) |
|
200 |
done |
|
201 |
qed |
|
202 |
||
203 |
lemma continuous_map_atin: |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
204 |
"continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. limitin Y f (f x) (atin X x))" |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
205 |
by (auto simp: limitin_def topcontinuous_at_atin continuous_map_eq_topcontinuous_at) |
69874 | 206 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
207 |
lemma limitin_continuous_map: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
208 |
"\<lbrakk>continuous_map X Y f; a \<in> topspace X; f a = b\<rbrakk> \<Longrightarrow> limitin Y f b (atin X a)" |
69874 | 209 |
by (auto simp: continuous_map_atin) |
210 |
||
211 |
||
212 |
subsection\<open>Combining theorems for continuous functions into the reals\<close> |
|
213 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
214 |
lemma continuous_map_canonical_const [continuous_intros]: |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
215 |
"continuous_map X euclidean (\<lambda>x. c)" |
69874 | 216 |
by simp |
217 |
||
218 |
lemma continuous_map_real_mult [continuous_intros]: |
|
219 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> |
|
220 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * g x)" |
|
221 |
by (simp add: continuous_map_atin tendsto_mult) |
|
222 |
||
223 |
lemma continuous_map_real_pow [continuous_intros]: |
|
224 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x ^ n)" |
|
225 |
by (induction n) (auto simp: continuous_map_real_mult) |
|
226 |
||
227 |
lemma continuous_map_real_mult_left: |
|
228 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. c * f x)" |
|
229 |
by (simp add: continuous_map_atin tendsto_mult) |
|
230 |
||
231 |
lemma continuous_map_real_mult_left_eq: |
|
232 |
"continuous_map X euclideanreal (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f" |
|
233 |
proof (cases "c = 0") |
|
234 |
case False |
|
235 |
have "continuous_map X euclideanreal (\<lambda>x. c * f x) \<Longrightarrow> continuous_map X euclideanreal f" |
|
236 |
apply (frule continuous_map_real_mult_left [where c="inverse c"]) |
|
237 |
apply (simp add: field_simps False) |
|
238 |
done |
|
239 |
with False show ?thesis |
|
240 |
using continuous_map_real_mult_left by blast |
|
241 |
qed simp |
|
242 |
||
243 |
lemma continuous_map_real_mult_right: |
|
244 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * c)" |
|
245 |
by (simp add: continuous_map_atin tendsto_mult) |
|
246 |
||
247 |
lemma continuous_map_real_mult_right_eq: |
|
248 |
"continuous_map X euclideanreal (\<lambda>x. f x * c) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f" |
|
249 |
by (simp add: mult.commute flip: continuous_map_real_mult_left_eq) |
|
250 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
251 |
lemma continuous_map_minus [continuous_intros]: |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
252 |
fixes f :: "'a\<Rightarrow>'b::real_normed_vector" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
253 |
shows "continuous_map X euclidean f \<Longrightarrow> continuous_map X euclidean (\<lambda>x. - f x)" |
69874 | 254 |
by (simp add: continuous_map_atin tendsto_minus) |
255 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
256 |
lemma continuous_map_minus_eq [simp]: |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
257 |
fixes f :: "'a\<Rightarrow>'b::real_normed_vector" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
258 |
shows "continuous_map X euclidean (\<lambda>x. - f x) \<longleftrightarrow> continuous_map X euclidean f" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
259 |
using continuous_map_minus add.inverse_inverse continuous_map_eq by fastforce |
69874 | 260 |
|
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
261 |
lemma continuous_map_add [continuous_intros]: |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
262 |
fixes f :: "'a\<Rightarrow>'b::real_normed_vector" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
263 |
shows "\<lbrakk>continuous_map X euclidean f; continuous_map X euclidean g\<rbrakk> \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x + g x)" |
69874 | 264 |
by (simp add: continuous_map_atin tendsto_add) |
265 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
266 |
lemma continuous_map_diff [continuous_intros]: |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
267 |
fixes f :: "'a\<Rightarrow>'b::real_normed_vector" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
268 |
shows "\<lbrakk>continuous_map X euclidean f; continuous_map X euclidean g\<rbrakk> \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x - g x)" |
69874 | 269 |
by (simp add: continuous_map_atin tendsto_diff) |
270 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
271 |
lemma continuous_map_norm [continuous_intros]: |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
272 |
fixes f :: "'a\<Rightarrow>'b::real_normed_vector" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
273 |
shows "continuous_map X euclidean f \<Longrightarrow> continuous_map X euclidean (\<lambda>x. norm(f x))" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
274 |
by (simp add: continuous_map_atin tendsto_norm) |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
275 |
|
69874 | 276 |
lemma continuous_map_real_abs [continuous_intros]: |
277 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. abs(f x))" |
|
278 |
by (simp add: continuous_map_atin tendsto_rabs) |
|
279 |
||
280 |
lemma continuous_map_real_max [continuous_intros]: |
|
281 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> |
|
282 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. max (f x) (g x))" |
|
283 |
by (simp add: continuous_map_atin tendsto_max) |
|
284 |
||
285 |
lemma continuous_map_real_min [continuous_intros]: |
|
286 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> |
|
287 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. min (f x) (g x))" |
|
288 |
by (simp add: continuous_map_atin tendsto_min) |
|
289 |
||
290 |
lemma continuous_map_sum [continuous_intros]: |
|
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
291 |
fixes f :: "'a\<Rightarrow>'b\<Rightarrow>'c::real_normed_vector" |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
292 |
shows "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> continuous_map X euclidean (\<lambda>x. f x i)\<rbrakk> |
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
293 |
\<Longrightarrow> continuous_map X euclidean (\<lambda>x. sum (f x) I)" |
69874 | 294 |
by (simp add: continuous_map_atin tendsto_sum) |
295 |
||
296 |
lemma continuous_map_prod [continuous_intros]: |
|
297 |
"\<lbrakk>finite I; |
|
298 |
\<And>i. i \<in> I \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)\<rbrakk> |
|
299 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. prod (f x) I)" |
|
300 |
by (simp add: continuous_map_atin tendsto_prod) |
|
301 |
||
302 |
lemma continuous_map_real_inverse [continuous_intros]: |
|
303 |
"\<lbrakk>continuous_map X euclideanreal f; \<And>x. x \<in> topspace X \<Longrightarrow> f x \<noteq> 0\<rbrakk> |
|
304 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. inverse(f x))" |
|
305 |
by (simp add: continuous_map_atin tendsto_inverse) |
|
306 |
||
307 |
lemma continuous_map_real_divide [continuous_intros]: |
|
308 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g; \<And>x. x \<in> topspace X \<Longrightarrow> g x \<noteq> 0\<rbrakk> |
|
309 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x / g x)" |
|
310 |
by (simp add: continuous_map_atin tendsto_divide) |
|
311 |
||
312 |
end |