@@ -139,8 +139,8 @@ def ghas_one.of_add_submonoids [semiring R] [has_zero ι]
139
139
ghas_one (λ i, carriers i) :=
140
140
{ one := ⟨1 , one_mem⟩ }
141
141
142
+ -- `@[simps]` doesn't generate a useful lemma, so we state one manually below.
142
143
/-- Build a `ghas_mul` instance for a collection of `add_submonoids`. -/
143
- @[simps mul]
144
144
def ghas_mul.of_add_submonoids [semiring R] [has_add ι]
145
145
(carriers : ι → add_submonoid R)
146
146
(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 ι]
153
153
map_add' := λ _ _, add_monoid_hom.ext $ λ _, subtype.ext (add_mul _ _ _),
154
154
map_zero' := add_monoid_hom.ext $ λ _, subtype.ext (zero_mul _) }, }
155
155
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
+
156
162
/-- Build a `gmonoid` instance for a collection of `add_submonoid`s. -/
157
163
@[simps to_ghas_one to_ghas_mul]
158
164
def gmonoid.of_add_submonoids [semiring R] [add_monoid ι]
@@ -187,14 +193,20 @@ def ghas_one.of_add_subgroups [ring R] [has_zero ι]
187
193
ghas_one (λ i, carriers i) :=
188
194
ghas_one.of_add_submonoids (λ i, (carriers i).to_add_submonoid) one_mem
189
195
196
+ -- `@[simps]` doesn't generate a useful lemma, so we state one manually below.
190
197
/-- Build a `ghas_mul` instance for a collection of `add_subgroup`s. -/
191
- @[simps mul]
192
198
def ghas_mul.of_add_subgroups [ring R] [has_add ι]
193
199
(carriers : ι → add_subgroup R)
194
200
(mul_mem : ∀ ⦃i j⦄ (gi : carriers i) (gj : carriers j), (gi * gj : R) ∈ carriers (i + j)) :
195
201
ghas_mul (λ i, carriers i) :=
196
202
ghas_mul.of_add_submonoids (λ i, (carriers i).to_add_submonoid) mul_mem
197
203
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
+
198
210
/-- Build a `gmonoid` instance for a collection of `add_subgroup`s. -/
199
211
@[simps to_ghas_one to_ghas_mul]
200
212
def gmonoid.of_add_subgroups [ring R] [add_monoid ι]
@@ -226,15 +238,22 @@ def ghas_one.of_submodules
226
238
ghas_one (λ i, carriers i) :=
227
239
ghas_one.of_add_submonoids (λ i, (carriers i).to_add_submonoid) one_mem
228
240
241
+ -- `@[simps]` doesn't generate a useful lemma, so we state one manually below.
229
242
/-- Build a `ghas_mul` instance for a collection of `submodule`s. -/
230
- @[simps mul]
231
243
def ghas_mul.of_submodules
232
244
[comm_semiring R] [semiring A] [algebra R A] [has_add ι]
233
245
(carriers : ι → submodule R A)
234
246
(mul_mem : ∀ ⦃i j⦄ (gi : carriers i) (gj : carriers j), (gi * gj : A) ∈ carriers (i + j)) :
235
247
ghas_mul (λ i, carriers i) :=
236
248
ghas_mul.of_add_submonoids (λ i, (carriers i).to_add_submonoid) mul_mem
237
249
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
+
238
257
/-- Build a `gmonoid` instance for a collection of `submodules`s. -/
239
258
@[simps to_ghas_one to_ghas_mul]
240
259
def gmonoid.of_submodules
0 commit comments