1
1
/-
2
2
Copyright (c) 2020 Sébastien Gouëzel. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
- Authors: Sébastien Gouëzel
4
+ Authors: Sébastien Gouëzel, Floris van Doorn
5
5
-/
6
6
import analysis.calculus.iterated_deriv
7
7
import analysis.inner_product_space.euclidean_dist
8
+ import measure_theory.function.locally_integrable
9
+ import measure_theory.integral.set_integral
8
10
9
11
/-!
10
12
# Infinitely smooth bump function
@@ -29,6 +31,9 @@ function cannot have:
29
31
function: real numbers `r`, `R`, and proofs of `0 < r < R`. The function itself is available
30
32
through `coe_fn`.
31
33
34
+ * If `f : cont_diff_bump_of_inner c` and `μ` is a measure on the domain of `f`, then `f.normed μ`
35
+ is a smooth bump function with integral `1` w.r.t. `μ`.
36
+
32
37
* `f : cont_diff_bump c`, where `c` is a point in a finite dimensional real vector space, is a
33
38
bundled smooth function such that
34
39
@@ -264,7 +269,7 @@ end smooth_transition
264
269
265
270
end real
266
271
267
- variable {E : Type *}
272
+ variables {E X : Type *}
268
273
269
274
/-- `f : cont_diff_bump_of_inner c`, where `c` is a point in an inner product space, is a
270
275
bundled smooth function such that
@@ -287,7 +292,8 @@ lemma R_pos {c : E} (f : cont_diff_bump_of_inner c) : 0 < f.R := f.r_pos.trans f
287
292
288
293
instance (c : E) : inhabited (cont_diff_bump_of_inner c) := ⟨⟨1 , 2 , zero_lt_one, one_lt_two⟩⟩
289
294
290
- variables [inner_product_space ℝ E] {c : E} (f : cont_diff_bump_of_inner c) {x : E}
295
+ variables [inner_product_space ℝ E] [normed_group X] [normed_space ℝ X]
296
+ variables {c : E} (f : cont_diff_bump_of_inner c) {x : E} {n : with_top ℕ}
291
297
292
298
/-- The function defined by `f : cont_diff_bump_of_inner c`. Use automatic coercion to
293
299
function instead. -/
@@ -296,6 +302,15 @@ def to_fun (f : cont_diff_bump_of_inner c) : E → ℝ :=
296
302
297
303
instance : has_coe_to_fun (cont_diff_bump_of_inner c) (λ _, E → ℝ) := ⟨to_fun⟩
298
304
305
+ protected lemma «def » (x : E) : f x = real.smooth_transition ((f.R - dist x c) / (f.R - f.r)) :=
306
+ rfl
307
+
308
+ protected lemma sub (x : E) : f (c - x) = f (c + x) :=
309
+ by simp_rw [f.def, dist_self_sub_left, dist_self_add_left]
310
+
311
+ protected lemma neg (f : cont_diff_bump_of_inner (0 : E)) (x : E) : f (- x) = f x :=
312
+ by simp_rw [← zero_sub, f.sub, zero_add]
313
+
299
314
open real (smooth_transition) real.smooth_transition metric
300
315
301
316
lemma one_of_mem_closed_ball (hx : x ∈ closed_ball c f.r) :
@@ -304,6 +319,10 @@ one_of_one_le $ (one_le_div (sub_pos.2 f.r_lt_R)).2 $ sub_le_sub_left hx _
304
319
305
320
lemma nonneg : 0 ≤ f x := nonneg _
306
321
322
+ /-- A version of `cont_diff_bump_of_inner.nonneg` with `x` explicit -/
323
+ lemma nonneg' (x : E) : 0 ≤ f x :=
324
+ f.nonneg
325
+
307
326
lemma le_one : f x ≤ 1 := le_one _
308
327
309
328
lemma pos_of_mem_ball (hx : x ∈ ball c f.R) : 0 < f x :=
@@ -324,6 +343,12 @@ begin
324
343
{ simp [hx.not_lt, f.zero_of_le_dist hx] }
325
344
end
326
345
346
+ lemma tsupport_eq : tsupport f = closed_ball c f.R :=
347
+ by simp_rw [tsupport, f.support_eq, closure_ball _ f.R_pos.ne']
348
+
349
+ protected lemma has_compact_support [finite_dimensional ℝ E] : has_compact_support f :=
350
+ by simp_rw [has_compact_support, f.tsupport_eq, is_compact_closed_ball]
351
+
327
352
lemma eventually_eq_one_of_mem_ball (h : x ∈ ball c f.r) :
328
353
f =ᶠ[𝓝 x] 1 :=
329
354
((is_open_lt (continuous_id.dist continuous_const) continuous_const).eventually_mem h).mono $
@@ -332,26 +357,105 @@ lemma eventually_eq_one_of_mem_ball (h : x ∈ ball c f.r) :
332
357
lemma eventually_eq_one : f =ᶠ[𝓝 c] 1 :=
333
358
f.eventually_eq_one_of_mem_ball (mem_ball_self f.r_pos)
334
359
335
- protected lemma cont_diff_at {n} :
336
- cont_diff_at ℝ n f x :=
360
+ /-- `cont_diff_bump` is `𝒞ⁿ` in all its arguments. -/
361
+ protected lemma _root_.cont_diff_at.cont_diff_bump {c g : X → E}
362
+ {f : ∀ x, cont_diff_bump_of_inner (c x)} {x : X}
363
+ (hc : cont_diff_at ℝ n c x) (hr : cont_diff_at ℝ n (λ x, (f x).r) x)
364
+ (hR : cont_diff_at ℝ n (λ x, (f x).R) x)
365
+ (hg : cont_diff_at ℝ n g x) : cont_diff_at ℝ n (λ x, f x (g x)) x :=
337
366
begin
338
- rcases em (x = c) with rfl|hx,
339
- { refine cont_diff_at.congr_of_eventually_eq _ f.eventually_eq_one,
340
- rw pi.one_def,
341
- exact cont_diff_at_const },
342
- { exact real.smooth_transition.cont_diff_at.comp x
343
- (cont_diff_at.div_const $ cont_diff_at_const.sub $
344
- cont_diff_at_id.dist cont_diff_at_const hx) }
367
+ rcases eq_or_ne (g x) (c x) with hx|hx,
368
+ { have : (λ x, f x (g x)) =ᶠ[𝓝 x] (λ x, 1 ),
369
+ { have : dist (g x) (c x) < (f x).r, { simp_rw [hx, dist_self, (f x).r_pos] },
370
+ have := continuous_at.eventually_lt (hg.continuous_at.dist hc.continuous_at) hr.continuous_at
371
+ this ,
372
+ exact eventually_of_mem this
373
+ (λ x hx, (f x).one_of_mem_closed_ball (mem_set_of_eq.mp hx).le) },
374
+ exact cont_diff_at_const.congr_of_eventually_eq this },
375
+ { refine real.smooth_transition.cont_diff_at.comp x _,
376
+ refine ((hR.sub $ hg.dist hc hx).div (hR.sub hr) (sub_pos.mpr (f x).r_lt_R).ne') }
345
377
end
346
378
347
- protected lemma cont_diff {n} :
348
- cont_diff ℝ n f :=
349
- cont_diff_iff_cont_diff_at.2 $ λ y, f.cont_diff_at
379
+ lemma _root_.cont_diff.cont_diff_bump {c g : X → E} {f : ∀ x, cont_diff_bump_of_inner (c x)}
380
+ (hc : cont_diff ℝ n c) (hr : cont_diff ℝ n (λ x, (f x).r)) (hR : cont_diff ℝ n (λ x, (f x).R))
381
+ (hg : cont_diff ℝ n g) : cont_diff ℝ n (λ x, f x (g x)) :=
382
+ by { rw [cont_diff_iff_cont_diff_at] at *, exact λ x, (hc x).cont_diff_bump (hr x) (hR x) (hg x) }
350
383
351
- protected lemma cont_diff_within_at {s n} :
352
- cont_diff_within_at ℝ n f s x :=
384
+ protected lemma cont_diff : cont_diff ℝ n f :=
385
+ cont_diff_const.cont_diff_bump cont_diff_const cont_diff_const cont_diff_id
386
+
387
+ protected lemma cont_diff_at : cont_diff_at ℝ n f x :=
388
+ f.cont_diff.cont_diff_at
389
+
390
+ protected lemma cont_diff_within_at {s : set E} : cont_diff_within_at ℝ n f s x :=
353
391
f.cont_diff_at.cont_diff_within_at
354
392
393
+ protected lemma continuous : continuous f :=
394
+ cont_diff_zero.mp f.cont_diff
395
+
396
+ open measure_theory
397
+ variables [measurable_space E] {μ : measure E}
398
+
399
+ /-- A bump function normed so that `∫ x, f.normed μ x ∂μ = 1`. -/
400
+ protected def normed (μ : measure E) : E → ℝ :=
401
+ λ x, f x / ∫ x, f x ∂μ
402
+
403
+ lemma normed_def {μ : measure E} (x : E) : f.normed μ x = f x / ∫ x, f x ∂μ :=
404
+ rfl
405
+
406
+ lemma nonneg_normed (x : E) : 0 ≤ f.normed μ x :=
407
+ div_nonneg f.nonneg $ integral_nonneg f.nonneg'
408
+
409
+ lemma cont_diff_normed {n : with_top ℕ} : cont_diff ℝ n (f.normed μ) :=
410
+ f.cont_diff.div_const
411
+
412
+ lemma continuous_normed : continuous (f.normed μ) :=
413
+ f.continuous.div_const
414
+
415
+ lemma normed_sub (x : E) : f.normed μ (c - x) = f.normed μ (c + x) :=
416
+ by simp_rw [f.normed_def, f.sub]
417
+
418
+ lemma normed_neg (f : cont_diff_bump_of_inner (0 : E)) (x : E) : f.normed μ (- x) = f.normed μ x :=
419
+ by simp_rw [f.normed_def, f.neg]
420
+
421
+ variables [borel_space E] [finite_dimensional ℝ E] [is_locally_finite_measure μ]
422
+
423
+ protected lemma integrable : integrable f μ :=
424
+ f.continuous.integrable_of_has_compact_support f.has_compact_support
425
+
426
+ protected lemma integrable_normed : integrable (f.normed μ) μ :=
427
+ f.integrable.div_const _
428
+
429
+ variables [μ .is_open_pos_measure]
430
+
431
+ lemma integral_pos : 0 < ∫ x, f x ∂μ :=
432
+ begin
433
+ refine (integral_pos_iff_support_of_nonneg f.nonneg' f.integrable).mpr _,
434
+ rw [f.support_eq],
435
+ refine is_open_ball.measure_pos _ (nonempty_ball.mpr f.R_pos)
436
+ end
437
+
438
+ lemma integral_normed : ∫ x, f.normed μ x ∂μ = 1 :=
439
+ begin
440
+ simp_rw [cont_diff_bump_of_inner.normed, div_eq_mul_inv, mul_comm (f _), ← smul_eq_mul,
441
+ integral_smul],
442
+ exact inv_mul_cancel (f.integral_pos.ne')
443
+ end
444
+
445
+ lemma support_normed_eq : support (f.normed μ) = metric.ball c f.R :=
446
+ by simp_rw [cont_diff_bump_of_inner.normed, support_div, f.support_eq,
447
+ support_const f.integral_pos.ne', inter_univ]
448
+
449
+ lemma tsupport_normed_eq : tsupport (f.normed μ) = metric.closed_ball c f.R :=
450
+ by simp_rw [tsupport, f.support_normed_eq, closure_ball _ f.R_pos.ne']
451
+
452
+ lemma has_compact_support_normed : has_compact_support (f.normed μ) :=
453
+ by simp_rw [has_compact_support, f.tsupport_normed_eq, is_compact_closed_ball]
454
+
455
+ variable (μ)
456
+ lemma integral_normed_smul (z : X) [complete_space X] : ∫ x, f.normed μ x • z ∂μ = z :=
457
+ by simp_rw [integral_smul_const, f.integral_normed, one_smul]
458
+
355
459
end cont_diff_bump_of_inner
356
460
357
461
/-- `f : cont_diff_bump c`, where `c` is a point in a finite dimensional real vector space, is
0 commit comments