Skip to content

Commit 4bf4f50

Browse files
committed
chore(MeasureSpace): move dirac and count to new files (#6116)
1 parent 2853253 commit 4bf4f50

File tree

7 files changed

+341
-329
lines changed

7 files changed

+341
-329
lines changed

Mathlib.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2371,6 +2371,8 @@ import Mathlib.MeasureTheory.Measure.AEDisjoint
23712371
import Mathlib.MeasureTheory.Measure.AEMeasurable
23722372
import Mathlib.MeasureTheory.Measure.Complex
23732373
import Mathlib.MeasureTheory.Measure.Content
2374+
import Mathlib.MeasureTheory.Measure.Count
2375+
import Mathlib.MeasureTheory.Measure.Dirac
23742376
import Mathlib.MeasureTheory.Measure.Doubling
23752377
import Mathlib.MeasureTheory.Measure.FiniteMeasure
23762378
import Mathlib.MeasureTheory.Measure.GiryMonad

Mathlib/MeasureTheory/Integral/Lebesgue.lean

Lines changed: 1 addition & 36 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Mario Carneiro, Johannes Hölzl
66
import Mathlib.Dynamics.Ergodic.MeasurePreserving
77
import Mathlib.MeasureTheory.Function.SimpleFunc
88
import Mathlib.MeasureTheory.Measure.MutuallySingular
9+
import Mathlib.MeasureTheory.Measure.Count
910

1011
#align_import measure_theory.integral.lebesgue from "leanprover-community/mathlib"@"c14c8fcde993801fca8946b0d80131a1a81d1520"
1112

@@ -41,42 +42,6 @@ open Classical Topology BigOperators NNReal ENNReal MeasureTheory
4142

4243
namespace MeasureTheory
4344

44-
section MoveThis
45-
46-
variable {m : MeasurableSpace α} {μ ν : Measure α} {f : α → ℝ≥0∞} {s : Set α}
47-
48-
-- todo after the port: move to measure_theory/measure/measure_space
49-
theorem restrict_dirac' (hs : MeasurableSet s) [Decidable (a ∈ s)] :
50-
(Measure.dirac a).restrict s = if a ∈ s then Measure.dirac a else 0 := by
51-
ext1 t ht
52-
classical
53-
simp only [Measure.restrict_apply ht, Measure.dirac_apply' _ (ht.inter hs), Set.indicator_apply,
54-
Set.mem_inter_iff, Pi.one_apply]
55-
by_cases has : a ∈ s
56-
· simp only [has, and_true_iff, if_true]
57-
split_ifs with hat
58-
· rw [Measure.dirac_apply_of_mem hat]
59-
· simp only [Measure.dirac_apply' _ ht, Set.indicator_apply, hat, if_false]
60-
· simp only [has, and_false_iff, if_false, Measure.coe_zero, Pi.zero_apply]
61-
#align measure_theory.restrict_dirac' MeasureTheory.restrict_dirac'
62-
63-
-- todo after the port: move to measure_theory/measure/measure_space
64-
theorem restrict_dirac [MeasurableSingletonClass α] [Decidable (a ∈ s)] :
65-
(Measure.dirac a).restrict s = if a ∈ s then Measure.dirac a else 0 := by
66-
ext1 t ht
67-
classical
68-
simp only [Measure.restrict_apply ht, Measure.dirac_apply a, Set.indicator_apply,
69-
Set.mem_inter_iff, Pi.one_apply]
70-
by_cases has : a ∈ s
71-
· simp only [has, and_true_iff, if_true]
72-
split_ifs with hat
73-
· rw [Measure.dirac_apply_of_mem hat]
74-
· simp only [Measure.dirac_apply' _ ht, Set.indicator_apply, hat, if_false]
75-
· simp only [has, and_false_iff, if_false, Measure.coe_zero, Pi.zero_apply]
76-
#align measure_theory.restrict_dirac MeasureTheory.restrict_dirac
77-
78-
end MoveThis
79-
8045
-- mathport name: «expr →ₛ »
8146
local infixr:25 " →ₛ " => SimpleFunc
8247

Lines changed: 188 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,188 @@
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
Lines changed: 148 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,148 @@
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.MeasureSpace
7+
/-!
8+
# Dirac measure
9+
10+
In this file we define the Dirac measure `MeasureTheory.Measure.dirac a`
11+
and prove some basic facts about it.
12+
-/
13+
14+
open Function Set
15+
open scoped ENNReal Classical
16+
17+
noncomputable section
18+
19+
variable [MeasurableSpace α] [MeasurableSpace β] {s : Set α}
20+
21+
namespace MeasureTheory
22+
23+
namespace Measure
24+
25+
/-- The dirac measure. -/
26+
def dirac (a : α) : Measure α := (OuterMeasure.dirac a).toMeasure (by simp)
27+
#align measure_theory.measure.dirac MeasureTheory.Measure.dirac
28+
29+
instance : MeasureSpace PUnit :=
30+
⟨dirac PUnit.unit⟩
31+
32+
theorem le_dirac_apply {a} : s.indicator 1 a ≤ dirac a s :=
33+
OuterMeasure.dirac_apply a s ▸ le_toMeasure_apply _ _ _
34+
#align measure_theory.measure.le_dirac_apply MeasureTheory.Measure.le_dirac_apply
35+
36+
@[simp]
37+
theorem dirac_apply' (a : α) (hs : MeasurableSet s) : dirac a s = s.indicator 1 a :=
38+
toMeasure_apply _ _ hs
39+
#align measure_theory.measure.dirac_apply' MeasureTheory.Measure.dirac_apply'
40+
41+
@[simp]
42+
theorem dirac_apply_of_mem {a : α} (h : a ∈ s) : dirac a s = 1 := by
43+
have : ∀ t : Set α, a ∈ t → t.indicator (1 : α → ℝ≥0∞) a = 1 := fun t ht => indicator_of_mem ht 1
44+
refine' le_antisymm (this univ trivial ▸ _) (this s h ▸ le_dirac_apply)
45+
rw [← dirac_apply' a MeasurableSet.univ]
46+
exact measure_mono (subset_univ s)
47+
#align measure_theory.measure.dirac_apply_of_mem MeasureTheory.Measure.dirac_apply_of_mem
48+
49+
@[simp]
50+
theorem dirac_apply [MeasurableSingletonClass α] (a : α) (s : Set α) :
51+
dirac a s = s.indicator 1 a := by
52+
by_cases h : a ∈ s; · rw [dirac_apply_of_mem h, indicator_of_mem h, Pi.one_apply]
53+
rw [indicator_of_not_mem h, ← nonpos_iff_eq_zero]
54+
calc
55+
dirac a s ≤ dirac a {a}ᶜ := measure_mono (subset_compl_comm.1 <| singleton_subset_iff.2 h)
56+
_ = 0 := by simp [dirac_apply' _ (measurableSet_singleton _).compl]
57+
58+
#align measure_theory.measure.dirac_apply MeasureTheory.Measure.dirac_apply
59+
60+
theorem map_dirac {f : α → β} (hf : Measurable f) (a : α) : (dirac a).map f = dirac (f a) :=
61+
ext fun s hs => by simp [hs, map_apply hf hs, hf hs, indicator_apply]
62+
#align measure_theory.measure.map_dirac MeasureTheory.Measure.map_dirac
63+
64+
@[simp]
65+
theorem restrict_singleton (μ : Measure α) (a : α) : μ.restrict {a} = μ {a} • dirac a := by
66+
ext1 s hs
67+
by_cases ha : a ∈ s
68+
· have : s ∩ {a} = {a} := by simpa
69+
simp [*]
70+
· have : s ∩ {a} = ∅ := inter_singleton_eq_empty.2 ha
71+
simp [*]
72+
#align measure_theory.measure.restrict_singleton MeasureTheory.Measure.restrict_singleton
73+
74+
/-- If `f` is a map with countable codomain, then `μ.map f` is a sum of Dirac measures. -/
75+
theorem map_eq_sum [Countable β] [MeasurableSingletonClass β] (μ : Measure α) (f : α → β)
76+
(hf : Measurable f) : μ.map f = sum fun b : β => μ (f ⁻¹' {b}) • dirac b := by
77+
ext1 s hs
78+
have : ∀ y ∈ s, MeasurableSet (f ⁻¹' {y}) := fun y _ => hf (measurableSet_singleton _)
79+
simp [← tsum_measure_preimage_singleton (to_countable s) this, *,
80+
tsum_subtype s fun b => μ (f ⁻¹' {b}), ← indicator_mul_right s fun b => μ (f ⁻¹' {b})]
81+
#align measure_theory.measure.map_eq_sum MeasureTheory.Measure.map_eq_sum
82+
83+
/-- A measure on a countable type is a sum of Dirac measures. -/
84+
@[simp]
85+
theorem sum_smul_dirac [Countable α] [MeasurableSingletonClass α] (μ : Measure α) :
86+
(sum fun a => μ {a} • dirac a) = μ := by simpa using (map_eq_sum μ id measurable_id).symm
87+
#align measure_theory.measure.sum_smul_dirac MeasureTheory.Measure.sum_smul_dirac
88+
89+
/-- Given that `α` is a countable, measurable space with all singleton sets measurable,
90+
write the measure of a set `s` as the sum of the measure of `{x}` for all `x ∈ s`. -/
91+
theorem tsum_indicator_apply_singleton [Countable α] [MeasurableSingletonClass α] (μ : Measure α)
92+
(s : Set α) (hs : MeasurableSet s) : (∑' x : α, s.indicator (fun x => μ {x}) x) = μ s :=
93+
calc
94+
(∑' x : α, s.indicator (fun x => μ {x}) x) =
95+
Measure.sum (fun a => μ {a} • Measure.dirac a) s := by
96+
simp only [Measure.sum_apply _ hs, Measure.smul_apply, smul_eq_mul, Measure.dirac_apply,
97+
Set.indicator_apply, mul_ite, Pi.one_apply, mul_one, mul_zero]
98+
_ = μ s := by rw [μ.sum_smul_dirac]
99+
#align measure_theory.measure.tsum_indicator_apply_singleton MeasureTheory.Measure.tsum_indicator_apply_singleton
100+
101+
end Measure
102+
103+
open Measure
104+
105+
theorem mem_ae_dirac_iff {a : α} (hs : MeasurableSet s) : s ∈ (dirac a).ae ↔ a ∈ s := by
106+
by_cases a ∈ s <;> simp [mem_ae_iff, dirac_apply', hs.compl, indicator_apply, *]
107+
#align measure_theory.mem_ae_dirac_iff MeasureTheory.mem_ae_dirac_iff
108+
109+
theorem ae_dirac_iff {a : α} {p : α → Prop} (hp : MeasurableSet { x | p x }) :
110+
(∀ᵐ x ∂dirac a, p x) ↔ p a :=
111+
mem_ae_dirac_iff hp
112+
#align measure_theory.ae_dirac_iff MeasureTheory.ae_dirac_iff
113+
114+
@[simp]
115+
theorem ae_dirac_eq [MeasurableSingletonClass α] (a : α) : (dirac a).ae = pure a := by
116+
ext s
117+
simp [mem_ae_iff, imp_false]
118+
#align measure_theory.ae_dirac_eq MeasureTheory.ae_dirac_eq
119+
120+
theorem ae_eq_dirac' [MeasurableSingletonClass β] {a : α} {f : α → β} (hf : Measurable f) :
121+
f =ᵐ[dirac a] const α (f a) :=
122+
(ae_dirac_iff <| show MeasurableSet (f ⁻¹' {f a}) from hf <| measurableSet_singleton _).2 rfl
123+
#align measure_theory.ae_eq_dirac' MeasureTheory.ae_eq_dirac'
124+
125+
theorem ae_eq_dirac [MeasurableSingletonClass α] {a : α} (f : α → δ) :
126+
f =ᵐ[dirac a] const α (f a) := by simp [Filter.EventuallyEq]
127+
#align measure_theory.ae_eq_dirac MeasureTheory.ae_eq_dirac
128+
129+
instance Measure.dirac.isProbabilityMeasure [MeasurableSpace α] {x : α} :
130+
IsProbabilityMeasure (dirac x) :=
131+
⟨dirac_apply_of_mem <| mem_univ x⟩
132+
#align measure_theory.measure.dirac.is_probability_measure MeasureTheory.Measure.dirac.isProbabilityMeasure
133+
134+
theorem restrict_dirac' (hs : MeasurableSet s) [Decidable (a ∈ s)] :
135+
(Measure.dirac a).restrict s = if a ∈ s then Measure.dirac a else 0 := by
136+
split_ifs with has
137+
· apply restrict_eq_self_of_ae_mem
138+
rw [ae_dirac_iff] <;> assumption
139+
· rw [restrict_eq_zero, dirac_apply' _ hs, indicator_of_not_mem has]
140+
#align measure_theory.restrict_dirac' MeasureTheory.restrict_dirac'
141+
142+
theorem restrict_dirac [MeasurableSingletonClass α] [Decidable (a ∈ s)] :
143+
(Measure.dirac a).restrict s = if a ∈ s then Measure.dirac a else 0 := by
144+
split_ifs with has
145+
· apply restrict_eq_self_of_ae_mem
146+
rwa [ae_dirac_eq]
147+
· rw [restrict_eq_zero, dirac_apply, indicator_of_not_mem has]
148+
#align measure_theory.restrict_dirac MeasureTheory.restrict_dirac

0 commit comments

Comments
 (0)