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

Commit e3de4e3

Browse files
fix(algebra/direct_sum_graded): replace badly-stated and slow simps lemmas with manual ones (#7403)
[Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/simps.20is.20very.20slow/near/236636962). The `simps mul` attribute on `direct_sum.ghas_mul.of_add_subgroups` was taking 44s, only to produce a lemma that wasn't very useful anyway. Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent aa37eee commit e3de4e3

File tree

1 file changed

+22
-3
lines changed

1 file changed

+22
-3
lines changed

src/algebra/direct_sum_graded.lean

Lines changed: 22 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -139,8 +139,8 @@ def ghas_one.of_add_submonoids [semiring R] [has_zero ι]
139139
ghas_one (λ i, carriers i) :=
140140
{ one := ⟨1, one_mem⟩ }
141141

142+
-- `@[simps]` doesn't generate a useful lemma, so we state one manually below.
142143
/-- Build a `ghas_mul` instance for a collection of `add_submonoids`. -/
143-
@[simps mul]
144144
def ghas_mul.of_add_submonoids [semiring R] [has_add ι]
145145
(carriers : ι → add_submonoid R)
146146
(mul_mem : ∀ ⦃i j⦄ (gi : carriers i) (gj : carriers j), (gi * gj : R) ∈ carriers (i + j)) :
@@ -153,6 +153,12 @@ def ghas_mul.of_add_submonoids [semiring R] [has_add ι]
153153
map_add' := λ _ _, add_monoid_hom.ext $ λ _, subtype.ext (add_mul _ _ _),
154154
map_zero' := add_monoid_hom.ext $ λ _, subtype.ext (zero_mul _) }, }
155155

156+
-- `@[simps]` doesn't generate this well
157+
@[simp] lemma ghas_mul.of_add_submonoids_mul [semiring R] [has_add ι]
158+
(carriers : ι → add_submonoid R) (mul_mem) {i j} (a : carriers i) (b : carriers j) :
159+
@ghas_mul.mul _ _ _ _ _ (ghas_mul.of_add_submonoids carriers mul_mem) i j a b =
160+
⟨a * b, mul_mem a b⟩ := rfl
161+
156162
/-- Build a `gmonoid` instance for a collection of `add_submonoid`s. -/
157163
@[simps to_ghas_one to_ghas_mul]
158164
def gmonoid.of_add_submonoids [semiring R] [add_monoid ι]
@@ -187,14 +193,20 @@ def ghas_one.of_add_subgroups [ring R] [has_zero ι]
187193
ghas_one (λ i, carriers i) :=
188194
ghas_one.of_add_submonoids (λ i, (carriers i).to_add_submonoid) one_mem
189195

196+
-- `@[simps]` doesn't generate a useful lemma, so we state one manually below.
190197
/-- Build a `ghas_mul` instance for a collection of `add_subgroup`s. -/
191-
@[simps mul]
192198
def ghas_mul.of_add_subgroups [ring R] [has_add ι]
193199
(carriers : ι → add_subgroup R)
194200
(mul_mem : ∀ ⦃i j⦄ (gi : carriers i) (gj : carriers j), (gi * gj : R) ∈ carriers (i + j)) :
195201
ghas_mul (λ i, carriers i) :=
196202
ghas_mul.of_add_submonoids (λ i, (carriers i).to_add_submonoid) mul_mem
197203

204+
-- `@[simps]` doesn't generate this well
205+
@[simp] lemma ghas_mul.of_add_subgroups_mul [ring R] [has_add ι]
206+
(carriers : ι → add_subgroup R) (mul_mem) {i j} (a : carriers i) (b : carriers j) :
207+
@ghas_mul.mul _ _ _ _ _ (ghas_mul.of_add_subgroups carriers mul_mem) i j a b =
208+
⟨a * b, mul_mem a b⟩ := rfl
209+
198210
/-- Build a `gmonoid` instance for a collection of `add_subgroup`s. -/
199211
@[simps to_ghas_one to_ghas_mul]
200212
def gmonoid.of_add_subgroups [ring R] [add_monoid ι]
@@ -226,15 +238,22 @@ def ghas_one.of_submodules
226238
ghas_one (λ i, carriers i) :=
227239
ghas_one.of_add_submonoids (λ i, (carriers i).to_add_submonoid) one_mem
228240

241+
-- `@[simps]` doesn't generate a useful lemma, so we state one manually below.
229242
/-- Build a `ghas_mul` instance for a collection of `submodule`s. -/
230-
@[simps mul]
231243
def ghas_mul.of_submodules
232244
[comm_semiring R] [semiring A] [algebra R A] [has_add ι]
233245
(carriers : ι → submodule R A)
234246
(mul_mem : ∀ ⦃i j⦄ (gi : carriers i) (gj : carriers j), (gi * gj : A) ∈ carriers (i + j)) :
235247
ghas_mul (λ i, carriers i) :=
236248
ghas_mul.of_add_submonoids (λ i, (carriers i).to_add_submonoid) mul_mem
237249

250+
-- `@[simps]` doesn't generate this well
251+
@[simp] lemma ghas_mul.of_submodules_mul
252+
[comm_semiring R] [semiring A] [algebra R A] [has_add ι]
253+
(carriers : ι → submodule R A) (mul_mem) {i j} (a : carriers i) (b : carriers j) :
254+
@ghas_mul.mul _ _ _ _ _ (ghas_mul.of_submodules carriers mul_mem) i j a b =
255+
⟨a * b, mul_mem a b⟩ := rfl
256+
238257
/-- Build a `gmonoid` instance for a collection of `submodules`s. -/
239258
@[simps to_ghas_one to_ghas_mul]
240259
def gmonoid.of_submodules

0 commit comments

Comments
 (0)