@@ -12,7 +12,7 @@ import Mathlib.Algebra.Category.MonCat.Basic
12
12
import Mathlib.CategoryTheory.Endomorphism
13
13
14
14
/-!
15
- # Category instances for group, add_group, comm_group , and add_comm_group .
15
+ # Category instances for Group, AddGroup, CommGroup , and AddCommGroup .
16
16
17
17
We introduce the bundled categories:
18
18
* `GroupCat`
@@ -81,7 +81,7 @@ set_option linter.uppercaseLean3 false in
81
81
set_option linter.uppercaseLean3 false in
82
82
#align AddGroup.of_hom AddGroupCat.ofHom
83
83
84
- /-- Typecheck a `add_monoid_hom ` as a morphism in `AddGroup`. -/
84
+ /-- Typecheck a `AddMonoidHom ` as a morphism in `AddGroup`. -/
85
85
add_decl_doc AddGroupCat.ofHom
86
86
87
87
-- porting note: this instance was not necessary in mathlib
@@ -373,9 +373,10 @@ lemma Hom.map_mul {X Y : CommGroupCat} (f : X ⟶ Y) (x y : X) : f (x * y) = f x
373
373
lemma Hom.map_one {X Y : CommGroupCat} (f : X ⟶ Y) : f (1 : X) = 1 := by
374
374
apply MonoidHom.map_one (show MonoidHom X Y from f)
375
375
376
+ -- Porting note: is this still relevant?
376
377
-- This example verifies an improvement possible in Lean 3.8.
377
- -- Before that, to have `monoid_hom .map_map` usable by `simp` here,
378
- -- we had to mark all the concrete category `has_coe_to_sort ` instances reducible.
378
+ -- Before that, to have `MonoidHom .map_map` usable by `simp` here,
379
+ -- we had to mark all the concrete category `CoeSort ` instances reducible.
379
380
-- Now, it just works.
380
381
@[to_additive]
381
382
example {R S : CommGroupCat} (i : R ⟶ S) (r : R) (h : r = 1 ) : i r = 1 := by simp [h]
@@ -386,7 +387,7 @@ namespace AddCommGroupCat
386
387
387
388
-- Note that because `ℤ : Type 0`, this forces `G : AddCommGroup.{0}`,
388
389
-- so we write this explicitly to be clear.
389
- -- TODO generalize this, requiring a `ulift_instances .lean` file
390
+ -- TODO generalize this, requiring a `ULiftInstances .lean` file
390
391
/-- Any element of an abelian group gives a unique morphism from `ℤ` sending
391
392
`1` to that element. -/
392
393
def asHom {G : AddCommGroupCat.{0 }} (g : G) : AddCommGroupCat.of ℤ ⟶ G :=
@@ -442,8 +443,7 @@ add_decl_doc AddEquiv.toAddGroupCatIso
442
443
/-- Build an isomorphism in the category `CommGroupCat` from a `MulEquiv`
443
444
between `CommGroup`s. -/
444
445
@[to_additive (attr := simps)]
445
- def MulEquiv.toCommGroupCatIso {X Y : CommGroupCat} (e : X ≃* Y) : X ≅ Y
446
- where
446
+ def MulEquiv.toCommGroupCatIso {X Y : CommGroupCat} (e : X ≃* Y) : X ≅ Y where
447
447
hom := CommGroupCat.Hom.mk e.toMonoidHom
448
448
inv := CommGroupCat.Hom.mk e.symm.toMonoidHom
449
449
set_option linter.uppercaseLean3 false in
@@ -501,8 +501,7 @@ add_decl_doc addEquivIsoAddGroupIso
501
501
/-- multiplicative equivalences between `comm_group`s are the same as (isomorphic to) isomorphisms
502
502
in `CommGroup` -/
503
503
@[to_additive]
504
- def mulEquivIsoCommGroupIso {X Y : CommGroupCat.{u}} : X ≃* Y ≅ X ≅ Y
505
- where
504
+ def mulEquivIsoCommGroupIso {X Y : CommGroupCat.{u}} : X ≃* Y ≅ X ≅ Y where
506
505
hom e := e.toCommGroupCatIso
507
506
inv i := i.commGroupIsoToMulEquiv
508
507
set_option linter.uppercaseLean3 false in
0 commit comments