Skip to content

Commit

Permalink
feat(group_theory/subgroup/basic): {multiset_,}noncomm_prod_mem (#12523)
Browse files Browse the repository at this point in the history
and same for submonoids.
  • Loading branch information
nomeata committed Mar 12, 2022
1 parent 6ee6203 commit 1eaf499
Show file tree
Hide file tree
Showing 2 changed files with 36 additions and 0 deletions.
13 changes: 13 additions & 0 deletions src/group_theory/subgroup/basic.lean
Expand Up @@ -292,6 +292,12 @@ is in the `add_subgroup`."]
lemma multiset_prod_mem {G} [comm_group G] (K : subgroup G) (g : multiset G) :
(∀ a ∈ g, a ∈ K) → g.prod ∈ K := K.to_submonoid.multiset_prod_mem g

@[to_additive]
lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G)
(comm : ∀ (x ∈ g) (y ∈ g), commute x y) :
(∀ a ∈ g, a ∈ K) → g.noncomm_prod comm ∈ K :=
K.to_submonoid.multiset_noncomm_prod_mem g comm

/-- Product of elements of a subgroup of a `comm_group` indexed by a `finset` is in the
subgroup. -/
@[to_additive "Sum of elements in an `add_subgroup` of an `add_comm_group` indexed by a `finset`
Expand All @@ -301,6 +307,13 @@ lemma prod_mem {G : Type*} [comm_group G] (K : subgroup G)
∏ c in t, f c ∈ K :=
K.to_submonoid.prod_mem h

@[to_additive]
lemma noncomm_prod_mem (K : subgroup G)
{ι : Type*} {t : finset ι} {f : ι → G} (comm : ∀ (x ∈ t) (y ∈ t), commute (f x) (f y)) :
(∀ c ∈ t, f c ∈ K) → t.noncomm_prod f comm ∈ K :=
K.to_submonoid.noncomm_prod_mem t f comm


@[to_additive add_subgroup.nsmul_mem]
lemma pow_mem {x : G} (hx : x ∈ K) : ∀ n : ℕ, x ^ n ∈ K := K.to_submonoid.pow_mem hx

Expand Down
23 changes: 23 additions & 0 deletions src/group_theory/submonoid/membership.lean
Expand Up @@ -7,6 +7,7 @@ Amelia Livingston, Yury Kudryashov
import group_theory.submonoid.operations
import algebra.big_operators.basic
import algebra.free_monoid
import data.finset.noncomm_prod

/-!
# Submonoids: membership criteria
Expand Down Expand Up @@ -65,6 +66,16 @@ lemma multiset_prod_mem {M} [comm_monoid M] (S : submonoid M) (m : multiset M)
(hm : ∀ a ∈ m, a ∈ S) : m.prod ∈ S :=
by { lift m to multiset S using hm, rw ← coe_multiset_prod, exact m.prod.coe_prop }

@[to_additive]
lemma multiset_noncomm_prod_mem (S : submonoid M) (m : multiset M)
(comm : ∀ (x ∈ m) (y ∈ m), commute x y) (h : ∀ (x ∈ m), x ∈ S) :
m.noncomm_prod comm ∈ S :=
begin
induction m using quotient.induction_on with l,
simp only [multiset.quot_mk_to_coe, multiset.noncomm_prod_coe],
exact submonoid.list_prod_mem _ h,
end

/-- Product of elements of a submonoid of a `comm_monoid` indexed by a `finset` is in the
submonoid. -/
@[to_additive "Sum of elements in an `add_submonoid` of an `add_comm_monoid` indexed by a `finset`
Expand All @@ -74,6 +85,18 @@ lemma prod_mem {M : Type*} [comm_monoid M] (S : submonoid M)
∏ c in t, f c ∈ S :=
S.multiset_prod_mem (t.1.map f) $ λ x hx, let ⟨i, hi, hix⟩ := multiset.mem_map.1 hx in hix ▸ h i hi

@[to_additive]
lemma noncomm_prod_mem (S : submonoid M) {ι : Type*} (t : finset ι) (f : ι → M)
(comm : ∀ (x ∈ t) (y ∈ t), commute (f x) (f y)) (h : ∀ c ∈ t, f c ∈ S) :
t.noncomm_prod f comm ∈ S :=
begin
apply multiset_noncomm_prod_mem,
intro y,
rw multiset.mem_map,
rintros ⟨x, ⟨hx, rfl⟩⟩,
exact h x hx,
end

end assoc

section non_assoc
Expand Down

0 comments on commit 1eaf499

Please sign in to comment.