@@ -207,6 +207,38 @@ lemma gpow_apply_eq_of_apply_apply_eq_self {f : perm α} {x : α} (hffx : f (f x
207
207
← pow_succ, eq_comm, inv_eq_iff_eq, ← mul_apply, ← pow_succ', @eq_comm _ x, or.comm],
208
208
exact pow_apply_eq_of_apply_apply_eq_self hffx _ }
209
209
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
+
210
242
variable [decidable_eq α]
211
243
212
244
/-- The `finset` of nonfixed points of a permutation. -/
0 commit comments