@@ -3,13 +3,11 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Yury Kudryashov
5
5
-/
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
9
9
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
13
11
import Mathlib.Dynamics.FixedPoints.Basic
14
12
15
13
/-!
@@ -43,6 +41,8 @@ is a periodic point of `f` of period `n` if and only if `minimalPeriod f x | n`.
43
41
44
42
-/
45
43
44
+ assert_not_exists MonoidWithZero
45
+
46
46
47
47
open Set
48
48
@@ -164,7 +164,7 @@ theorem eq_of_apply_eq_same (hx : IsPeriodicPt f n x) (hy : IsPeriodicPt f n y)
164
164
then `x = y`. -/
165
165
theorem eq_of_apply_eq (hx : IsPeriodicPt f m x) (hy : IsPeriodicPt f n y) (hm : 0 < m) (hn : 0 < n)
166
166
(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
168
168
169
169
end IsPeriodicPt
170
170
@@ -187,9 +187,6 @@ theorem bijOn_ptsOfPeriod (f : α → α) {n : ℕ} (hn : 0 < n) :
187
187
⟨f^[n.pred] x, hx.apply_iterate _, by
188
188
rw [← comp_apply (f := f), comp_iterate_pred_of_pos f hn, hx.eq]⟩⟩
189
189
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
-
193
190
/-- The set of periodic points of a map `f : α → α`. -/
194
191
def periodicPts (f : α → α) : Set α :=
195
192
{ x : α | ∃ n > 0 , IsPeriodicPt f n x }
@@ -208,7 +205,7 @@ theorem isPeriodicPt_of_mem_periodicPts_of_isPeriodicPt_iterate (hx : x ∈ peri
208
205
change _ = _
209
206
convert (hm.apply_iterate ((n / r + 1 ) * r - n)).eq <;>
210
207
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]
212
209
exact (Nat.lt_div_mul_add hr).le
213
210
214
211
variable (f)
@@ -219,10 +216,6 @@ theorem bUnion_ptsOfPeriod : ⋃ n > 0, ptsOfPeriod f n = periodicPts f :=
219
216
theorem iUnion_pNat_ptsOfPeriod : ⋃ n : ℕ+, ptsOfPeriod f n = periodicPts f :=
220
217
iSup_subtype.trans <| bUnion_ptsOfPeriod f
221
218
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
-
226
219
variable {f}
227
220
228
221
theorem Semiconj.mapsTo_periodicPts {g : α → β} (h : Semiconj g fa fb) :
@@ -339,7 +332,7 @@ theorem not_isPeriodicPt_of_pos_of_lt_minimalPeriod :
339
332
| _ + 1 , _, hn => fun hp => Nat.succ_ne_zero _ (hp.eq_zero_of_lt_minimalPeriod hn)
340
333
341
334
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 =>
343
336
-- Porting note: `Nat.dvd_iff_mod_eq_zero` gained explicit arguments
344
337
Nat.dvd_iff_mod_eq_zero.2 <|
345
338
(hx.mod <| isPeriodicPt_minimalPeriod f x).eq_zero_of_lt_minimalPeriod <|
@@ -354,40 +347,11 @@ theorem minimalPeriod_eq_minimalPeriod_iff {g : β → β} {y : β} :
354
347
minimalPeriod f x = minimalPeriod g y ↔ ∀ n, IsPeriodicPt f n x ↔ IsPeriodicPt g n y := by
355
348
simp_rw [isPeriodicPt_iff_minimalPeriod_dvd, dvd_right_iff_eq]
356
349
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
-
367
350
theorem Commute.minimalPeriod_of_comp_dvd_lcm {g : α → α} (h : Commute f g) :
368
351
minimalPeriod (f ∘ g) x ∣ Nat.lcm (minimalPeriod f x) (minimalPeriod g x) := by
369
352
rw [← isPeriodicPt_iff_minimalPeriod_dvd]
370
353
exact (isPeriodicPt_minimalPeriod f x).comp_lcm h (isPeriodicPt_minimalPeriod g x)
371
354
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
-
391
355
private theorem minimalPeriod_iterate_eq_div_gcd_aux (h : 0 < gcd (minimalPeriod f x) n) :
392
356
minimalPeriod f^[n] x = minimalPeriod f x / Nat.gcd (minimalPeriod f x) n := by
393
357
apply Nat.dvd_antisymm
@@ -521,16 +485,6 @@ theorem isPeriodicPt_prod_map (x : α × β) :
521
485
IsPeriodicPt (Prod.map f g) n x ↔ IsPeriodicPt f n x.1 ∧ IsPeriodicPt g n x.2 := by
522
486
simp [IsPeriodicPt]
523
487
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
-
534
488
end Function
535
489
536
490
namespace MulAction
@@ -584,7 +538,7 @@ theorem zpow_smul_eq_iff_period_dvd {j : ℤ} {g : G} {a : α} :
584
538
g ^ j • a = a ↔ (period g a : ℤ) ∣ j := by
585
539
rcases j with n | n
586
540
· 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,
588
542
Int.natCast_dvd_natCast, pow_smul_eq_iff_period_dvd]
589
543
590
544
@[to_additive (attr := simp)]
0 commit comments