Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 0d77f29

Browse files
committed
feat(analysis/calculus/specific_functions): define normed bump functions (#13463)
* Normed bump functions have integral 1 w.r.t. the specified measure. * Also add a few more properties of bump functions, including its smoothness in all arguments (including midpoint and the two radii). * From the sphere eversion project * Required for convolutions
1 parent 06a6044 commit 0d77f29

File tree

5 files changed

+159
-17
lines changed

5 files changed

+159
-17
lines changed

src/analysis/calculus/specific_functions.lean

Lines changed: 121 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -1,10 +1,12 @@
11
/-
22
Copyright (c) 2020 Sébastien Gouëzel. All rights reserved.
33
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
55
-/
66
import analysis.calculus.iterated_deriv
77
import analysis.inner_product_space.euclidean_dist
8+
import measure_theory.function.locally_integrable
9+
import measure_theory.integral.set_integral
810

911
/-!
1012
# Infinitely smooth bump function
@@ -29,6 +31,9 @@ function cannot have:
2931
function: real numbers `r`, `R`, and proofs of `0 < r < R`. The function itself is available
3032
through `coe_fn`.
3133
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+
3237
* `f : cont_diff_bump c`, where `c` is a point in a finite dimensional real vector space, is a
3338
bundled smooth function such that
3439
@@ -264,7 +269,7 @@ end smooth_transition
264269

265270
end real
266271

267-
variable {E : Type*}
272+
variables {E X : Type*}
268273

269274
/-- `f : cont_diff_bump_of_inner c`, where `c` is a point in an inner product space, is a
270275
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
287292

288293
instance (c : E) : inhabited (cont_diff_bump_of_inner c) := ⟨⟨1, 2, zero_lt_one, one_lt_two⟩⟩
289294

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 ℕ}
291297

292298
/-- The function defined by `f : cont_diff_bump_of_inner c`. Use automatic coercion to
293299
function instead. -/
@@ -296,6 +302,15 @@ def to_fun (f : cont_diff_bump_of_inner c) : E → ℝ :=
296302

297303
instance : has_coe_to_fun (cont_diff_bump_of_inner c) (λ _, E → ℝ) := ⟨to_fun⟩
298304

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+
299314
open real (smooth_transition) real.smooth_transition metric
300315

301316
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 _
304319

305320
lemma nonneg : 0 ≤ f x := nonneg _
306321

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+
307326
lemma le_one : f x ≤ 1 := le_one _
308327

309328
lemma pos_of_mem_ball (hx : x ∈ ball c f.R) : 0 < f x :=
@@ -324,6 +343,12 @@ begin
324343
{ simp [hx.not_lt, f.zero_of_le_dist hx] }
325344
end
326345

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+
327352
lemma eventually_eq_one_of_mem_ball (h : x ∈ ball c f.r) :
328353
f =ᶠ[𝓝 x] 1 :=
329354
((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) :
332357
lemma eventually_eq_one : f =ᶠ[𝓝 c] 1 :=
333358
f.eventually_eq_one_of_mem_ball (mem_ball_self f.r_pos)
334359

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 :=
337366
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') }
345377
end
346378

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) }
350383

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 :=
353391
f.cont_diff_at.cont_diff_within_at
354392

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+
355459
end cont_diff_bump_of_inner
356460

357461
/-- `f : cont_diff_bump c`, where `c` is a point in a finite dimensional real vector space, is

src/analysis/normed/group/basic.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -158,6 +158,12 @@ by rw [← dist_zero_left, ← dist_add_left g 0 h, add_zero]
158158
@[simp] theorem dist_self_add_left (g h : E) : dist (g + h) g = ∥h∥ :=
159159
by rw [dist_comm, dist_self_add_right]
160160

161+
@[simp] theorem dist_self_sub_right (g h : E) : dist g (g - h) = ∥h∥ :=
162+
by rw [sub_eq_add_neg, dist_self_add_right, norm_neg]
163+
164+
@[simp] theorem dist_self_sub_left (g h : E) : dist (g - h) g = ∥h∥ :=
165+
by rw [dist_comm, dist_self_sub_right]
166+
161167
/-- **Triangle inequality** for the norm. -/
162168
lemma norm_add_le (g h : E) : ∥g + h∥ ≤ ∥g∥ + ∥h∥ :=
163169
by simpa [dist_eq_norm] using dist_triangle g 0 (-h)

src/measure_theory/function/l1_space.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -833,6 +833,10 @@ lemma integrable.mul_const {f : α → ℝ} (h : integrable f μ) (c : ℝ) :
833833
integrable (λ x, f x * c) μ :=
834834
by simp_rw [mul_comm, h.const_mul _]
835835

836+
lemma integrable.div_const {f : α → ℝ} (h : integrable f μ) (c : ℝ) :
837+
integrable (λ x, f x / c) μ :=
838+
by simp_rw [div_eq_mul_inv, h.mul_const]
839+
836840
end normed_space
837841

838842
section normed_space_over_complete_field

src/order/cover.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -224,6 +224,18 @@ h.wcovby.Icc_eq
224224

225225
end partial_order
226226

227+
section linear_order
228+
229+
variables [linear_order α] {a b : α}
230+
231+
lemma covby.Ioi_eq (h : a ⋖ b) : Ioi a = Ici b :=
232+
by rw [← Ioo_union_Ici_eq_Ioi h.lt, h.Ioo_eq, empty_union]
233+
234+
lemma covby.Iio_eq (h : a ⋖ b) : Iio b = Iic a :=
235+
by rw [← Iic_union_Ioo_eq_Iio h.lt, h.Ioo_eq, union_empty]
236+
237+
end linear_order
238+
227239
namespace set
228240

229241
lemma wcovby_insert (x : α) (s : set α) : s ⩿ insert x s :=

src/topology/algebra/order/basic.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -559,6 +559,22 @@ lemma continuous.if_le [topological_space γ] [Π x, decidable (f x ≤ g x)] {f
559559
continuous (λ x, if f x ≤ g x then f' x else g' x) :=
560560
continuous_if_le hf hg hf'.continuous_on hg'.continuous_on hfg
561561

562+
lemma tendsto.eventually_lt {l : filter γ} {f g : γ → α} {y z : α}
563+
(hf : tendsto f l (𝓝 y)) (hg : tendsto g l (𝓝 z)) (hyz : y < z) : ∀ᶠ x in l, f x < g x :=
564+
begin
565+
by_cases h : y ⋖ z,
566+
{ filter_upwards [hf (Iio_mem_nhds hyz), hg (Ioi_mem_nhds hyz)],
567+
rw [h.Iio_eq],
568+
exact λ x hfx hgx, lt_of_le_of_lt hfx hgx },
569+
{ obtain ⟨w, hyw, hwz⟩ := (not_covby_iff hyz).mp h,
570+
filter_upwards [hf (Iio_mem_nhds hyw), hg (Ioi_mem_nhds hwz)],
571+
exact λ x, lt_trans },
572+
end
573+
574+
lemma continuous_at.eventually_lt {x₀ : β} (hf : continuous_at f x₀)
575+
(hg : continuous_at g x₀) (hfg : f x₀ < g x₀) : ∀ᶠ x in 𝓝 x₀, f x < g x :=
576+
tendsto.eventually_lt hf hg hfg
577+
562578
@[continuity] lemma continuous.min (hf : continuous f) (hg : continuous g) :
563579
continuous (λb, min (f b) (g b)) :=
564580
by { simp only [min_def], exact hf.if_le hg hf hg (λ x, id) }

0 commit comments

Comments
 (0)