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

Commit 45ef183

Browse files
committed
feat(group_theory/subgroup/basic): The center can be written as the intersection of centralizers (#16837)
This PRs adds a lemma that writes the center as the intersection of centralizers. I kept `S` explicit so that you can `rw center_eq_infi S` (specifying `S` without having to immediately supply a proof of `hS`).
1 parent 0fbd703 commit 45ef183

File tree

1 file changed

+8
-0
lines changed

1 file changed

+8
-0
lines changed

src/group_theory/subgroup/basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2772,6 +2772,14 @@ le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ su
27722772
$ le_centralizer_iff.1 $ (closure_le _).2
27732773
$ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _
27742774

2775+
@[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) :
2776+
center G = ⨅ g ∈ S, centralizer (zpowers g) :=
2777+
by rw [←centralizer_top, ←hS, centralizer_closure]
2778+
2779+
@[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) :
2780+
center G = ⨅ g : S, centralizer (zpowers g) :=
2781+
by rw [center_eq_infi S hS, ←infi_subtype'']
2782+
27752783
end subgroup
27762784

27772785
namespace monoid_hom

0 commit comments

Comments
 (0)