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

Commit 13b0d72

Browse files
committed
feat(measure_theory/function/l1_space): generalize multiplicative results to is_unit (#19073)
This generalization lets us use these lemmas for left- and right- multiplication by non-commutative normed rings. This doesn't make it all the way to `integral_mul_const`, but it's a step in that direction.
1 parent 5a2df4c commit 13b0d72

File tree

2 files changed

+36
-28
lines changed

2 files changed

+36
-28
lines changed

src/measure_theory/function/l1_space.lean

Lines changed: 35 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -352,16 +352,16 @@ end pos_part
352352
section normed_space
353353
variables {𝕜 : Type*}
354354

355-
lemma has_finite_integral.smul' [has_smul 𝕜 β] [has_nnnorm 𝕜] (c : 𝕜) {f : α → β}
356-
(h : ∀ (k : 𝕜) (b : β), ‖k • b‖₊ ≤ ‖k‖₊ * ‖b‖₊) :
355+
lemma has_finite_integral.smul
356+
[normed_add_comm_group 𝕜] [smul_zero_class 𝕜 β] [has_bounded_smul 𝕜 β] (c : 𝕜) {f : α → β} :
357357
has_finite_integral f μ → has_finite_integral (c • f) μ :=
358358
begin
359359
simp only [has_finite_integral], assume hfi,
360360
calc
361361
∫⁻ (a : α), ‖c • f a‖₊ ∂μ ≤ ∫⁻ (a : α), (‖c‖₊) * ‖f a‖₊ ∂μ : begin
362362
refine lintegral_mono _,
363363
intro i,
364-
exact_mod_cast h c (f i),
364+
exact_mod_cast (nnnorm_smul_le c (f i) : _),
365365
end
366366
... < ∞ :
367367
begin
@@ -370,30 +370,28 @@ begin
370370
end
371371
end
372372

373-
lemma has_finite_integral.smul [normed_field 𝕜] [normed_space 𝕜 β] (c : 𝕜) {f : α → β} :
374-
has_finite_integral f μ → has_finite_integral (c • f) μ :=
375-
has_finite_integral.smul' _ $ λ a b, nnnorm_smul_le a b
376-
377-
lemma has_finite_integral_smul_iff [normed_field 𝕜] [normed_space 𝕜 β] {c : 𝕜} (hc : c ≠ 0)
373+
lemma has_finite_integral_smul_iff
374+
[normed_ring 𝕜] [mul_action_with_zero 𝕜 β] [has_bounded_smul 𝕜 β]
375+
{c : 𝕜} (hc : is_unit c)
378376
(f : α → β) :
379377
has_finite_integral (c • f) μ ↔ has_finite_integral f μ :=
380378
begin
379+
obtain ⟨c, rfl⟩ := hc,
381380
split,
382381
{ assume h,
383-
simpa only [smul_smul, inv_mul_cancel hc, one_smul] using h.smul c⁻¹ },
382+
simpa only [smul_smul, units.inv_mul, one_smul] using h.smul (↑c⁻¹ : 𝕜) },
384383
exact has_finite_integral.smul _
385384
end
386385

387386
lemma has_finite_integral.const_mul [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ)
388387
(c : 𝕜) :
389388
has_finite_integral (λ x, c * f x) μ :=
390-
(has_finite_integral.smul' c nnnorm_mul_le h : _)
389+
h.smul c
391390

392391
lemma has_finite_integral.mul_const [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ)
393392
(c : 𝕜) :
394393
has_finite_integral (λ x, f x * c) μ :=
395-
(has_finite_integral.smul' (mul_opposite.op c)
396-
(λ a b, (nnnorm_mul_le b a.unop).trans_eq $ mul_comm _ _) h : _)
394+
h.smul (mul_opposite.op c)
397395

398396
end normed_space
399397

@@ -673,7 +671,8 @@ end
673671

674672
/-- Hölder's inequality for integrable functions: the scalar multiplication of an integrable
675673
scalar-valued function by a vector-value function with finite essential supremum is integrable. -/
676-
lemma integrable.smul_ess_sup {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β] {f : α → 𝕜}
674+
lemma integrable.smul_ess_sup {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
675+
{f : α → 𝕜}
677676
(hf : integrable f μ) {g : α → β} (g_ae_strongly_measurable : ae_strongly_measurable g μ)
678677
(ess_sup_g : ess_sup (λ x, (‖g x‖₊ : ℝ≥0∞)) μ ≠ ∞) :
679678
integrable (λ (x : α), f x • g x) μ :=
@@ -932,16 +931,25 @@ hf.neg.pos_part
932931

933932
end pos_part
934933

935-
section normed_space
936-
variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β]
934+
section has_bounded_smul
935+
variables {𝕜 : Type*}
937936

938-
lemma integrable.smul (c : 𝕜) {f : α → β}
937+
lemma integrable.smul [normed_add_comm_group 𝕜] [smul_zero_class 𝕜 β] [has_bounded_smul 𝕜 β]
938+
(c : 𝕜) {f : α → β}
939939
(hf : integrable f μ) : integrable (c • f) μ :=
940940
⟨hf.ae_strongly_measurable.const_smul c, hf.has_finite_integral.smul c⟩
941941

942-
lemma integrable_smul_iff {c : 𝕜} (hc : c ≠ 0) (f : α → β) :
942+
lemma is_unit.integrable_smul_iff [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
943+
{c : 𝕜} (hc : is_unit c) (f : α → β) :
944+
integrable (c • f) μ ↔ integrable f μ :=
945+
and_congr (hc.ae_strongly_measurable_const_smul_iff) (has_finite_integral_smul_iff hc f)
946+
947+
lemma integrable_smul_iff [normed_division_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
948+
{c : 𝕜} (hc : c ≠ 0) (f : α → β) :
943949
integrable (c • f) μ ↔ integrable f μ :=
944-
and_congr (ae_strongly_measurable_const_smul_iff₀ hc) (has_finite_integral_smul_iff hc f)
950+
(is_unit.mk0 _ hc).integrable_smul_iff f
951+
952+
variables [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
945953

946954
lemma integrable.smul_of_top_right {f : α → β} {φ : α → 𝕜}
947955
(hf : integrable f μ) (hφ : mem_ℒp φ ∞ μ) :
@@ -957,7 +965,7 @@ lemma integrable.smul_const {f : α → 𝕜} (hf : integrable f μ) (c : β) :
957965
integrable (λ x, f x • c) μ :=
958966
hf.smul_of_top_left (mem_ℒp_top_const c)
959967

960-
end normed_space
968+
end has_bounded_smul
961969

962970
section normed_space_over_complete_field
963971
variables {𝕜 : Type*} [nontrivially_normed_field 𝕜] [complete_space 𝕜]
@@ -980,27 +988,27 @@ variables {𝕜 : Type*} [normed_ring 𝕜] {f : α → 𝕜}
980988

981989
lemma integrable.const_mul {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
982990
integrable (λ x, c * f x) μ :=
983-
⟨h.ae_strongly_measurable.const_smul c, h.has_finite_integral.const_mul c⟩
991+
h.smul c
984992

985993
lemma integrable.const_mul' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
986994
integrable ((λ (x : α), c) * f) μ :=
987995
integrable.const_mul h c
988996

989997
lemma integrable.mul_const {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
990998
integrable (λ x, f x * c) μ :=
991-
⟨h.ae_strongly_measurable.const_smul (mul_opposite.op c), h.has_finite_integral.mul_const c⟩
999+
h.smul (mul_opposite.op c)
9921000

9931001
lemma integrable.mul_const' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
9941002
integrable (f * (λ (x : α), c)) μ :=
9951003
integrable.mul_const h c
9961004

9971005
lemma integrable_const_mul_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) :
9981006
integrable (λ x, c * f x) μ ↔ integrable f μ :=
999-
let ⟨u, hc⟩ := hc in hc ▸ ⟨λ h, by simpa using h.const_mul ↑(u⁻¹), λ h, h.const_mul _⟩
1007+
hc.integrable_smul_iff f
10001008

10011009
lemma integrable_mul_const_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) :
10021010
integrable (λ x, f x * c) μ ↔ integrable f μ :=
1003-
let ⟨u, hc⟩ := hc in hc ▸ ⟨λ h, by simpa using h.mul_const ↑(u⁻¹), λ h, h.mul_const _⟩
1011+
hc.op.integrable_smul_iff f
10041012

10051013
lemma integrable.bdd_mul' {f g : α → 𝕜} {c : ℝ} (hg : integrable g μ)
10061014
(hf : ae_strongly_measurable f μ) (hf_bound : ∀ᵐ x ∂μ, ‖f x‖ ≤ c) :
@@ -1136,13 +1144,13 @@ lemma integrable.sub {f g : α →ₘ[μ] β} (hf : integrable f) (hg : integrab
11361144

11371145
end
11381146

1139-
section normed_space
1140-
variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β]
1147+
section has_bounded_smul
1148+
variables {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
11411149

11421150
lemma integrable.smul {c : 𝕜} {f : α →ₘ[μ] β} : integrable f → integrable (c • f) :=
11431151
induction_on f $ λ f hfm hfi, (integrable_mk _).2 $ ((integrable_mk hfm).1 hfi).smul _
11441152

1145-
end normed_space
1153+
end has_bounded_smul
11461154

11471155
end
11481156

@@ -1263,7 +1271,7 @@ by { simp [integrable.to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm_sub] }
12631271
edist (hf.to_L1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ :=
12641272
by { simp [integrable.to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm] }
12651273

1266-
variables {𝕜 : Type*} [normed_field 𝕜] [normed_space 𝕜 β]
1274+
variables {𝕜 : Type*} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
12671275

12681276
lemma to_L1_smul (f : α → β) (hf : integrable f μ) (k : 𝕜) :
12691277
to_L1 (λ a, k • f a) (hf.smul k) = k • to_L1 f hf := rfl

src/measure_theory/integral/peak_function.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -58,7 +58,7 @@ begin
5858
filter_upwards [tendsto_uniformly_on_iff.1 (hlφ u u_open x₀u) 1 zero_lt_one, hiφ]
5959
with i hi h'i,
6060
have A : integrable_on (λ x, φ i x • g x) (s \ u) μ,
61-
{ apply integrable.smul_of_top_right (hmg.mono (diff_subset _ _) le_rfl),
61+
{ refine integrable.smul_of_top_right (hmg.mono (diff_subset _ _) le_rfl) _,
6262
apply mem_ℒp_top_of_bound
6363
((integrable_of_integral_eq_one h'i).ae_strongly_measurable.mono_set ((diff_subset _ _))) 1,
6464
filter_upwards [self_mem_ae_restrict (hs.diff u_open.measurable_set)] with x hx,

0 commit comments

Comments
 (0)