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

Commit 28c902d

Browse files
committed
fix(algebra/group/pi): Fix apply-simp-lemmas for monoid_hom.single (#12474)
so that the simp-normal form really is `pi.mul_single`. While adjusting related lemmas in `group_theory.subgroup.basic`, add a few missing `to_additive` attributes.
1 parent 64d953a commit 28c902d

File tree

2 files changed

+24
-11
lines changed

2 files changed

+24
-11
lines changed

src/algebra/group/pi.lean

Lines changed: 10 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -189,23 +189,31 @@ This is the `one_hom` version of `pi.mul_single`. -/
189189
@[to_additive zero_hom.single "The zero-preserving homomorphism including a single value
190190
into a dependent family of values, as functions supported at a point.
191191
192-
This is the `zero_hom` version of `pi.single`.", simps]
192+
This is the `zero_hom` version of `pi.single`."]
193193
def one_hom.single [Π i, has_one $ f i] (i : I) : one_hom (f i) (Π i, f i) :=
194194
{ to_fun := mul_single i,
195195
map_one' := mul_single_one i }
196196

197+
@[to_additive, simp]
198+
lemma one_hom.single_apply [Π i, has_one $ f i] (i : I) (x : f i) :
199+
one_hom.single f i x = mul_single i x := rfl
200+
197201
/-- The monoid homomorphism including a single monoid into a dependent family of additive monoids,
198202
as functions supported at a point.
199203
200204
This is the `monoid_hom` version of `pi.mul_single`. -/
201205
@[to_additive "The additive monoid homomorphism including a single additive
202206
monoid into a dependent family of additive monoids, as functions supported at a point.
203207
204-
This is the `add_monoid_hom` version of `pi.single`.", simps]
208+
This is the `add_monoid_hom` version of `pi.single`."]
205209
def monoid_hom.single [Π i, mul_one_class $ f i] (i : I) : f i →* Π i, f i :=
206210
{ map_mul' := mul_single_op₂ (λ _, (*)) (λ _, one_mul _) _,
207211
.. (one_hom.single f i) }
208212

213+
@[to_additive, simp]
214+
lemma monoid_hom.single_apply [Π i, mul_one_class $ f i] (i : I) (x : f i) :
215+
monoid_hom.single f i x = mul_single i x := rfl
216+
209217
/-- The multiplicative homomorphism including a single `mul_zero_class`
210218
into a dependent family of `mul_zero_class`es, as functions supported at a point.
211219

src/group_theory/subgroup/basic.lean

Lines changed: 14 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -1197,10 +1197,10 @@ begin
11971197
{ intros h x hx i hi, refine h i hi ⟨_, hx, rfl⟩, }
11981198
end
11991199

1200-
@[simp]
1201-
lemma single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)}
1200+
@[simp, to_additive]
1201+
lemma mul_single_mem_pi [decidable_eq η] {I : set η} {H : Π i, subgroup (f i)}
12021202
(i : η) (x : f i) :
1203-
monoid_hom.single f i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) :=
1203+
pi.mul_single i x ∈ pi I H ↔ (i ∈ I → x ∈ H i) :=
12041204
begin
12051205
split,
12061206
{ intros h hi, simpa using h i hi, },
@@ -1210,7 +1210,8 @@ begin
12101210
{ simp [heq, one_mem], }, }
12111211
end
12121212

1213-
lemma pi_mem_of_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) }
1213+
@[to_additive]
1214+
lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) }
12141215
(x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) :
12151216
x ∈ H :=
12161217
begin
@@ -1237,19 +1238,23 @@ begin
12371238
{ apply h2, simp, } }
12381239
end
12391240

1240-
lemma pi_mem_of_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i) } (x : Π i, f i)
1241-
(h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H :=
1242-
pi_mem_of_single_mem_aux finset.univ x (by simp) (λ i _, h i)
1241+
@[to_additive]
1242+
lemma pi_mem_of_mul_single_mem [fintype η] [decidable_eq η] {H : subgroup (Π i, f i)}
1243+
(x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H :=
1244+
pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i)
12431245

12441246
/-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/
1247+
@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the
1248+
additive groups."]
12451249
lemma pi_le_iff [decidable_eq η] [fintype η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} :
12461250
pi univ H ≤ J ↔ (∀ i : η, map (monoid_hom.single f i) (H i) ≤ J) :=
12471251
begin
12481252
split,
12491253
{ rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx },
1250-
{ exact λ h x hx, pi_mem_of_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), }
1254+
{ exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), }
12511255
end
12521256

1257+
@[to_additive]
12531258
lemma pi_eq_bot_iff (H : Π i, subgroup (f i)) :
12541259
pi set.univ H = ⊥ ↔ ∀ i, H i = ⊥ :=
12551260
begin
@@ -1258,7 +1263,7 @@ begin
12581263
split,
12591264
{ intros h i x hx,
12601265
have : monoid_hom.single f i x = 1 :=
1261-
h (monoid_hom.single f i x) ((single_mem_pi i x).mpr (λ _, hx)),
1266+
h (monoid_hom.single f i x) ((mul_single_mem_pi i x).mpr (λ _, hx)),
12621267
simpa using congr_fun this i, },
12631268
{ exact λ h x hx, funext (λ i, h _ _ (hx i trivial)), },
12641269
end

0 commit comments

Comments
 (0)