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

Commit 21cc806

Browse files
committed
feat(data/equiv/basic): perm_congr group lemmas (#6982)
Co-authored-by: Yakov Pechersky <pechersky@users.noreply.github.com>
1 parent 3365c44 commit 21cc806

File tree

2 files changed

+24
-5
lines changed

2 files changed

+24
-5
lines changed

src/data/equiv/basic.lean

Lines changed: 19 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -317,22 +317,36 @@ by { ext, refl }
317317
@[simp] lemma equiv_congr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) :
318318
ab.equiv_congr cd e x = cd (e (ab.symm x)) := rfl
319319

320+
section perm_congr
321+
322+
variables {α' β' : Type*} (e : α' ≃ β')
323+
320324
/-- If `α` is equivalent to `β`, then `perm α` is equivalent to `perm β`. -/
321-
def perm_congr {α : Type*} {β : Type*} (e : α ≃ β) : perm α ≃ perm β :=
325+
def perm_congr : perm α' ≃ perm β' :=
322326
equiv_congr e e
323327

324-
lemma perm_congr_def {α β : Type*} (e : α ≃ β) (p : equiv.perm α) :
328+
lemma perm_congr_def (p : equiv.perm α') :
325329
e.perm_congr p = (e.symm.trans p).trans e := rfl
326330

327-
@[simp] lemma perm_congr_symm {α β : Type*} (e : α ≃ β) :
331+
@[simp] lemma perm_congr_refl :
332+
e.perm_congr (equiv.refl _) = equiv.refl _ :=
333+
by simp [perm_congr_def]
334+
335+
@[simp] lemma perm_congr_symm :
328336
e.perm_congr.symm = e.symm.perm_congr := rfl
329337

330-
@[simp] lemma perm_congr_apply {α β : Type*} (e : α ≃ β) (p : equiv.perm α) (x) :
338+
@[simp] lemma perm_congr_apply (p : equiv.perm α') (x) :
331339
e.perm_congr p x = e (p (e.symm x)) := rfl
332340

333-
lemma perm_congr_symm_apply {α β : Type*} (e : α ≃ β) (p : equiv.perm β) (x) :
341+
lemma perm_congr_symm_apply (p : equiv.perm β') (x) :
334342
e.perm_congr.symm p x = e.symm (p (e x)) := rfl
335343

344+
lemma perm_congr_trans (p p' : equiv.perm α') :
345+
(e.perm_congr p).trans (e.perm_congr p') = e.perm_congr (p.trans p') :=
346+
by { ext, simp }
347+
348+
end perm_congr
349+
336350
protected lemma image_eq_preimage {α β} (e : α ≃ β) (s : set α) : e '' s = e.symm ⁻¹' s :=
337351
set.ext $ assume x, set.mem_image_iff_of_inverse e.left_inv e.right_inv
338352

src/group_theory/perm/basic.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -176,6 +176,11 @@ begin
176176
simpa using equiv.congr_fun h i
177177
end
178178

179+
/-- If `e` is also a permutation, we can write `perm_congr`
180+
completely in terms of the group structure. -/
181+
@[simp] lemma perm_congr_eq_mul (e p : perm α) :
182+
e.perm_congr p = e * p * e⁻¹ := rfl
183+
179184
/-- If the permutation `f` fixes the subtype `{x // p x}`, then this returns the permutation
180185
on `{x // p x}` induced by `f`. -/
181186
def subtype_perm (f : perm α) {p : α → Prop} (h : ∀ x, p x ↔ p (f x)) : perm {x // p x} :=

0 commit comments

Comments
 (0)