author | haftmann |
Mon, 13 Sep 2021 14:18:24 +0000 | |
changeset 74309 | 42523fbf643b |
parent 71172 | 575b3a818de5 |
child 77934 | 01c88cf514fc |
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 |
|
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
50 |
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
|
51 |
by (auto simp: limitin_def tendsto_def) |
69874 | 52 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
53 |
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
|
54 |
by (simp add: limitin_def) |
69874 | 55 |
|
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
56 |
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
|
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: "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
|
60 |
by simp |
69874 | 61 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
62 |
lemma limitin_eventually: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
63 |
"\<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
|
64 |
by (auto simp: limitin_def eventually_mono) |
69874 | 65 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
66 |
lemma limitin_subsequence: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
67 |
"\<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
|
68 |
unfolding limitin_def using eventually_subseq by fastforce |
69874 | 69 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
70 |
lemma limitin_subtopology: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
71 |
"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
|
72 |
\<longleftrightarrow> l \<in> S \<and> eventually (\<lambda>a. f a \<in> S) F \<and> limitin X f l F" (is "?lhs = ?rhs") |
69874 | 73 |
proof (cases "l \<in> S \<inter> topspace X") |
74 |
case True |
|
75 |
show ?thesis |
|
76 |
proof |
|
77 |
assume L: ?lhs |
|
78 |
with True |
|
79 |
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
|
80 |
by (metis (no_types) limitin_def openin_topspace topspace_subtopology) |
69874 | 81 |
with L show ?rhs |
71172 | 82 |
apply (clarsimp simp add: limitin_def eventually_mono openin_subtopology_alt) |
69874 | 83 |
apply (drule_tac x="S \<inter> U" in spec, force simp: elim: eventually_mono) |
84 |
done |
|
85 |
next |
|
86 |
assume ?rhs |
|
87 |
then show ?lhs |
|
88 |
using eventually_elim2 |
|
71172 | 89 |
by (fastforce simp add: limitin_def openin_subtopology_alt) |
69874 | 90 |
qed |
71172 | 91 |
qed (auto simp: limitin_def) |
69874 | 92 |
|
93 |
||
70065
cc89a395b5a3
Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents:
70019
diff
changeset
|
94 |
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
|
95 |
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
|
96 |
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
|
97 |
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
|
98 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
99 |
lemma limitin_sequentially: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
100 |
"limitin X S l sequentially \<longleftrightarrow> |
69874 | 101 |
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
|
102 |
by (simp add: limitin_def eventually_sequentially) |
69874 | 103 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
104 |
lemma limitin_sequentially_offset: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
105 |
"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
|
106 |
unfolding limitin_sequentially |
69874 | 107 |
by (metis add.commute le_add2 order_trans) |
108 |
||
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
109 |
lemma limitin_sequentially_offset_rev: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
110 |
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
|
111 |
shows "limitin X f l sequentially" |
69874 | 112 |
proof - |
113 |
have "\<exists>N. \<forall>n\<ge>N. f n \<in> U" if U: "openin X U" "l \<in> U" for U |
|
114 |
proof - |
|
115 |
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
|
116 |
using assms U unfolding limitin_sequentially by blast |
69874 | 117 |
then have "\<forall>n\<ge>N+k. f n \<in> U" |
118 |
by (metis add_leD2 le_add_diff_inverse ordered_cancel_comm_monoid_diff_class.le_diff_conv2 add.commute) |
|
119 |
then show ?thesis .. |
|
120 |
qed |
|
121 |
with assms show ?thesis |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
122 |
unfolding limitin_sequentially |
69874 | 123 |
by simp |
124 |
qed |
|
125 |
||
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
126 |
lemma limitin_atin: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
127 |
"limitin Y f y (atin X x) \<longleftrightarrow> |
69874 | 128 |
y \<in> topspace Y \<and> |
129 |
(x \<in> topspace X |
|
130 |
\<longrightarrow> (\<forall>V. openin Y V \<and> y \<in> V |
|
131 |
\<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
|
132 |
by (auto simp: limitin_def eventually_atin image_subset_iff) |
69874 | 133 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
134 |
lemma limitin_atin_self: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
135 |
"limitin Y f (f a) (atin X a) \<longleftrightarrow> |
69874 | 136 |
f a \<in> topspace Y \<and> |
137 |
(a \<in> topspace X |
|
138 |
\<longrightarrow> (\<forall>V. openin Y V \<and> f a \<in> V |
|
139 |
\<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
|
140 |
unfolding limitin_atin by fastforce |
69874 | 141 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
142 |
lemma limitin_trivial: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
143 |
"\<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
|
144 |
by (simp add: limitin_def) |
69874 | 145 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
146 |
lemma limitin_transform_eventually: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
147 |
"\<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
|
148 |
unfolding limitin_def using eventually_elim2 by fastforce |
69874 | 149 |
|
150 |
lemma continuous_map_limit: |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
151 |
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
|
152 |
shows "limitin Y (g \<circ> f) (g l) F" |
69874 | 153 |
proof - |
154 |
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
|
155 |
by (meson assms continuous_map_def limitin_topspace) |
69874 | 156 |
moreover |
157 |
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> |
|
158 |
\<Longrightarrow> \<forall>\<^sub>F x in F. g (f x) \<in> U" |
|
159 |
using assms eventually_mono |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
160 |
by (fastforce simp: limitin_def dest!: openin_continuous_map_preimage) |
69874 | 161 |
ultimately show ?thesis |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
162 |
using f by (fastforce simp add: limitin_def) |
69874 | 163 |
qed |
164 |
||
165 |
||
166 |
subsection\<open>Pointwise continuity in topological spaces\<close> |
|
167 |
||
168 |
definition topcontinuous_at where |
|
169 |
"topcontinuous_at X Y f x \<longleftrightarrow> |
|
170 |
x \<in> topspace X \<and> |
|
171 |
(\<forall>x \<in> topspace X. f x \<in> topspace Y) \<and> |
|
172 |
(\<forall>V. openin Y V \<and> f x \<in> V |
|
173 |
\<longrightarrow> (\<exists>U. openin X U \<and> x \<in> U \<and> (\<forall>y \<in> U. f y \<in> V)))" |
|
174 |
||
175 |
lemma topcontinuous_at_atin: |
|
176 |
"topcontinuous_at X Y f x \<longleftrightarrow> |
|
177 |
x \<in> topspace X \<and> |
|
178 |
(\<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
|
179 |
limitin Y f (f x) (atin X x)" |
69874 | 180 |
unfolding topcontinuous_at_def |
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
181 |
by (fastforce simp add: limitin_atin)+ |
69874 | 182 |
|
183 |
lemma continuous_map_eq_topcontinuous_at: |
|
184 |
"continuous_map X Y f \<longleftrightarrow> (\<forall>x \<in> topspace X. topcontinuous_at X Y f x)" |
|
185 |
(is "?lhs = ?rhs") |
|
186 |
proof |
|
187 |
assume ?lhs |
|
188 |
then show ?rhs |
|
189 |
by (auto simp: continuous_map_def topcontinuous_at_def) |
|
190 |
next |
|
191 |
assume R: ?rhs |
|
192 |
then show ?lhs |
|
193 |
apply (auto simp: continuous_map_def topcontinuous_at_def) |
|
194 |
apply (subst openin_subopen, safe) |
|
195 |
apply (drule bspec, assumption) |
|
196 |
using openin_subset[of X] apply (auto simp: subset_iff dest!: spec) |
|
197 |
done |
|
198 |
qed |
|
199 |
||
200 |
lemma continuous_map_atin: |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
201 |
"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
|
202 |
by (auto simp: limitin_def topcontinuous_at_atin continuous_map_eq_topcontinuous_at) |
69874 | 203 |
|
69875
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
204 |
lemma limitin_continuous_map: |
03bc14eab432
renamed the constant "limit" as it is too "generic"
paulson <lp15@cam.ac.uk>
parents:
69874
diff
changeset
|
205 |
"\<lbrakk>continuous_map X Y f; a \<in> topspace X; f a = b\<rbrakk> \<Longrightarrow> limitin Y f b (atin X a)" |
69874 | 206 |
by (auto simp: continuous_map_atin) |
207 |
||
208 |
||
209 |
subsection\<open>Combining theorems for continuous functions into the reals\<close> |
|
210 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
211 |
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
|
212 |
"continuous_map X euclidean (\<lambda>x. c)" |
69874 | 213 |
by simp |
214 |
||
215 |
lemma continuous_map_real_mult [continuous_intros]: |
|
216 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> |
|
217 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * g x)" |
|
218 |
by (simp add: continuous_map_atin tendsto_mult) |
|
219 |
||
220 |
lemma continuous_map_real_pow [continuous_intros]: |
|
221 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x ^ n)" |
|
222 |
by (induction n) (auto simp: continuous_map_real_mult) |
|
223 |
||
224 |
lemma continuous_map_real_mult_left: |
|
225 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. c * f x)" |
|
226 |
by (simp add: continuous_map_atin tendsto_mult) |
|
227 |
||
228 |
lemma continuous_map_real_mult_left_eq: |
|
229 |
"continuous_map X euclideanreal (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f" |
|
230 |
proof (cases "c = 0") |
|
231 |
case False |
|
232 |
have "continuous_map X euclideanreal (\<lambda>x. c * f x) \<Longrightarrow> continuous_map X euclideanreal f" |
|
233 |
apply (frule continuous_map_real_mult_left [where c="inverse c"]) |
|
234 |
apply (simp add: field_simps False) |
|
235 |
done |
|
236 |
with False show ?thesis |
|
237 |
using continuous_map_real_mult_left by blast |
|
238 |
qed simp |
|
239 |
||
240 |
lemma continuous_map_real_mult_right: |
|
241 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x * c)" |
|
242 |
by (simp add: continuous_map_atin tendsto_mult) |
|
243 |
||
244 |
lemma continuous_map_real_mult_right_eq: |
|
245 |
"continuous_map X euclideanreal (\<lambda>x. f x * c) \<longleftrightarrow> c = 0 \<or> continuous_map X euclideanreal f" |
|
246 |
by (simp add: mult.commute flip: continuous_map_real_mult_left_eq) |
|
247 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
248 |
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
|
249 |
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
|
250 |
shows "continuous_map X euclidean f \<Longrightarrow> continuous_map X euclidean (\<lambda>x. - f x)" |
69874 | 251 |
by (simp add: continuous_map_atin tendsto_minus) |
252 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
253 |
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
|
254 |
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
|
255 |
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
|
256 |
using continuous_map_minus add.inverse_inverse continuous_map_eq by fastforce |
69874 | 257 |
|
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
258 |
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
|
259 |
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
|
260 |
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 | 261 |
by (simp add: continuous_map_atin tendsto_add) |
262 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
263 |
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
|
264 |
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
|
265 |
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 | 266 |
by (simp add: continuous_map_atin tendsto_diff) |
267 |
||
70019
095dce9892e8
A few results in Algebra, and bits for Analysis
paulson <lp15@cam.ac.uk>
parents:
69875
diff
changeset
|
268 |
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
|
269 |
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
|
270 |
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
|
271 |
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
|
272 |
|
69874 | 273 |
lemma continuous_map_real_abs [continuous_intros]: |
274 |
"continuous_map X euclideanreal f \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. abs(f x))" |
|
275 |
by (simp add: continuous_map_atin tendsto_rabs) |
|
276 |
||
277 |
lemma continuous_map_real_max [continuous_intros]: |
|
278 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> |
|
279 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. max (f x) (g x))" |
|
280 |
by (simp add: continuous_map_atin tendsto_max) |
|
281 |
||
282 |
lemma continuous_map_real_min [continuous_intros]: |
|
283 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g\<rbrakk> |
|
284 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. min (f x) (g x))" |
|
285 |
by (simp add: continuous_map_atin tendsto_min) |
|
286 |
||
287 |
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
|
288 |
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
|
289 |
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
|
290 |
\<Longrightarrow> continuous_map X euclidean (\<lambda>x. sum (f x) I)" |
69874 | 291 |
by (simp add: continuous_map_atin tendsto_sum) |
292 |
||
293 |
lemma continuous_map_prod [continuous_intros]: |
|
294 |
"\<lbrakk>finite I; |
|
295 |
\<And>i. i \<in> I \<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x i)\<rbrakk> |
|
296 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. prod (f x) I)" |
|
297 |
by (simp add: continuous_map_atin tendsto_prod) |
|
298 |
||
299 |
lemma continuous_map_real_inverse [continuous_intros]: |
|
300 |
"\<lbrakk>continuous_map X euclideanreal f; \<And>x. x \<in> topspace X \<Longrightarrow> f x \<noteq> 0\<rbrakk> |
|
301 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. inverse(f x))" |
|
302 |
by (simp add: continuous_map_atin tendsto_inverse) |
|
303 |
||
304 |
lemma continuous_map_real_divide [continuous_intros]: |
|
305 |
"\<lbrakk>continuous_map X euclideanreal f; continuous_map X euclideanreal g; \<And>x. x \<in> topspace X \<Longrightarrow> g x \<noteq> 0\<rbrakk> |
|
306 |
\<Longrightarrow> continuous_map X euclideanreal (\<lambda>x. f x / g x)" |
|
307 |
by (simp add: continuous_map_atin tendsto_divide) |
|
308 |
||
309 |
end |