Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat: add MeasureTheory.MeasurePreserving.measure_symmDiff_preimage_iterate_le #6175

Closed
wants to merge 3 commits into from
Closed
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
6 changes: 3 additions & 3 deletions Mathlib/Algebra/Ring/BooleanRing.lean
Original file line number Diff line number Diff line change
Expand Up @@ -306,7 +306,7 @@ theorem ofBoolAlg_sdiff (a b : AsBoolAlg α) : ofBoolAlg (a \ b) = ofBoolAlg a *
rfl
#align of_boolalg_sdiff ofBoolAlg_sdiff

private theorem of_boolalg_symm_diff_aux (a b : α) : (a + b + a * b) * (1 + a * b) = a + b :=
private theorem of_boolalg_symmDiff_aux (a b : α) : (a + b + a * b) * (1 + a * b) = a + b :=
calc
(a + b + a * b) * (1 + a * b) = a + b + (a * b + a * b * (a * b)) + (a * (b * b) + a * a * b) :=
by ring
Expand All @@ -315,7 +315,7 @@ private theorem of_boolalg_symm_diff_aux (a b : α) : (a + b + a * b) * (1 + a *
@[simp]
theorem ofBoolAlg_symmDiff (a b : AsBoolAlg α) : ofBoolAlg (a ∆ b) = ofBoolAlg a + ofBoolAlg b := by
rw [symmDiff_eq_sup_sdiff_inf]
exact of_boolalg_symm_diff_aux _ _
exact of_boolalg_symmDiff_aux _ _
#align of_boolalg_symm_diff ofBoolAlg_symmDiff

@[simp]
Expand Down Expand Up @@ -553,7 +553,7 @@ protected def BoundedLatticeHom.asBoolRing (f : BoundedLatticeHom α β) :
toFun := toBoolRing ∘ f ∘ ofBoolRing
map_zero' := f.map_bot'
map_one' := f.map_top'
map_add' := map_symm_diff' f
map_add' := map_symmDiff' f
map_mul' := f.map_inf'
#align bounded_lattice_hom.as_boolring BoundedLatticeHom.asBoolRing

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Data/Finset/Image.lean
Original file line number Diff line number Diff line change
Expand Up @@ -536,7 +536,7 @@ theorem image_symmDiff [DecidableEq α] {f : α → β} (s t : Finset α) (hf :
(s ∆ t).image f = s.image f ∆ t.image f :=
coe_injective <| by
push_cast
exact Set.image_symm_diff hf _ _
exact Set.image_symmDiff hf _ _
#align finset.image_symm_diff Finset.image_symmDiff

@[simp]
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Data/Finset/Pointwise.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2093,9 +2093,9 @@ theorem smul_finset_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t
image_sdiff _ _ <| MulAction.injective₀ ha
#align finset.smul_finset_sdiff₀ Finset.smul_finset_sdiff₀

theorem smul_finset_symm_diff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) :=
theorem smul_finset_symmDiff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) :=
image_symmDiff _ _ <| MulAction.injective₀ ha
#align finset.smul_finset_symm_diff₀ Finset.smul_finset_symm_diff
#align finset.smul_finset_symm_diff₀ Finset.smul_finset_symmDiff

theorem smul_univ₀ [Fintype β] {s : Finset α} (hs : ¬s ⊆ 0) : s • (univ : Finset β) = univ :=
coe_injective <| by
Expand Down
5 changes: 5 additions & 0 deletions Mathlib/Data/Real/ENNReal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -944,6 +944,9 @@ theorem le_of_top_imp_top_of_toNNReal_le {a b : ℝ≥0∞} (h : a = ⊤ → b =
simpa using h_nnreal
#align ennreal.le_of_top_imp_top_of_to_nnreal_le ENNReal.le_of_top_imp_top_of_toNNReal_le

@[simp]
theorem abs_toReal {x : ℝ≥0∞} : |x.toReal| = x.toReal := by cases x <;> simp

end Order

section CompleteLattice
Expand Down Expand Up @@ -2272,6 +2275,8 @@ theorem toReal_mul : (a * b).toReal = a.toReal * b.toReal :=
toRealHom.map_mul a b
#align ennreal.to_real_mul ENNReal.toReal_mul

theorem toReal_nsmul (a : ℝ≥0∞) (n : ℕ) : (n • a).toReal = n • a.toReal := by simp

@[simp]
theorem toReal_pow (a : ℝ≥0∞) (n : ℕ) : (a ^ n).toReal = a.toReal ^ n :=
toRealHom.map_pow a n
Expand Down
12 changes: 8 additions & 4 deletions Mathlib/Data/Set/Image.lean
Original file line number Diff line number Diff line change
Expand Up @@ -102,6 +102,10 @@ theorem preimage_diff (f : α → β) (s t : Set β) : f ⁻¹' (s \ t) = f ⁻
rfl
#align set.preimage_diff Set.preimage_diff

@[simp]
lemma preimage_symmDiff {f : α → β} (s t : Set β) : f ⁻¹' (s ∆ t) = (f ⁻¹' s) ∆ (f ⁻¹' t) :=
rfl

@[simp]
theorem preimage_ite (f : α → β) (s t₁ t₂ : Set β) :
f ⁻¹' s.ite t₁ t₂ = (f ⁻¹' s).ite (f ⁻¹' t₁) (f ⁻¹' t₂) :=
Expand Down Expand Up @@ -432,20 +436,20 @@ theorem subset_image_diff (f : α → β) (s t : Set α) : f '' s \ f '' t ⊆ f
exact image_subset f (subset_union_right t s)
#align set.subset_image_diff Set.subset_image_diff

theorem subset_image_symm_diff : (f '' s) ∆ (f '' t) ⊆ f '' s ∆ t :=
theorem subset_image_symmDiff : (f '' s) ∆ (f '' t) ⊆ f '' s ∆ t :=
(union_subset_union (subset_image_diff _ _ _) <| subset_image_diff _ _ _).trans
(superset_of_eq (image_union _ _ _))
#align set.subset_image_symm_diff Set.subset_image_symm_diff
#align set.subset_image_symm_diff Set.subset_image_symmDiff

theorem image_diff {f : α → β} (hf : Injective f) (s t : Set α) : f '' (s \ t) = f '' s \ f '' t :=
Subset.antisymm
(Subset.trans (image_inter_subset _ _ _) <| inter_subset_inter_right _ <| image_compl_subset hf)
(subset_image_diff f s t)
#align set.image_diff Set.image_diff

theorem image_symm_diff (hf : Injective f) (s t : Set α) : f '' s ∆ t = (f '' s) ∆ (f '' t) := by
theorem image_symmDiff (hf : Injective f) (s t : Set α) : f '' s ∆ t = (f '' s) ∆ (f '' t) := by
simp_rw [Set.symmDiff_def, image_union, image_diff hf]
#align set.image_symm_diff Set.image_symm_diff
#align set.image_symm_diff Set.image_symmDiff

theorem Nonempty.image (f : α → β) {s : Set α} : s.Nonempty → (f '' s).Nonempty
| ⟨x, hx⟩ => ⟨f x, mem_image_of_mem f hx⟩
Expand Down
14 changes: 7 additions & 7 deletions Mathlib/Data/Set/Pointwise/SMul.lean
Original file line number Diff line number Diff line change
Expand Up @@ -941,10 +941,10 @@ theorem smul_set_sdiff : a • (s \ t) = a • s \ a • t :=
#align set.vadd_set_sdiff Set.vadd_set_sdiff

@[to_additive]
theorem smul_set_symm_diff : a • s ∆ t = (a • s) ∆ (a • t) :=
image_symm_diff (MulAction.injective a) _ _
#align set.smul_set_symm_diff Set.smul_set_symm_diff
#align set.vadd_set_symm_diff Set.vadd_set_symm_diff
theorem smul_set_symmDiff : a • s ∆ t = (a • s) ∆ (a • t) :=
image_symmDiff (MulAction.injective a) _ _
#align set.smul_set_symm_diff Set.smul_set_symmDiff
#align set.vadd_set_symm_diff Set.vadd_set_symmDiff

@[to_additive (attr := simp)]
theorem smul_set_univ : a • (univ : Set β) = univ :=
Expand Down Expand Up @@ -1053,9 +1053,9 @@ theorem smul_set_sdiff₀ (ha : a ≠ 0) : a • (s \ t) = a • s \ a • t :=
image_diff (MulAction.injective₀ ha) _ _
#align set.smul_set_sdiff₀ Set.smul_set_sdiff₀

theorem smul_set_symm_diff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) :=
image_symm_diff (MulAction.injective₀ ha) _ _
#align set.smul_set_symm_diff₀ Set.smul_set_symm_diff
theorem smul_set_symmDiff₀ (ha : a ≠ 0) : a • s ∆ t = (a • s) ∆ (a • t) :=
image_symmDiff (MulAction.injective₀ ha) _ _
#align set.smul_set_symm_diff₀ Set.smul_set_symmDiff

theorem smul_set_univ₀ (ha : a ≠ 0) : a • (univ : Set β) = univ :=
image_univ_of_surjective <| MulAction.surjective₀ ha
Expand Down
9 changes: 9 additions & 0 deletions Mathlib/Dynamics/Ergodic/MeasurePreserving.lean
Original file line number Diff line number Diff line change
Expand Up @@ -135,6 +135,15 @@ protected theorem iterate {f : α → α} (hf : MeasurePreserving f μa μa) :

variable {μ : Measure α} {f : α → α} {s : Set α}

lemma measure_symmDiff_preimage_iterate_le
(hf : MeasurePreserving f μ μ) (hs : MeasurableSet s) (n : ℕ) :
μ (s ∆ (f^[n] ⁻¹' s)) ≤ n • μ (s ∆ (f ⁻¹' s)) := by
induction' n with n ih; simp
simp only [add_smul, one_smul, ← n.add_one]
refine' le_trans (measure_symmDiff_le s (f^[n] ⁻¹' s) (f^[n+1] ⁻¹' s)) (add_le_add ih _)
replace hs : MeasurableSet (s ∆ (f ⁻¹' s)) := hs.symmDiff $ hf.measurable hs
rw [iterate_succ', preimage_comp, ← preimage_symmDiff, (hf.iterate n).measure_preimage hs]

/-- If `μ univ < n * μ s` and `f` is a map preserving measure `μ`,
then for some `x ∈ s` and `0 < m < n`, `f^[m] x ∈ s`. -/
theorem exists_mem_image_mem_of_volume_lt_mul_volume (hf : MeasurePreserving f μ μ)
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/MeasureTheory/Measure/Lebesgue/EqHaar.lean
Original file line number Diff line number Diff line change
Expand Up @@ -378,7 +378,7 @@ theorem NullMeasurableSet.const_smul (hs : NullMeasurableSet s μ) (r : ℝ) :
obtain ⟨t, ht, hst⟩ := hs
refine' ⟨_, ht.const_smul_of_ne_zero hr, _⟩
rw [← measure_symmDiff_eq_zero_iff] at hst ⊢
rw [← smul_set_symm_diff₀ hr, addHaar_smul μ, hst, mul_zero]
rw [← smul_set_symmDiff₀ hr, addHaar_smul μ, hst, mul_zero]
#align measure_theory.measure.null_measurable_set.const_smul MeasureTheory.Measure.NullMeasurableSet.const_smul

variable (μ)
Expand Down
27 changes: 27 additions & 0 deletions Mathlib/MeasureTheory/Measure/MeasureSpace.lean
Original file line number Diff line number Diff line change
Expand Up @@ -149,6 +149,14 @@ theorem measure_union_add_inter' (hs : MeasurableSet s) (t : Set α) :
rw [union_comm, inter_comm, measure_union_add_inter t hs, add_comm]
#align measure_theory.measure_union_add_inter' MeasureTheory.measure_union_add_inter'

lemma measure_symmDiff_eq (hs : MeasurableSet s) (ht : MeasurableSet t) :
μ (s ∆ t) = μ (s \ t) + μ (t \ s) := by
simpa only [symmDiff_def, sup_eq_union] using measure_union disjoint_sdiff_sdiff (ht.diff hs)

lemma measure_symmDiff_le (s t u : Set α) :
μ (s ∆ u) ≤ μ (s ∆ t) + μ (t ∆ u) :=
le_trans (μ.mono $ symmDiff_triangle s t u) (measure_union_le (s ∆ t) (t ∆ u))

theorem measure_add_measure_compl (h : MeasurableSet s) : μ s + μ sᶜ = μ univ :=
measure_add_measure_compl₀ h.nullMeasurableSet
#align measure_theory.measure_add_measure_compl MeasureTheory.measure_add_measure_compl
Expand Down Expand Up @@ -2906,6 +2914,25 @@ theorem ae_mem_iff_measure_eq [IsFiniteMeasure μ] {s : Set α} (hs : NullMeasur
ae_iff_measure_eq hs
#align measure_theory.ae_mem_iff_measure_eq MeasureTheory.ae_mem_iff_measure_eq

theorem abs_toReal_measure_sub_le_measure_symmDiff'
(hs : MeasurableSet s) (ht : MeasurableSet t) (hs' : μ s ≠ ∞) (ht' : μ t ≠ ∞) :
|(μ s).toReal - (μ t).toReal| ≤ (μ (s ∆ t)).toReal := by
have hst : μ (s \ t) ≠ ∞ := (measure_lt_top_of_subset (diff_subset s t) hs').ne
have hts : μ (t \ s) ≠ ∞ := (measure_lt_top_of_subset (diff_subset t s) ht').ne
suffices : (μ s).toReal - (μ t).toReal = (μ (s \ t)).toReal - (μ (t \ s)).toReal
· rw [this, measure_symmDiff_eq hs ht, ENNReal.toReal_add hst hts]
convert abs_sub (μ (s \ t)).toReal (μ (t \ s)).toReal <;> simp
rw [measure_diff' s ht ht', measure_diff' t hs hs',
ENNReal.toReal_sub_of_le measure_le_measure_union_right (measure_union_ne_top hs' ht'),
ENNReal.toReal_sub_of_le measure_le_measure_union_right (measure_union_ne_top ht' hs'),
union_comm t s]
abel

theorem abs_toReal_measure_sub_le_measure_symmDiff [IsFiniteMeasure μ]
(hs : MeasurableSet s) (ht : MeasurableSet t) :
|(μ s).toReal - (μ t).toReal| ≤ (μ (s ∆ t)).toReal :=
abs_toReal_measure_sub_le_measure_symmDiff' hs ht (measure_ne_top μ s) (measure_ne_top μ t)

end IsFiniteMeasure

section IsProbabilityMeasure
Expand Down
13 changes: 11 additions & 2 deletions Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean
Original file line number Diff line number Diff line change
Expand Up @@ -201,6 +201,12 @@ theorem measure_mono_top (h : s₁ ⊆ s₂) (h₁ : μ s₁ = ∞) : μ s₂ =
top_unique <| h₁ ▸ measure_mono h
#align measure_theory.measure_mono_top MeasureTheory.measure_mono_top

@[simp, mono]
theorem measure_le_measure_union_left : μ s ≤ μ (s ∪ t) := μ.mono $ subset_union_left s t

@[simp, mono]
theorem measure_le_measure_union_right : μ t ≤ μ (s ∪ t) := μ.mono $ subset_union_right s t

/-- For every set there exists a measurable superset of the same measure. -/
theorem exists_measurable_superset (μ : Measure α) (s : Set α) :
∃ t, s ⊆ t ∧ MeasurableSet t ∧ μ t = μ s := by
Expand Down Expand Up @@ -329,12 +335,15 @@ theorem exists_measure_pos_of_not_measure_iUnion_null [Countable β] {s : β →
exact measure_iUnion_null fun n => nonpos_iff_eq_zero.1 (hs n)
#align measure_theory.exists_measure_pos_of_not_measure_Union_null MeasureTheory.exists_measure_pos_of_not_measure_iUnion_null

theorem measure_lt_top_of_subset (hst : t ⊆ s) (hs : μ s ≠ ∞) : μ t < ∞ :=
lt_of_le_of_lt (μ.mono hst) hs.lt_top

theorem measure_inter_lt_top_of_left_ne_top (hs_finite : μ s ≠ ∞) : μ (s ∩ t) < ∞ :=
(measure_mono (Set.inter_subset_left s t)).trans_lt hs_finite.lt_top
measure_lt_top_of_subset (inter_subset_left s t) hs_finite
#align measure_theory.measure_inter_lt_top_of_left_ne_top MeasureTheory.measure_inter_lt_top_of_left_ne_top

theorem measure_inter_lt_top_of_right_ne_top (ht_finite : μ t ≠ ∞) : μ (s ∩ t) < ∞ :=
inter_comm t s ▸ measure_inter_lt_top_of_left_ne_top ht_finite
measure_lt_top_of_subset (inter_subset_right s t) ht_finite
#align measure_theory.measure_inter_lt_top_of_right_ne_top MeasureTheory.measure_inter_lt_top_of_right_ne_top

theorem measure_inter_null_of_null_right (S : Set α) {T : Set α} (h : μ T = 0) : μ (S ∩ T) = 0 :=
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Order/Hom/Lattice.lean
Original file line number Diff line number Diff line change
Expand Up @@ -292,9 +292,9 @@ theorem map_sdiff' (a b : α) : f (a \ b) = f a \ f b := by
#align map_sdiff' map_sdiff'

/-- Special case of `map_symmDiff` for boolean algebras. -/
theorem map_symm_diff' (a b : α) : f (a ∆ b) = f a ∆ f b := by
theorem map_symmDiff' (a b : α) : f (a ∆ b) = f a ∆ f b := by
rw [symmDiff, symmDiff, map_sup, map_sdiff', map_sdiff']
#align map_symm_diff' map_symm_diff'
#align map_symm_diff' map_symmDiff'

end BooleanAlgebra

Expand Down
Loading