|
| 1 | +/- |
| 2 | +Copyright (c) 2025 Lua Viana Reis. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Lua Viana Reis |
| 5 | +-/ |
| 6 | +module |
| 7 | + |
| 8 | +public import Mathlib.MeasureTheory.Integral.Bochner.Basic |
| 9 | +import Mathlib.MeasureTheory.Integral.Bochner.Set |
| 10 | +import Mathlib.Algebra.Order.Group.PartialSups |
| 11 | +import Mathlib.Algebra.Order.Ring.Star |
| 12 | +import Mathlib.Analysis.InnerProductSpace.Basic |
| 13 | +import Mathlib.Analysis.RCLike.Lemmas |
| 14 | +import Mathlib.Data.Real.StarOrdered |
| 15 | +import Mathlib.Dynamics.BirkhoffSum.Average |
| 16 | + |
| 17 | +/-! |
| 18 | +# Maximal ergodic theorem. |
| 19 | +
|
| 20 | +We define the set `birkhoffAverageSupSet f g a` of points `x` where the supremum of |
| 21 | +`birkhoffAverage ℝ f g n x` for varying `n` is greater than `a`. Then, we show its |
| 22 | +measure multiplied by `a` is bounded above by the integral of `g` on this set. In particular, |
| 23 | +for a positive `g`, this integral is bounded above by the `L1` norm of `g`, so for `g` |
| 24 | +taking values in a `NormedAddCommGroup`, we get an inequality involving the norm. |
| 25 | +
|
| 26 | +## Main results |
| 27 | +
|
| 28 | +* `meas_birkhoffAverageSupSet_smul_const_le_integral`: the measure of `birkhoffAverageSupSet f g a` |
| 29 | + multiplied by `a` is bounded above by the integral of `g` on the same set. |
| 30 | +* `meas_birkhoffAverageSupSet_smul_const_le_norm`: the measure of `birkhoffAverageSupSet f ‖g‖ a` |
| 31 | + multiplied by `a` is bounded above by the `L1` norm of `g`. |
| 32 | +-/ |
| 33 | + |
| 34 | +variable {α : Type*} {f : α → α} |
| 35 | + |
| 36 | +open MeasureTheory Measure MeasurableSpace Filter Topology |
| 37 | + |
| 38 | +section BirkhoffMax |
| 39 | + |
| 40 | +/-- The maximum of `birkhoffSum f g i` for `i` ranging from `0` to `n`. -/ |
| 41 | +def birkhoffMax (f : α → α) (g : α → ℝ) : ℕ →o (α → ℝ) := |
| 42 | + partialSups (birkhoffSum f g) |
| 43 | + |
| 44 | +lemma birkhoffMax_nonneg {g n} : |
| 45 | + 0 ≤ birkhoffMax f g n := by |
| 46 | + apply (le_partialSups_of_le _ n.zero_le).trans' |
| 47 | + rfl |
| 48 | + |
| 49 | +lemma birkhoffMax_succ {g n} : |
| 50 | + birkhoffMax f g (n + 1) = 0 ⊔ (g + birkhoffMax f g n ∘ f) := by |
| 51 | + have : birkhoffSum f g ∘ Nat.succ = fun k ↦ g + birkhoffSum f g k ∘ f := by |
| 52 | + funext |
| 53 | + exact birkhoffSum_succ' .. |
| 54 | + erw [partialSups_succ', this, partialSups_const_add, birkhoffSum_zero'] |
| 55 | + funext |
| 56 | + simp [birkhoffMax, partialSups] |
| 57 | + |
| 58 | +lemma birkhoffMax_succ' {g n x} (hpos : 0 < birkhoffMax f g (n + 1) x) : |
| 59 | + birkhoffMax f g (n + 1) x = g x + birkhoffMax f g n (f x) := by |
| 60 | + erw [birkhoffMax_succ, lt_sup_iff] at hpos |
| 61 | + cases hpos with |
| 62 | + | inl h => absurd h; exact lt_irrefl 0 |
| 63 | + | inr h => |
| 64 | + erw [birkhoffMax_succ, Pi.sup_apply, sup_of_le_right h.le] |
| 65 | + rfl |
| 66 | + |
| 67 | +lemma birkhoffMax_comp_le_succ {g n} : |
| 68 | + birkhoffMax f g n ≤ birkhoffMax f g (n + 1) := by |
| 69 | + gcongr |
| 70 | + exact n.le_succ |
| 71 | + |
| 72 | +lemma birkhoffMax_le_birkhoffMax {g n x} (hpos : 0 < birkhoffMax f g n x) : |
| 73 | + birkhoffMax f g n x ≤ g x + birkhoffMax f g n (f x) := by |
| 74 | + match n with |
| 75 | + | 0 => absurd hpos; exact lt_irrefl 0 |
| 76 | + | n + 1 => |
| 77 | + apply le_of_eq_of_le (birkhoffMax_succ' hpos) |
| 78 | + apply add_le_add_right |
| 79 | + exact birkhoffMax_comp_le_succ (f x) |
| 80 | + |
| 81 | +lemma birkhoffMax_pos_of_mem_support {g n x} |
| 82 | + (hx : x ∈ (birkhoffMax f g n).support) : 0 < birkhoffMax f g n x := by |
| 83 | + apply lt_or_gt_of_ne at hx |
| 84 | + cases hx with |
| 85 | + | inl h => |
| 86 | + absurd h; exact (birkhoffMax_nonneg x).not_gt |
| 87 | + | inr h => exact h |
| 88 | + |
| 89 | +-- TODO: move elsewhere |
| 90 | +@[measurability] |
| 91 | +lemma birkhoffSum_measurable [MeasurableSpace α] {f : α → α} (hf : Measurable f) {g : α → ℝ} |
| 92 | + (hg : Measurable g) {n} : Measurable (birkhoffSum f g n) := by |
| 93 | + apply Finset.measurable_sum |
| 94 | + measurability |
| 95 | + |
| 96 | +-- TODO: move elsewhere |
| 97 | +@[measurability] |
| 98 | +lemma birkhoffMax_measurable [MeasurableSpace α] (hf : Measurable f) {g : α → ℝ} |
| 99 | + (hg : Measurable g) {n} : Measurable (birkhoffMax f g n) := by |
| 100 | + unfold birkhoffMax |
| 101 | + induction n <;> measurability |
| 102 | + |
| 103 | +section MeasurePreserving |
| 104 | + |
| 105 | +variable {f : α → α} [MeasurableSpace α] (μ : Measure α := by volume_tac) {g : α → ℝ} {n} |
| 106 | + (hf : MeasurePreserving f μ μ) (hg : Integrable g μ) |
| 107 | + |
| 108 | +include hf |
| 109 | + |
| 110 | +-- todo: move elsewhere |
| 111 | +@[measurability] |
| 112 | +lemma birkhoffSum_aestronglyMeasurable (hg : AEStronglyMeasurable g μ) : |
| 113 | + AEStronglyMeasurable (birkhoffSum f g n) μ := by |
| 114 | + apply Finset.aestronglyMeasurable_fun_sum |
| 115 | + exact fun i _ => hg.comp_measurePreserving (hf.iterate i) |
| 116 | + |
| 117 | +-- todo: move elsewhere |
| 118 | +@[measurability] |
| 119 | +lemma birkhoffMax_aestronglyMeasurable (hg : AEStronglyMeasurable g μ) : |
| 120 | + AEStronglyMeasurable (birkhoffMax f g n) μ := by |
| 121 | + unfold birkhoffMax |
| 122 | + induction n <;> measurability |
| 123 | + |
| 124 | +include hg |
| 125 | + |
| 126 | +-- todo: move elsewhere |
| 127 | +lemma birkhoffSum_integrable : Integrable (birkhoffSum f g n) μ := |
| 128 | + integrable_finset_sum _ fun _ _ ↦ (hf.iterate _).integrable_comp_of_integrable hg |
| 129 | + |
| 130 | +-- todo: move elsewhere |
| 131 | +lemma birkhoffMax_integrable : Integrable (birkhoffMax f g n) μ := by |
| 132 | + unfold birkhoffMax |
| 133 | + induction n with |
| 134 | + | zero => exact integrable_zero .. |
| 135 | + | succ n hn => simpa using Integrable.sup hn (birkhoffSum_integrable μ hf hg) |
| 136 | + |
| 137 | +lemma birkhoffMax_integral_le : |
| 138 | + ∫ x, birkhoffMax f g n x ∂μ ≤ |
| 139 | + ∫ x in (birkhoffMax f g n).support, g x ∂μ + |
| 140 | + ∫ x in (birkhoffMax f g n).support, birkhoffMax f g n (f x) ∂μ := by |
| 141 | + have := hf.integrable_comp_of_integrable (birkhoffMax_integrable μ hf hg (n := n)) |
| 142 | + rw [←integral_add hg.restrict, ←setIntegral_support] |
| 143 | + · apply setIntegral_mono_on₀ |
| 144 | + · exact (birkhoffMax_integrable μ hf hg).restrict |
| 145 | + · exact .add hg.restrict this.restrict |
| 146 | + · exact AEStronglyMeasurable.nullMeasurableSet_support (by measurability) |
| 147 | + · intro x hx |
| 148 | + exact birkhoffMax_le_birkhoffMax (birkhoffMax_pos_of_mem_support hx) |
| 149 | + · exact this.restrict |
| 150 | + |
| 151 | +lemma setIntegral_nonneg_on_birkhoffMax_support : |
| 152 | + 0 ≤ ∫ x in (birkhoffMax f g n).support, g x ∂μ := by |
| 153 | + have hg₁ : AEStronglyMeasurable (birkhoffMax f g n) μ := by measurability |
| 154 | + have hg₂ : Integrable (birkhoffMax f g n) μ := birkhoffMax_integrable μ hf hg |
| 155 | + have hg₃ : Integrable (birkhoffMax f g n ∘ f) μ := hf.integrable_comp_of_integrable hg₂ |
| 156 | + calc |
| 157 | + 0 ≤ ∫ x in (birkhoffMax f g n).supportᶜ, birkhoffMax f g n (f x) ∂μ := by |
| 158 | + exact integral_nonneg (fun x => birkhoffMax_nonneg (f x)) |
| 159 | + _ = ∫ x, birkhoffMax f g n (f x) ∂μ - |
| 160 | + ∫ x in (birkhoffMax f g n).support, birkhoffMax f g n (f x) ∂μ := by |
| 161 | + exact setIntegral_compl₀ hg₁.nullMeasurableSet_support hg₃ |
| 162 | + _ = ∫ x, birkhoffMax f g n x ∂μ - |
| 163 | + ∫ x in (birkhoffMax f g n).support, birkhoffMax f g n (f x) ∂μ := by |
| 164 | + rw [←integral_map hf.aemeasurable (hf.map_eq.symm ▸ hg₁), hf.map_eq] |
| 165 | + _ ≤ ∫ x in (birkhoffMax f g n).support, g x ∂μ := by |
| 166 | + grw [birkhoffMax_integral_le μ hf hg] |
| 167 | + grind |
| 168 | + |
| 169 | +end MeasurePreserving |
| 170 | + |
| 171 | +end BirkhoffMax |
| 172 | + |
| 173 | +def birkhoffSupSet (f : α → α) (g : α → ℝ) : Set α := |
| 174 | + {x | ∃ n : ℕ, 0 < birkhoffSum f g n x} |
| 175 | + |
| 176 | +lemma birkhoffSupSet_eq_iSup_birkhoffMax_support {f : α → α} {g : α → ℝ} : |
| 177 | + birkhoffSupSet f g = ⋃ n : ℕ, (birkhoffMax f g n).support := by |
| 178 | + ext x |
| 179 | + simp only [birkhoffSupSet, Set.mem_setOf_eq, Set.mem_iUnion, Function.mem_support] |
| 180 | + constructor |
| 181 | + · refine fun ⟨n, hn⟩ => ⟨n, ?_⟩ |
| 182 | + apply ne_of_gt (hn.trans_le _) |
| 183 | + exact le_partialSups (birkhoffSum f g) _ _ |
| 184 | + · rintro ⟨n, hn⟩ |
| 185 | + apply lt_or_gt_of_ne at hn |
| 186 | + cases hn with |
| 187 | + | inl h => absurd h; exact not_lt_of_ge (birkhoffMax_nonneg x) |
| 188 | + | inr h => |
| 189 | + rw [birkhoffMax, Pi.partialSups_apply] at h |
| 190 | + rcases exists_partialSups_eq (birkhoffSum f g · x) n with ⟨m, _, hm₂⟩ |
| 191 | + exact ⟨m, hm₂ ▸ h⟩ |
| 192 | + |
| 193 | +public def birkhoffAverageSupSet (f : α → α) (g : α → ℝ) (a : ℝ) : Set α := |
| 194 | + {x | ∃ n : ℕ, a < birkhoffAverage ℝ f g n x} |
| 195 | + |
| 196 | +theorem birkhoffAverage_iff_birkhoffSum {f : α → α} {x n g} {a : ℝ} (hn : 0 < n) : |
| 197 | + a < birkhoffAverage ℝ f g n x ↔ 0 < birkhoffSum f (g - fun _ ↦ a) n x := by |
| 198 | + nth_rw 2 [←smul_lt_smul_iff_of_pos_left (a := (↑n : ℝ)⁻¹) (by positivity)] |
| 199 | + rw [smul_zero, ←birkhoffAverage, birkhoffAverage_sub] |
| 200 | + simp only [Pi.sub_apply, sub_pos] |
| 201 | + nth_rw 2 [birkhoffAverage_of_comp_eq rfl hn.ne'] |
| 202 | + |
| 203 | +theorem birkhoffAverageSupSet_eq_birkhoffSupSet {f : α → α} {g a} (ha : 0 < a) : |
| 204 | + birkhoffAverageSupSet f g a = birkhoffSupSet f (g - fun _ ↦ a) := by |
| 205 | + unfold birkhoffAverageSupSet birkhoffSupSet |
| 206 | + have {n x} : a < birkhoffAverage ℝ f g n x ↔ 0 < birkhoffSum f (g - fun _ ↦ a) n x := by |
| 207 | + cases n with |
| 208 | + | zero => |
| 209 | + refine ⟨fun h => ?_, fun h => ?_⟩ |
| 210 | + · exfalso; rw [birkhoffAverage_zero] at h; exact lt_asymm ha h |
| 211 | + · exfalso; rw [birkhoffSum_zero] at h; exact lt_irrefl 0 h |
| 212 | + | succ n => exact birkhoffAverage_iff_birkhoffSum (by positivity) |
| 213 | + conv => |
| 214 | + enter [1, 1, x, 1, n] |
| 215 | + rw [this] |
| 216 | + |
| 217 | +section MeasurePreserving |
| 218 | + |
| 219 | +variable {f : α → α} [MeasurableSpace α] (μ : Measure α := by volume_tac) {n} |
| 220 | + (hf : MeasurePreserving f μ μ) |
| 221 | + |
| 222 | +include hf |
| 223 | + |
| 224 | +section Real |
| 225 | + |
| 226 | +variable {g : α → ℝ} (hg : Integrable g μ) |
| 227 | + |
| 228 | +include hg |
| 229 | + |
| 230 | +lemma tendsto_setIntegral_on_birkhoffMax_support_birkhoffSupSet : |
| 231 | + Tendsto (fun n ↦ ∫ x in (birkhoffMax f g n).support, g x ∂μ) atTop |
| 232 | + (𝓝 <| ∫ x in birkhoffSupSet f g, g x ∂ μ) := by |
| 233 | + rw [birkhoffSupSet_eq_iSup_birkhoffMax_support] |
| 234 | + apply tendsto_setIntegral_of_monotone₀ _ _ hg.integrableOn |
| 235 | + · intros |
| 236 | + exact AEStronglyMeasurable.nullMeasurableSet_support (by measurability) |
| 237 | + · intro i j hij x |
| 238 | + have : 0 ≤ birkhoffMax f g i x := birkhoffMax_nonneg x |
| 239 | + have := (birkhoffMax f g).mono hij x |
| 240 | + grind [Function.mem_support] |
| 241 | +theorem setIntegral_nonneg_on_birkhoffSupSet : |
| 242 | + 0 ≤ ∫ x in birkhoffSupSet f g, g x ∂μ := by |
| 243 | + apply ge_of_tendsto' (tendsto_setIntegral_on_birkhoffMax_support_birkhoffSupSet μ hf hg) |
| 244 | + intro n |
| 245 | + exact setIntegral_nonneg_on_birkhoffMax_support μ hf hg |
| 246 | + |
| 247 | +variable [IsFiniteMeasure μ] |
| 248 | + |
| 249 | +/-- **Maximal ergodic theorem**: The measure of the set where the supremum of the Birkhoff |
| 250 | +averages of `g` is greater than `a`, multiplied by `a`, is bounded above by the integral of |
| 251 | +`g` on this set. -/ |
| 252 | +public theorem meas_birkhoffAverageSupSet_smul_const_le_integral (a : ℝ) (ha : 0 < a) : |
| 253 | + μ.real (birkhoffAverageSupSet f g a) • a ≤ ∫ x in birkhoffAverageSupSet f g a, g x ∂μ := by |
| 254 | + have p₁ := Integrable.sub hg (integrable_const a) |
| 255 | + calc |
| 256 | + _ = ∫ x in birkhoffSupSet f (g - fun _ ↦ a), a ∂μ := by |
| 257 | + rw [setIntegral_const, birkhoffAverageSupSet_eq_birkhoffSupSet ha] |
| 258 | + _ ≤ ∫ x in birkhoffSupSet f (g - fun _ ↦ a), a ∂μ + |
| 259 | + ∫ x in birkhoffSupSet f (g - fun _ ↦ a), g x - a ∂μ := by |
| 260 | + exact le_add_of_nonneg_right (setIntegral_nonneg_on_birkhoffSupSet μ hf p₁) |
| 261 | + _ = ∫ x in birkhoffAverageSupSet f g a, g x ∂μ := by |
| 262 | + rw [← integral_add, birkhoffAverageSupSet_eq_birkhoffSupSet ha] |
| 263 | + · rcongr; grind |
| 264 | + · exact (integrable_const a).restrict |
| 265 | + · exact p₁.restrict |
| 266 | + |
| 267 | +end Real |
| 268 | + |
| 269 | +section NormedAddCommGroup |
| 270 | + |
| 271 | +variable {E : Type*} [NormedAddCommGroup E] {g : α → E} (hg : Integrable g μ) [IsFiniteMeasure μ] |
| 272 | + |
| 273 | +include hg |
| 274 | + |
| 275 | +/-- **Maximal ergodic theorem** for group-valued functions: The measure of the set where |
| 276 | +the supremum of the Birkhoff averages of `‖g‖` is greater than `a`, multiplied by `a`, is |
| 277 | +bounded above by the norm of `g`. -/ |
| 278 | +public theorem meas_birkhoffAverageSupSet_smul_const_le_norm (a : ℝ) (ha : 0 < a) : |
| 279 | + μ.real (birkhoffAverageSupSet f (fun x ↦ ‖g x‖) a) • a ≤ ∫ x, ‖g x‖ ∂μ := |
| 280 | + calc |
| 281 | + _ ≤ ∫ x in birkhoffAverageSupSet f (fun x ↦ ‖g x‖) a, ‖g x‖ ∂μ := by |
| 282 | + exact meas_birkhoffAverageSupSet_smul_const_le_integral μ hf hg.norm a ha |
| 283 | + _ ≤ ∫ x, ‖g x‖ ∂μ := by |
| 284 | + exact setIntegral_le_integral hg.norm (ae_of_all _ (norm_nonneg <| g ·)) |
| 285 | + |
| 286 | +end NormedAddCommGroup |
| 287 | + |
| 288 | +end MeasurePreserving |
0 commit comments