@@ -24,6 +24,12 @@ and `α → (E)NNReal` functions in two cases,
24
24
* `ENNReal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0∞ functions,
25
25
* `NNReal.lintegral_mul_le_Lp_mul_Lq` : ℝ≥0 functions.
26
26
27
+ `ENNReal.lintegral_mul_norm_pow_le` is a variant where the exponents are not reciprocals:
28
+ `∫ (f ^ p * g ^ q) ∂μ ≤ (∫ f ∂μ) ^ p * (∫ g ∂μ) ^ q` where `p, q ≥ 0` and `p + q = 1`.
29
+ `ENNReal.lintegral_prod_norm_pow_le` generalizes this to a finite family of functions:
30
+ `∫ (∏ i, f i ^ p i) ∂μ ≤ ∏ i, (∫ f i ∂μ) ^ p i` when the `p` is a collection
31
+ of nonnegative weights with sum 1.
32
+
27
33
Minkowski's inequality for the Lebesgue integral of measurable functions with `ℝ≥0∞` values:
28
34
we prove `(∫ (f + g)^p ∂μ) ^ (1/p) ≤ (∫ f^p ∂μ) ^ (1/p) + (∫ g^p ∂μ) ^ (1/p)` for `1 ≤ p`.
29
35
-/
@@ -48,7 +54,7 @@ only to prove the more general results:
48
54
49
55
noncomputable section
50
56
51
- open Classical BigOperators NNReal ENNReal MeasureTheory
57
+ open Classical BigOperators NNReal ENNReal MeasureTheory Finset
52
58
53
59
set_option linter.uppercaseLean3 false
54
60
@@ -175,6 +181,101 @@ theorem lintegral_mul_le_Lp_mul_Lq (μ : Measure α) {p q : ℝ} (hpq : p.IsConj
175
181
exact ENNReal.lintegral_mul_le_Lp_mul_Lq_of_ne_zero_of_ne_top hpq hf hf_top hg_top hf_zero hg_zero
176
182
#align ennreal.lintegral_mul_le_Lp_mul_Lq ENNReal.lintegral_mul_le_Lp_mul_Lq
177
183
184
+ /-- A different formulation of Hölder's inequality for two functions, with two exponents that sum to
185
+ 1, instead of reciprocals of -/
186
+ theorem lintegral_mul_norm_pow_le {α} [MeasurableSpace α] {μ : Measure α}
187
+ {f g : α → ℝ≥0 ∞} (hf : AEMeasurable f μ) (hg : AEMeasurable g μ)
188
+ {p q : ℝ} (hp : 0 ≤ p) (hq : 0 ≤ q) (hpq : p + q = 1 ) :
189
+ ∫⁻ a, f a ^ p * g a ^ q ∂μ ≤ (∫⁻ a, f a ∂μ) ^ p * (∫⁻ a, g a ∂μ) ^ q := by
190
+ rcases hp.eq_or_lt with rfl|hp
191
+ · rw [zero_add] at hpq
192
+ simp [hpq]
193
+ rcases hq.eq_or_lt with rfl|hq
194
+ · rw [add_zero] at hpq
195
+ simp [hpq]
196
+ have h2p : 1 < 1 / p := by
197
+ rw [one_div]
198
+ apply one_lt_inv hp
199
+ linarith
200
+ have h2pq : 1 / (1 / p) + 1 / (1 / q) = 1 := by
201
+ simp [hp.ne', hq.ne', hpq]
202
+ have := ENNReal.lintegral_mul_le_Lp_mul_Lq μ ⟨h2p, h2pq⟩ (hf.pow_const p) (hg.pow_const q)
203
+ simpa [← ENNReal.rpow_mul, hp.ne', hq.ne'] using this
204
+
205
+ /-- A version of Hölder with multiple arguments -/
206
+ theorem lintegral_prod_norm_pow_le {α ι : Type *} [MeasurableSpace α] {μ : Measure α}
207
+ (s : Finset ι) {f : ι → α → ℝ≥0 ∞} (hf : ∀ i ∈ s, AEMeasurable (f i) μ)
208
+ {p : ι → ℝ} (hp : ∑ i in s, p i = 1 ) (h2p : ∀ i ∈ s, 0 ≤ p i) :
209
+ ∫⁻ a, ∏ i in s, f i a ^ p i ∂μ ≤ ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by
210
+ induction s using Finset.induction generalizing p
211
+ case empty =>
212
+ simp at hp
213
+ case insert i₀ s hi₀ ih =>
214
+ rcases eq_or_ne (p i₀) 1 with h2i₀|h2i₀
215
+ · simp [hi₀]
216
+ have h2p : ∀ i ∈ s, p i = 0 := by
217
+ simpa [hi₀, h2i₀, sum_eq_zero_iff_of_nonneg (fun i hi ↦ h2p i <| mem_insert_of_mem hi)]
218
+ using hp
219
+ calc ∫⁻ a, f i₀ a ^ p i₀ * ∏ i in s, f i a ^ p i ∂μ
220
+ = ∫⁻ a, f i₀ a ^ p i₀ * ∏ i in s, 1 ∂μ := by
221
+ congr! 3 with x
222
+ apply prod_congr rfl fun i hi ↦ by rw [h2p i hi, ENNReal.rpow_zero]
223
+ _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i in s, 1 := by simp [h2i₀]
224
+ _ = (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by
225
+ congr 1
226
+ apply prod_congr rfl fun i hi ↦ by rw [h2p i hi, ENNReal.rpow_zero]
227
+ · have hpi₀ : 0 ≤ 1 - p i₀ := by
228
+ simp_rw [sub_nonneg, ← hp, single_le_sum h2p (mem_insert_self ..)]
229
+ have h2pi₀ : 1 - p i₀ ≠ 0 := by
230
+ rwa [sub_ne_zero, ne_comm]
231
+ let q := fun i ↦ p i / (1 - p i₀)
232
+ have hq : ∑ i in s, q i = 1 := by
233
+ rw [← Finset.sum_div, ← sum_insert_sub hi₀, hp, div_self h2pi₀]
234
+ have h2q : ∀ i ∈ s, 0 ≤ q i :=
235
+ fun i hi ↦ div_nonneg (h2p i <| mem_insert_of_mem hi) hpi₀
236
+ calc ∫⁻ a, ∏ i in insert i₀ s, f i a ^ p i ∂μ
237
+ = ∫⁻ a, f i₀ a ^ p i₀ * ∏ i in s, f i a ^ p i ∂μ := by simp [hi₀]
238
+ _ = ∫⁻ a, f i₀ a ^ p i₀ * (∏ i in s, f i a ^ q i) ^ (1 - p i₀) ∂μ := by
239
+ simp [← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul,
240
+ div_mul_cancel (h := h2pi₀)]
241
+ _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * (∫⁻ a, ∏ i in s, f i a ^ q i ∂μ) ^ (1 - p i₀) := by
242
+ apply ENNReal.lintegral_mul_norm_pow_le
243
+ · exact hf i₀ <| mem_insert_self ..
244
+ · exact s.aemeasurable_prod <| fun i hi ↦ (hf i <| mem_insert_of_mem hi).pow_const _
245
+ · exact h2p i₀ <| mem_insert_self ..
246
+ · exact hpi₀
247
+ · apply add_sub_cancel'_right
248
+ _ ≤ (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * (∏ i in s, (∫⁻ a, f i a ∂μ) ^ q i) ^ (1 - p i₀) := by
249
+ gcongr -- behavior of gcongr is heartbeat-dependent, which makes code really fragile...
250
+ exact ih (fun i hi ↦ hf i <| mem_insert_of_mem hi) hq h2q
251
+ _ = (∫⁻ a, f i₀ a ∂μ) ^ p i₀ * ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by
252
+ simp [← ENNReal.prod_rpow_of_nonneg hpi₀, ← ENNReal.rpow_mul,
253
+ div_mul_cancel (h := h2pi₀)]
254
+ _ = ∏ i in insert i₀ s, (∫⁻ a, f i a ∂μ) ^ p i := by simp [hi₀]
255
+
256
+ /-- A version of Hölder with multiple arguments, one of which plays a distinguished role. -/
257
+ theorem lintegral_mul_prod_norm_pow_le {α ι : Type *} [MeasurableSpace α] {μ : Measure α}
258
+ (s : Finset ι) {g : α → ℝ≥0 ∞} {f : ι → α → ℝ≥0 ∞} (hg : AEMeasurable g μ)
259
+ (hf : ∀ i ∈ s, AEMeasurable (f i) μ) (q : ℝ) {p : ι → ℝ} (hpq : q + ∑ i in s, p i = 1 )
260
+ (hq : 0 ≤ q) (hp : ∀ i ∈ s, 0 ≤ p i) :
261
+ ∫⁻ a, g a ^ q * ∏ i in s, f i a ^ p i ∂μ ≤
262
+ (∫⁻ a, g a ∂μ) ^ q * ∏ i in s, (∫⁻ a, f i a ∂μ) ^ p i := by
263
+ suffices
264
+ ∫⁻ t, ∏ j in insertNone s, Option.elim j (g t) (fun j ↦ f j t) ^ Option.elim j q p ∂μ
265
+ ≤ ∏ j in insertNone s, (∫⁻ t, Option.elim j (g t) (fun j ↦ f j t) ∂μ) ^ Option.elim j q p by
266
+ simpa using this
267
+ refine ENNReal.lintegral_prod_norm_pow_le _ ?_ ?_ ?_
268
+ · rintro (_|i) hi
269
+ · exact hg
270
+ · refine hf i ?_
271
+ simpa using hi
272
+ · simp_rw [sum_insertNone, Option.elim]
273
+ exact hpq
274
+ · rintro (_|i) hi
275
+ · exact hq
276
+ · refine hp i ?_
277
+ simpa using hi
278
+
178
279
theorem lintegral_rpow_add_lt_top_of_lintegral_rpow_lt_top {p : ℝ} {f g : α → ℝ≥0 ∞}
179
280
(hf : AEMeasurable f μ) (hf_top : (∫⁻ a, f a ^ p ∂μ) < ⊤) (hg_top : (∫⁻ a, g a ^ p ∂μ) < ⊤)
180
281
(hp1 : 1 ≤ p) : (∫⁻ a, (f + g) a ^ p ∂μ) < ⊤ := by
0 commit comments