| author | paulson <lp15@cam.ac.uk> |
| Tue, 13 Oct 2015 12:42:08 +0100 | |
| changeset 61426 | d53db136e8fd |
| parent 61260 | e6f03fae14d5 |
| child 61808 | fc1556774cfe |
| permissions | -rw-r--r-- |
| 60420 | 1 |
section \<open>Bounded Continuous Functions\<close> |
| 60421 | 2 |
|
| 59453 | 3 |
theory Bounded_Continuous_Function |
4 |
imports Integration |
|
5 |
begin |
|
6 |
||
| 60421 | 7 |
subsection \<open>Definition\<close> |
| 59453 | 8 |
|
| 60421 | 9 |
definition bcontfun :: "('a::topological_space \<Rightarrow> 'b::metric_space) set"
|
10 |
where "bcontfun = {f. continuous_on UNIV f \<and> bounded (range f)}"
|
|
| 59453 | 11 |
|
| 61260 | 12 |
typedef (overloaded) ('a, 'b) bcontfun =
|
13 |
"bcontfun :: ('a::topological_space \<Rightarrow> 'b::metric_space) set"
|
|
| 59453 | 14 |
by (auto simp: bcontfun_def intro: continuous_intros simp: bounded_def) |
15 |
||
16 |
lemma bcontfunE: |
|
17 |
assumes "f \<in> bcontfun" |
|
18 |
obtains y where "continuous_on UNIV f" "\<And>x. dist (f x) u \<le> y" |
|
19 |
using assms unfolding bcontfun_def |
|
20 |
by (metis (lifting) bounded_any_center dist_commute mem_Collect_eq rangeI) |
|
21 |
||
22 |
lemma bcontfunE': |
|
23 |
assumes "f \<in> bcontfun" |
|
24 |
obtains y where "continuous_on UNIV f" "\<And>x. dist (f x) undefined \<le> y" |
|
25 |
using assms bcontfunE |
|
26 |
by metis |
|
27 |
||
| 60421 | 28 |
lemma bcontfunI: "continuous_on UNIV f \<Longrightarrow> (\<And>x. dist (f x) u \<le> b) \<Longrightarrow> f \<in> bcontfun" |
| 59453 | 29 |
unfolding bcontfun_def |
30 |
by (metis (lifting, no_types) bounded_def dist_commute mem_Collect_eq rangeE) |
|
31 |
||
| 60421 | 32 |
lemma bcontfunI': "continuous_on UNIV f \<Longrightarrow> (\<And>x. dist (f x) undefined \<le> b) \<Longrightarrow> f \<in> bcontfun" |
| 59453 | 33 |
using bcontfunI by metis |
34 |
||
35 |
lemma continuous_on_Rep_bcontfun[intro, simp]: "continuous_on T (Rep_bcontfun x)" |
|
36 |
using Rep_bcontfun[of x] |
|
37 |
by (auto simp: bcontfun_def intro: continuous_on_subset) |
|
38 |
||
39 |
instantiation bcontfun :: (topological_space, metric_space) metric_space |
|
40 |
begin |
|
41 |
||
| 60421 | 42 |
definition dist_bcontfun :: "('a, 'b) bcontfun \<Rightarrow> ('a, 'b) bcontfun \<Rightarrow> real"
|
43 |
where "dist_bcontfun f g = (SUP x. dist (Rep_bcontfun f x) (Rep_bcontfun g x))" |
|
| 59453 | 44 |
|
| 60421 | 45 |
definition open_bcontfun :: "('a, 'b) bcontfun set \<Rightarrow> bool"
|
46 |
where "open_bcontfun S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" |
|
| 59453 | 47 |
|
48 |
lemma dist_bounded: |
|
| 60421 | 49 |
fixes f :: "('a, 'b) bcontfun"
|
| 59453 | 50 |
shows "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f g" |
51 |
proof - |
|
| 60421 | 52 |
have "Rep_bcontfun f \<in> bcontfun" by (rule Rep_bcontfun) |
| 59453 | 53 |
from bcontfunE'[OF this] obtain y where y: |
54 |
"continuous_on UNIV (Rep_bcontfun f)" |
|
55 |
"\<And>x. dist (Rep_bcontfun f x) undefined \<le> y" |
|
56 |
by auto |
|
| 60421 | 57 |
have "Rep_bcontfun g \<in> bcontfun" by (rule Rep_bcontfun) |
| 59453 | 58 |
from bcontfunE'[OF this] obtain z where z: |
59 |
"continuous_on UNIV (Rep_bcontfun g)" |
|
60 |
"\<And>x. dist (Rep_bcontfun g x) undefined \<le> z" |
|
61 |
by auto |
|
| 60421 | 62 |
show ?thesis |
63 |
unfolding dist_bcontfun_def |
|
| 59453 | 64 |
proof (intro cSUP_upper bdd_aboveI2) |
65 |
fix x |
|
| 60421 | 66 |
have "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> |
67 |
dist (Rep_bcontfun f x) undefined + dist (Rep_bcontfun g x) undefined" |
|
| 59453 | 68 |
by (rule dist_triangle2) |
| 60421 | 69 |
also have "\<dots> \<le> y + z" |
70 |
using y(2)[of x] z(2)[of x] by (rule add_mono) |
|
| 59453 | 71 |
finally show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> y + z" . |
72 |
qed simp |
|
73 |
qed |
|
74 |
||
75 |
lemma dist_bound: |
|
| 60421 | 76 |
fixes f :: "('a, 'b) bcontfun"
|
| 59453 | 77 |
assumes "\<And>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> b" |
78 |
shows "dist f g \<le> b" |
|
79 |
using assms by (auto simp: dist_bcontfun_def intro: cSUP_least) |
|
80 |
||
81 |
lemma dist_bounded_Abs: |
|
| 60421 | 82 |
fixes f g :: "'a \<Rightarrow> 'b" |
| 59453 | 83 |
assumes "f \<in> bcontfun" "g \<in> bcontfun" |
84 |
shows "dist (f x) (g x) \<le> dist (Abs_bcontfun f) (Abs_bcontfun g)" |
|
85 |
by (metis Abs_bcontfun_inverse assms dist_bounded) |
|
86 |
||
87 |
lemma const_bcontfun: "(\<lambda>x::'a. b::'b) \<in> bcontfun" |
|
88 |
by (auto intro: bcontfunI continuous_on_const) |
|
89 |
||
90 |
lemma dist_fun_lt_imp_dist_val_lt: |
|
91 |
assumes "dist f g < e" |
|
92 |
shows "dist (Rep_bcontfun f x) (Rep_bcontfun g x) < e" |
|
93 |
using dist_bounded assms by (rule le_less_trans) |
|
94 |
||
95 |
lemma dist_val_lt_imp_dist_fun_le: |
|
96 |
assumes "\<forall>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) < e" |
|
97 |
shows "dist f g \<le> e" |
|
| 60421 | 98 |
unfolding dist_bcontfun_def |
| 59453 | 99 |
proof (intro cSUP_least) |
100 |
fix x |
|
101 |
show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> e" |
|
102 |
using assms[THEN spec[where x=x]] by (simp add: dist_norm) |
|
| 60421 | 103 |
qed simp |
| 59453 | 104 |
|
105 |
instance |
|
106 |
proof |
|
| 60421 | 107 |
fix f g h :: "('a, 'b) bcontfun"
|
| 59453 | 108 |
show "dist f g = 0 \<longleftrightarrow> f = g" |
109 |
proof |
|
| 60421 | 110 |
have "\<And>x. dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f g" |
111 |
by (rule dist_bounded) |
|
| 59453 | 112 |
also assume "dist f g = 0" |
| 60421 | 113 |
finally show "f = g" |
114 |
by (auto simp: Rep_bcontfun_inject[symmetric] Abs_bcontfun_inverse) |
|
| 59453 | 115 |
qed (auto simp: dist_bcontfun_def SUP_def simp del: Sup_image_eq intro!: cSup_eq) |
116 |
show "dist f g \<le> dist f h + dist g h" |
|
117 |
proof (subst dist_bcontfun_def, safe intro!: cSUP_least) |
|
118 |
fix x |
|
119 |
have "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> |
|
120 |
dist (Rep_bcontfun f x) (Rep_bcontfun h x) + dist (Rep_bcontfun g x) (Rep_bcontfun h x)" |
|
121 |
by (rule dist_triangle2) |
|
| 60421 | 122 |
also have "dist (Rep_bcontfun f x) (Rep_bcontfun h x) \<le> dist f h" |
123 |
by (rule dist_bounded) |
|
124 |
also have "dist (Rep_bcontfun g x) (Rep_bcontfun h x) \<le> dist g h" |
|
125 |
by (rule dist_bounded) |
|
126 |
finally show "dist (Rep_bcontfun f x) (Rep_bcontfun g x) \<le> dist f h + dist g h" |
|
127 |
by simp |
|
| 59453 | 128 |
qed |
129 |
qed (simp add: open_bcontfun_def) |
|
| 60421 | 130 |
|
| 59453 | 131 |
end |
132 |
||
133 |
lemma closed_Pi_bcontfun: |
|
| 60421 | 134 |
fixes I :: "'a::metric_space set" |
135 |
and X :: "'a \<Rightarrow> 'b::complete_space set" |
|
| 59453 | 136 |
assumes "\<And>i. i \<in> I \<Longrightarrow> closed (X i)" |
137 |
shows "closed (Abs_bcontfun ` (Pi I X \<inter> bcontfun))" |
|
138 |
unfolding closed_sequential_limits |
|
139 |
proof safe |
|
140 |
fix f l |
|
141 |
assume seq: "\<forall>n. f n \<in> Abs_bcontfun ` (Pi I X \<inter> bcontfun)" and lim: "f ----> l" |
|
142 |
have lim_fun: "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (Rep_bcontfun l x) < e" |
|
143 |
using LIMSEQ_imp_Cauchy[OF lim, simplified Cauchy_def] metric_LIMSEQ_D[OF lim] |
|
| 60421 | 144 |
by (intro uniformly_cauchy_imp_uniformly_convergent[where P="\<lambda>_. True", simplified]) |
| 59453 | 145 |
(metis dist_fun_lt_imp_dist_val_lt)+ |
146 |
show "l \<in> Abs_bcontfun ` (Pi I X \<inter> bcontfun)" |
|
147 |
proof (rule, safe) |
|
148 |
fix x assume "x \<in> I" |
|
| 60421 | 149 |
then have "closed (X x)" |
150 |
using assms by simp |
|
| 59453 | 151 |
moreover have "eventually (\<lambda>xa. Rep_bcontfun (f xa) x \<in> X x) sequentially" |
152 |
proof (rule always_eventually, safe) |
|
153 |
fix i |
|
| 60420 | 154 |
from seq[THEN spec, of i] \<open>x \<in> I\<close> |
| 59453 | 155 |
show "Rep_bcontfun (f i) x \<in> X x" |
156 |
by (auto simp: Abs_bcontfun_inverse) |
|
157 |
qed |
|
158 |
moreover note sequentially_bot |
|
159 |
moreover have "(\<lambda>n. Rep_bcontfun (f n) x) ----> Rep_bcontfun l x" |
|
160 |
using lim_fun by (blast intro!: metric_LIMSEQ_I) |
|
161 |
ultimately show "Rep_bcontfun l x \<in> X x" |
|
162 |
by (rule Lim_in_closed_set) |
|
163 |
qed (auto simp: Rep_bcontfun Rep_bcontfun_inverse) |
|
164 |
qed |
|
165 |
||
| 60421 | 166 |
|
| 60420 | 167 |
subsection \<open>Complete Space\<close> |
| 59453 | 168 |
|
169 |
instance bcontfun :: (metric_space, complete_space) complete_space |
|
170 |
proof |
|
| 60421 | 171 |
fix f :: "nat \<Rightarrow> ('a, 'b) bcontfun"
|
172 |
assume "Cauchy f" -- \<open>Cauchy equals uniform convergence\<close> |
|
| 59453 | 173 |
then obtain g where limit_function: |
174 |
"\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e" |
|
175 |
using uniformly_convergent_eq_cauchy[of "\<lambda>_. True" |
|
176 |
"\<lambda>n. Rep_bcontfun (f n)"] |
|
| 60421 | 177 |
unfolding Cauchy_def |
178 |
by (metis dist_fun_lt_imp_dist_val_lt) |
|
| 59453 | 179 |
|
| 60421 | 180 |
then obtain N where fg_dist: -- \<open>for an upper bound on @{term g}\<close>
|
| 59453 | 181 |
"\<forall>n\<ge>N. \<forall>x. dist (g x) ( Rep_bcontfun (f n) x) < 1" |
182 |
by (force simp add: dist_commute) |
|
183 |
from bcontfunE'[OF Rep_bcontfun, of "f N"] obtain b where |
|
| 60421 | 184 |
f_bound: "\<forall>x. dist (Rep_bcontfun (f N) x) undefined \<le> b" |
185 |
by force |
|
186 |
have "g \<in> bcontfun" -- \<open>The limit function is bounded and continuous\<close> |
|
| 59453 | 187 |
proof (intro bcontfunI) |
188 |
show "continuous_on UNIV g" |
|
189 |
using bcontfunE[OF Rep_bcontfun] limit_function |
|
| 60421 | 190 |
by (intro continuous_uniform_limit[where f="\<lambda>n. Rep_bcontfun (f n)" and F="sequentially"]) |
191 |
(auto simp add: eventually_sequentially trivial_limit_def dist_norm) |
|
| 59453 | 192 |
next |
193 |
fix x |
|
194 |
from fg_dist have "dist (g x) (Rep_bcontfun (f N) x) < 1" |
|
195 |
by (simp add: dist_norm norm_minus_commute) |
|
196 |
with dist_triangle[of "g x" undefined "Rep_bcontfun (f N) x"] |
|
197 |
show "dist (g x) undefined \<le> 1 + b" using f_bound[THEN spec, of x] |
|
198 |
by simp |
|
199 |
qed |
|
200 |
show "convergent f" |
|
|
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:
59865
diff
changeset
|
201 |
proof (rule convergentI, subst lim_sequentially, safe) |
| 60421 | 202 |
-- \<open>The limit function converges according to its norm\<close> |
203 |
fix e :: real |
|
204 |
assume "e > 0" |
|
205 |
then have "e/2 > 0" by simp |
|
| 59453 | 206 |
with limit_function[THEN spec, of"e/2"] |
207 |
have "\<exists>N. \<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e/2" |
|
208 |
by simp |
|
209 |
then obtain N where N: "\<forall>n\<ge>N. \<forall>x. dist (Rep_bcontfun (f n) x) (g x) < e / 2" by auto |
|
210 |
show "\<exists>N. \<forall>n\<ge>N. dist (f n) (Abs_bcontfun g) < e" |
|
211 |
proof (rule, safe) |
|
212 |
fix n |
|
213 |
assume "N \<le> n" |
|
214 |
with N show "dist (f n) (Abs_bcontfun g) < e" |
|
215 |
using dist_val_lt_imp_dist_fun_le[of |
|
216 |
"f n" "Abs_bcontfun g" "e/2"] |
|
| 60420 | 217 |
Abs_bcontfun_inverse[OF \<open>g \<in> bcontfun\<close>] \<open>e > 0\<close> by simp |
| 59453 | 218 |
qed |
219 |
qed |
|
220 |
qed |
|
221 |
||
| 60421 | 222 |
|
223 |
subsection \<open>Supremum norm for a normed vector space\<close> |
|
| 59453 | 224 |
|
225 |
instantiation bcontfun :: (topological_space, real_normed_vector) real_vector |
|
226 |
begin |
|
227 |
||
228 |
definition "-f = Abs_bcontfun (\<lambda>x. -(Rep_bcontfun f x))" |
|
229 |
||
230 |
definition "f + g = Abs_bcontfun (\<lambda>x. Rep_bcontfun f x + Rep_bcontfun g x)" |
|
231 |
||
232 |
definition "f - g = Abs_bcontfun (\<lambda>x. Rep_bcontfun f x - Rep_bcontfun g x)" |
|
233 |
||
234 |
definition "0 = Abs_bcontfun (\<lambda>x. 0)" |
|
235 |
||
236 |
definition "scaleR r f = Abs_bcontfun (\<lambda>x. r *\<^sub>R Rep_bcontfun f x)" |
|
237 |
||
238 |
lemma plus_cont: |
|
| 60421 | 239 |
fixes f g :: "'a \<Rightarrow> 'b" |
240 |
assumes f: "f \<in> bcontfun" |
|
241 |
and g: "g \<in> bcontfun" |
|
| 59453 | 242 |
shows "(\<lambda>x. f x + g x) \<in> bcontfun" |
243 |
proof - |
|
244 |
from bcontfunE'[OF f] obtain y where "continuous_on UNIV f" "\<And>x. dist (f x) undefined \<le> y" |
|
245 |
by auto |
|
246 |
moreover |
|
247 |
from bcontfunE'[OF g] obtain z where "continuous_on UNIV g" "\<And>x. dist (g x) undefined \<le> z" |
|
248 |
by auto |
|
249 |
ultimately show ?thesis |
|
250 |
proof (intro bcontfunI) |
|
251 |
fix x |
|
| 60421 | 252 |
have "dist (f x + g x) 0 = dist (f x + g x) (0 + 0)" |
253 |
by simp |
|
254 |
also have "\<dots> \<le> dist (f x) 0 + dist (g x) 0" |
|
255 |
by (rule dist_triangle_add) |
|
| 59453 | 256 |
also have "\<dots> \<le> dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" |
257 |
unfolding zero_bcontfun_def using assms |
|
258 |
by (auto intro!: add_mono dist_bounded_Abs const_bcontfun) |
|
| 60421 | 259 |
finally show "dist (f x + g x) 0 \<le> dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" . |
| 59453 | 260 |
qed (simp add: continuous_on_add) |
261 |
qed |
|
262 |
||
263 |
lemma Rep_bcontfun_plus[simp]: "Rep_bcontfun (f + g) x = Rep_bcontfun f x + Rep_bcontfun g x" |
|
264 |
by (simp add: plus_bcontfun_def Abs_bcontfun_inverse plus_cont Rep_bcontfun) |
|
265 |
||
266 |
lemma uminus_cont: |
|
| 60421 | 267 |
fixes f :: "'a \<Rightarrow> 'b" |
| 59453 | 268 |
assumes "f \<in> bcontfun" |
269 |
shows "(\<lambda>x. - f x) \<in> bcontfun" |
|
270 |
proof - |
|
| 60421 | 271 |
from bcontfunE[OF assms, of 0] obtain y |
272 |
where "continuous_on UNIV f" "\<And>x. dist (f x) 0 \<le> y" |
|
| 59453 | 273 |
by auto |
| 60421 | 274 |
then show ?thesis |
| 59453 | 275 |
proof (intro bcontfunI) |
276 |
fix x |
|
277 |
assume "\<And>x. dist (f x) 0 \<le> y" |
|
| 60421 | 278 |
then show "dist (- f x) 0 \<le> y" |
279 |
by (subst dist_minus[symmetric]) simp |
|
| 59453 | 280 |
qed (simp add: continuous_on_minus) |
281 |
qed |
|
282 |
||
| 60421 | 283 |
lemma Rep_bcontfun_uminus[simp]: "Rep_bcontfun (- f) x = - Rep_bcontfun f x" |
| 59453 | 284 |
by (simp add: uminus_bcontfun_def Abs_bcontfun_inverse uminus_cont Rep_bcontfun) |
285 |
||
286 |
lemma minus_cont: |
|
| 60421 | 287 |
fixes f g :: "'a \<Rightarrow> 'b" |
288 |
assumes f: "f \<in> bcontfun" |
|
289 |
and g: "g \<in> bcontfun" |
|
| 59453 | 290 |
shows "(\<lambda>x. f x - g x) \<in> bcontfun" |
| 60421 | 291 |
using plus_cont [of f "- g"] assms |
292 |
by (simp add: uminus_cont fun_Compl_def) |
|
| 59453 | 293 |
|
| 60421 | 294 |
lemma Rep_bcontfun_minus[simp]: "Rep_bcontfun (f - g) x = Rep_bcontfun f x - Rep_bcontfun g x" |
| 59453 | 295 |
by (simp add: minus_bcontfun_def Abs_bcontfun_inverse minus_cont Rep_bcontfun) |
296 |
||
297 |
lemma scaleR_cont: |
|
| 60421 | 298 |
fixes a :: real |
299 |
and f :: "'a \<Rightarrow> 'b" |
|
| 59453 | 300 |
assumes "f \<in> bcontfun" |
301 |
shows " (\<lambda>x. a *\<^sub>R f x) \<in> bcontfun" |
|
302 |
proof - |
|
| 60421 | 303 |
from bcontfunE[OF assms, of 0] obtain y |
304 |
where "continuous_on UNIV f" "\<And>x. dist (f x) 0 \<le> y" |
|
| 59453 | 305 |
by auto |
| 60421 | 306 |
then show ?thesis |
| 59453 | 307 |
proof (intro bcontfunI) |
| 60421 | 308 |
fix x |
309 |
assume "\<And>x. dist (f x) 0 \<le> y" |
|
|
59554
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
59453
diff
changeset
|
310 |
then show "dist (a *\<^sub>R f x) 0 \<le> abs a * y" |
|
4044f53326c9
inlined rules to free user-space from technical names
haftmann
parents:
59453
diff
changeset
|
311 |
by (metis norm_cmul_rule_thm norm_conv_dist) |
| 59453 | 312 |
qed (simp add: continuous_intros) |
313 |
qed |
|
314 |
||
| 60421 | 315 |
lemma Rep_bcontfun_scaleR[simp]: "Rep_bcontfun (a *\<^sub>R g) x = a *\<^sub>R Rep_bcontfun g x" |
| 59453 | 316 |
by (simp add: scaleR_bcontfun_def Abs_bcontfun_inverse scaleR_cont Rep_bcontfun) |
317 |
||
318 |
instance |
|
| 61169 | 319 |
by standard |
| 60421 | 320 |
(simp_all add: plus_bcontfun_def zero_bcontfun_def minus_bcontfun_def scaleR_bcontfun_def |
321 |
Abs_bcontfun_inverse Rep_bcontfun_inverse Rep_bcontfun algebra_simps |
|
322 |
plus_cont const_bcontfun minus_cont scaleR_cont) |
|
323 |
||
| 59453 | 324 |
end |
325 |
||
326 |
instantiation bcontfun :: (topological_space, real_normed_vector) real_normed_vector |
|
327 |
begin |
|
328 |
||
| 60421 | 329 |
definition norm_bcontfun :: "('a, 'b) bcontfun \<Rightarrow> real"
|
330 |
where "norm_bcontfun f = dist f 0" |
|
| 59453 | 331 |
|
332 |
definition "sgn (f::('a,'b) bcontfun) = f /\<^sub>R norm f"
|
|
333 |
||
334 |
instance |
|
335 |
proof |
|
| 60421 | 336 |
fix a :: real |
337 |
fix f g :: "('a, 'b) bcontfun"
|
|
| 59453 | 338 |
show "dist f g = norm (f - g)" |
339 |
by (simp add: norm_bcontfun_def dist_bcontfun_def zero_bcontfun_def |
|
| 60421 | 340 |
Abs_bcontfun_inverse const_bcontfun norm_conv_dist[symmetric] dist_norm) |
| 59453 | 341 |
show "norm (f + g) \<le> norm f + norm g" |
342 |
unfolding norm_bcontfun_def |
|
343 |
proof (subst dist_bcontfun_def, safe intro!: cSUP_least) |
|
344 |
fix x |
|
345 |
have "dist (Rep_bcontfun (f + g) x) (Rep_bcontfun 0 x) \<le> |
|
346 |
dist (Rep_bcontfun f x) 0 + dist (Rep_bcontfun g x) 0" |
|
347 |
by (metis (hide_lams, no_types) Rep_bcontfun_minus Rep_bcontfun_plus diff_0_right dist_norm |
|
348 |
le_less_linear less_irrefl norm_triangle_lt) |
|
349 |
also have "dist (Rep_bcontfun f x) 0 \<le> dist f 0" |
|
350 |
using dist_bounded[of f x 0] |
|
351 |
by (simp add: Abs_bcontfun_inverse const_bcontfun zero_bcontfun_def) |
|
352 |
also have "dist (Rep_bcontfun g x) 0 \<le> dist g 0" using dist_bounded[of g x 0] |
|
353 |
by (simp add: Abs_bcontfun_inverse const_bcontfun zero_bcontfun_def) |
|
354 |
finally show "dist (Rep_bcontfun (f + g) x) (Rep_bcontfun 0 x) \<le> dist f 0 + dist g 0" by simp |
|
355 |
qed |
|
356 |
show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f" |
|
357 |
proof - |
|
358 |
have "\<bar>a\<bar> * Sup (range (\<lambda>x. dist (Rep_bcontfun f x) 0)) = |
|
359 |
(SUP i:range (\<lambda>x. dist (Rep_bcontfun f x) 0). \<bar>a\<bar> * i)" |
|
360 |
proof (intro continuous_at_Sup_mono bdd_aboveI2) |
|
361 |
fix x |
|
362 |
show "dist (Rep_bcontfun f x) 0 \<le> norm f" using dist_bounded[of f x 0] |
|
363 |
by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def |
|
364 |
const_bcontfun) |
|
365 |
qed (auto intro!: monoI mult_left_mono continuous_intros) |
|
366 |
moreover |
|
| 60421 | 367 |
have "range (\<lambda>x. dist (Rep_bcontfun (a *\<^sub>R f) x) 0) = |
| 59453 | 368 |
(\<lambda>x. \<bar>a\<bar> * x) ` (range (\<lambda>x. dist (Rep_bcontfun f x) 0))" |
369 |
by (auto simp: norm_conv_dist[symmetric]) |
|
370 |
ultimately |
|
371 |
show "norm (a *\<^sub>R f) = \<bar>a\<bar> * norm f" |
|
372 |
by (simp add: norm_bcontfun_def dist_bcontfun_def norm_conv_dist Abs_bcontfun_inverse |
|
373 |
zero_bcontfun_def const_bcontfun SUP_def del: Sup_image_eq) |
|
374 |
qed |
|
375 |
qed (auto simp: norm_bcontfun_def sgn_bcontfun_def) |
|
376 |
||
377 |
end |
|
378 |
||
| 60421 | 379 |
lemma bcontfun_normI: "continuous_on UNIV f \<Longrightarrow> (\<And>x. norm (f x) \<le> b) \<Longrightarrow> f \<in> bcontfun" |
| 59453 | 380 |
unfolding norm_conv_dist |
381 |
by (auto intro: bcontfunI) |
|
382 |
||
383 |
lemma norm_bounded: |
|
| 60421 | 384 |
fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun"
|
| 59453 | 385 |
shows "norm (Rep_bcontfun f x) \<le> norm f" |
386 |
using dist_bounded[of f x 0] |
|
387 |
by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def |
|
388 |
const_bcontfun) |
|
389 |
||
390 |
lemma norm_bound: |
|
| 60421 | 391 |
fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun"
|
| 59453 | 392 |
assumes "\<And>x. norm (Rep_bcontfun f x) \<le> b" |
393 |
shows "norm f \<le> b" |
|
394 |
using dist_bound[of f 0 b] assms |
|
395 |
by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def |
|
396 |
const_bcontfun) |
|
397 |
||
398 |
||
| 60421 | 399 |
subsection \<open>Continuously Extended Functions\<close> |
400 |
||
401 |
definition clamp :: "'a::euclidean_space \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a" where |
|
| 59453 | 402 |
"clamp a b x = (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i)" |
403 |
||
| 60421 | 404 |
definition ext_cont :: "('a::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> ('a, 'b) bcontfun"
|
| 59453 | 405 |
where "ext_cont f a b = Abs_bcontfun ((\<lambda>x. f (clamp a b x)))" |
406 |
||
407 |
lemma ext_cont_def': |
|
408 |
"ext_cont f a b = Abs_bcontfun (\<lambda>x. |
|
409 |
f (\<Sum>i\<in>Basis. (if x\<bullet>i < a\<bullet>i then a\<bullet>i else if x\<bullet>i \<le> b\<bullet>i then x\<bullet>i else b\<bullet>i) *\<^sub>R i))" |
|
| 60421 | 410 |
unfolding ext_cont_def clamp_def .. |
| 59453 | 411 |
|
412 |
lemma clamp_in_interval: |
|
413 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" |
|
414 |
shows "clamp a b x \<in> cbox a b" |
|
415 |
unfolding clamp_def |
|
416 |
using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) |
|
417 |
||
418 |
lemma dist_clamps_le_dist_args: |
|
| 60421 | 419 |
fixes x :: "'a::euclidean_space" |
| 59453 | 420 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" |
421 |
shows "dist (clamp a b y) (clamp a b x) \<le> dist y x" |
|
422 |
proof - |
|
| 60421 | 423 |
from box_ne_empty(1)[of a b] assms have "(\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" |
424 |
by (simp add: cbox_def) |
|
425 |
then have "(\<Sum>i\<in>Basis. (dist (clamp a b y \<bullet> i) (clamp a b x \<bullet> i))\<^sup>2) \<le> |
|
426 |
(\<Sum>i\<in>Basis. (dist (y \<bullet> i) (x \<bullet> i))\<^sup>2)" |
|
427 |
by (auto intro!: setsum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) |
|
428 |
then show ?thesis |
|
429 |
by (auto intro: real_sqrt_le_mono |
|
430 |
simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] setL2_def) |
|
| 59453 | 431 |
qed |
432 |
||
433 |
lemma clamp_continuous_at: |
|
| 60421 | 434 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" |
435 |
and x :: 'a |
|
| 59453 | 436 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" |
| 60421 | 437 |
and f_cont: "continuous_on (cbox a b) f" |
| 59453 | 438 |
shows "continuous (at x) (\<lambda>x. f (clamp a b x))" |
| 60421 | 439 |
unfolding continuous_at_eps_delta |
440 |
proof safe |
|
441 |
fix x :: 'a |
|
442 |
fix e :: real |
|
443 |
assume "e > 0" |
|
444 |
moreover have "clamp a b x \<in> cbox a b" |
|
445 |
by (simp add: clamp_in_interval assms) |
|
446 |
moreover note f_cont[simplified continuous_on_iff] |
|
| 59453 | 447 |
ultimately |
448 |
obtain d where d: "0 < d" |
|
449 |
"\<And>x'. x' \<in> cbox a b \<Longrightarrow> dist x' (clamp a b x) < d \<Longrightarrow> dist (f x') (f (clamp a b x)) < e" |
|
450 |
by force |
|
451 |
show "\<exists>d>0. \<forall>x'. dist x' x < d \<longrightarrow> |
|
452 |
dist (f (clamp a b x')) (f (clamp a b x)) < e" |
|
453 |
by (auto intro!: d clamp_in_interval assms dist_clamps_le_dist_args[THEN le_less_trans]) |
|
454 |
qed |
|
455 |
||
456 |
lemma clamp_continuous_on: |
|
| 60421 | 457 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::metric_space" |
| 59453 | 458 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" |
| 60421 | 459 |
and f_cont: "continuous_on (cbox a b) f" |
| 59453 | 460 |
shows "continuous_on UNIV (\<lambda>x. f (clamp a b x))" |
461 |
using assms |
|
462 |
by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) |
|
463 |
||
464 |
lemma clamp_bcontfun: |
|
| 60421 | 465 |
fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" |
| 59453 | 466 |
assumes "\<And>i. i \<in> Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i" |
| 60421 | 467 |
and continuous: "continuous_on (cbox a b) f" |
| 59453 | 468 |
shows "(\<lambda>x. f (clamp a b x)) \<in> bcontfun" |
469 |
proof - |
|
| 60421 | 470 |
have "bounded (f ` (cbox a b))" |
471 |
by (rule compact_continuous_image[OF continuous compact_cbox[of a b], THEN compact_imp_bounded]) |
|
472 |
then obtain c where f_bound: "\<forall>x\<in>f ` cbox a b. norm x \<le> c" |
|
473 |
by (auto simp add: bounded_pos) |
|
| 59453 | 474 |
show "(\<lambda>x. f (clamp a b x)) \<in> bcontfun" |
475 |
proof (intro bcontfun_normI) |
|
476 |
fix x |
|
| 60421 | 477 |
show "norm (f (clamp a b x)) \<le> c" |
478 |
using clamp_in_interval[OF assms(1), of x] f_bound |
|
479 |
by (simp add: ext_cont_def) |
|
| 59453 | 480 |
qed (simp add: clamp_continuous_on assms) |
481 |
qed |
|
482 |
||
483 |
lemma integral_clamp: |
|
484 |
"integral {t0::real..clamp t0 t1 x} f =
|
|
485 |
(if x < t0 then 0 else if x \<le> t1 then integral {t0..x} f else integral {t0..t1} f)"
|
|
486 |
by (auto simp: clamp_def) |
|
487 |
||
488 |
||
489 |
declare [[coercion Rep_bcontfun]] |
|
490 |
||
491 |
lemma ext_cont_cancel[simp]: |
|
| 60421 | 492 |
fixes x a b :: "'a::euclidean_space" |
| 59453 | 493 |
assumes x: "x \<in> cbox a b" |
| 60421 | 494 |
and "continuous_on (cbox a b) f" |
| 59453 | 495 |
shows "ext_cont f a b x = f x" |
496 |
using assms |
|
497 |
unfolding ext_cont_def |
|
498 |
proof (subst Abs_bcontfun_inverse[OF clamp_bcontfun]) |
|
499 |
show "f (clamp a b x) = f x" |
|
500 |
using x unfolding clamp_def mem_box |
|
501 |
by (intro arg_cong[where f=f] euclidean_eqI[where 'a='a]) (simp add: not_less) |
|
502 |
qed (auto simp: cbox_def) |
|
503 |
||
504 |
lemma ext_cont_cong: |
|
505 |
assumes "t0 = s0" |
|
| 60421 | 506 |
and "t1 = s1" |
507 |
and "\<And>t. t \<in> (cbox t0 t1) \<Longrightarrow> f t = g t" |
|
508 |
and "continuous_on (cbox t0 t1) f" |
|
509 |
and "continuous_on (cbox s0 s1) g" |
|
510 |
and ord: "\<And>i. i \<in> Basis \<Longrightarrow> t0 \<bullet> i \<le> t1 \<bullet> i" |
|
| 59453 | 511 |
shows "ext_cont f t0 t1 = ext_cont g s0 s1" |
512 |
unfolding assms ext_cont_def |
|
513 |
using assms clamp_in_interval[OF ord] |
|
514 |
by (subst Rep_bcontfun_inject[symmetric]) simp |
|
515 |
||
516 |
end |