Skip to content

Commit bdee03d

Browse files
committed
chore(Dynamics/PeriodicPts): don't import MonoidWithZero (#20765)
For this, split the file into `.Defs` and `.Lemmas`.
1 parent b10900e commit bdee03d

File tree

8 files changed

+93
-62
lines changed

8 files changed

+93
-62
lines changed

Archive/Imo/Imo2006Q5.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Violeta Hernández Palacios
55
-/
66
import Mathlib.Algebra.Polynomial.Roots
7-
import Mathlib.Dynamics.PeriodicPts
7+
import Mathlib.Dynamics.PeriodicPts.Lemmas
88

99
/-!
1010
# IMO 2006 Q5

Mathlib.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3071,7 +3071,8 @@ import Mathlib.Dynamics.Flow
30713071
import Mathlib.Dynamics.Minimal
30723072
import Mathlib.Dynamics.Newton
30733073
import Mathlib.Dynamics.OmegaLimit
3074-
import Mathlib.Dynamics.PeriodicPts
3074+
import Mathlib.Dynamics.PeriodicPts.Defs
3075+
import Mathlib.Dynamics.PeriodicPts.Lemmas
30753076
import Mathlib.Dynamics.TopologicalEntropy.CoverEntropy
30763077
import Mathlib.Dynamics.TopologicalEntropy.DynamicalEntourage
30773078
import Mathlib.Dynamics.TopologicalEntropy.NetEntropy

Mathlib/Dynamics/PeriodicPts.lean renamed to Mathlib/Dynamics/PeriodicPts/Defs.lean

Lines changed: 10 additions & 56 deletions
Original file line numberDiff line numberDiff line change
@@ -3,13 +3,11 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury Kudryashov
55
-/
6-
import Mathlib.Algebra.GroupPower.IterateHom
7-
import Mathlib.Algebra.Ring.Divisibility.Basic
8-
import Mathlib.Algebra.Ring.Int.Defs
6+
import Batteries.Data.Nat.Gcd
7+
import Mathlib.Algebra.Order.Group.Nat
8+
import Mathlib.Algebra.Order.Sub.Basic
99
import Mathlib.Data.List.Cycle
10-
import Mathlib.Data.Nat.GCD.Basic
11-
import Mathlib.Data.Nat.Prime.Basic
12-
import Mathlib.Data.PNat.Basic
10+
import Mathlib.Data.PNat.Notation
1311
import Mathlib.Dynamics.FixedPoints.Basic
1412

1513
/-!
@@ -43,6 +41,8 @@ is a periodic point of `f` of period `n` if and only if `minimalPeriod f x | n`.
4341
4442
-/
4543

44+
assert_not_exists MonoidWithZero
45+
4646

4747
open Set
4848

@@ -164,7 +164,7 @@ theorem eq_of_apply_eq_same (hx : IsPeriodicPt f n x) (hy : IsPeriodicPt f n y)
164164
then `x = y`. -/
165165
theorem eq_of_apply_eq (hx : IsPeriodicPt f m x) (hy : IsPeriodicPt f n y) (hm : 0 < m) (hn : 0 < n)
166166
(h : f x = f y) : x = y :=
167-
(hx.mul_const n).eq_of_apply_eq_same (hy.const_mul m) (mul_pos hm hn) h
167+
(hx.mul_const n).eq_of_apply_eq_same (hy.const_mul m) (Nat.mul_pos hm hn) h
168168

169169
end IsPeriodicPt
170170

@@ -187,9 +187,6 @@ theorem bijOn_ptsOfPeriod (f : α → α) {n : ℕ} (hn : 0 < n) :
187187
⟨f^[n.pred] x, hx.apply_iterate _, by
188188
rw [← comp_apply (f := f), comp_iterate_pred_of_pos f hn, hx.eq]⟩⟩
189189

190-
theorem directed_ptsOfPeriod_pNat (f : α → α) : Directed (· ⊆ ·) fun n : ℕ+ => ptsOfPeriod f n :=
191-
fun m n => ⟨m * n, fun _ hx => hx.mul_const n, fun _ hx => hx.const_mul m⟩
192-
193190
/-- The set of periodic points of a map `f : α → α`. -/
194191
def periodicPts (f : α → α) : Set α :=
195192
{ x : α | ∃ n > 0, IsPeriodicPt f n x }
@@ -208,7 +205,7 @@ theorem isPeriodicPt_of_mem_periodicPts_of_isPeriodicPt_iterate (hx : x ∈ peri
208205
change _ = _
209206
convert (hm.apply_iterate ((n / r + 1) * r - n)).eq <;>
210207
rw [← iterate_add_apply, Nat.sub_add_cancel this, iterate_mul, (hr'.iterate _).eq]
211-
rw [add_mul, one_mul]
208+
rw [Nat.add_mul, one_mul]
212209
exact (Nat.lt_div_mul_add hr).le
213210

214211
variable (f)
@@ -219,10 +216,6 @@ theorem bUnion_ptsOfPeriod : ⋃ n > 0, ptsOfPeriod f n = periodicPts f :=
219216
theorem iUnion_pNat_ptsOfPeriod : ⋃ n : ℕ+, ptsOfPeriod f n = periodicPts f :=
220217
iSup_subtype.trans <| bUnion_ptsOfPeriod f
221218

222-
theorem bijOn_periodicPts : BijOn f (periodicPts f) (periodicPts f) :=
223-
iUnion_pNat_ptsOfPeriod f ▸
224-
bijOn_iUnion_of_directed (directed_ptsOfPeriod_pNat f) fun i => bijOn_ptsOfPeriod f i.pos
225-
226219
variable {f}
227220

228221
theorem Semiconj.mapsTo_periodicPts {g : α → β} (h : Semiconj g fa fb) :
@@ -339,7 +332,7 @@ theorem not_isPeriodicPt_of_pos_of_lt_minimalPeriod :
339332
| _ + 1, _, hn => fun hp => Nat.succ_ne_zero _ (hp.eq_zero_of_lt_minimalPeriod hn)
340333

341334
theorem IsPeriodicPt.minimalPeriod_dvd (hx : IsPeriodicPt f n x) : minimalPeriod f x ∣ n :=
342-
(eq_or_lt_of_le <| n.zero_le).elim (fun hn0 => hn0 ▸ dvd_zero _) fun hn0 =>
335+
(eq_or_lt_of_le <| n.zero_le).elim (fun hn0 => hn0 ▸ Nat.dvd_zero _) fun hn0 =>
343336
-- Porting note: `Nat.dvd_iff_mod_eq_zero` gained explicit arguments
344337
Nat.dvd_iff_mod_eq_zero.2 <|
345338
(hx.mod <| isPeriodicPt_minimalPeriod f x).eq_zero_of_lt_minimalPeriod <|
@@ -354,40 +347,11 @@ theorem minimalPeriod_eq_minimalPeriod_iff {g : β → β} {y : β} :
354347
minimalPeriod f x = minimalPeriod g y ↔ ∀ n, IsPeriodicPt f n x ↔ IsPeriodicPt g n y := by
355348
simp_rw [isPeriodicPt_iff_minimalPeriod_dvd, dvd_right_iff_eq]
356349

357-
theorem minimalPeriod_eq_prime {p : ℕ} [hp : Fact p.Prime] (hper : IsPeriodicPt f p x)
358-
(hfix : ¬IsFixedPt f x) : minimalPeriod f x = p :=
359-
(hp.out.eq_one_or_self_of_dvd _ hper.minimalPeriod_dvd).resolve_left
360-
(mt minimalPeriod_eq_one_iff_isFixedPt.1 hfix)
361-
362-
theorem minimalPeriod_eq_prime_pow {p k : ℕ} [hp : Fact p.Prime] (hk : ¬IsPeriodicPt f (p ^ k) x)
363-
(hk1 : IsPeriodicPt f (p ^ (k + 1)) x) : minimalPeriod f x = p ^ (k + 1) := by
364-
apply Nat.eq_prime_pow_of_dvd_least_prime_pow hp.out <;>
365-
rwa [← isPeriodicPt_iff_minimalPeriod_dvd]
366-
367350
theorem Commute.minimalPeriod_of_comp_dvd_lcm {g : α → α} (h : Commute f g) :
368351
minimalPeriod (f ∘ g) x ∣ Nat.lcm (minimalPeriod f x) (minimalPeriod g x) := by
369352
rw [← isPeriodicPt_iff_minimalPeriod_dvd]
370353
exact (isPeriodicPt_minimalPeriod f x).comp_lcm h (isPeriodicPt_minimalPeriod g x)
371354

372-
theorem Commute.minimalPeriod_of_comp_dvd_mul {g : α → α} (h : Commute f g) :
373-
minimalPeriod (f ∘ g) x ∣ minimalPeriod f x * minimalPeriod g x :=
374-
dvd_trans h.minimalPeriod_of_comp_dvd_lcm (lcm_dvd_mul _ _)
375-
376-
theorem Commute.minimalPeriod_of_comp_eq_mul_of_coprime {g : α → α} (h : Commute f g)
377-
(hco : Coprime (minimalPeriod f x) (minimalPeriod g x)) :
378-
minimalPeriod (f ∘ g) x = minimalPeriod f x * minimalPeriod g x := by
379-
apply h.minimalPeriod_of_comp_dvd_mul.antisymm
380-
suffices
381-
∀ {f g : α → α},
382-
Commute f g →
383-
Coprime (minimalPeriod f x) (minimalPeriod g x) →
384-
minimalPeriod f x ∣ minimalPeriod (f ∘ g) x from
385-
hco.mul_dvd_of_dvd_of_dvd (this h hco) (h.comp_eq.symm ▸ this h.symm hco.symm)
386-
intro f g h hco
387-
refine hco.dvd_of_dvd_mul_left (IsPeriodicPt.left_of_comp h ?_ ?_).minimalPeriod_dvd
388-
· exact (isPeriodicPt_minimalPeriod _ _).const_mul _
389-
· exact (isPeriodicPt_minimalPeriod _ _).mul_const _
390-
391355
private theorem minimalPeriod_iterate_eq_div_gcd_aux (h : 0 < gcd (minimalPeriod f x) n) :
392356
minimalPeriod f^[n] x = minimalPeriod f x / Nat.gcd (minimalPeriod f x) n := by
393357
apply Nat.dvd_antisymm
@@ -521,16 +485,6 @@ theorem isPeriodicPt_prod_map (x : α × β) :
521485
IsPeriodicPt (Prod.map f g) n x ↔ IsPeriodicPt f n x.1 ∧ IsPeriodicPt g n x.2 := by
522486
simp [IsPeriodicPt]
523487

524-
theorem minimalPeriod_prod_map (f : α → α) (g : β → β) (x : α × β) :
525-
minimalPeriod (Prod.map f g) x = (minimalPeriod f x.1).lcm (minimalPeriod g x.2) :=
526-
eq_of_forall_dvd <| by cases x; simp [← isPeriodicPt_iff_minimalPeriod_dvd, Nat.lcm_dvd_iff]
527-
528-
theorem minimalPeriod_fst_dvd : minimalPeriod f x.1 ∣ minimalPeriod (Prod.map f g) x := by
529-
rw [minimalPeriod_prod_map]; exact Nat.dvd_lcm_left _ _
530-
531-
theorem minimalPeriod_snd_dvd : minimalPeriod g x.2 ∣ minimalPeriod (Prod.map f g) x := by
532-
rw [minimalPeriod_prod_map]; exact Nat.dvd_lcm_right _ _
533-
534488
end Function
535489

536490
namespace MulAction
@@ -584,7 +538,7 @@ theorem zpow_smul_eq_iff_period_dvd {j : ℤ} {g : G} {a : α} :
584538
g ^ j • a = a ↔ (period g a : ℤ) ∣ j := by
585539
rcases j with n | n
586540
· rw [Int.ofNat_eq_coe, zpow_natCast, Int.natCast_dvd_natCast, pow_smul_eq_iff_period_dvd]
587-
· rw [Int.negSucc_coe, zpow_neg, zpow_natCast, inv_smul_eq_iff, eq_comm, dvd_neg,
541+
· rw [Int.negSucc_coe, zpow_neg, zpow_natCast, inv_smul_eq_iff, eq_comm, Int.dvd_neg,
588542
Int.natCast_dvd_natCast, pow_smul_eq_iff_period_dvd]
589543

590544
@[to_additive (attr := simp)]
Lines changed: 75 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,75 @@
1+
/-
2+
Copyright (c) 2020 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.Data.Nat.GCD.Basic
7+
import Mathlib.Data.Nat.Prime.Basic
8+
import Mathlib.Data.PNat.Basic
9+
import Mathlib.Dynamics.PeriodicPts.Defs
10+
11+
/-!
12+
# Extra lemmas about periodic points
13+
-/
14+
15+
open Nat Set
16+
17+
namespace Function
18+
variable {α : Type*} {f : α → α} {x y : α}
19+
20+
open Function (Commute)
21+
22+
theorem directed_ptsOfPeriod_pNat (f : α → α) : Directed (· ⊆ ·) fun n : ℕ+ => ptsOfPeriod f n :=
23+
fun m n => ⟨m * n, fun _ hx => hx.mul_const n, fun _ hx => hx.const_mul m⟩
24+
25+
variable (f) in
26+
theorem bijOn_periodicPts : BijOn f (periodicPts f) (periodicPts f) :=
27+
iUnion_pNat_ptsOfPeriod f ▸
28+
bijOn_iUnion_of_directed (directed_ptsOfPeriod_pNat f) fun i => bijOn_ptsOfPeriod f i.pos
29+
30+
theorem minimalPeriod_eq_prime {p : ℕ} [hp : Fact p.Prime] (hper : IsPeriodicPt f p x)
31+
(hfix : ¬IsFixedPt f x) : minimalPeriod f x = p :=
32+
(hp.out.eq_one_or_self_of_dvd _ hper.minimalPeriod_dvd).resolve_left
33+
(mt minimalPeriod_eq_one_iff_isFixedPt.1 hfix)
34+
35+
theorem minimalPeriod_eq_prime_pow {p k : ℕ} [hp : Fact p.Prime] (hk : ¬IsPeriodicPt f (p ^ k) x)
36+
(hk1 : IsPeriodicPt f (p ^ (k + 1)) x) : minimalPeriod f x = p ^ (k + 1) := by
37+
apply Nat.eq_prime_pow_of_dvd_least_prime_pow hp.out <;>
38+
rwa [← isPeriodicPt_iff_minimalPeriod_dvd]
39+
40+
theorem Commute.minimalPeriod_of_comp_dvd_mul {g : α → α} (h : Commute f g) :
41+
minimalPeriod (f ∘ g) x ∣ minimalPeriod f x * minimalPeriod g x :=
42+
dvd_trans h.minimalPeriod_of_comp_dvd_lcm (lcm_dvd_mul _ _)
43+
44+
theorem Commute.minimalPeriod_of_comp_eq_mul_of_coprime {g : α → α} (h : Commute f g)
45+
(hco : Coprime (minimalPeriod f x) (minimalPeriod g x)) :
46+
minimalPeriod (f ∘ g) x = minimalPeriod f x * minimalPeriod g x := by
47+
apply h.minimalPeriod_of_comp_dvd_mul.antisymm
48+
suffices
49+
∀ {f g : α → α},
50+
Commute f g →
51+
Coprime (minimalPeriod f x) (minimalPeriod g x) →
52+
minimalPeriod f x ∣ minimalPeriod (f ∘ g) x from
53+
hco.mul_dvd_of_dvd_of_dvd (this h hco) (h.comp_eq.symm ▸ this h.symm hco.symm)
54+
intro f g h hco
55+
refine hco.dvd_of_dvd_mul_left (IsPeriodicPt.left_of_comp h ?_ ?_).minimalPeriod_dvd
56+
· exact (isPeriodicPt_minimalPeriod _ _).const_mul _
57+
· exact (isPeriodicPt_minimalPeriod _ _).mul_const _
58+
59+
end Function
60+
61+
namespace Function
62+
63+
variable {α β : Type*} {f : α → α} {g : β → β} {x : α × β} {a : α} {b : β} {m n : ℕ}
64+
65+
theorem minimalPeriod_prod_map (f : α → α) (g : β → β) (x : α × β) :
66+
minimalPeriod (Prod.map f g) x = (minimalPeriod f x.1).lcm (minimalPeriod g x.2) :=
67+
eq_of_forall_dvd <| by cases x; simp [← isPeriodicPt_iff_minimalPeriod_dvd, Nat.lcm_dvd_iff]
68+
69+
theorem minimalPeriod_fst_dvd : minimalPeriod f x.1 ∣ minimalPeriod (Prod.map f g) x := by
70+
rw [minimalPeriod_prod_map]; exact Nat.dvd_lcm_left _ _
71+
72+
theorem minimalPeriod_snd_dvd : minimalPeriod g x.2 ∣ minimalPeriod (Prod.map f g) x := by
73+
rw [minimalPeriod_prod_map]; exact Nat.dvd_lcm_right _ _
74+
75+
end Function

Mathlib/GroupTheory/GroupAction/FixedPoints.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ Authors: Emilie Burgun
55
-/
66
import Mathlib.Algebra.Group.Commute.Basic
77
import Mathlib.Data.Set.Pointwise.SMul
8-
import Mathlib.Dynamics.PeriodicPts
8+
import Mathlib.Dynamics.PeriodicPts.Defs
99
import Mathlib.GroupTheory.GroupAction.Defs
1010

1111
/-!
@@ -83,7 +83,7 @@ theorem fixedBy_subset_fixedBy_zpow (g : G) (j : ℤ) :
8383
fixedBy α g ⊆ fixedBy α (g ^ j) := by
8484
intro a a_in_fixedBy
8585
rw [mem_fixedBy, zpow_smul_eq_iff_minimalPeriod_dvd,
86-
minimalPeriod_eq_one_iff_fixedBy.mpr a_in_fixedBy, Nat.cast_one]
86+
minimalPeriod_eq_one_iff_fixedBy.mpr a_in_fixedBy, Int.natCast_one]
8787
exact one_dvd j
8888

8989
variable (M α) in

Mathlib/GroupTheory/GroupAction/Period.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Emilie Burgun
55
-/
66

7-
import Mathlib.Dynamics.PeriodicPts
7+
import Mathlib.Dynamics.PeriodicPts.Lemmas
88
import Mathlib.GroupTheory.Exponent
99
import Mathlib.GroupTheory.GroupAction.Basic
1010

Mathlib/GroupTheory/GroupAction/Quotient.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ Authors: Chris Hughes, Thomas Browning
66
import Mathlib.Algebra.Group.Subgroup.Actions
77
import Mathlib.Algebra.Group.Subgroup.ZPowers.Lemmas
88
import Mathlib.Data.Fintype.BigOperators
9-
import Mathlib.Dynamics.PeriodicPts
9+
import Mathlib.Dynamics.PeriodicPts.Defs
1010
import Mathlib.GroupTheory.Commutator.Basic
1111
import Mathlib.GroupTheory.Coset.Basic
1212
import Mathlib.GroupTheory.GroupAction.Basic

Mathlib/GroupTheory/OrderOfElement.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import Mathlib.Algebra.Group.Subgroup.Finite
88
import Mathlib.Algebra.Module.NatInt
99
import Mathlib.Algebra.Order.Group.Action
1010
import Mathlib.Algebra.Order.Ring.Abs
11+
import Mathlib.Dynamics.PeriodicPts.Lemmas
1112
import Mathlib.GroupTheory.Index
1213
import Mathlib.Order.Interval.Set.Infinite
1314
import Mathlib.Tactic.Positivity

0 commit comments

Comments
 (0)