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

Commit b868323

Browse files
committed
feat(group_theory/perm/cycle/basic): Cycle on a set (#17899)
Define `equiv.perm.is_cycle_on`, a predicate for a permutation to be a cycle on a set. This has two discrepancies with the existing `equiv.perm.is_cycle_on` (apart from the obvious one that `is_cycle_on` is restricted to a set): * `is_cycle_on` forbids fixed points (on `s`) while `is_cycle` allows them. * `is_cycle` forbids the identity while `is_cycle_on` allows it. I think the first one isn't so bad given that `is_cycle` is meant to be used over the entire type (so you can't just rule out fixed points). The second one is more of a worry as I had to case-split on `s.subsingleton ∨ s.nontrivial` in several `is_cycle_on` lemmas to derive them from the corresponding `is_cycle` ones, while of course the `is_cycle` ones would also have held for a weaker version of `is_cycle` that allows the identity.
1 parent 6f413f3 commit b868323

File tree

6 files changed

+284
-101
lines changed

6 files changed

+284
-101
lines changed

src/data/set/function.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1449,6 +1449,12 @@ lemma inv_on : inv_on e e.symm t s :=
14491449
lemma bij_on_image : bij_on e s (e '' s) := (e.injective.inj_on _).bij_on_image
14501450
lemma bij_on_symm_image : bij_on e.symm (e '' s) s := e.bij_on_image.symm e.inv_on
14511451

1452+
variables {e}
1453+
1454+
@[simp] lemma bij_on_symm : bij_on e.symm t s ↔ bij_on e s t := bij_on_comm e.symm.inv_on
1455+
1456+
alias bij_on_symm ↔ _root_.set.bij_on.of_equiv_symm _root_.set.bij_on.equiv_symm
1457+
14521458
variables [decidable_eq α] {a b : α}
14531459

14541460
lemma bij_on_swap (ha : a ∈ s) (hb : b ∈ s) : bij_on (swap a b) s s :=

src/dynamics/fixed_points/basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -87,7 +87,7 @@ h.to_left_inverse e.left_inverse_symm
8787
protected lemma perm_inv (h : is_fixed_pt e x) : is_fixed_pt ⇑(e⁻¹) x := h.equiv_symm
8888

8989
protected lemma perm_pow (h : is_fixed_pt e x) (n : ℕ) : is_fixed_pt ⇑(e ^ n) x :=
90-
by { rw equiv.perm.iterate_eq_pow, exact h.iterate _ }
90+
by { rw equiv.perm.coe_pow, exact h.iterate _ }
9191

9292
protected lemma perm_zpow (h : is_fixed_pt e x) : ∀ n : ℤ, is_fixed_pt ⇑(e ^ n) x
9393
| (int.of_nat n) := h.perm_pow _

src/group_theory/perm/basic.lean

Lines changed: 27 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Leonardo de Moura, Mario Carneiro
66
import algebra.group.pi
77
import algebra.group.prod
88
import algebra.hom.iterate
9+
import logic.equiv.set
910

1011
/-!
1112
# The group of permutations (self-equivalences) of a type `α`
@@ -67,9 +68,9 @@ lemma inv_def (f : perm α) : f⁻¹ = f.symm := rfl
6768

6869
@[simp, norm_cast] lemma coe_one : ⇑(1 : perm α) = id := rfl
6970
@[simp, norm_cast] lemma coe_mul (f g : perm α) : ⇑(f * g) = f ∘ g := rfl
70-
7171
@[norm_cast] lemma coe_pow (f : perm α) (n : ℕ) : ⇑(f ^ n) = (f^[n]) :=
7272
hom_coe_pow _ rfl (λ _ _, rfl) _ _
73+
@[simp] lemma iterate_eq_pow (f : perm α) (n : ℕ) : (f^[n]) = ⇑(f ^ n) := (coe_pow _ _).symm
7374

7475
lemma eq_inv_iff_eq {f : perm α} {x y : α} : x = f⁻¹ y ↔ f x = y := f.eq_symm_apply
7576

@@ -79,9 +80,9 @@ lemma zpow_apply_comm {α : Type*} (σ : perm α) (m n : ℤ) {x : α} :
7980
(σ ^ m) ((σ ^ n) x) = (σ ^ n) ((σ ^ m) x) :=
8081
by rw [←equiv.perm.mul_apply, ←equiv.perm.mul_apply, zpow_mul_comm]
8182

82-
@[simp] lemma iterate_eq_pow (f : perm α) : ∀ n, f^[n] = ⇑(f ^ n)
83-
| 0 := rfl
84-
| (n + 1) := by { rw [function.iterate_succ, pow_add, iterate_eq_pow], refl }
83+
@[simp] lemma image_inv (f : perm α) (s : set α) : ⇑f⁻¹ '' s = f ⁻¹' s := f⁻¹.image_eq_preimage _
84+
@[simp] lemma preimage_inv (f : perm α) (s : set α) : ⇑f⁻¹ ⁻¹' s = f '' s :=
85+
(f.image_eq_preimage _).symm
8586

8687
/-! Lemmas about mixing `perm` with `equiv`. Because we have multiple ways to express
8788
`equiv.refl`, `equiv.symm`, and `equiv.trans`, we want simp lemmas for every combination.
@@ -494,3 +495,25 @@ lemma zpow_mul_right : ∀ n : ℤ, equiv.mul_right a ^ n = equiv.mul_right (a ^
494495

495496
end group
496497
end equiv
498+
499+
open equiv function
500+
501+
namespace set
502+
variables {α : Type*} {f : perm α} {s t : set α}
503+
504+
@[simp] lemma bij_on_perm_inv : bij_on ⇑f⁻¹ t s ↔ bij_on f s t := equiv.bij_on_symm
505+
506+
alias bij_on_perm_inv ↔ bij_on.of_perm_inv bij_on.perm_inv
507+
508+
lemma maps_to.perm_pow : maps_to f s s → ∀ n : ℕ, maps_to ⇑(f ^ n) s s :=
509+
by { simp_rw equiv.perm.coe_pow, exact maps_to.iterate }
510+
lemma surj_on.perm_pow : surj_on f s s → ∀ n : ℕ, surj_on ⇑(f ^ n) s s :=
511+
by { simp_rw equiv.perm.coe_pow, exact surj_on.iterate }
512+
lemma bij_on.perm_pow : bij_on f s s → ∀ n : ℕ, bij_on ⇑(f ^ n) s s :=
513+
by { simp_rw equiv.perm.coe_pow, exact bij_on.iterate }
514+
515+
lemma bij_on.perm_zpow (hf : bij_on f s s) : ∀ n : ℤ, bij_on ⇑(f ^ n) s s
516+
| (int.of_nat n) := hf.perm_pow _
517+
| (int.neg_succ_of_nat n) := by { rw zpow_neg_succ_of_nat, exact (hf.perm_pow _).perm_inv }
518+
519+
end set

0 commit comments

Comments
 (0)