|
| 1 | +/- |
| 2 | +Copyright (c) 2018 Johannes Hölzl. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Johannes Hölzl |
| 5 | +-/ |
| 6 | +import Mathlib.MeasureTheory.Measure.Dirac |
| 7 | + |
| 8 | +/-! |
| 9 | +# Counting measure |
| 10 | +
|
| 11 | +In this file we define the counting measure `MeasurTheory.Measure.count` |
| 12 | +as `MeasureTheory.Measure.sum MeasureTheory.Measure.dirac` |
| 13 | +and prove basic properties of this measure. |
| 14 | +-/ |
| 15 | +open Set |
| 16 | +open scoped ENNReal BigOperators Classical |
| 17 | + |
| 18 | +variable [MeasurableSpace α] [MeasurableSpace β] {s : Set α} |
| 19 | + |
| 20 | +noncomputable section |
| 21 | + |
| 22 | +namespace MeasureTheory.Measure |
| 23 | + |
| 24 | +/-- Counting measure on any measurable space. -/ |
| 25 | +def count : Measure α := |
| 26 | + sum dirac |
| 27 | +#align measure_theory.measure.count MeasureTheory.Measure.count |
| 28 | + |
| 29 | +theorem le_count_apply : ∑' _ : s, (1 : ℝ≥0∞) ≤ count s := |
| 30 | + calc |
| 31 | + (∑' _ : s, 1 : ℝ≥0∞) = ∑' i, indicator s 1 i := tsum_subtype s 1 |
| 32 | + _ ≤ ∑' i, dirac i s := (ENNReal.tsum_le_tsum fun _ => le_dirac_apply) |
| 33 | + _ ≤ count s := le_sum_apply _ _ |
| 34 | +#align measure_theory.measure.le_count_apply MeasureTheory.Measure.le_count_apply |
| 35 | + |
| 36 | +theorem count_apply (hs : MeasurableSet s) : count s = ∑' i : s, 1 := by |
| 37 | + simp only [count, sum_apply, hs, dirac_apply', ← tsum_subtype s (1 : α → ℝ≥0∞), Pi.one_apply] |
| 38 | +#align measure_theory.measure.count_apply MeasureTheory.Measure.count_apply |
| 39 | + |
| 40 | +-- @[simp] -- Porting note: simp can prove this |
| 41 | +theorem count_empty : count (∅ : Set α) = 0 := by rw [count_apply MeasurableSet.empty, tsum_empty] |
| 42 | +#align measure_theory.measure.count_empty MeasureTheory.Measure.count_empty |
| 43 | + |
| 44 | +@[simp] |
| 45 | +theorem count_apply_finset' {s : Finset α} (s_mble : MeasurableSet (s : Set α)) : |
| 46 | + count (↑s : Set α) = s.card := |
| 47 | + calc |
| 48 | + count (↑s : Set α) = ∑' i : (↑s : Set α), 1 := count_apply s_mble |
| 49 | + _ = ∑ i in s, 1 := (s.tsum_subtype 1) |
| 50 | + _ = s.card := by simp |
| 51 | + |
| 52 | +#align measure_theory.measure.count_apply_finset' MeasureTheory.Measure.count_apply_finset' |
| 53 | + |
| 54 | +@[simp] |
| 55 | +theorem count_apply_finset [MeasurableSingletonClass α] (s : Finset α) : |
| 56 | + count (↑s : Set α) = s.card := |
| 57 | + count_apply_finset' s.measurableSet |
| 58 | +#align measure_theory.measure.count_apply_finset MeasureTheory.Measure.count_apply_finset |
| 59 | + |
| 60 | +theorem count_apply_finite' {s : Set α} (s_fin : s.Finite) (s_mble : MeasurableSet s) : |
| 61 | + count s = s_fin.toFinset.card := by |
| 62 | + simp [← |
| 63 | + @count_apply_finset' _ _ s_fin.toFinset (by simpa only [Finite.coe_toFinset] using s_mble)] |
| 64 | +#align measure_theory.measure.count_apply_finite' MeasureTheory.Measure.count_apply_finite' |
| 65 | + |
| 66 | +theorem count_apply_finite [MeasurableSingletonClass α] (s : Set α) (hs : s.Finite) : |
| 67 | + count s = hs.toFinset.card := by rw [← count_apply_finset, Finite.coe_toFinset] |
| 68 | +#align measure_theory.measure.count_apply_finite MeasureTheory.Measure.count_apply_finite |
| 69 | + |
| 70 | +/-- `count` measure evaluates to infinity at infinite sets. -/ |
| 71 | +theorem count_apply_infinite (hs : s.Infinite) : count s = ∞ := by |
| 72 | + refine' top_unique (le_of_tendsto' ENNReal.tendsto_nat_nhds_top fun n => _) |
| 73 | + rcases hs.exists_subset_card_eq n with ⟨t, ht, rfl⟩ |
| 74 | + calc |
| 75 | + (t.card : ℝ≥0∞) = ∑ i in t, 1 := by simp |
| 76 | + _ = ∑' i : (t : Set α), 1 := (t.tsum_subtype 1).symm |
| 77 | + _ ≤ count (t : Set α) := le_count_apply |
| 78 | + _ ≤ count s := measure_mono ht |
| 79 | + |
| 80 | +#align measure_theory.measure.count_apply_infinite MeasureTheory.Measure.count_apply_infinite |
| 81 | + |
| 82 | +@[simp] |
| 83 | +theorem count_apply_eq_top' (s_mble : MeasurableSet s) : count s = ∞ ↔ s.Infinite := by |
| 84 | + by_cases hs : s.Finite |
| 85 | + · simp [Set.Infinite, hs, count_apply_finite' hs s_mble] |
| 86 | + · change s.Infinite at hs |
| 87 | + simp [hs, count_apply_infinite] |
| 88 | +#align measure_theory.measure.count_apply_eq_top' MeasureTheory.Measure.count_apply_eq_top' |
| 89 | + |
| 90 | +@[simp] |
| 91 | +theorem count_apply_eq_top [MeasurableSingletonClass α] : count s = ∞ ↔ s.Infinite := by |
| 92 | + by_cases hs : s.Finite |
| 93 | + · exact count_apply_eq_top' hs.measurableSet |
| 94 | + · change s.Infinite at hs |
| 95 | + simp [hs, count_apply_infinite] |
| 96 | +#align measure_theory.measure.count_apply_eq_top MeasureTheory.Measure.count_apply_eq_top |
| 97 | + |
| 98 | +@[simp] |
| 99 | +theorem count_apply_lt_top' (s_mble : MeasurableSet s) : count s < ∞ ↔ s.Finite := |
| 100 | + calc |
| 101 | + count s < ∞ ↔ count s ≠ ∞ := lt_top_iff_ne_top |
| 102 | + _ ↔ ¬s.Infinite := (not_congr (count_apply_eq_top' s_mble)) |
| 103 | + _ ↔ s.Finite := Classical.not_not |
| 104 | + |
| 105 | +#align measure_theory.measure.count_apply_lt_top' MeasureTheory.Measure.count_apply_lt_top' |
| 106 | + |
| 107 | +@[simp] |
| 108 | +theorem count_apply_lt_top [MeasurableSingletonClass α] : count s < ∞ ↔ s.Finite := |
| 109 | + calc |
| 110 | + count s < ∞ ↔ count s ≠ ∞ := lt_top_iff_ne_top |
| 111 | + _ ↔ ¬s.Infinite := (not_congr count_apply_eq_top) |
| 112 | + _ ↔ s.Finite := Classical.not_not |
| 113 | + |
| 114 | +#align measure_theory.measure.count_apply_lt_top MeasureTheory.Measure.count_apply_lt_top |
| 115 | + |
| 116 | +theorem empty_of_count_eq_zero' (s_mble : MeasurableSet s) (hsc : count s = 0) : s = ∅ := by |
| 117 | + have hs : s.Finite := by |
| 118 | + rw [← count_apply_lt_top' s_mble, hsc] |
| 119 | + exact WithTop.zero_lt_top |
| 120 | + simpa [count_apply_finite' hs s_mble] using hsc |
| 121 | +#align measure_theory.measure.empty_of_count_eq_zero' MeasureTheory.Measure.empty_of_count_eq_zero' |
| 122 | + |
| 123 | +theorem empty_of_count_eq_zero [MeasurableSingletonClass α] (hsc : count s = 0) : s = ∅ := by |
| 124 | + have hs : s.Finite := by |
| 125 | + rw [← count_apply_lt_top, hsc] |
| 126 | + exact WithTop.zero_lt_top |
| 127 | + simpa [count_apply_finite _ hs] using hsc |
| 128 | +#align measure_theory.measure.empty_of_count_eq_zero MeasureTheory.Measure.empty_of_count_eq_zero |
| 129 | + |
| 130 | +@[simp] |
| 131 | +theorem count_eq_zero_iff' (s_mble : MeasurableSet s) : count s = 0 ↔ s = ∅ := |
| 132 | + ⟨empty_of_count_eq_zero' s_mble, fun h => h.symm ▸ count_empty⟩ |
| 133 | +#align measure_theory.measure.count_eq_zero_iff' MeasureTheory.Measure.count_eq_zero_iff' |
| 134 | + |
| 135 | +@[simp] |
| 136 | +theorem count_eq_zero_iff [MeasurableSingletonClass α] : count s = 0 ↔ s = ∅ := |
| 137 | + ⟨empty_of_count_eq_zero, fun h => h.symm ▸ count_empty⟩ |
| 138 | +#align measure_theory.measure.count_eq_zero_iff MeasureTheory.Measure.count_eq_zero_iff |
| 139 | + |
| 140 | +theorem count_ne_zero' (hs' : s.Nonempty) (s_mble : MeasurableSet s) : count s ≠ 0 := by |
| 141 | + rw [Ne.def, count_eq_zero_iff' s_mble] |
| 142 | + exact hs'.ne_empty |
| 143 | +#align measure_theory.measure.count_ne_zero' MeasureTheory.Measure.count_ne_zero' |
| 144 | + |
| 145 | +theorem count_ne_zero [MeasurableSingletonClass α] (hs' : s.Nonempty) : count s ≠ 0 := by |
| 146 | + rw [Ne.def, count_eq_zero_iff] |
| 147 | + exact hs'.ne_empty |
| 148 | +#align measure_theory.measure.count_ne_zero MeasureTheory.Measure.count_ne_zero |
| 149 | + |
| 150 | +@[simp] |
| 151 | +theorem count_singleton' {a : α} (ha : MeasurableSet ({a} : Set α)) : count ({a} : Set α) = 1 := by |
| 152 | + rw [count_apply_finite' (Set.finite_singleton a) ha, Set.Finite.toFinset] |
| 153 | + simp [@toFinset_card _ _ (Set.finite_singleton a).fintype, |
| 154 | + @Fintype.card_unique _ _ (Set.finite_singleton a).fintype] |
| 155 | +#align measure_theory.measure.count_singleton' MeasureTheory.Measure.count_singleton' |
| 156 | + |
| 157 | +-- @[simp] -- Porting note: simp can prove this |
| 158 | +theorem count_singleton [MeasurableSingletonClass α] (a : α) : count ({a} : Set α) = 1 := |
| 159 | + count_singleton' (measurableSet_singleton a) |
| 160 | +#align measure_theory.measure.count_singleton MeasureTheory.Measure.count_singleton |
| 161 | + |
| 162 | +theorem count_injective_image' {f : β → α} (hf : Function.Injective f) {s : Set β} |
| 163 | + (s_mble : MeasurableSet s) (fs_mble : MeasurableSet (f '' s)) : count (f '' s) = count s := by |
| 164 | + by_cases hs : s.Finite |
| 165 | + · lift s to Finset β using hs |
| 166 | + rw [← Finset.coe_image, count_apply_finset' _, count_apply_finset' s_mble, |
| 167 | + s.card_image_of_injective hf] |
| 168 | + simpa only [Finset.coe_image] using fs_mble |
| 169 | + · rw [count_apply_infinite hs] |
| 170 | + rw [← finite_image_iff <| hf.injOn _] at hs |
| 171 | + rw [count_apply_infinite hs] |
| 172 | +#align measure_theory.measure.count_injective_image' MeasureTheory.Measure.count_injective_image' |
| 173 | + |
| 174 | +theorem count_injective_image [MeasurableSingletonClass α] [MeasurableSingletonClass β] {f : β → α} |
| 175 | + (hf : Function.Injective f) (s : Set β) : count (f '' s) = count s := by |
| 176 | + by_cases hs : s.Finite |
| 177 | + · exact count_injective_image' hf hs.measurableSet (Finite.image f hs).measurableSet |
| 178 | + rw [count_apply_infinite hs] |
| 179 | + rw [← finite_image_iff <| hf.injOn _] at hs |
| 180 | + rw [count_apply_infinite hs] |
| 181 | +#align measure_theory.measure.count_injective_image MeasureTheory.Measure.count_injective_image |
| 182 | + |
| 183 | +instance count.isFiniteMeasure [Finite α] [MeasurableSpace α] : |
| 184 | + IsFiniteMeasure (Measure.count : Measure α) := |
| 185 | + ⟨by |
| 186 | + cases nonempty_fintype α |
| 187 | + simpa [Measure.count_apply, tsum_fintype] using (ENNReal.nat_ne_top _).lt_top⟩ |
| 188 | +#align measure_theory.measure.count.is_finite_measure MeasureTheory.Measure.count.isFiniteMeasure |
0 commit comments