Skip to content

Commit cf60565

Browse files
committed
feat(Dynamics/Birkhoff): define Birkhoff sum and Birkhoff average (#6131)
1 parent bafa842 commit cf60565

File tree

3 files changed

+166
-0
lines changed

3 files changed

+166
-0
lines changed

Mathlib.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1834,6 +1834,8 @@ import Mathlib.Deprecated.Subfield
18341834
import Mathlib.Deprecated.Subgroup
18351835
import Mathlib.Deprecated.Submonoid
18361836
import Mathlib.Deprecated.Subring
1837+
import Mathlib.Dynamics.BirkhoffSum.Average
1838+
import Mathlib.Dynamics.BirkhoffSum.Basic
18371839
import Mathlib.Dynamics.Circle.RotationNumber.TranslationNumber
18381840
import Mathlib.Dynamics.Ergodic.AddCircle
18391841
import Mathlib.Dynamics.Ergodic.Conservative
Lines changed: 84 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,84 @@
1+
/-
2+
Copyright (c) 2023 Yury Kudryashov. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yury Kudryashov
5+
-/
6+
import Mathlib.Dynamics.BirkhoffSum.Basic
7+
import Mathlib.Algebra.Module.Basic
8+
9+
/-!
10+
# Birkhoff average
11+
12+
In this file we define `birkhoffAverage f g n x` to be
13+
$$
14+
\frac{1}{n}\sum_{k=0}^{n-1}g(f^{[k]}(x)),
15+
$$
16+
where `f : α → α` is a self-map on some type `α`,
17+
`g : α → M` is a function from `α` to a module over a division semiring `R`,
18+
and `R` is used to formalize division by `n` as `(n : R)⁻¹ • _`.
19+
20+
While we need an auxiliary division semiring `R` to define `birkhoffAverage`,
21+
the definition does not depend on the choice of `R`,
22+
see `birkhoffAverage_congr_ring`.
23+
24+
-/
25+
26+
section birkhoffAverage
27+
28+
variable (R : Type _) {α M : Type _} [DivisionSemiring R] [AddCommMonoid M] [Module R M]
29+
30+
/-- The average value of `g` on the first `n` points of the orbit of `x` under `f`,
31+
i.e. the Birkhoff sum `∑ k in Finset.range n, g (f^[k] x)` divided by `n`.
32+
33+
This average appears in many ergodic theorems
34+
which say that `(birkhoffAverage R f g · x)`
35+
converges to the "space average" `⨍ x, g x ∂μ` as `n → ∞`.
36+
37+
We use an auxiliary `[DivisionSemiring R]` to define division by `n`.
38+
However, the definition does not depend on the choice of `R`,
39+
see `birkhoffAverage_congr_ring`. -/
40+
def birkhoffAverage (f : α → α) (g : α → M) (n : ℕ) (x : α) : M := (n : R)⁻¹ • birkhoffSum f g n x
41+
42+
theorem birkhoffAverage_zero (f : α → α) (g : α → M) (x : α) :
43+
birkhoffAverage R f g 0 x = 0 := by simp [birkhoffAverage]
44+
45+
@[simp] theorem birkhoffAverage_zero' (f : α → α) (g : α → M) : birkhoffAverage R f g 0 = 0 :=
46+
funext <| birkhoffAverage_zero _ _ _
47+
48+
theorem birkhoffAverage_one (f : α → α) (g : α → M) (x : α) :
49+
birkhoffAverage R f g 1 x = g x := by simp [birkhoffAverage]
50+
51+
@[simp]
52+
theorem birkhoffAverage_one' (f : α → α) (g : α → M) : birkhoffAverage R f g 1 = g :=
53+
funext <| birkhoffAverage_one R f g
54+
55+
theorem map_birkhoffAverage (S : Type _) {F N : Type _}
56+
[DivisionSemiring S] [AddCommMonoid N] [Module S N]
57+
[AddMonoidHomClass F M N] (g' : F) (f : α → α) (g : α → M) (n : ℕ) (x : α) :
58+
g' (birkhoffAverage R f g n x) = birkhoffAverage S f (g' ∘ g) n x := by
59+
simp only [birkhoffAverage, map_inv_nat_cast_smul g' R S, map_birkhoffSum]
60+
61+
theorem birkhoffAverage_congr_ring (S : Type _) [DivisionSemiring S] [Module S M]
62+
(f : α → α) (g : α → M) (n : ℕ) (x : α) :
63+
birkhoffAverage R f g n x = birkhoffAverage S f g n x :=
64+
map_birkhoffAverage R S (AddMonoidHom.id M) f g n x
65+
66+
theorem birkhoffAverage_congr_ring' (S : Type _) [DivisionSemiring S] [Module S M] :
67+
birkhoffAverage (α := α) (M := M) R = birkhoffAverage S := by
68+
ext; apply birkhoffAverage_congr_ring
69+
70+
theorem Function.IsFixedPt.birkhoffAverage_eq [CharZero R] {f : α → α} {x : α} (h : IsFixedPt f x)
71+
(g : α → M) {n : ℕ} (hn : n ≠ 0) : birkhoffAverage R f g n x = g x := by
72+
rw [birkhoffAverage, h.birkhoffSum_eq, nsmul_eq_smul_cast R, inv_smul_smul₀]
73+
rwa [Nat.cast_ne_zero]
74+
75+
end birkhoffAverage
76+
77+
/-- Birkhoff average is "almost invariant" under `f`:
78+
the difference between `birkhoffAverage R f g n (f x)` and `birkhoffAverage R f g n x`
79+
is equal to `(n : R)⁻¹ • (g (f^[n] x) - g x)`. -/
80+
theorem birkhoffAverage_apply_sub_birkhoffAverage (R : Type _) [DivisionRing R]
81+
[AddCommGroup M] [Module R M] (f : α → α) (g : α → M) (n : ℕ) (x : α) :
82+
birkhoffAverage R f g n (f x) - birkhoffAverage R f g n x =
83+
(n : R)⁻¹ • (g (f^[n] x) - g x) := by
84+
simp only [birkhoffAverage, birkhoffSum_apply_sub_birkhoffSum, ← smul_sub]
Lines changed: 80 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,80 @@
1+
/-
2+
Copyright (c) 2023 Yury Kudryashov. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yury Kudryashov
5+
-/
6+
import Mathlib.Algebra.BigOperators.Basic
7+
import Mathlib.Dynamics.FixedPoints.Basic
8+
9+
/-!
10+
# Birkhoff sums
11+
12+
In this file we define `birkhoffSum f g n x` to be the sum `∑ k in Finset.range n, g (f^[k] x)`.
13+
This sum (more precisely, the corresponding average `n⁻¹ • birkhoffSum f g n x`)
14+
appears in various ergodic theorems
15+
saying that these averages converge to the "space average" `⨍ x, g x ∂μ` in some sense.
16+
17+
See also `birkhoffAverage` defined in `Dynamics/BirkhoffSum/Average`.
18+
-/
19+
20+
open Finset Function
21+
open scoped BigOperators
22+
23+
section AddCommMonoid
24+
25+
variable {α M : Type _} [AddCommMonoid M]
26+
27+
/-- The sum of values of `g` on the first `n` points of the orbit of `x` under `f`. -/
28+
def birkhoffSum (f : α → α) (g : α → M) (n : ℕ) (x : α) : M := ∑ k in range n, g (f^[k] x)
29+
30+
theorem birkhoffSum_zero (f : α → α) (g : α → M) (x : α) : birkhoffSum f g 0 x = 0 :=
31+
sum_range_zero _
32+
33+
@[simp]
34+
theorem birkhoffSum_zero' (f : α → α) (g : α → M) : birkhoffSum f g 0 = 0 :=
35+
funext <| birkhoffSum_zero _ _
36+
37+
theorem birkhoffSum_one (f : α → α) (g : α → M) (x : α) : birkhoffSum f g 1 x = g x :=
38+
sum_range_one _
39+
40+
@[simp]
41+
theorem birkhoffSum_one' (f : α → α) (g : α → M) : birkhoffSum f g 1 = g :=
42+
funext <| birkhoffSum_one f g
43+
44+
theorem birkhoffSum_succ (f : α → α) (g : α → M) (n : ℕ) (x : α) :
45+
birkhoffSum f g (n + 1) x = birkhoffSum f g n x + g (f^[n] x) :=
46+
sum_range_succ _ _
47+
48+
theorem birkhoffSum_succ' (f : α → α) (g : α → M) (n : ℕ) (x : α) :
49+
birkhoffSum f g (n + 1) x = g x + birkhoffSum f g n (f x) :=
50+
(sum_range_succ' _ _).trans (add_comm _ _)
51+
52+
theorem birkhoffSum_add (f : α → α) (g : α → M) (m n : ℕ) (x : α) :
53+
birkhoffSum f g (m + n) x = birkhoffSum f g m x + birkhoffSum f g n (f^[m] x) := by
54+
simp_rw [birkhoffSum, sum_range_add, add_comm m, iterate_add_apply]
55+
56+
theorem Function.IsFixedPt.birkhoffSum_eq {f : α → α} {x : α} (h : IsFixedPt f x) (g : α → M)
57+
(n : ℕ) : birkhoffSum f g n x = n • g x := by
58+
simp [birkhoffSum, (h.iterate _).eq]
59+
60+
theorem map_birkhoffSum {F N : Type _} [AddCommMonoid N] [AddMonoidHomClass F M N]
61+
(g' : F) (f : α → α) (g : α → M) (n : ℕ) (x : α) :
62+
g' (birkhoffSum f g n x) = birkhoffSum f (g' ∘ g) n x :=
63+
map_sum g' _ _
64+
65+
end AddCommMonoid
66+
67+
section AddCommGroup
68+
69+
variable {α G : Type _} [AddCommGroup G]
70+
71+
/-- Birkhoff sum is "almost invariant" under `f`:
72+
the difference between `birkhoffSum f g n (f x)` and `birkhoffSum f g n x`
73+
is equal to `g (f^[n] x) - g x`. -/
74+
theorem birkhoffSum_apply_sub_birkhoffSum (f : α → α) (g : α → G) (n : ℕ) (x : α) :
75+
birkhoffSum f g n (f x) - birkhoffSum f g n x = g (f^[n] x) - g x := by
76+
rw [← sub_eq_iff_eq_add.2 (birkhoffSum_succ f g n x),
77+
← sub_eq_iff_eq_add.2 (birkhoffSum_succ' f g n x),
78+
← sub_add, ← sub_add, sub_add_comm]
79+
80+
end AddCommGroup

0 commit comments

Comments
 (0)