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

Commit 18f0093

Browse files
committed
feat(measure_theory/measure/measure_space): add measure_Union_of_null_inter (#9307)
From #2819
1 parent 7e3256b commit 18f0093

File tree

1 file changed

+34
-0
lines changed

1 file changed

+34
-0
lines changed

src/measure_theory/measure/measure_space.lean

Lines changed: 34 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -217,6 +217,40 @@ begin
217217
exact supr_le (λ s, sum_measure_le_measure_univ (λ i hi, hs i) (λ i hi j hj hij, H i j hij))
218218
end
219219

220+
lemma measure_Union_of_null_inter [encodable ι] {f : ι → set α} (h : ∀ i, measurable_set (f i))
221+
(hn : pairwise ((λ S T, μ (S ∩ T) = 0) on f)) : μ (⋃ i, f i) = ∑' i, μ (f i) :=
222+
begin
223+
have h_null : μ (⋃ (ij : ι × ι) (hij : ij.fst ≠ ij.snd), f ij.fst ∩ f ij.snd) = 0,
224+
{ rw measure_Union_null_iff,
225+
rintro ⟨i, j⟩,
226+
by_cases hij : i = j,
227+
{ simp [hij], },
228+
{ suffices : μ (f i ∩ f j) = 0,
229+
{ simpa [hij], },
230+
apply hn i j hij, }, },
231+
have h_pair : pairwise (disjoint on
232+
(λ i, f i \ (⋃ (ij : ι × ι) (hij : ij.fst ≠ ij.snd), f ij.fst ∩ f ij.snd))),
233+
{ intros i j hij x hx,
234+
simp only [not_exists, exists_prop, mem_Union, mem_inter_eq, not_and,
235+
inf_eq_inter, ne.def, mem_diff, prod.exists] at hx,
236+
simp only [mem_empty_eq, bot_eq_empty],
237+
rcases hx with ⟨⟨hx_left_left, hx_left_right⟩, hx_right_left, hx_right_right⟩,
238+
exact hx_left_right _ _ hij hx_left_left hx_right_left, },
239+
have h_meas :
240+
∀ i, measurable_set (f i \ (⋃ (ij : ι × ι) (hij : ij.fst ≠ ij.snd), f ij.fst ∩ f ij.snd)),
241+
{ intro w,
242+
apply (h w).diff,
243+
apply measurable_set.Union,
244+
rintro ⟨i, j⟩,
245+
by_cases hij : i = j,
246+
{ simp [hij], },
247+
{ simp [hij, measurable_set.inter (h i) (h j)], }, },
248+
have : μ _ = _ := measure_Union h_pair h_meas,
249+
rw ← Union_diff at this,
250+
simp_rw measure_diff_null h_null at this,
251+
exact this,
252+
end
253+
220254
/-- Pigeonhole principle for measure spaces: if `∑' i, μ (s i) > μ univ`, then
221255
one of the intersections `s i ∩ s j` is not empty. -/
222256
lemma exists_nonempty_inter_of_measure_univ_lt_tsum_measure {m : measurable_space α} (μ : measure α)

0 commit comments

Comments
 (0)