@@ -352,16 +352,16 @@ end pos_part
352
352
section normed_space
353
353
variables {𝕜 : Type *}
354
354
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 : α → β} :
357
357
has_finite_integral f μ → has_finite_integral (c • f) μ :=
358
358
begin
359
359
simp only [has_finite_integral], assume hfi,
360
360
calc
361
361
∫⁻ (a : α), ‖c • f a‖₊ ∂μ ≤ ∫⁻ (a : α), (‖c‖₊) * ‖f a‖₊ ∂μ : begin
362
362
refine lintegral_mono _,
363
363
intro i,
364
- exact_mod_cast h c (f i),
364
+ exact_mod_cast (nnnorm_smul_le c (f i) : _ ),
365
365
end
366
366
... < ∞ :
367
367
begin
@@ -370,30 +370,28 @@ begin
370
370
end
371
371
end
372
372
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)
378
376
(f : α → β) :
379
377
has_finite_integral (c • f) μ ↔ has_finite_integral f μ :=
380
378
begin
379
+ obtain ⟨c, rfl⟩ := hc,
381
380
split,
382
381
{ 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⁻¹ : 𝕜) },
384
383
exact has_finite_integral.smul _
385
384
end
386
385
387
386
lemma has_finite_integral.const_mul [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ)
388
387
(c : 𝕜) :
389
388
has_finite_integral (λ x, c * f x) μ :=
390
- (has_finite_integral .smul' c nnnorm_mul_le h : _)
389
+ h .smul c
391
390
392
391
lemma has_finite_integral.mul_const [normed_ring 𝕜] {f : α → 𝕜} (h : has_finite_integral f μ)
393
392
(c : 𝕜) :
394
393
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)
397
395
398
396
end normed_space
399
397
673
671
674
672
/-- Hölder's inequality for integrable functions: the scalar multiplication of an integrable
675
673
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 : α → 𝕜}
677
676
(hf : integrable f μ) {g : α → β} (g_ae_strongly_measurable : ae_strongly_measurable g μ)
678
677
(ess_sup_g : ess_sup (λ x, (‖g x‖₊ : ℝ≥0 ∞)) μ ≠ ∞) :
679
678
integrable (λ (x : α), f x • g x) μ :=
@@ -932,16 +931,25 @@ hf.neg.pos_part
932
931
933
932
end pos_part
934
933
935
- section normed_space
936
- variables {𝕜 : Type *} [normed_field 𝕜] [normed_space 𝕜 β]
934
+ section has_bounded_smul
935
+ variables {𝕜 : Type *}
937
936
938
- lemma integrable.smul (c : 𝕜) {f : α → β}
937
+ lemma integrable.smul [normed_add_comm_group 𝕜] [smul_zero_class 𝕜 β] [has_bounded_smul 𝕜 β]
938
+ (c : 𝕜) {f : α → β}
939
939
(hf : integrable f μ) : integrable (c • f) μ :=
940
940
⟨hf.ae_strongly_measurable.const_smul c, hf.has_finite_integral.smul c⟩
941
941
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 : α → β) :
943
949
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 𝕜 β]
945
953
946
954
lemma integrable.smul_of_top_right {f : α → β} {φ : α → 𝕜}
947
955
(hf : integrable f μ) (hφ : mem_ℒp φ ∞ μ) :
@@ -957,7 +965,7 @@ lemma integrable.smul_const {f : α → 𝕜} (hf : integrable f μ) (c : β) :
957
965
integrable (λ x, f x • c) μ :=
958
966
hf.smul_of_top_left (mem_ℒp_top_const c)
959
967
960
- end normed_space
968
+ end has_bounded_smul
961
969
962
970
section normed_space_over_complete_field
963
971
variables {𝕜 : Type *} [nontrivially_normed_field 𝕜] [complete_space 𝕜]
@@ -980,27 +988,27 @@ variables {𝕜 : Type*} [normed_ring 𝕜] {f : α → 𝕜}
980
988
981
989
lemma integrable.const_mul {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
982
990
integrable (λ x, c * f x) μ :=
983
- ⟨h.ae_strongly_measurable.const_smul c, h.has_finite_integral.const_mul c⟩
991
+ h.smul c
984
992
985
993
lemma integrable.const_mul' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
986
994
integrable ((λ (x : α), c) * f) μ :=
987
995
integrable.const_mul h c
988
996
989
997
lemma integrable.mul_const {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
990
998
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)
992
1000
993
1001
lemma integrable.mul_const' {f : α → 𝕜} (h : integrable f μ) (c : 𝕜) :
994
1002
integrable (f * (λ (x : α), c)) μ :=
995
1003
integrable.mul_const h c
996
1004
997
1005
lemma integrable_const_mul_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) :
998
1006
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
1000
1008
1001
1009
lemma integrable_mul_const_iff {c : 𝕜} (hc : is_unit c) (f : α → 𝕜) :
1002
1010
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
1004
1012
1005
1013
lemma integrable.bdd_mul' {f g : α → 𝕜} {c : ℝ} (hg : integrable g μ)
1006
1014
(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
1136
1144
1137
1145
end
1138
1146
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 𝕜 β]
1141
1149
1142
1150
lemma integrable.smul {c : 𝕜} {f : α →ₘ[μ] β} : integrable f → integrable (c • f) :=
1143
1151
induction_on f $ λ f hfm hfi, (integrable_mk _).2 $ ((integrable_mk hfm).1 hfi).smul _
1144
1152
1145
- end normed_space
1153
+ end has_bounded_smul
1146
1154
1147
1155
end
1148
1156
@@ -1263,7 +1271,7 @@ by { simp [integrable.to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm_sub] }
1263
1271
edist (hf.to_L1 f) 0 = ∫⁻ a, edist (f a) 0 ∂μ :=
1264
1272
by { simp [integrable.to_L1, snorm, snorm'], simp [edist_eq_coe_nnnorm] }
1265
1273
1266
- variables {𝕜 : Type *} [normed_field 𝕜] [normed_space 𝕜 β]
1274
+ variables {𝕜 : Type *} [normed_ring 𝕜] [module 𝕜 β] [has_bounded_smul 𝕜 β]
1267
1275
1268
1276
lemma to_L1_smul (f : α → β) (hf : integrable f μ) (k : 𝕜) :
1269
1277
to_L1 (λ a, k • f a) (hf.smul k) = k • to_L1 f hf := rfl
0 commit comments