Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit d07245f

Browse files
committed
feat(group_theory/order_of_element): Order in α × β (#18719)
The order of `(a, b)` is the lcm of the orders of `a` and `b`. Match `pow` and `zpow` lemmas. Also some `variables` noise because I could not use `x` to mean what I wanted, and incidentally the type `A` was mostly unused.
1 parent d608fc5 commit d07245f

File tree

3 files changed

+95
-22
lines changed

3 files changed

+95
-22
lines changed

src/data/prod/basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -97,7 +97,7 @@ funext (λ p, ext (map_fst f g p) (map_snd f g p))
9797
lemma id_prod : (λ (p : α × β), (p.1, p.2)) = id :=
9898
funext $ λ ⟨a, b⟩, rfl
9999

100-
lemma map_id : (prod.map (@id α) (@id β)) = id :=
100+
@[simp] lemma map_id : (prod.map (@id α) (@id β)) = id :=
101101
id_prod
102102

103103
lemma fst_surjective [h : nonempty β] : function.surjective (@fst α β) :=

src/dynamics/periodic_pts.lean

Lines changed: 25 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -519,6 +519,31 @@ end
519519

520520
end function
521521

522+
namespace function
523+
variables {α β : Type*} {f : α → α} {g : β → β} {x : α × β} {a : α} {b : β} {m n : ℕ}
524+
525+
@[simp] lemma iterate_prod_map (f : α → α) (g : β → β) (n : ℕ) :
526+
(prod.map f g)^[n] = prod.map (f^[n]) (g^[n]) := by induction n; simp [*, prod.map_comp_map]
527+
528+
@[simp] lemma is_fixed_pt_prod_map (x : α × β) :
529+
is_fixed_pt (prod.map f g) x ↔ is_fixed_pt f x.1 ∧ is_fixed_pt g x.2 := prod.ext_iff
530+
531+
@[simp] lemma is_periodic_pt_prod_map (x : α × β) :
532+
is_periodic_pt (prod.map f g) n x ↔ is_periodic_pt f n x.1 ∧ is_periodic_pt g n x.2 :=
533+
by simp [is_periodic_pt]
534+
535+
lemma minimal_period_prod_map (f : α → α) (g : β → β) (x : α × β) :
536+
minimal_period (prod.map f g) x = (minimal_period f x.1).lcm (minimal_period g x.2) :=
537+
eq_of_forall_dvd $ by cases x; simp [←is_periodic_pt_iff_minimal_period_dvd, nat.lcm_dvd_iff]
538+
539+
lemma minimal_period_fst_dvd : minimal_period f x.1 ∣ minimal_period (prod.map f g) x :=
540+
by { rw minimal_period_prod_map, exact nat.dvd_lcm_left _ _ }
541+
542+
lemma minimal_period_snd_dvd : minimal_period g x.2 ∣ minimal_period (prod.map f g) x :=
543+
by { rw minimal_period_prod_map, exact nat.dvd_lcm_right _ _ }
544+
545+
end function
546+
522547
namespace mul_action
523548

524549
open function

src/group_theory/order_of_element.lean

Lines changed: 69 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -3,8 +3,9 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Julian Kuelshammer
55
-/
6+
import algebra.gcd_monoid.finset
67
import algebra.hom.iterate
7-
import data.nat.modeq
8+
import data.int.modeq
89
import data.set.pointwise.basic
910
import data.set.intervals.infinite
1011
import dynamics.periodic_pts
@@ -35,14 +36,11 @@ order of an element
3536
open function nat
3637
open_locale pointwise
3738

38-
universes u v
39-
40-
variables {G : Type u} {A : Type v}
41-
variables {x y : G} {a b : A} {n m : ℕ}
39+
variables {G H A α β : Type*}
4240

4341
section monoid_add_monoid
4442

45-
variables [monoid G] [add_monoid A]
43+
variables [monoid G] [add_monoid A] {x y : G} {a b : A} {n m : ℕ}
4644

4745
section is_of_fin_order
4846

@@ -95,8 +93,7 @@ by { rw [is_of_fin_order_iff_pow_eq_one, is_of_fin_order_iff_pow_eq_one], norm_c
9593
/-- The image of an element of finite order has finite order. -/
9694
@[to_additive add_monoid_hom.is_of_fin_order
9795
"The image of an element of finite additive order has finite additive order."]
98-
lemma monoid_hom.is_of_fin_order
99-
{H : Type v} [monoid H] (f : G →* H) {x : G} (h : is_of_fin_order x) :
96+
lemma monoid_hom.is_of_fin_order [monoid H] (f : G →* H) {x : G} (h : is_of_fin_order x) :
10097
is_of_fin_order $ f x :=
10198
(is_of_fin_order_iff_pow_eq_one _).mpr $ begin
10299
rcases (is_of_fin_order_iff_pow_eq_one _).mp h with ⟨n, npos, hn⟩,
@@ -171,6 +168,11 @@ end
171168
lemma order_of_pos_iff : 0 < order_of x ↔ is_of_fin_order x :=
172169
by rwa [iff_not_comm.mp order_of_eq_zero_iff, pos_iff_ne_zero]
173170

171+
@[to_additive is_of_fin_add_order.mono]
172+
lemma is_of_fin_order.mono [monoid β] {y : β} (hx : is_of_fin_order x)
173+
(h : order_of y ∣ order_of x) : is_of_fin_order y :=
174+
by { rw ←order_of_pos_iff at ⊢ hx, exact nat.pos_of_dvd_of_pos h hx }
175+
174176
@[to_additive nsmul_ne_zero_of_lt_add_order_of']
175177
lemma pow_ne_one_of_lt_order_of' (n0 : n ≠ 0) (h : n < order_of x) : x ^ n ≠ 1 :=
176178
λ j, not_is_periodic_pt_of_pos_of_lt_minimal_period n0 h
@@ -311,11 +313,11 @@ end
311313
lemma order_of_dvd_lcm_mul : order_of y ∣ nat.lcm (order_of x) (order_of (x * y)) :=
312314
begin
313315
by_cases h0 : order_of x = 0,
314-
{ rw [h0, lcm_zero_left], apply dvd_zero },
316+
{ rw [h0, nat.lcm_zero_left], apply dvd_zero },
315317
conv_lhs { rw [← one_mul y, ← pow_order_of_eq_one x,
316318
← succ_pred_eq_of_pos (nat.pos_of_ne_zero h0), pow_succ', mul_assoc] },
317319
exact (((commute.refl x).mul_right h).pow_left _).order_of_mul_dvd_lcm.trans
318-
(lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩),
320+
(nat.lcm_dvd_iff.2 ⟨trans (order_of_pow_dvd _) (dvd_lcm_left _ _), dvd_lcm_right _ _⟩),
319321
end
320322

321323
@[to_additive add_order_of_add_dvd_mul_add_order_of]
@@ -394,7 +396,7 @@ end p_prime
394396
end monoid_add_monoid
395397

396398
section cancel_monoid
397-
variables [left_cancel_monoid G] (x y)
399+
variables [left_cancel_monoid G] (x y : G) {m n : ℕ}
398400

399401
@[to_additive nsmul_injective_of_lt_add_order_of]
400402
lemma pow_injective_of_lt_order_of
@@ -451,7 +453,7 @@ end
451453
end cancel_monoid
452454

453455
section group
454-
variables [group G] [add_group A] {x a} {i : ℤ}
456+
variables [group G] {x y : G} {i : ℤ}
455457

456458
/-- Inverses of elements of finite order have finite order. -/
457459
@[to_additive "Inverses of elements of finite additive order have finite additive order."]
@@ -560,7 +562,7 @@ end group
560562

561563
section comm_monoid
562564

563-
variables [comm_monoid G]
565+
variables [comm_monoid G] {x y : G}
564566

565567
/-- Elements of finite order are closed under multiplication. -/
566568
@[to_additive "Elements of finite additive order are closed under addition."]
@@ -571,7 +573,7 @@ lemma is_of_fin_order.mul (hx : is_of_fin_order x) (hy : is_of_fin_order y) :
571573
end comm_monoid
572574

573575
section finite_monoid
574-
variables [monoid G]
576+
variables [monoid G] {n : ℕ}
575577
open_locale big_operators
576578

577579
@[to_additive sum_card_add_order_of_eq_card_nsmul_eq_zero]
@@ -592,7 +594,7 @@ end finite_monoid
592594

593595
section finite_cancel_monoid
594596
-- TODO: Of course everything also works for right_cancel_monoids.
595-
variables [left_cancel_monoid G] [add_left_cancel_monoid A]
597+
variables [left_cancel_monoid G] {x y : G} {n : ℕ}
596598

597599
-- TODO: Use this to show that a finite left cancellative monoid is a group.
598600
@[to_additive]
@@ -682,7 +684,7 @@ lemma order_eq_card_powers [fintype G] : order_of x = fintype.card (submonoid.po
682684
end finite_cancel_monoid
683685

684686
section finite_group
685-
variables [group G] [add_group A]
687+
variables [group G] {x y : G} {n : ℕ}
686688

687689
@[to_additive]
688690
lemma exists_zpow_eq_one [finite G] (x : G) : ∃ (i : ℤ) (H : i ≠ 0), x ^ (i : ℤ) = 1 :=
@@ -712,6 +714,23 @@ lemma mem_zpowers_iff_mem_range_order_of [finite G] [decidable_eq G] :
712714
y ∈ subgroup.zpowers x ↔ y ∈ (finset.range (order_of x)).image ((^) x : ℕ → G) :=
713715
by rw [← mem_powers_iff_mem_zpowers, mem_powers_iff_mem_range_order_of]
714716

717+
@[to_additive] lemma zpow_eq_one_iff_modeq {n : ℤ} : x ^ n = 1 ↔ n ≡ 0 [ZMOD (order_of x)] :=
718+
by rw [int.modeq_zero_iff_dvd, order_of_dvd_iff_zpow_eq_one]
719+
720+
@[to_additive] lemma zpow_eq_zpow_iff_modeq {m n : ℤ} : x ^ m = x ^ n ↔ m ≡ n [ZMOD (order_of x)] :=
721+
by rw [←mul_inv_eq_one, ←zpow_sub, zpow_eq_one_iff_modeq, int.modeq_iff_dvd, int.modeq_iff_dvd,
722+
zero_sub, neg_sub]
723+
724+
@[simp, to_additive] lemma injective_zpow_iff_not_is_of_fin_order :
725+
injective (λ n : ℤ, x ^ n) ↔ ¬ is_of_fin_order x :=
726+
begin
727+
refine ⟨_, λ h n m hnm, _⟩,
728+
{ simp_rw is_of_fin_order_iff_pow_eq_one,
729+
rintro h ⟨n, hn, hx⟩,
730+
exact nat.cast_ne_zero.2 hn.ne' (h $ by simpa using hx) },
731+
rwa [zpow_eq_zpow_iff_modeq, order_of_eq_zero_iff.2 h, nat.cast_zero, int.modeq_zero_iff] at hnm,
732+
end
733+
715734
@[to_additive decidable_zmultiples]
716735
noncomputable instance decidable_zpowers : decidable_pred (∈ subgroup.zpowers x) :=
717736
classical.dec_pred _
@@ -755,8 +774,8 @@ end
755774

756775
variables [fintype G]
757776

758-
/-- See also `order_eq_card_zpowers'`. -/
759-
@[to_additive add_order_eq_card_zmultiples "See also `add_order_eq_card_zmultiples'`."]
777+
/-- See also `nat.card_zpowers'`. -/
778+
@[to_additive add_order_eq_card_zmultiples "See also `nat.card_zmultiples`."]
760779
lemma order_eq_card_zpowers : order_of x = fintype.card (zpowers x) :=
761780
(fintype.card_fin (order_of x)).symm.trans (fintype.card_eq.2 ⟨fin_equiv_zpowers x⟩)
762781

@@ -846,8 +865,6 @@ begin
846865
(congr_arg (∣ fintype.card K) (order_of_subgroup ⟨x, hx.2⟩)).mpr order_of_dvd_card_univ⟩,
847866
end
848867

849-
variable (a)
850-
851868
/-- TODO: Generalise to `submonoid.powers`.-/
852869
@[to_additive image_range_add_order_of, nolint to_additive_doc]
853870
lemma image_range_order_of [decidable_eq G] :
@@ -909,7 +926,7 @@ end pow_is_subgroup
909926

910927
section linear_ordered_ring
911928

912-
variable [linear_ordered_ring G]
929+
variables [linear_ordered_ring G] {x : G}
913930

914931
lemma order_of_abs_ne_one (h : |x| ≠ 1) : order_of x = 0 :=
915932
begin
@@ -931,3 +948,34 @@ begin
931948
end
932949

933950
end linear_ordered_ring
951+
952+
section prod
953+
variables [monoid α] [monoid β] {x : α × β} {a : α} {b : β}
954+
955+
@[to_additive prod.add_order_of] protected lemma prod.order_of (x : α × β) :
956+
order_of x = (order_of x.1).lcm (order_of x.2) :=
957+
minimal_period_prod_map _ _ _
958+
959+
@[to_additive add_order_of_fst_dvd_add_order_of] lemma order_of_fst_dvd_order_of :
960+
order_of x.1 ∣ order_of x :=
961+
minimal_period_fst_dvd
962+
963+
@[to_additive add_order_of_snd_dvd_add_order_of] lemma order_of_snd_dvd_order_of :
964+
order_of x.2 ∣ order_of x :=
965+
minimal_period_snd_dvd
966+
967+
@[to_additive is_of_fin_add_order.fst]
968+
lemma is_of_fin_order.fst {x : α × β} (hx : is_of_fin_order x) : is_of_fin_order x.1 :=
969+
hx.mono order_of_fst_dvd_order_of
970+
971+
@[to_additive is_of_fin_add_order.snd]
972+
lemma is_of_fin_order.snd {x : α × β} (hx : is_of_fin_order x) : is_of_fin_order x.2 :=
973+
hx.mono order_of_snd_dvd_order_of
974+
975+
@[to_additive is_of_fin_add_order.prod_mk]
976+
lemma is_of_fin_order.prod_mk : is_of_fin_order a → is_of_fin_order b → is_of_fin_order (a, b) :=
977+
by simpa only [←order_of_pos_iff, prod.order_of] using nat.lcm_pos
978+
979+
end prod
980+
981+
-- TODO: Corresponding `pi` lemmas. We cannot currently state them here because of import cycles

0 commit comments

Comments
 (0)