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

Commit 030a704

Browse files
committed
feat(group_theory/perm/sign): Order of product of disjoint permutations (#6998)
The order of the product of disjoint permutations is the lcm of the orders.
1 parent fe16235 commit 030a704

File tree

2 files changed

+37
-0
lines changed

2 files changed

+37
-0
lines changed

src/group_theory/order_of_element.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -293,6 +293,11 @@ lemma add_order_of_dvd_iff_nsmul_eq_zero {n : ℕ} : add_order_of x ∣ n ↔ n
293293
lemma order_of_dvd_iff_pow_eq_one {n : ℕ} : order_of a ∣ n ↔ a ^ n = 1 :=
294294
⟨λ h, by rw [pow_eq_mod_order_of, nat.mod_eq_zero_of_dvd h, pow_zero], order_of_dvd_of_pow_eq_one⟩
295295

296+
lemma commute.order_of_mul_dvd_lcm (h : commute a b) :
297+
order_of (a * b) ∣ nat.lcm (order_of a) (order_of b) :=
298+
by rw [order_of_dvd_iff_pow_eq_one, h.mul_pow, order_of_dvd_iff_pow_eq_one.mp
299+
(nat.dvd_lcm_left _ _), order_of_dvd_iff_pow_eq_one.mp (nat.dvd_lcm_right _ _), one_mul]
300+
296301
lemma add_order_of_eq_prime {p : ℕ} [hp : fact p.prime]
297302
(hg : p •ℕ x = 0) (hg1 : x ≠ 0) : add_order_of x = p :=
298303
(hp.out.2 _ (add_order_of_dvd_of_nsmul_eq_zero hg)).resolve_left (mt add_order_of_eq_one_iff.1 hg1)

src/group_theory/perm/sign.lean

Lines changed: 32 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -207,6 +207,38 @@ lemma gpow_apply_eq_of_apply_apply_eq_self {f : perm α} {x : α} (hffx : f (f x
207207
← pow_succ, eq_comm, inv_eq_iff_eq, ← mul_apply, ← pow_succ', @eq_comm _ x, or.comm],
208208
exact pow_apply_eq_of_apply_apply_eq_self hffx _ }
209209

210+
lemma disjoint.mul_apply_eq_iff {σ τ : perm α} (hστ : disjoint σ τ) {a : α} :
211+
(σ * τ) a = a ↔ σ a = a ∧ τ a = a :=
212+
begin
213+
refine ⟨λ h, _, λ h, by rw [mul_apply, h.2, h.1]⟩,
214+
cases hστ a with hσ hτ,
215+
{ exact ⟨hσ, σ.injective (h.trans hσ.symm)⟩ },
216+
{ exact ⟨(congr_arg σ hτ).symm.trans h, hτ⟩ },
217+
end
218+
219+
lemma disjoint.mul_eq_one_iff {σ τ : perm α} (hστ : disjoint σ τ) :
220+
σ * τ = 1 ↔ σ = 1 ∧ τ = 1 :=
221+
by simp_rw [ext_iff, one_apply, hστ.mul_apply_eq_iff, forall_and_distrib]
222+
223+
lemma disjoint.gpow_disjoint_gpow {σ τ : perm α} (hστ : disjoint σ τ) (m n : ℤ) :
224+
disjoint (σ ^ m) (τ ^ n) :=
225+
λ x, or.imp (λ h, gpow_apply_eq_self_of_apply_eq_self h m)
226+
(λ h, gpow_apply_eq_self_of_apply_eq_self h n) (hστ x)
227+
228+
lemma disjoint.pow_disjoint_pow {σ τ : perm α} (hστ : disjoint σ τ) (m n : ℕ) :
229+
disjoint (σ ^ m) (τ ^ n) :=
230+
hστ.gpow_disjoint_gpow m n
231+
232+
lemma disjoint.order_of {σ τ : perm α} (hστ : disjoint σ τ) :
233+
order_of (σ * τ) = nat.lcm (order_of σ) (order_of τ) :=
234+
begin
235+
have h : ∀ n : ℕ, (σ * τ) ^ n = 1 ↔ σ ^ n = 1 ∧ τ ^ n = 1 :=
236+
λ n, by rw [commute.mul_pow hστ.mul_comm, disjoint.mul_eq_one_iff (hστ.pow_disjoint_pow n n)],
237+
exact nat.dvd_antisymm (commute.order_of_mul_dvd_lcm hστ.mul_comm) (nat.lcm_dvd
238+
(order_of_dvd_of_pow_eq_one ((h (order_of (σ * τ))).mp (pow_order_of_eq_one (σ * τ))).1)
239+
(order_of_dvd_of_pow_eq_one ((h (order_of (σ * τ))).mp (pow_order_of_eq_one (σ * τ))).2)),
240+
end
241+
210242
variable [decidable_eq α]
211243

212244
/-- The `finset` of nonfixed points of a permutation. -/

0 commit comments

Comments
 (0)