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

Commit 242d687

Browse files
committed
feat(algebra/hom/group and *): introduce mul_hom M N notation M →ₙ* N (#13526)
The discussion and poll related to this new notation can be found in this [Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/ring_hom.20notation.20and.20friends/near/279313301)
1 parent 7bfaa5c commit 242d687

File tree

18 files changed

+101
-96
lines changed

18 files changed

+101
-96
lines changed

src/algebra/category/Semigroup/basic.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -56,13 +56,13 @@ def of (M : Type u) [has_mul M] : Magma := bundled.of M
5656
add_decl_doc AddMagma.of
5757

5858
/-- Typecheck a `mul_hom` as a morphism in `Magma`. -/
59-
@[to_additive] def of_hom {X Y : Type u} [has_mul X] [has_mul Y] (f : mul_hom X Y) :
59+
@[to_additive] def of_hom {X Y : Type u} [has_mul X] [has_mul Y] (f : X →ₙ* Y) :
6060
of X ⟶ of Y := f
6161

6262
/-- Typecheck a `add_hom` as a morphism in `AddMagma`. -/
6363
add_decl_doc AddMagma.of_hom
6464

65-
@[simp, to_additive] lemma of_hom_apply {X Y : Type u} [has_mul X] [has_mul Y] (f : mul_hom X Y)
65+
@[simp, to_additive] lemma of_hom_apply {X Y : Type u} [has_mul X] [has_mul Y] (f : X →ₙ* Y)
6666
(x : X) : of_hom f x = f x := rfl
6767

6868
@[to_additive]
@@ -100,13 +100,13 @@ def of (M : Type u) [semigroup M] : Semigroup := bundled.of M
100100
add_decl_doc AddSemigroup.of
101101

102102
/-- Typecheck a `mul_hom` as a morphism in `Semigroup`. -/
103-
@[to_additive] def of_hom {X Y : Type u} [semigroup X] [semigroup Y] (f : mul_hom X Y) :
103+
@[to_additive] def of_hom {X Y : Type u} [semigroup X] [semigroup Y] (f : X →ₙ* Y) :
104104
of X ⟶ of Y := f
105105

106106
/-- Typecheck a `add_hom` as a morphism in `AddSemigroup`. -/
107107
add_decl_doc AddSemigroup.of_hom
108108

109-
@[simp, to_additive] lemma of_hom_apply {X Y : Type u} [semigroup X] [semigroup Y] (f : mul_hom X Y)
109+
@[simp, to_additive] lemma of_hom_apply {X Y : Type u} [semigroup X] [semigroup Y] (f : X →ₙ* Y)
110110
(x : X) : of_hom f x = f x := rfl
111111

112112
@[to_additive]

src/algebra/divisibility.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -77,7 +77,7 @@ variables {M N : Type*} [monoid M] [monoid N]
7777
lemma map_dvd {F : Type*} [mul_hom_class F M N] (f : F) {a b} : a ∣ b → f a ∣ f b
7878
| ⟨c, h⟩ := ⟨f c, h.symm ▸ map_mul f a c⟩
7979

80-
lemma mul_hom.map_dvd (f : mul_hom M N) {a b} : a ∣ b → f a ∣ f b := map_dvd f
80+
lemma mul_hom.map_dvd (f : M →ₙ* N) {a b} : a ∣ b → f a ∣ f b := map_dvd f
8181

8282
lemma monoid_hom.map_dvd (f : M →* N) {a b} : a ∣ b → f a ∣ f b := map_dvd f
8383

src/algebra/free.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -82,13 +82,13 @@ namespace free_magma
8282
variables {α : Type u} {β : Type v} [has_mul β] (f : α → β)
8383

8484
@[to_additive]
85-
theorem lift_aux_unique (F : mul_hom (free_magma α) β) : ⇑F = lift_aux (F ∘ of) :=
85+
theorem lift_aux_unique (F : free_magma α →ₙ* β) : ⇑F = lift_aux (F ∘ of) :=
8686
funext $ λ x, free_magma.rec_on x (λ x, rfl) $ λ x y ih1 ih2,
8787
(F.map_mul x y).trans $ congr (congr_arg _ ih1) ih2
8888

8989
/-- The universal property of the free magma expressing its adjointness. -/
9090
@[to_additive "The universal property of the free additive magma expressing its adjointness."]
91-
def lift : (α → β) ≃ mul_hom (free_magma α) β :=
91+
def lift : (α → β) ≃ (free_magma α →ₙ* β) :=
9292
{ to_fun := λ f,
9393
{ to_fun := lift_aux f,
9494
map_mul' := λ x y, rfl, },

src/algebra/group/pi.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -126,7 +126,7 @@ end pi
126126
namespace mul_hom
127127

128128
@[to_additive] lemma coe_mul {M N} {mM : has_mul M} {mN : comm_semigroup N}
129-
(f g : mul_hom M N) :
129+
(f g : M →ₙ* N) :
130130
(f * g : M → N) = λ x, f x * g x := rfl
131131

132132
end mul_hom
@@ -219,7 +219,7 @@ lemma monoid_hom.single_apply [Π i, mul_one_class $ f i] (i : I) (x : f i) :
219219
into a dependent family of `mul_zero_class`es, as functions supported at a point.
220220
221221
This is the `mul_hom` version of `pi.single`. -/
222-
@[simps] def mul_hom.single [Π i, mul_zero_class $ f i] (i : I) : mul_hom (f i) (Π i, f i) :=
222+
@[simps] def mul_hom.single [Π i, mul_zero_class $ f i] (i : I) : (f i) →ₙ* (Π i, f i) :=
223223
{ to_fun := single i,
224224
map_mul' := pi.single_op₂ (λ _, (*)) (λ _, zero_mul _) _, }
225225

src/algebra/group/prod.lean

Lines changed: 18 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -194,42 +194,42 @@ variables (M N) [has_mul M] [has_mul N] [has_mul P]
194194
/-- Given magmas `M`, `N`, the natural projection homomorphism from `M × N` to `M`.-/
195195
@[to_additive "Given additive magmas `A`, `B`, the natural projection homomorphism
196196
from `A × B` to `A`"]
197-
def fst : mul_hom (M × N) M := ⟨prod.fst, λ _ _, rfl⟩
197+
def fst : (M × N) →ₙ* M := ⟨prod.fst, λ _ _, rfl⟩
198198

199199
/-- Given magmas `M`, `N`, the natural projection homomorphism from `M × N` to `N`.-/
200200
@[to_additive "Given additive magmas `A`, `B`, the natural projection homomorphism
201201
from `A × B` to `B`"]
202-
def snd : mul_hom (M × N) N := ⟨prod.snd, λ _ _, rfl⟩
202+
def snd : (M × N) →ₙ* N := ⟨prod.snd, λ _ _, rfl⟩
203203

204204
variables {M N}
205205

206206
@[simp, to_additive] lemma coe_fst : ⇑(fst M N) = prod.fst := rfl
207207
@[simp, to_additive] lemma coe_snd : ⇑(snd M N) = prod.snd := rfl
208208

209-
/-- Combine two `monoid_hom`s `f : mul_hom M N`, `g : mul_hom M P` into
210-
`f.prod g : mul_hom M (N × P)` given by `(f.prod g) x = (f x, g x)`. -/
209+
/-- Combine two `monoid_hom`s `f : M →ₙ* N`, `g : M →ₙ* P` into
210+
`f.prod g : M →ₙ* (N × P)` given by `(f.prod g) x = (f x, g x)`. -/
211211
@[to_additive prod "Combine two `add_monoid_hom`s `f : add_hom M N`, `g : add_hom M P` into
212212
`f.prod g : add_hom M (N × P)` given by `(f.prod g) x = (f x, g x)`"]
213-
protected def prod (f : mul_hom M N) (g : mul_hom M P) : mul_hom M (N × P) :=
213+
protected def prod (f : M →ₙ* N) (g : M →ₙ* P) : M →ₙ* (N × P) :=
214214
{ to_fun := pi.prod f g,
215215
map_mul' := λ x y, prod.ext (f.map_mul x y) (g.map_mul x y) }
216216

217217
@[to_additive coe_prod]
218-
lemma coe_prod (f : mul_hom M N) (g : mul_hom M P) : ⇑(f.prod g) = pi.prod f g := rfl
218+
lemma coe_prod (f : M →ₙ* N) (g : M →ₙ* P) : ⇑(f.prod g) = pi.prod f g := rfl
219219

220220
@[simp, to_additive prod_apply]
221-
lemma prod_apply (f : mul_hom M N) (g : mul_hom M P) (x) : f.prod g x = (f x, g x) := rfl
221+
lemma prod_apply (f : M →ₙ* N) (g : M →ₙ* P) (x) : f.prod g x = (f x, g x) := rfl
222222

223223
@[simp, to_additive fst_comp_prod]
224-
lemma fst_comp_prod (f : mul_hom M N) (g : mul_hom M P) : (fst N P).comp (f.prod g) = f :=
224+
lemma fst_comp_prod (f : M →ₙ* N) (g : M →ₙ* P) : (fst N P).comp (f.prod g) = f :=
225225
ext $ λ x, rfl
226226

227227
@[simp, to_additive snd_comp_prod]
228-
lemma snd_comp_prod (f : mul_hom M N) (g : mul_hom M P) : (snd N P).comp (f.prod g) = g :=
228+
lemma snd_comp_prod (f : M →ₙ* N) (g : M →ₙ* P) : (snd N P).comp (f.prod g) = g :=
229229
ext $ λ x, rfl
230230

231231
@[simp, to_additive prod_unique]
232-
lemma prod_unique (f : mul_hom M (N × P)) :
232+
lemma prod_unique (f : M →ₙ* (N × P)) :
233233
((fst N P).comp f).prod ((snd N P).comp f) = f :=
234234
ext $ λ x, by simp only [prod_apply, coe_fst, coe_snd, comp_apply, prod.mk.eta]
235235

@@ -238,11 +238,11 @@ end prod
238238
section prod_map
239239

240240
variables {M' : Type*} {N' : Type*} [has_mul M] [has_mul N] [has_mul M'] [has_mul N'] [has_mul P]
241-
(f : mul_hom M M') (g : mul_hom N N')
241+
(f : M →ₙ* M') (g : N →ₙ* N')
242242

243243
/-- `prod.map` as a `monoid_hom`. -/
244244
@[to_additive prod_map "`prod.map` as an `add_monoid_hom`"]
245-
def prod_map : mul_hom (M × N) (M' × N') := (f.comp (fst M N)).prod (g.comp (snd M N))
245+
def prod_map : (M × N) →ₙ* (M' × N') := (f.comp (fst M N)).prod (g.comp (snd M N))
246246

247247
@[to_additive prod_map_def]
248248
lemma prod_map_def : prod_map f g = (f.comp (fst M N)).prod (g.comp (snd M N)) := rfl
@@ -251,29 +251,29 @@ lemma prod_map_def : prod_map f g = (f.comp (fst M N)).prod (g.comp (snd M N)) :
251251
lemma coe_prod_map : ⇑(prod_map f g) = prod.map f g := rfl
252252

253253
@[to_additive prod_comp_prod_map]
254-
lemma prod_comp_prod_map (f : mul_hom P M) (g : mul_hom P N)
255-
(f' : mul_hom M M') (g' : mul_hom N N') :
254+
lemma prod_comp_prod_map (f : P →ₙ* M) (g : P →ₙ* N)
255+
(f' : M →ₙ* M') (g' : N →ₙ* N') :
256256
(f'.prod_map g').comp (f.prod g) = (f'.comp f).prod (g'.comp g) :=
257257
rfl
258258

259259
end prod_map
260260

261261
section coprod
262262

263-
variables [has_mul M] [has_mul N] [comm_semigroup P] (f : mul_hom M P) (g : mul_hom N P)
263+
variables [has_mul M] [has_mul N] [comm_semigroup P] (f : M →ₙ* P) (g : N →ₙ* P)
264264

265265
/-- Coproduct of two `mul_hom`s with the same codomain:
266266
`f.coprod g (p : M × N) = f p.1 * g p.2`. -/
267267
@[to_additive "Coproduct of two `add_hom`s with the same codomain:
268268
`f.coprod g (p : M × N) = f p.1 + g p.2`."]
269-
def coprod : mul_hom (M × N) P := f.comp (fst M N) * g.comp (snd M N)
269+
def coprod : (M × N) →ₙ* P := f.comp (fst M N) * g.comp (snd M N)
270270

271271
@[simp, to_additive]
272272
lemma coprod_apply (p : M × N) : f.coprod g p = f p.1 * g p.2 := rfl
273273

274274
@[to_additive]
275275
lemma comp_coprod {Q : Type*} [comm_semigroup Q]
276-
(h : mul_hom P Q) (f : mul_hom M P) (g : mul_hom N P) :
276+
(h : P →ₙ* Q) (f : M →ₙ* P) (g : N →ₙ* P) :
277277
h.comp (f.coprod g) = (h.comp f).coprod (h.comp g) :=
278278
ext $ λ x, by simp
279279

@@ -475,7 +475,7 @@ variables {α : Type*}
475475

476476
/-- Multiplication as a multiplicative homomorphism. -/
477477
@[to_additive "Addition as an additive homomorphism.", simps]
478-
def mul_mul_hom [comm_semigroup α] : mul_hom (α × α) α :=
478+
def mul_mul_hom [comm_semigroup α] : (α × α) →ₙ* α :=
479479
{ to_fun := λ a, a.1 * a.2,
480480
map_mul' := λ a b, mul_mul_mul_comm _ _ _ _ }
481481

src/algebra/group/with_one.lean

Lines changed: 7 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -116,7 +116,7 @@ section
116116
local attribute [irreducible] with_one with_zero
117117
/-- `coe` as a bundled morphism -/
118118
@[to_additive "`coe` as a bundled morphism", simps apply]
119-
def coe_mul_hom [has_mul α] : mul_hom α (with_one α) :=
119+
def coe_mul_hom [has_mul α] : α →ₙ* (with_one α) :=
120120
{ to_fun := coe, map_mul' := λ x y, rfl }
121121

122122
end
@@ -127,7 +127,7 @@ variables [has_mul α] [mul_one_class β]
127127

128128
/-- Lift a semigroup homomorphism `f` to a bundled monoid homorphism. -/
129129
@[to_additive "Lift an add_semigroup homomorphism `f` to a bundled add_monoid homorphism."]
130-
def lift : mul_hom α β ≃ (with_one α →* β) :=
130+
def lift : (α →ₙ* β) ≃ (with_one α →* β) :=
131131
{ to_fun := λ f,
132132
{ to_fun := λ x, option.cases_on x 1 f,
133133
map_one' := rfl,
@@ -139,7 +139,7 @@ def lift : mul_hom α β ≃ (with_one α →* β) :=
139139
left_inv := λ f, mul_hom.ext $ λ x, rfl,
140140
right_inv := λ F, monoid_hom.ext $ λ x, with_one.cases_on x F.map_one.symm $ λ x, rfl }
141141

142-
variables (f : mul_hom α β)
142+
variables (f : α →ₙ* β)
143143

144144
@[simp, to_additive]
145145
lemma lift_coe (x : α) : lift f x = f x := rfl
@@ -163,23 +163,23 @@ variables [has_mul α] [has_mul β] [has_mul γ]
163163
from `with_one α` to `with_one β` -/
164164
@[to_additive "Given an additive map from `α → β` returns an add_monoid homomorphism
165165
from `with_zero α` to `with_zero β`"]
166-
def map (f : mul_hom α β) : with_one α →* with_one β :=
166+
def map (f : α →ₙ* β) : with_one α →* with_one β :=
167167
lift (coe_mul_hom.comp f)
168168

169-
@[simp, to_additive] lemma map_coe (f : mul_hom α β) (a : α) : map f (a : with_one α) = f a :=
169+
@[simp, to_additive] lemma map_coe (f : α →ₙ* β) (a : α) : map f (a : with_one α) = f a :=
170170
lift_coe _ _
171171

172172
@[simp, to_additive]
173173
lemma map_id : map (mul_hom.id α) = monoid_hom.id (with_one α) :=
174174
by { ext, induction x using with_one.cases_on; refl }
175175

176176
@[to_additive]
177-
lemma map_map (f : mul_hom α β) (g : mul_hom β γ) (x) :
177+
lemma map_map (f : α →ₙ* β) (g : β →ₙ* γ) (x) :
178178
map g (map f x) = map (g.comp f) x :=
179179
by { induction x using with_one.cases_on; refl }
180180

181181
@[simp, to_additive]
182-
lemma map_comp (f : mul_hom α β) (g : mul_hom β γ) :
182+
lemma map_comp (f : α →ₙ* β) (g : β →ₙ* γ) :
183183
map (g.comp f) = (map g).comp (map f) :=
184184
monoid_hom.ext $ λ x, (map_map f g x).symm
185185

src/algebra/hom/equiv.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -35,8 +35,8 @@ variables {F α β A B M N P Q G H : Type*}
3535

3636
/-- Makes a multiplicative inverse from a bijection which preserves multiplication. -/
3737
@[to_additive "Makes an additive inverse from a bijection which preserves addition."]
38-
def mul_hom.inverse [has_mul M] [has_mul N] (f : mul_hom M N) (g : N → M)
39-
(h₁ : function.left_inverse g f) (h₂ : function.right_inverse g f) : mul_hom N M :=
38+
def mul_hom.inverse [has_mul M] [has_mul N] (f : M →ₙ* N) (g : N → M)
39+
(h₁ : function.left_inverse g f) (h₂ : function.right_inverse g f) : N →ₙ* M :=
4040
{ to_fun := g,
4141
map_mul' := λ x y,
4242
calc g (x * y) = g (f (g x) * f (g y)) : by rw [h₂ x, h₂ y]
@@ -50,7 +50,7 @@ def monoid_hom.inverse {A B : Type*} [monoid A] [monoid B] (f : A →* B) (g : B
5050
B →* A :=
5151
{ to_fun := g,
5252
map_one' := by rw [← f.map_one, h₁],
53-
.. (f : mul_hom A B).inverse g h₁ h₂, }
53+
.. (f : A →ₙ* B).inverse g h₁ h₂, }
5454

5555
set_option old_structure_cmd true
5656

@@ -71,7 +71,7 @@ add_decl_doc add_equiv.to_add_hom
7171

7272
/-- `mul_equiv α β` is the type of an equiv `α ≃ β` which preserves multiplication. -/
7373
@[ancestor equiv mul_hom, to_additive]
74-
structure mul_equiv (M N : Type*) [has_mul M] [has_mul N] extends M ≃ N, mul_hom M N
74+
structure mul_equiv (M N : Type*) [has_mul M] [has_mul N] extends M ≃ N, M →ₙ* N
7575

7676
/-- The `equiv` underlying a `mul_equiv`. -/
7777
add_decl_doc mul_equiv.to_equiv

0 commit comments

Comments
 (0)