diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 1210abee0c080..73bc6a93fd4f2 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -971,23 +971,23 @@ begin exact snorm'_lt_top_of_snorm'_lt_top_of_exponent_le hfq_m hfq_lt_top (le_of_lt hp_pos) hpq_real, end -lemma snorm'_sum_le [second_countable_topology E] {ι} {f : ι → α → E} {s : finset ι} +section has_measurable_add +variable [has_measurable_add₂ E] + +lemma snorm'_sum_le {ι} {f : ι → α → E} {s : finset ι} (hfs : ∀ i, i ∈ s → ae_measurable (f i) μ) (hq1 : 1 ≤ q) : snorm' (∑ i in s, f i) q μ ≤ ∑ i in s, snorm' (f i) q μ := finset.le_sum_of_subadditive_on_pred (λ (f : α → E), snorm' f q μ) (λ f, ae_measurable f μ) (snorm'_zero (zero_lt_one.trans_le hq1)) (λ f g hf hg, snorm'_add_le hf hg hq1) (λ x y, ae_measurable.add) _ hfs -lemma snorm_sum_le [second_countable_topology E] {ι} {f : ι → α → E} {s : finset ι} +lemma snorm_sum_le {ι} {f : ι → α → E} {s : finset ι} (hfs : ∀ i, i ∈ s → ae_measurable (f i) μ) (hp1 : 1 ≤ p) : snorm (∑ i in s, f i) p μ ≤ ∑ i in s, snorm (f i) p μ := finset.le_sum_of_subadditive_on_pred (λ (f : α → E), snorm f p μ) (λ f, ae_measurable f μ) snorm_zero (λ f g hf hg, snorm_add_le hf hg hp1) (λ x y, ae_measurable.add) _ hfs -section second_countable_topology -variable [second_countable_topology E] - lemma mem_ℒp.add {f g : α → E} (hf : mem_ℒp f p μ) (hg : mem_ℒp g p μ) : mem_ℒp (f + g) p μ := ⟨ae_measurable.add hf.1 hg.1, snorm_add_lt_top hf hg⟩ @@ -1006,7 +1006,7 @@ begin exact (hf i (s.mem_insert_self i)).add (ih (λ j hj, hf j (finset.mem_insert_of_mem hj))), }, end -end second_countable_topology +end has_measurable_add end borel_space