From 2e6f6dff0e3e5ea6b97bca1a10ebab73bd8131a7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 20 Jun 2024 02:11:25 +0000 Subject: [PATCH] chore: Rename to `Grp` (#3731) This is a proposal to rename what was in mathlib `Group` and in mathlib4 `GroupCat` to its literature name `Grp`. This has the advantage not to conflict with `group` that has been capitalised to `Group` and to be shorter. --- Mathlib.lean | 46 +-- .../GroupCat/EquivalenceGroupAddGroup.lean | 94 ----- .../Category/{GroupCat => Grp}/Abelian.lean | 28 +- .../{GroupCat => Grp}/Adjunctions.lean | 84 ++-- .../Category/{GroupCat => Grp}/Basic.lean | 370 +++++++++--------- .../{GroupCat => Grp}/Biproducts.lean | 80 ++-- .../Category/{GroupCat => Grp}/Colimits.lean | 88 ++--- .../{GroupCat => Grp}/EnoughInjectives.lean | 22 +- .../Category/{GroupCat => Grp}/EpiMono.lean | 164 ++++---- .../Grp/EquivalenceGroupAddGroup.lean | 94 +++++ .../{GroupCat => Grp}/FilteredColimits.lean | 190 ++++----- .../ForgetCorepresentable.lean | 56 +-- .../Category/{GroupCat => Grp}/Images.lean | 46 +-- .../Category/{GroupCat => Grp}/Injective.lean | 20 +- .../Category/{GroupCat => Grp}/Kernels.lean | 10 +- .../Category/{GroupCat => Grp}/Limits.lean | 368 ++++++++--------- .../{GroupCat => Grp}/Preadditive.lean | 14 +- .../Category/{GroupCat => Grp}/Subobject.lean | 12 +- .../{GroupCat => Grp}/ZModuleEquivalence.lean | 12 +- .../Category/{GroupCat => Grp}/Zero.lean | 30 +- ...GroupWithZeroCat.lean => GrpWithZero.lean} | 48 +-- .../Algebra/Category/ModuleCat/Abelian.lean | 4 +- Mathlib/Algebra/Category/ModuleCat/Basic.lean | 52 +-- .../Category/ModuleCat/ChangeOfRings.lean | 14 +- .../Algebra/Category/ModuleCat/Colimits.lean | 58 +-- .../Category/ModuleCat/FilteredColimits.lean | 36 +- .../Algebra/Category/ModuleCat/Injective.lean | 6 +- .../Algebra/Category/ModuleCat/Limits.lean | 24 +- .../Algebra/Category/ModuleCat/Presheaf.lean | 12 +- .../Category/ModuleCat/Presheaf/Colimits.lean | 6 +- .../Category/ModuleCat/Presheaf/Limits.lean | 2 +- .../ModuleCat/Presheaf/Sheafification.lean | 20 +- .../Category/ModuleCat/Presheaf/Sheafify.lean | 10 +- Mathlib/Algebra/Category/ModuleCat/Sheaf.lean | 14 +- .../Category/ModuleCat/Sheaf/Abelian.lean | 10 +- .../Category/ModuleCat/Sheaf/Limits.lean | 6 +- .../Algebra/Category/ModuleCat/Tannaka.lean | 8 +- .../Algebra/Category/MonCat/Adjunctions.lean | 14 +- Mathlib/Algebra/Category/Ring/Basic.lean | 10 +- .../Category/Ring/FilteredColimits.lean | 6 +- Mathlib/Algebra/Category/Ring/Limits.lean | 20 +- .../{SemigroupCat => Semigrp}/Basic.lean | 120 +++--- Mathlib/Algebra/HierarchyDesign.lean | 6 +- .../Homology/HomotopyCategory/HomComplex.lean | 10 +- Mathlib/Algebra/Homology/LocalCohomology.lean | 2 +- Mathlib/Algebra/Homology/ShortComplex/Ab.lean | 22 +- .../ShortComplex/ConcreteCategory.lean | 4 +- Mathlib/Algebra/Module/CharacterModule.lean | 2 +- .../Normed/Group/SemiNormedGroupCat.lean | 270 ------------- .../Analysis/Normed/Group/SemiNormedGrp.lean | 270 +++++++++++++ .../Completion.lean | 68 ++-- .../Kernels.lean | 194 ++++----- .../Galois/Prorepresentability.lean | 12 +- Mathlib/CategoryTheory/Linear/Yoneda.lean | 4 +- Mathlib/CategoryTheory/Monoidal/Mon_.lean | 2 +- .../CategoryTheory/Preadditive/Generator.lean | 8 +- .../Preadditive/Yoneda/Basic.lean | 30 +- .../Preadditive/Yoneda/Injective.lean | 4 +- .../Preadditive/Yoneda/Projective.lean | 6 +- .../Sites/NonabelianCohomology/H1.lean | 10 +- Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 26 +- .../RepresentationTheory/Action/Basic.lean | 10 +- .../RepresentationTheory/Action/Monoidal.lean | 2 +- Mathlib/RepresentationTheory/Character.lean | 2 +- Mathlib/RepresentationTheory/FdRep.lean | 2 +- Mathlib/RepresentationTheory/Invariants.lean | 4 +- Mathlib/Tactic/ToAdditive.lean | 2 + .../Sheaves/SheafCondition/UniqueGluing.lean | 2 +- Mathlib/Topology/VectorBundle/Basic.lean | 2 +- scripts/nolints.json | 4 +- 70 files changed, 1655 insertions(+), 1653 deletions(-) delete mode 100644 Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean rename Mathlib/Algebra/Category/{GroupCat => Grp}/Abelian.lean (80%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Adjunctions.lean (72%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Basic.lean (52%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Biproducts.lean (61%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Colimits.lean (77%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/EnoughInjectives.lean (61%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/EpiMono.lean (70%) create mode 100644 Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean rename Mathlib/Algebra/Category/{GroupCat => Grp}/FilteredColimits.lean (52%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/ForgetCorepresentable.lean (65%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Images.lean (65%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Injective.lean (76%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Kernels.lean (89%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Limits.lean (54%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Preadditive.lean (70%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Subobject.lean (56%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/ZModuleEquivalence.lean (82%) rename Mathlib/Algebra/Category/{GroupCat => Grp}/Zero.lean (64%) rename Mathlib/Algebra/Category/{GroupWithZeroCat.lean => GrpWithZero.lean} (56%) rename Mathlib/Algebra/Category/{SemigroupCat => Semigrp}/Basic.lean (72%) delete mode 100644 Mathlib/Analysis/Normed/Group/SemiNormedGroupCat.lean create mode 100644 Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean rename Mathlib/Analysis/Normed/Group/{SemiNormedGroupCat => SemiNormedGrp}/Completion.lean (64%) rename Mathlib/Analysis/Normed/Group/{SemiNormedGroupCat => SemiNormedGrp}/Kernels.lean (66%) diff --git a/Mathlib.lean b/Mathlib.lean index 91199d3aa3b68..46217cd6ef43c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -56,25 +56,25 @@ import Mathlib.Algebra.Category.BoolRing import Mathlib.Algebra.Category.CoalgebraCat.Basic import Mathlib.Algebra.Category.FGModuleCat.Basic import Mathlib.Algebra.Category.FGModuleCat.Limits -import Mathlib.Algebra.Category.GroupCat.Abelian -import Mathlib.Algebra.Category.GroupCat.Adjunctions -import Mathlib.Algebra.Category.GroupCat.Basic -import Mathlib.Algebra.Category.GroupCat.Biproducts -import Mathlib.Algebra.Category.GroupCat.Colimits -import Mathlib.Algebra.Category.GroupCat.EnoughInjectives -import Mathlib.Algebra.Category.GroupCat.EpiMono -import Mathlib.Algebra.Category.GroupCat.EquivalenceGroupAddGroup -import Mathlib.Algebra.Category.GroupCat.FilteredColimits -import Mathlib.Algebra.Category.GroupCat.ForgetCorepresentable -import Mathlib.Algebra.Category.GroupCat.Images -import Mathlib.Algebra.Category.GroupCat.Injective -import Mathlib.Algebra.Category.GroupCat.Kernels -import Mathlib.Algebra.Category.GroupCat.Limits -import Mathlib.Algebra.Category.GroupCat.Preadditive -import Mathlib.Algebra.Category.GroupCat.Subobject -import Mathlib.Algebra.Category.GroupCat.ZModuleEquivalence -import Mathlib.Algebra.Category.GroupCat.Zero -import Mathlib.Algebra.Category.GroupWithZeroCat +import Mathlib.Algebra.Category.Grp.Abelian +import Mathlib.Algebra.Category.Grp.Adjunctions +import Mathlib.Algebra.Category.Grp.Basic +import Mathlib.Algebra.Category.Grp.Biproducts +import Mathlib.Algebra.Category.Grp.Colimits +import Mathlib.Algebra.Category.Grp.EnoughInjectives +import Mathlib.Algebra.Category.Grp.EpiMono +import Mathlib.Algebra.Category.Grp.EquivalenceGroupAddGroup +import Mathlib.Algebra.Category.Grp.FilteredColimits +import Mathlib.Algebra.Category.Grp.ForgetCorepresentable +import Mathlib.Algebra.Category.Grp.Images +import Mathlib.Algebra.Category.Grp.Injective +import Mathlib.Algebra.Category.Grp.Kernels +import Mathlib.Algebra.Category.Grp.Limits +import Mathlib.Algebra.Category.Grp.Preadditive +import Mathlib.Algebra.Category.Grp.Subobject +import Mathlib.Algebra.Category.Grp.ZModuleEquivalence +import Mathlib.Algebra.Category.Grp.Zero +import Mathlib.Algebra.Category.GrpWithZero import Mathlib.Algebra.Category.ModuleCat.Abelian import Mathlib.Algebra.Category.ModuleCat.Adjunctions import Mathlib.Algebra.Category.ModuleCat.Algebra @@ -122,7 +122,7 @@ import Mathlib.Algebra.Category.Ring.Constructions import Mathlib.Algebra.Category.Ring.FilteredColimits import Mathlib.Algebra.Category.Ring.Instances import Mathlib.Algebra.Category.Ring.Limits -import Mathlib.Algebra.Category.SemigroupCat.Basic +import Mathlib.Algebra.Category.Semigrp.Basic import Mathlib.Algebra.CharP.Algebra import Mathlib.Algebra.CharP.Basic import Mathlib.Algebra.CharP.CharAndCard @@ -1041,9 +1041,9 @@ import Mathlib.Analysis.Normed.Group.InfiniteSum import Mathlib.Analysis.Normed.Group.Lemmas import Mathlib.Analysis.Normed.Group.Pointwise import Mathlib.Analysis.Normed.Group.Quotient -import Mathlib.Analysis.Normed.Group.SemiNormedGroupCat -import Mathlib.Analysis.Normed.Group.SemiNormedGroupCat.Completion -import Mathlib.Analysis.Normed.Group.SemiNormedGroupCat.Kernels +import Mathlib.Analysis.Normed.Group.SemiNormedGrp +import Mathlib.Analysis.Normed.Group.SemiNormedGrp.Completion +import Mathlib.Analysis.Normed.Group.SemiNormedGrp.Kernels import Mathlib.Analysis.Normed.Group.Seminorm import Mathlib.Analysis.Normed.Group.Tannery import Mathlib.Analysis.Normed.Group.ZeroAtInfty diff --git a/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean b/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean deleted file mode 100644 index 47ef4b1479dfe..0000000000000 --- a/Mathlib/Algebra/Category/GroupCat/EquivalenceGroupAddGroup.lean +++ /dev/null @@ -1,94 +0,0 @@ -/- -Copyright (c) 2022 Jujian Zhang. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Jujian Zhang --/ -import Mathlib.Algebra.Category.GroupCat.Basic -import Mathlib.Algebra.Group.Equiv.TypeTags - -#align_import algebra.category.Group.equivalence_Group_AddGroup from "leanprover-community/mathlib"@"47b51515e69f59bca5cf34ef456e6000fe205a69" - -/-! -# Equivalence between `Group` and `AddGroup` - -This file contains two equivalences: -* `groupAddGroupEquivalence` : the equivalence between `GroupCat` and `AddGroupCat` by sending - `X : GroupCat` to `Additive X` and `Y : AddGroupCat` to `Multiplicative Y`. -* `commGroupAddCommGroupEquivalence` : the equivalence between `CommGroupCat` and `AddCommGroupCat` - by sending `X : CommGroupCat` to `Additive X` and `Y : AddCommGroupCat` to `Multiplicative Y`. --/ - --- Porting note: everything is upper case -set_option linter.uppercaseLean3 false - -open CategoryTheory - -namespace GroupCat - --- Porting note: Lean cannot find these now -private instance (X : GroupCat) : MulOneClass X.α := X.str.toMulOneClass -private instance (X : CommGroupCat) : MulOneClass X.α := X.str.toMulOneClass -private instance (X : AddGroupCat) : AddZeroClass X.α := X.str.toAddZeroClass -private instance (X : AddCommGroupCat) : AddZeroClass X.α := X.str.toAddZeroClass - -/-- The functor `Group ⥤ AddGroup` by sending `X ↦ additive X` and `f ↦ f`. --/ -@[simps] -def toAddGroupCat : GroupCat ⥤ AddGroupCat where - obj X := AddGroupCat.of (Additive X) - map {X} {Y} := MonoidHom.toAdditive -#align Group.to_AddGroup GroupCat.toAddGroupCat - -end GroupCat - -namespace CommGroupCat - -/-- The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ additive X` and `f ↦ f`. --/ -@[simps] -def toAddCommGroupCat : CommGroupCat ⥤ AddCommGroupCat where - obj X := AddCommGroupCat.of (Additive X) - map {X} {Y} := MonoidHom.toAdditive -#align CommGroup.to_AddCommGroup CommGroupCat.toAddCommGroupCat - -end CommGroupCat - -namespace AddGroupCat - -/-- The functor `AddGroup ⥤ Group` by sending `X ↦ multiplicative Y` and `f ↦ f`. --/ -@[simps] -def toGroupCat : AddGroupCat ⥤ GroupCat where - obj X := GroupCat.of (Multiplicative X) - map {X} {Y} := AddMonoidHom.toMultiplicative -#align AddGroup.to_Group AddGroupCat.toGroupCat - -end AddGroupCat - -namespace AddCommGroupCat - -/-- The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ multiplicative Y` and `f ↦ f`. --/ -@[simps] -def toCommGroupCat : AddCommGroupCat ⥤ CommGroupCat where - obj X := CommGroupCat.of (Multiplicative X) - map {X} {Y} := AddMonoidHom.toMultiplicative -#align AddCommGroup.to_CommGroup AddCommGroupCat.toCommGroupCat - -end AddCommGroupCat - -/-- The equivalence of categories between `Group` and `AddGroup` --/ -def groupAddGroupEquivalence : GroupCat ≌ AddGroupCat := - CategoryTheory.Equivalence.mk GroupCat.toAddGroupCat AddGroupCat.toGroupCat - (NatIso.ofComponents fun X => MulEquiv.toGroupCatIso (MulEquiv.multiplicativeAdditive X)) - (NatIso.ofComponents fun X => AddEquiv.toAddGroupCatIso (AddEquiv.additiveMultiplicative X)) -#align Group_AddGroup_equivalence groupAddGroupEquivalence - -/-- The equivalence of categories between `CommGroup` and `AddCommGroup`. --/ -def commGroupAddCommGroupEquivalence : CommGroupCat ≌ AddCommGroupCat := - CategoryTheory.Equivalence.mk CommGroupCat.toAddCommGroupCat AddCommGroupCat.toCommGroupCat - (NatIso.ofComponents fun X => MulEquiv.toCommGroupCatIso (MulEquiv.multiplicativeAdditive X)) - (NatIso.ofComponents fun X => AddEquiv.toAddCommGroupCatIso (AddEquiv.additiveMultiplicative X)) -#align CommGroup_AddCommGroup_equivalence commGroupAddCommGroupEquivalence diff --git a/Mathlib/Algebra/Category/GroupCat/Abelian.lean b/Mathlib/Algebra/Category/Grp/Abelian.lean similarity index 80% rename from Mathlib/Algebra/Category/GroupCat/Abelian.lean rename to Mathlib/Algebra/Category/Grp/Abelian.lean index 78c6ee40ca846..88a7b9791a0f4 100644 --- a/Mathlib/Algebra/Category/GroupCat/Abelian.lean +++ b/Mathlib/Algebra/Category/Grp/Abelian.lean @@ -3,11 +3,11 @@ Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ -import Mathlib.Algebra.Category.GroupCat.Colimits -import Mathlib.Algebra.Category.GroupCat.FilteredColimits -import Mathlib.Algebra.Category.GroupCat.Kernels -import Mathlib.Algebra.Category.GroupCat.Limits -import Mathlib.Algebra.Category.GroupCat.ZModuleEquivalence +import Mathlib.Algebra.Category.Grp.Colimits +import Mathlib.Algebra.Category.Grp.FilteredColimits +import Mathlib.Algebra.Category.Grp.Kernels +import Mathlib.Algebra.Category.Grp.Limits +import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.Category.ModuleCat.Abelian import Mathlib.CategoryTheory.Abelian.FunctorCategory import Mathlib.CategoryTheory.Limits.ConcreteCategory @@ -24,26 +24,26 @@ universe u noncomputable section -namespace AddCommGroupCat +namespace AddCommGrp -variable {X Y Z : AddCommGroupCat.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) +variable {X Y Z : AddCommGrp.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) /-- In the category of abelian groups, every monomorphism is normal. -/ def normalMono (_ : Mono f) : NormalMono f := - equivalenceReflectsNormalMono (forget₂ (ModuleCat.{u} ℤ) AddCommGroupCat.{u}).inv <| + equivalenceReflectsNormalMono (forget₂ (ModuleCat.{u} ℤ) AddCommGrp.{u}).inv <| ModuleCat.normalMono _ inferInstance set_option linter.uppercaseLean3 false in -#align AddCommGroup.normal_mono AddCommGroupCat.normalMono +#align AddCommGroup.normal_mono AddCommGrp.normalMono /-- In the category of abelian groups, every epimorphism is normal. -/ def normalEpi (_ : Epi f) : NormalEpi f := - equivalenceReflectsNormalEpi (forget₂ (ModuleCat.{u} ℤ) AddCommGroupCat.{u}).inv <| + equivalenceReflectsNormalEpi (forget₂ (ModuleCat.{u} ℤ) AddCommGrp.{u}).inv <| ModuleCat.normalEpi _ inferInstance set_option linter.uppercaseLean3 false in -#align AddCommGroup.normal_epi AddCommGroupCat.normalEpi +#align AddCommGroup.normal_epi AddCommGrp.normalEpi /-- The category of abelian groups is abelian. -/ -instance : Abelian AddCommGroupCat.{u} where +instance : Abelian AddCommGrp.{u} where has_finite_products := ⟨HasFiniteProducts.out⟩ normalMonoOfMono := normalMono normalEpiOfEpi := normalEpi @@ -58,7 +58,7 @@ theorem exact_iff : Exact f g ↔ f.range = g.ker := by /-- The category of abelian groups satisfies Grothedieck's Axiom AB5. -/ instance {J : Type u} [SmallCategory J] [IsFiltered J] : - PreservesFiniteLimits <| colim (J := J) (C := AddCommGroupCat.{u}) := by + PreservesFiniteLimits <| colim (J := J) (C := AddCommGrp.{u}) := by refine Functor.preservesFiniteLimitsOfMapExact _ fun F G H η γ h => (exact_iff _ _).mpr (le_antisymm ?_ ?_) all_goals replace h : ∀ j : J, Exact (η.app j) (γ.app j) := @@ -77,4 +77,4 @@ instance {J : Type u} [SmallCategory J] [IsFiltered J] : erw [← comp_apply, colimit.ι_map, comp_apply, ht] exact colimit.w_apply G e₁ y -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean b/Mathlib/Algebra/Category/Grp/Adjunctions.lean similarity index 72% rename from Mathlib/Algebra/Category/GroupCat/Adjunctions.lean rename to Mathlib/Algebra/Category/Grp/Adjunctions.lean index af27c8e461069..19b3515edc05c 100644 --- a/Mathlib/Algebra/Category/GroupCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/Grp/Adjunctions.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Johannes Hölzl -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.GroupTheory.FreeAbelianGroup #align_import algebra.category.Group.adjunctions from "leanprover-community/mathlib"@"ecef68622cf98f6d42c459e5b5a079aeecdd9842" @@ -17,17 +17,17 @@ category of abelian groups. ## Main definitions -* `AddCommGroupCat.free`: constructs the functor associating to a type `X` the free abelian group +* `AddCommGrp.free`: constructs the functor associating to a type `X` the free abelian group with generators `x : X`. -* `GroupCat.free`: constructs the functor associating to a type `X` the free group with +* `Grp.free`: constructs the functor associating to a type `X` the free group with generators `x : X`. * `abelianize`: constructs the functor which associates to a group `G` its abelianization `Gᵃᵇ`. ## Main statements -* `AddCommGroupCat.adj`: proves that `AddCommGroupCat.free` is the left adjoint +* `AddCommGrp.adj`: proves that `AddCommGrp.free` is the left adjoint of the forgetful functor from abelian groups to types. -* `GroupCat.adj`: proves that `GroupCat.free` is the left adjoint of the forgetful functor +* `Grp.adj`: proves that `Grp.free` is the left adjoint of the forgetful functor from groups to types. * `abelianizeAdj`: proves that `abelianize` is left adjoint to the forgetful functor from abelian groups to groups. @@ -41,24 +41,24 @@ universe u open CategoryTheory -namespace AddCommGroupCat +namespace AddCommGrp open scoped Classical /-- The free functor `Type u ⥤ AddCommGroup` sending a type `X` to the free abelian group with generators `x : X`. -/ -def free : Type u ⥤ AddCommGroupCat where +def free : Type u ⥤ AddCommGrp where obj α := of (FreeAbelianGroup α) map := FreeAbelianGroup.map map_id _ := AddMonoidHom.ext FreeAbelianGroup.map_id_apply map_comp _ _ := AddMonoidHom.ext FreeAbelianGroup.map_comp_apply -#align AddCommGroup.free AddCommGroupCat.free +#align AddCommGroup.free AddCommGrp.free @[simp] theorem free_obj_coe {α : Type u} : (free.obj α : Type u) = FreeAbelianGroup α := rfl -#align AddCommGroup.free_obj_coe AddCommGroupCat.free_obj_coe +#align AddCommGroup.free_obj_coe AddCommGrp.free_obj_coe -- This currently can't be a `simp` lemma, -- because `free_obj_coe` will simplify implicit arguments in the LHS. @@ -66,11 +66,11 @@ theorem free_obj_coe {α : Type u} : (free.obj α : Type u) = FreeAbelianGroup theorem free_map_coe {α β : Type u} {f : α → β} (x : FreeAbelianGroup α) : (free.map f) x = f <$> x := rfl -#align AddCommGroup.free_map_coe AddCommGroupCat.free_map_coe +#align AddCommGroup.free_map_coe AddCommGrp.free_map_coe /-- The free-forgetful adjunction for abelian groups. -/ -def adj : free ⊣ forget AddCommGroupCat.{u} := +def adj : free ⊣ forget AddCommGrp.{u} := Adjunction.mkOfHomEquiv { homEquiv := fun X G => FreeAbelianGroup.lift.symm -- Porting note: used to be just `by intros; ext; rfl`. @@ -79,9 +79,9 @@ def adj : free ⊣ forget AddCommGroupCat.{u} := ext simp only [Equiv.symm_symm] apply FreeAbelianGroup.lift_comp } -#align AddCommGroup.adj AddCommGroupCat.adj +#align AddCommGroup.adj AddCommGrp.adj -instance : (forget AddCommGroupCat.{u}).IsRightAdjoint := +instance : (forget AddCommGrp.{u}).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ /-- As an example, we now give a high-powered proof that @@ -89,17 +89,17 @@ the monomorphisms in `AddCommGroup` are just the injective functions. (This proof works in all universes.) -/ -example {G H : AddCommGroupCat.{u}} (f : G ⟶ H) [Mono f] : Function.Injective f := - (mono_iff_injective (f : G → H)).mp (Functor.map_mono (forget AddCommGroupCat) f) +example {G H : AddCommGrp.{u}} (f : G ⟶ H) [Mono f] : Function.Injective f := + (mono_iff_injective (f : G → H)).mp (Functor.map_mono (forget AddCommGrp) f) -end AddCommGroupCat +end AddCommGrp -namespace GroupCat +namespace Grp /-- The free functor `Type u ⥤ Group` sending a type `X` to the free group with generators `x : X`. -/ -def free : Type u ⥤ GroupCat where +def free : Type u ⥤ Grp where obj α := of (FreeGroup α) map := FreeGroup.map map_id := by @@ -108,11 +108,11 @@ def free : Type u ⥤ GroupCat where map_comp := by -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 intros; ext1; erw [← FreeGroup.map.unique] <;> intros <;> rfl -#align Group.free GroupCat.free +#align Group.free Grp.free /-- The free-forgetful adjunction for groups. -/ -def adj : free ⊣ forget GroupCat.{u} := +def adj : free ⊣ forget Grp.{u} := Adjunction.mkOfHomEquiv { homEquiv := fun X G => FreeGroup.lift.symm -- Porting note: used to be just `by intros; ext1; rfl`. @@ -124,17 +124,17 @@ def adj : free ⊣ forget GroupCat.{u} := apply FreeGroup.lift.unique intros apply FreeGroup.lift.of } -#align Group.adj GroupCat.adj +#align Group.adj Grp.adj -instance : (forget GroupCat.{u}).IsRightAdjoint := +instance : (forget Grp.{u}).IsRightAdjoint := ⟨_, ⟨adj⟩⟩ section Abelianization /-- The abelianization functor `Group ⥤ CommGroup` sending a group `G` to its abelianization `Gᵃᵇ`. -/ -def abelianize : GroupCat.{u} ⥤ CommGroupCat.{u} where - obj G := CommGroupCat.of (Abelianization G) +def abelianize : Grp.{u} ⥤ CommGrp.{u} where + obj G := CommGrp.of (Abelianization G) map f := Abelianization.lift (Abelianization.of.comp f) map_id := by intros; simp only [coe_id] @@ -142,10 +142,10 @@ def abelianize : GroupCat.{u} ⥤ CommGroupCat.{u} where map_comp := by intros; simp only [coe_comp] apply (Equiv.apply_eq_iff_eq_symm_apply Abelianization.lift).mpr; rfl -#align abelianize GroupCat.abelianize +#align abelianize Grp.abelianize /-- The abelianization-forgetful adjuction from `Group` to `CommGroup`. -/ -def abelianizeAdj : abelianize ⊣ forget₂ CommGroupCat.{u} GroupCat.{u} := +def abelianizeAdj : abelianize ⊣ forget₂ CommGrp.{u} Grp.{u} := Adjunction.mkOfHomEquiv { homEquiv := fun G A => Abelianization.lift.symm -- Porting note: used to be just `by intros; ext1; rfl`. @@ -157,23 +157,23 @@ def abelianizeAdj : abelianize ⊣ forget₂ CommGroupCat.{u} GroupCat.{u} := apply Abelianization.lift.unique intros apply Abelianization.lift.of } -#align abelianize_adj GroupCat.abelianizeAdj +#align abelianize_adj Grp.abelianizeAdj end Abelianization -end GroupCat +end Grp /-- The functor taking a monoid to its subgroup of units. -/ @[simps] -def MonCat.units : MonCat.{u} ⥤ GroupCat.{u} where - obj R := GroupCat.of Rˣ - map f := GroupCat.ofHom <| Units.map f +def MonCat.units : MonCat.{u} ⥤ Grp.{u} where + obj R := Grp.of Rˣ + map f := Grp.ofHom <| Units.map f map_id _ := MonoidHom.ext fun _ => Units.ext rfl map_comp _ _ := MonoidHom.ext fun _ => Units.ext rfl #align Mon.units MonCat.units -/-- The forgetful-units adjunction between `GroupCat` and `MonCat`. -/ -def GroupCat.forget₂MonAdj : forget₂ GroupCat MonCat ⊣ MonCat.units.{u} where +/-- The forgetful-units adjunction between `Grp` and `MonCat`. -/ +def Grp.forget₂MonAdj : forget₂ Grp MonCat ⊣ MonCat.units.{u} where homEquiv X Y := { toFun := fun f => MonoidHom.toHomUnits f invFun := fun f => (Units.coeHom Y).comp f @@ -187,22 +187,22 @@ def GroupCat.forget₂MonAdj : forget₂ GroupCat MonCat ⊣ MonCat.units.{u} wh naturality := by intros; exact MonoidHom.ext fun x => rfl } homEquiv_unit := MonoidHom.ext fun _ => Units.ext rfl homEquiv_counit := MonoidHom.ext fun _ => rfl -#align Group.forget₂_Mon_adj GroupCat.forget₂MonAdj +#align Group.forget₂_Mon_adj Grp.forget₂MonAdj instance : MonCat.units.{u}.IsRightAdjoint := - ⟨_, ⟨GroupCat.forget₂MonAdj⟩⟩ + ⟨_, ⟨Grp.forget₂MonAdj⟩⟩ /-- The functor taking a monoid to its subgroup of units. -/ @[simps] -def CommMonCat.units : CommMonCat.{u} ⥤ CommGroupCat.{u} where - obj R := CommGroupCat.of Rˣ - map f := CommGroupCat.ofHom <| Units.map f +def CommMonCat.units : CommMonCat.{u} ⥤ CommGrp.{u} where + obj R := CommGrp.of Rˣ + map f := CommGrp.ofHom <| Units.map f map_id _ := MonoidHom.ext fun _ => Units.ext rfl map_comp _ _ := MonoidHom.ext fun _ => Units.ext rfl #align CommMon.units CommMonCat.units -/-- The forgetful-units adjunction between `CommGroupCat` and `CommMonCat`. -/ -def CommGroupCat.forget₂CommMonAdj : forget₂ CommGroupCat CommMonCat ⊣ CommMonCat.units.{u} where +/-- The forgetful-units adjunction between `CommGrp` and `CommMonCat`. -/ +def CommGrp.forget₂CommMonAdj : forget₂ CommGrp CommMonCat ⊣ CommMonCat.units.{u} where homEquiv X Y := { toFun := fun f => MonoidHom.toHomUnits f invFun := fun f => (Units.coeHom Y).comp f @@ -216,7 +216,7 @@ def CommGroupCat.forget₂CommMonAdj : forget₂ CommGroupCat CommMonCat ⊣ Com naturality := by intros; exact MonoidHom.ext fun x => rfl } homEquiv_unit := MonoidHom.ext fun _ => Units.ext rfl homEquiv_counit := MonoidHom.ext fun _ => rfl -#align CommGroup.forget₂_CommMon_adj CommGroupCat.forget₂CommMonAdj +#align CommGroup.forget₂_CommMon_adj CommGrp.forget₂CommMonAdj instance : CommMonCat.units.{u}.IsRightAdjoint := - ⟨_, ⟨CommGroupCat.forget₂CommMonAdj⟩⟩ + ⟨_, ⟨CommGrp.forget₂CommMonAdj⟩⟩ diff --git a/Mathlib/Algebra/Category/GroupCat/Basic.lean b/Mathlib/Algebra/Category/Grp/Basic.lean similarity index 52% rename from Mathlib/Algebra/Category/GroupCat/Basic.lean rename to Mathlib/Algebra/Category/Grp/Basic.lean index 0fb65b73eda80..3ee44e2ab9c6d 100644 --- a/Mathlib/Algebra/Category/GroupCat/Basic.lean +++ b/Mathlib/Algebra/Category/Grp/Basic.lean @@ -13,10 +13,10 @@ import Mathlib.CategoryTheory.Endomorphism # Category instances for Group, AddGroup, CommGroup, and AddCommGroup. We introduce the bundled categories: -* `GroupCat` -* `AddGroupCat` -* `CommGroupCat` -* `AddCommGroupCat` +* `Grp` +* `AddGrp` +* `CommGrp` +* `AddCommGrp` along with the relevant forgetful functors between them, and to the bundled monoid categories. -/ @@ -26,424 +26,424 @@ open CategoryTheory /-- The category of groups and group morphisms. -/ @[to_additive] -def GroupCat : Type (u + 1) := +def Grp : Type (u + 1) := Bundled Group set_option linter.uppercaseLean3 false in -#align Group GroupCat +#align Group Grp set_option linter.uppercaseLean3 false in -#align AddGroup AddGroupCat +#align AddGroup AddGrp /-- The category of additive groups and group morphisms -/ -add_decl_doc AddGroupCat +add_decl_doc AddGrp -namespace GroupCat +namespace Grp @[to_additive] instance : BundledHom.ParentProjection (fun {α : Type*} (h : Group α) => h.toDivInvMonoid.toMonoid) := ⟨⟩ -deriving instance LargeCategory for GroupCat -attribute [to_additive] instGroupCatLargeCategory +deriving instance LargeCategory for Grp +attribute [to_additive] instGrpLargeCategory @[to_additive] -instance concreteCategory : ConcreteCategory GroupCat := by - dsimp only [GroupCat] +instance concreteCategory : ConcreteCategory Grp := by + dsimp only [Grp] infer_instance @[to_additive] -instance : CoeSort GroupCat Type* where +instance : CoeSort Grp Type* where coe X := X.α @[to_additive] -instance (X : GroupCat) : Group X := X.str +instance (X : Grp) : Group X := X.str -- porting note (#10670): this instance was not necessary in mathlib @[to_additive] -instance {X Y : GroupCat} : CoeFun (X ⟶ Y) fun _ => X → Y where +instance {X Y : Grp} : CoeFun (X ⟶ Y) fun _ => X → Y where coe (f : X →* Y) := f @[to_additive] -instance instFunLike (X Y : GroupCat) : FunLike (X ⟶ Y) X Y := +instance instFunLike (X Y : Grp) : FunLike (X ⟶ Y) X Y := show FunLike (X →* Y) X Y from inferInstance @[to_additive (attr := simp)] -lemma coe_id {X : GroupCat} : (𝟙 X : X → X) = id := rfl +lemma coe_id {X : Grp} : (𝟙 X : X → X) = id := rfl @[to_additive (attr := simp)] -lemma coe_comp {X Y Z : GroupCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +lemma coe_comp {X Y Z : Grp} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl @[to_additive] -lemma comp_def {X Y Z : GroupCat} {f : X ⟶ Y} {g : Y ⟶ Z} : f ≫ g = g.comp f := rfl +lemma comp_def {X Y Z : Grp} {f : X ⟶ Y} {g : Y ⟶ Z} : f ≫ g = g.comp f := rfl -@[simp] lemma forget_map {X Y : GroupCat} (f : X ⟶ Y) : (forget GroupCat).map f = (f : X → Y) := rfl +@[simp] lemma forget_map {X Y : Grp} (f : X ⟶ Y) : (forget Grp).map f = (f : X → Y) := rfl @[to_additive (attr := ext)] -lemma ext {X Y : GroupCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := +lemma ext {X Y : Grp} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := MonoidHom.ext w /-- Construct a bundled `Group` from the underlying type and typeclass. -/ @[to_additive] -def of (X : Type u) [Group X] : GroupCat := +def of (X : Type u) [Group X] : Grp := Bundled.of X set_option linter.uppercaseLean3 false in -#align Group.of GroupCat.of +#align Group.of Grp.of set_option linter.uppercaseLean3 false in -#align AddGroup.of AddGroupCat.of +#align AddGroup.of AddGrp.of /-- Construct a bundled `AddGroup` from the underlying type and typeclass. -/ -add_decl_doc AddGroupCat.of +add_decl_doc AddGrp.of @[to_additive (attr := simp)] -theorem coe_of (R : Type u) [Group R] : ↑(GroupCat.of R) = R := +theorem coe_of (R : Type u) [Group R] : ↑(Grp.of R) = R := rfl set_option linter.uppercaseLean3 false in -#align Group.coe_of GroupCat.coe_of +#align Group.coe_of Grp.coe_of set_option linter.uppercaseLean3 false in -#align AddGroup.coe_of AddGroupCat.coe_of +#align AddGroup.coe_of AddGrp.coe_of @[to_additive (attr := simp)] theorem coe_comp' {G H K : Type _} [Group G] [Group H] [Group K] (f : G →* H) (g : H →* K) : @DFunLike.coe (G →* K) G (fun _ ↦ K) MonoidHom.instFunLike (CategoryStruct.comp - (X := GroupCat.of G) (Y := GroupCat.of H) (Z := GroupCat.of K) f g) = g ∘ f := + (X := Grp.of G) (Y := Grp.of H) (Z := Grp.of K) f g) = g ∘ f := rfl @[to_additive (attr := simp)] theorem coe_id' {G : Type _} [Group G] : @DFunLike.coe (G →* G) G (fun _ ↦ G) MonoidHom.instFunLike - (CategoryStruct.id (X := GroupCat.of G)) = id := + (CategoryStruct.id (X := Grp.of G)) = id := rfl @[to_additive] -instance : Inhabited GroupCat := - ⟨GroupCat.of PUnit⟩ +instance : Inhabited Grp := + ⟨Grp.of PUnit⟩ @[to_additive hasForgetToAddMonCat] -instance hasForgetToMonCat : HasForget₂ GroupCat MonCat := +instance hasForgetToMonCat : HasForget₂ Grp MonCat := BundledHom.forget₂ _ _ set_option linter.uppercaseLean3 false in -#align Group.has_forget_to_Mon GroupCat.hasForgetToMonCat +#align Group.has_forget_to_Mon Grp.hasForgetToMonCat set_option linter.uppercaseLean3 false in -#align AddGroup.has_forget_to_AddMon AddGroupCat.hasForgetToAddMonCat +#align AddGroup.has_forget_to_AddMon AddGrp.hasForgetToAddMonCat @[to_additive] -instance : Coe GroupCat.{u} MonCat.{u} where coe := (forget₂ GroupCat MonCat).obj +instance : Coe Grp.{u} MonCat.{u} where coe := (forget₂ Grp MonCat).obj -- porting note (#10670): this instance was not necessary in mathlib @[to_additive] -instance (G H : GroupCat) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) +instance (G H : Grp) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) @[to_additive (attr := simp)] -theorem one_apply (G H : GroupCat) (g : G) : ((1 : G ⟶ H) : G → H) g = 1 := +theorem one_apply (G H : Grp) (g : G) : ((1 : G ⟶ H) : G → H) g = 1 := rfl set_option linter.uppercaseLean3 false in -#align Group.one_apply GroupCat.one_apply +#align Group.one_apply Grp.one_apply set_option linter.uppercaseLean3 false in -#align AddGroup.zero_apply AddGroupCat.zero_apply +#align AddGroup.zero_apply AddGrp.zero_apply -/-- Typecheck a `MonoidHom` as a morphism in `GroupCat`. -/ +/-- Typecheck a `MonoidHom` as a morphism in `Grp`. -/ @[to_additive] def ofHom {X Y : Type u} [Group X] [Group Y] (f : X →* Y) : of X ⟶ of Y := f set_option linter.uppercaseLean3 false in -#align Group.of_hom GroupCat.ofHom +#align Group.of_hom Grp.ofHom set_option linter.uppercaseLean3 false in -#align AddGroup.of_hom AddGroupCat.ofHom +#align AddGroup.of_hom AddGrp.ofHom /-- Typecheck an `AddMonoidHom` as a morphism in `AddGroup`. -/ -add_decl_doc AddGroupCat.ofHom +add_decl_doc AddGrp.ofHom @[to_additive] theorem ofHom_apply {X Y : Type _} [Group X] [Group Y] (f : X →* Y) (x : X) : (ofHom f) x = f x := rfl set_option linter.uppercaseLean3 false in -#align Group.of_hom_apply GroupCat.ofHom_apply +#align Group.of_hom_apply Grp.ofHom_apply set_option linter.uppercaseLean3 false in -#align AddGroup.of_hom_apply AddGroupCat.ofHom_apply +#align AddGroup.of_hom_apply AddGrp.ofHom_apply @[to_additive] -instance ofUnique (G : Type*) [Group G] [i : Unique G] : Unique (GroupCat.of G) := i +instance ofUnique (G : Type*) [Group G] [i : Unique G] : Unique (Grp.of G) := i set_option linter.uppercaseLean3 false in -#align Group.of_unique GroupCat.ofUnique +#align Group.of_unique Grp.ofUnique set_option linter.uppercaseLean3 false in -#align AddGroup.of_unique AddGroupCat.ofUnique +#align AddGroup.of_unique AddGrp.ofUnique -- We verify that simp lemmas apply when coercing morphisms to functions. @[to_additive] -example {R S : GroupCat} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [h] +example {R S : Grp} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [h] /-- Universe lift functor for groups. -/ @[to_additive (attr := simps) "Universe lift functor for additive groups."] -def uliftFunctor : GroupCat.{u} ⥤ GroupCat.{max u v} where - obj X := GroupCat.of (ULift.{v, u} X) - map {X Y} f := GroupCat.ofHom <| +def uliftFunctor : Grp.{u} ⥤ Grp.{max u v} where + obj X := Grp.of (ULift.{v, u} X) + map {X Y} f := Grp.ofHom <| MulEquiv.ulift.symm.toMonoidHom.comp <| f.comp MulEquiv.ulift.toMonoidHom map_id X := by rfl map_comp {X Y Z} f g := by rfl -end GroupCat +end Grp /-- The category of commutative groups and group morphisms. -/ @[to_additive] -def CommGroupCat : Type (u + 1) := +def CommGrp : Type (u + 1) := Bundled CommGroup set_option linter.uppercaseLean3 false in -#align CommGroup CommGroupCat +#align CommGroup CommGrp set_option linter.uppercaseLean3 false in -#align AddCommGroup AddCommGroupCat +#align AddCommGroup AddCommGrp /-- The category of additive commutative groups and group morphisms. -/ -add_decl_doc AddCommGroupCat +add_decl_doc AddCommGrp /-- `Ab` is an abbreviation for `AddCommGroup`, for the sake of mathematicians' sanity. -/ -abbrev Ab := AddCommGroupCat +abbrev Ab := AddCommGrp set_option linter.uppercaseLean3 false in #align Ab Ab -namespace CommGroupCat +namespace CommGrp @[to_additive] instance : BundledHom.ParentProjection @CommGroup.toGroup := ⟨⟩ -deriving instance LargeCategory for CommGroupCat -attribute [to_additive] instCommGroupCatLargeCategory +deriving instance LargeCategory for CommGrp +attribute [to_additive] instCommGrpLargeCategory @[to_additive] -instance concreteCategory : ConcreteCategory CommGroupCat := by - dsimp only [CommGroupCat] +instance concreteCategory : ConcreteCategory CommGrp := by + dsimp only [CommGrp] infer_instance @[to_additive] -instance : CoeSort CommGroupCat Type* where +instance : CoeSort CommGrp Type* where coe X := X.α @[to_additive] -instance commGroupInstance (X : CommGroupCat) : CommGroup X := X.str +instance commGroupInstance (X : CommGrp) : CommGroup X := X.str set_option linter.uppercaseLean3 false in -#align CommGroup.comm_group_instance CommGroupCat.commGroupInstance +#align CommGroup.comm_group_instance CommGrp.commGroupInstance set_option linter.uppercaseLean3 false in -#align AddCommGroup.add_comm_group_instance AddCommGroupCat.addCommGroupInstance +#align AddCommGroup.add_comm_group_instance AddCommGrp.addCommGroupInstance -- porting note (#10670): this instance was not necessary in mathlib @[to_additive] -instance {X Y : CommGroupCat} : CoeFun (X ⟶ Y) fun _ => X → Y where +instance {X Y : CommGrp} : CoeFun (X ⟶ Y) fun _ => X → Y where coe (f : X →* Y) := f @[to_additive] -instance instFunLike (X Y : CommGroupCat) : FunLike (X ⟶ Y) X Y := +instance instFunLike (X Y : CommGrp) : FunLike (X ⟶ Y) X Y := show FunLike (X →* Y) X Y from inferInstance @[to_additive (attr := simp)] -lemma coe_id {X : CommGroupCat} : (𝟙 X : X → X) = id := rfl +lemma coe_id {X : CommGrp} : (𝟙 X : X → X) = id := rfl @[to_additive (attr := simp)] -lemma coe_comp {X Y Z : CommGroupCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +lemma coe_comp {X Y Z : CommGrp} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl @[to_additive] -lemma comp_def {X Y Z : CommGroupCat} {f : X ⟶ Y} {g : Y ⟶ Z} : f ≫ g = g.comp f := rfl +lemma comp_def {X Y Z : CommGrp} {f : X ⟶ Y} {g : Y ⟶ Z} : f ≫ g = g.comp f := rfl @[to_additive (attr := simp)] -lemma forget_map {X Y : CommGroupCat} (f : X ⟶ Y) : - (forget CommGroupCat).map f = (f : X → Y) := +lemma forget_map {X Y : CommGrp} (f : X ⟶ Y) : + (forget CommGrp).map f = (f : X → Y) := rfl @[to_additive (attr := ext)] -lemma ext {X Y : CommGroupCat} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := +lemma ext {X Y : CommGrp} {f g : X ⟶ Y} (w : ∀ x : X, f x = g x) : f = g := MonoidHom.ext w /-- Construct a bundled `CommGroup` from the underlying type and typeclass. -/ @[to_additive] -def of (G : Type u) [CommGroup G] : CommGroupCat := +def of (G : Type u) [CommGroup G] : CommGrp := Bundled.of G set_option linter.uppercaseLean3 false in -#align CommGroup.of CommGroupCat.of +#align CommGroup.of CommGrp.of set_option linter.uppercaseLean3 false in -#align AddCommGroup.of AddCommGroupCat.of +#align AddCommGroup.of AddCommGrp.of /-- Construct a bundled `AddCommGroup` from the underlying type and typeclass. -/ -add_decl_doc AddCommGroupCat.of +add_decl_doc AddCommGrp.of @[to_additive] -instance : Inhabited CommGroupCat := - ⟨CommGroupCat.of PUnit⟩ +instance : Inhabited CommGrp := + ⟨CommGrp.of PUnit⟩ @[to_additive (attr := simp)] -theorem coe_of (R : Type u) [CommGroup R] : (CommGroupCat.of R : Type u) = R := +theorem coe_of (R : Type u) [CommGroup R] : (CommGrp.of R : Type u) = R := rfl set_option linter.uppercaseLean3 false in -#align CommGroup.coe_of CommGroupCat.coe_of +#align CommGroup.coe_of CommGrp.coe_of set_option linter.uppercaseLean3 false in -#align AddCommGroup.coe_of AddCommGroupCat.coe_of +#align AddCommGroup.coe_of AddCommGrp.coe_of @[to_additive (attr := simp)] theorem coe_comp' {G H K : Type _} [CommGroup G] [CommGroup H] [CommGroup K] (f : G →* H) (g : H →* K) : @DFunLike.coe (G →* K) G (fun _ ↦ K) MonoidHom.instFunLike (CategoryStruct.comp - (X := CommGroupCat.of G) (Y := CommGroupCat.of H) (Z := CommGroupCat.of K) f g) = g ∘ f := + (X := CommGrp.of G) (Y := CommGrp.of H) (Z := CommGrp.of K) f g) = g ∘ f := rfl @[to_additive (attr := simp)] theorem coe_id' {G : Type _} [CommGroup G] : @DFunLike.coe (G →* G) G (fun _ ↦ G) MonoidHom.instFunLike - (CategoryStruct.id (X := CommGroupCat.of G)) = id := + (CategoryStruct.id (X := CommGrp.of G)) = id := rfl @[to_additive] -instance ofUnique (G : Type*) [CommGroup G] [i : Unique G] : Unique (CommGroupCat.of G) := +instance ofUnique (G : Type*) [CommGroup G] [i : Unique G] : Unique (CommGrp.of G) := i set_option linter.uppercaseLean3 false in -#align CommGroup.of_unique CommGroupCat.ofUnique +#align CommGroup.of_unique CommGrp.ofUnique set_option linter.uppercaseLean3 false in -#align AddCommGroup.of_unique AddCommGroupCat.ofUnique +#align AddCommGroup.of_unique AddCommGrp.ofUnique @[to_additive] -instance hasForgetToGroup : HasForget₂ CommGroupCat GroupCat := +instance hasForgetToGroup : HasForget₂ CommGrp Grp := BundledHom.forget₂ _ _ set_option linter.uppercaseLean3 false in -#align CommGroup.has_forget_to_Group CommGroupCat.hasForgetToGroup +#align CommGroup.has_forget_to_Group CommGrp.hasForgetToGroup set_option linter.uppercaseLean3 false in -#align AddCommGroup.has_forget_to_AddGroup AddCommGroupCat.hasForgetToAddGroup +#align AddCommGroup.has_forget_to_AddGroup AddCommGrp.hasForgetToAddGroup @[to_additive] -instance : Coe CommGroupCat.{u} GroupCat.{u} where coe := (forget₂ CommGroupCat GroupCat).obj +instance : Coe CommGrp.{u} Grp.{u} where coe := (forget₂ CommGrp Grp).obj @[to_additive hasForgetToAddCommMonCat] -instance hasForgetToCommMonCat : HasForget₂ CommGroupCat CommMonCat := - InducedCategory.hasForget₂ fun G : CommGroupCat => CommMonCat.of G +instance hasForgetToCommMonCat : HasForget₂ CommGrp CommMonCat := + InducedCategory.hasForget₂ fun G : CommGrp => CommMonCat.of G set_option linter.uppercaseLean3 false in -#align CommGroup.has_forget_to_CommMon CommGroupCat.hasForgetToCommMonCat +#align CommGroup.has_forget_to_CommMon CommGrp.hasForgetToCommMonCat set_option linter.uppercaseLean3 false in -#align AddCommGroup.has_forget_to_AddCommMon AddCommGroupCat.hasForgetToAddCommMonCat +#align AddCommGroup.has_forget_to_AddCommMon AddCommGrp.hasForgetToAddCommMonCat @[to_additive] -instance : Coe CommGroupCat.{u} CommMonCat.{u} where coe := (forget₂ CommGroupCat CommMonCat).obj +instance : Coe CommGrp.{u} CommMonCat.{u} where coe := (forget₂ CommGrp CommMonCat).obj -- porting note (#10670): this instance was not necessary in mathlib @[to_additive] -instance (G H : CommGroupCat) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) +instance (G H : CommGrp) : One (G ⟶ H) := (inferInstance : One (MonoidHom G H)) @[to_additive (attr := simp)] -theorem one_apply (G H : CommGroupCat) (g : G) : ((1 : G ⟶ H) : G → H) g = 1 := +theorem one_apply (G H : CommGrp) (g : G) : ((1 : G ⟶ H) : G → H) g = 1 := rfl set_option linter.uppercaseLean3 false in -#align CommGroup.one_apply CommGroupCat.one_apply +#align CommGroup.one_apply CommGrp.one_apply set_option linter.uppercaseLean3 false in -#align AddCommGroup.zero_apply AddCommGroupCat.zero_apply +#align AddCommGroup.zero_apply AddCommGrp.zero_apply /-- Typecheck a `MonoidHom` as a morphism in `CommGroup`. -/ @[to_additive] def ofHom {X Y : Type u} [CommGroup X] [CommGroup Y] (f : X →* Y) : of X ⟶ of Y := f set_option linter.uppercaseLean3 false in -#align CommGroup.of_hom CommGroupCat.ofHom +#align CommGroup.of_hom CommGrp.ofHom set_option linter.uppercaseLean3 false in -#align AddCommGroup.of_hom AddCommGroupCat.ofHom +#align AddCommGroup.of_hom AddCommGrp.ofHom /-- Typecheck an `AddMonoidHom` as a morphism in `AddCommGroup`. -/ -add_decl_doc AddCommGroupCat.ofHom +add_decl_doc AddCommGrp.ofHom @[to_additive (attr := simp)] theorem ofHom_apply {X Y : Type _} [CommGroup X] [CommGroup Y] (f : X →* Y) (x : X) : @DFunLike.coe (X →* Y) X (fun _ ↦ Y) _ (ofHom f) x = f x := rfl set_option linter.uppercaseLean3 false in -#align CommGroup.of_hom_apply CommGroupCat.ofHom_apply +#align CommGroup.of_hom_apply CommGrp.ofHom_apply set_option linter.uppercaseLean3 false in -#align AddCommGroup.of_hom_apply AddCommGroupCat.ofHom_apply +#align AddCommGroup.of_hom_apply AddCommGrp.ofHom_apply -- We verify that simp lemmas apply when coercing morphisms to functions. @[to_additive] -example {R S : CommGroupCat} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [h] +example {R S : CommGrp} (i : R ⟶ S) (r : R) (h : r = 1) : i r = 1 := by simp [h] /-- Universe lift functor for commutative groups. -/ @[to_additive (attr := simps) "Universe lift functor for additive commutative groups."] -def uliftFunctor : CommGroupCat.{u} ⥤ CommGroupCat.{max u v} where - obj X := CommGroupCat.of (ULift.{v, u} X) - map {X Y} f := CommGroupCat.ofHom <| +def uliftFunctor : CommGrp.{u} ⥤ CommGrp.{max u v} where + obj X := CommGrp.of (ULift.{v, u} X) + map {X Y} f := CommGrp.ofHom <| MulEquiv.ulift.symm.toMonoidHom.comp <| f.comp MulEquiv.ulift.toMonoidHom map_id X := by rfl map_comp {X Y Z} f g := by rfl -end CommGroupCat +end CommGrp -namespace AddCommGroupCat +namespace AddCommGrp -- Note that because `ℤ : Type 0`, this forces `G : AddCommGroup.{0}`, -- so we write this explicitly to be clear. -- TODO generalize this, requiring a `ULiftInstances.lean` file /-- Any element of an abelian group gives a unique morphism from `ℤ` sending `1` to that element. -/ -def asHom {G : AddCommGroupCat.{0}} (g : G) : AddCommGroupCat.of ℤ ⟶ G := +def asHom {G : AddCommGrp.{0}} (g : G) : AddCommGrp.of ℤ ⟶ G := zmultiplesHom G g set_option linter.uppercaseLean3 false in -#align AddCommGroup.as_hom AddCommGroupCat.asHom +#align AddCommGroup.as_hom AddCommGrp.asHom @[simp] -theorem asHom_apply {G : AddCommGroupCat.{0}} (g : G) (i : ℤ) : +theorem asHom_apply {G : AddCommGrp.{0}} (g : G) (i : ℤ) : @DFunLike.coe (ℤ →+ ↑G) ℤ (fun _ ↦ ↑G) _ (asHom g) i = i • g := rfl set_option linter.uppercaseLean3 false in -#align AddCommGroup.as_hom_apply AddCommGroupCat.asHom_apply +#align AddCommGroup.as_hom_apply AddCommGrp.asHom_apply -theorem asHom_injective {G : AddCommGroupCat.{0}} : Function.Injective (@asHom G) := fun h k w => by - convert congr_arg (fun k : AddCommGroupCat.of ℤ ⟶ G => (k : ℤ → G) (1 : ℤ)) w <;> simp +theorem asHom_injective {G : AddCommGrp.{0}} : Function.Injective (@asHom G) := fun h k w => by + convert congr_arg (fun k : AddCommGrp.of ℤ ⟶ G => (k : ℤ → G) (1 : ℤ)) w <;> simp set_option linter.uppercaseLean3 false in -#align AddCommGroup.as_hom_injective AddCommGroupCat.asHom_injective +#align AddCommGroup.as_hom_injective AddCommGrp.asHom_injective @[ext] -theorem int_hom_ext {G : AddCommGroupCat.{0}} (f g : AddCommGroupCat.of ℤ ⟶ G) +theorem int_hom_ext {G : AddCommGrp.{0}} (f g : AddCommGrp.of ℤ ⟶ G) (w : f (1 : ℤ) = g (1 : ℤ)) : f = g := @AddMonoidHom.ext_int G _ f g w set_option linter.uppercaseLean3 false in -#align AddCommGroup.int_hom_ext AddCommGroupCat.int_hom_ext +#align AddCommGroup.int_hom_ext AddCommGrp.int_hom_ext -- TODO: this argument should be generalised to the situation where -- the forgetful functor is representable. -theorem injective_of_mono {G H : AddCommGroupCat.{0}} (f : G ⟶ H) [Mono f] : Function.Injective f := +theorem injective_of_mono {G H : AddCommGrp.{0}} (f : G ⟶ H) [Mono f] : Function.Injective f := fun g₁ g₂ h => by have t0 : asHom g₁ ≫ f = asHom g₂ ≫ f := by aesop_cat have t1 : asHom g₁ = asHom g₂ := (cancel_mono _).1 t0 apply asHom_injective t1 set_option linter.uppercaseLean3 false in -#align AddCommGroup.injective_of_mono AddCommGroupCat.injective_of_mono +#align AddCommGroup.injective_of_mono AddCommGrp.injective_of_mono -end AddCommGroupCat +end AddCommGrp -/-- Build an isomorphism in the category `GroupCat` from a `MulEquiv` between `Group`s. -/ +/-- Build an isomorphism in the category `Grp` from a `MulEquiv` between `Group`s. -/ @[to_additive (attr := simps)] -def MulEquiv.toGroupCatIso {X Y : GroupCat} (e : X ≃* Y) : X ≅ Y where +def MulEquiv.toGrpIso {X Y : Grp} (e : X ≃* Y) : X ≅ Y where hom := e.toMonoidHom inv := e.symm.toMonoidHom set_option linter.uppercaseLean3 false in -#align mul_equiv.to_Group_iso MulEquiv.toGroupCatIso +#align mul_equiv.to_Group_iso MulEquiv.toGrpIso set_option linter.uppercaseLean3 false in -#align add_equiv.to_AddGroup_iso AddEquiv.toAddGroupCatIso +#align add_equiv.to_AddGroup_iso AddEquiv.toAddGrpIso /-- Build an isomorphism in the category `AddGroup` from an `AddEquiv` between `AddGroup`s. -/ -add_decl_doc AddEquiv.toAddGroupCatIso +add_decl_doc AddEquiv.toAddGrpIso -/-- Build an isomorphism in the category `CommGroupCat` from a `MulEquiv` +/-- Build an isomorphism in the category `CommGrp` from a `MulEquiv` between `CommGroup`s. -/ @[to_additive (attr := simps)] -def MulEquiv.toCommGroupCatIso {X Y : CommGroupCat} (e : X ≃* Y) : X ≅ Y where +def MulEquiv.toCommGrpIso {X Y : CommGrp} (e : X ≃* Y) : X ≅ Y where hom := e.toMonoidHom inv := e.symm.toMonoidHom set_option linter.uppercaseLean3 false in -#align mul_equiv.to_CommGroup_iso MulEquiv.toCommGroupCatIso +#align mul_equiv.to_CommGroup_iso MulEquiv.toCommGrpIso set_option linter.uppercaseLean3 false in -#align add_equiv.to_AddCommGroup_iso AddEquiv.toAddCommGroupCatIso +#align add_equiv.to_AddCommGroup_iso AddEquiv.toAddCommGrpIso -/-- Build an isomorphism in the category `AddCommGroupCat` from an `AddEquiv` +/-- Build an isomorphism in the category `AddCommGrp` from an `AddEquiv` between `AddCommGroup`s. -/ -add_decl_doc AddEquiv.toAddCommGroupCatIso +add_decl_doc AddEquiv.toAddCommGrpIso namespace CategoryTheory.Iso -/-- Build a `MulEquiv` from an isomorphism in the category `GroupCat`. -/ +/-- Build a `MulEquiv` from an isomorphism in the category `Grp`. -/ @[to_additive (attr := simp)] -def groupIsoToMulEquiv {X Y : GroupCat} (i : X ≅ Y) : X ≃* Y := +def groupIsoToMulEquiv {X Y : Grp} (i : X ≅ Y) : X ≃* Y := MonoidHom.toMulEquiv i.hom i.inv i.hom_inv_id i.inv_hom_id set_option linter.uppercaseLean3 false in #align category_theory.iso.Group_iso_to_mul_equiv CategoryTheory.Iso.groupIsoToMulEquiv @@ -455,7 +455,7 @@ add_decl_doc addGroupIsoToAddEquiv /-- Build a `MulEquiv` from an isomorphism in the category `CommGroup`. -/ @[to_additive (attr := simps!)] -def commGroupIsoToMulEquiv {X Y : CommGroupCat} (i : X ≅ Y) : X ≃* Y := +def commGroupIsoToMulEquiv {X Y : CommGrp} (i : X ≅ Y) : X ≃* Y := MonoidHom.toMulEquiv i.hom i.inv i.hom_inv_id i.inv_hom_id set_option linter.uppercaseLean3 false in #align category_theory.iso.CommGroup_iso_to_mul_equiv CategoryTheory.Iso.commGroupIsoToMulEquiv @@ -468,10 +468,10 @@ add_decl_doc addCommGroupIsoToAddEquiv end CategoryTheory.Iso /-- multiplicative equivalences between `Group`s are the same as (isomorphic to) isomorphisms -in `GroupCat` -/ +in `Grp` -/ @[to_additive] -def mulEquivIsoGroupIso {X Y : GroupCat.{u}} : X ≃* Y ≅ X ≅ Y where - hom e := e.toGroupCatIso +def mulEquivIsoGroupIso {X Y : Grp.{u}} : X ≃* Y ≅ X ≅ Y where + hom e := e.toGrpIso inv i := i.groupIsoToMulEquiv set_option linter.uppercaseLean3 false in #align mul_equiv_iso_Group_iso mulEquivIsoGroupIso @@ -479,14 +479,14 @@ set_option linter.uppercaseLean3 false in #align add_equiv_iso_AddGroup_iso addEquivIsoAddGroupIso /-- Additive equivalences between `AddGroup`s are the same -as (isomorphic to) isomorphisms in `AddGroupCat`. -/ +as (isomorphic to) isomorphisms in `AddGrp`. -/ add_decl_doc addEquivIsoAddGroupIso /-- Multiplicative equivalences between `CommGroup`s are the same as (isomorphic to) isomorphisms -in `CommGroupCat`. -/ +in `CommGrp`. -/ @[to_additive] -def mulEquivIsoCommGroupIso {X Y : CommGroupCat.{u}} : X ≃* Y ≅ X ≅ Y where - hom e := e.toCommGroupCatIso +def mulEquivIsoCommGroupIso {X Y : CommGrp.{u}} : X ≃* Y ≅ X ≅ Y where + hom e := e.toCommGrpIso inv i := i.commGroupIsoToMulEquiv set_option linter.uppercaseLean3 false in #align mul_equiv_iso_CommGroup_iso mulEquivIsoCommGroupIso @@ -494,14 +494,14 @@ set_option linter.uppercaseLean3 false in #align add_equiv_iso_AddCommGroup_iso addEquivIsoAddCommGroupIso /-- Additive equivalences between `AddCommGroup`s are -the same as (isomorphic to) isomorphisms in `AddCommGroupCat`. -/ +the same as (isomorphic to) isomorphisms in `AddCommGrp`. -/ add_decl_doc addEquivIsoAddCommGroupIso namespace CategoryTheory.Aut /-- The (bundled) group of automorphisms of a type is isomorphic to the (bundled) group of permutations. -/ -def isoPerm {α : Type u} : GroupCat.of (Aut α) ≅ GroupCat.of (Equiv.Perm α) where +def isoPerm {α : Type u} : Grp.of (Aut α) ≅ Grp.of (Equiv.Perm α) where hom := { toFun := fun g => g.toEquiv map_one' := by aesop @@ -523,61 +523,61 @@ set_option linter.uppercaseLean3 false in end CategoryTheory.Aut @[to_additive] -instance GroupCat.forget_reflects_isos : (forget GroupCat.{u}).ReflectsIsomorphisms where +instance Grp.forget_reflects_isos : (forget Grp.{u}).ReflectsIsomorphisms where reflects {X Y} f _ := by - let i := asIso ((forget GroupCat).map f) + let i := asIso ((forget Grp).map f) let e : X ≃* Y := { i.toEquiv with map_mul' := map_mul _ } - exact e.toGroupCatIso.isIso_hom + exact e.toGrpIso.isIso_hom set_option linter.uppercaseLean3 false in -#align Group.forget_reflects_isos GroupCat.forget_reflects_isos +#align Group.forget_reflects_isos Grp.forget_reflects_isos set_option linter.uppercaseLean3 false in -#align AddGroup.forget_reflects_isos AddGroupCat.forget_reflects_isos +#align AddGroup.forget_reflects_isos AddGrp.forget_reflects_isos @[to_additive] -instance CommGroupCat.forget_reflects_isos : (forget CommGroupCat.{u}).ReflectsIsomorphisms where +instance CommGrp.forget_reflects_isos : (forget CommGrp.{u}).ReflectsIsomorphisms where reflects {X Y} f _ := by - let i := asIso ((forget CommGroupCat).map f) + let i := asIso ((forget CommGrp).map f) let e : X ≃* Y := { i.toEquiv with map_mul' := map_mul _} - exact e.toCommGroupCatIso.isIso_hom + exact e.toCommGrpIso.isIso_hom set_option linter.uppercaseLean3 false in -#align CommGroup.forget_reflects_isos CommGroupCat.forget_reflects_isos +#align CommGroup.forget_reflects_isos CommGrp.forget_reflects_isos set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget_reflects_isos AddCommGroupCat.forget_reflects_isos +#align AddCommGroup.forget_reflects_isos AddCommGrp.forget_reflects_isos -- note: in the following definitions, there is a problem with `@[to_additive]` -- as the `Category` instance is not found on the additive variant -- this variant is then renamed with a `Aux` suffix -/-- An alias for `GroupCat.{max u v}`, to deal around unification issues. -/ -@[to_additive (attr := nolint checkUnivs) GroupCatMaxAux - "An alias for `AddGroupCat.{max u v}`, to deal around unification issues."] -abbrev GroupCatMax.{u1, u2} := GroupCat.{max u1 u2} -/-- An alias for `AddGroupCat.{max u v}`, to deal around unification issues. -/ +/-- An alias for `Grp.{max u v}`, to deal around unification issues. -/ +@[to_additive (attr := nolint checkUnivs) GrpMaxAux + "An alias for `AddGrp.{max u v}`, to deal around unification issues."] +abbrev GrpMax.{u1, u2} := Grp.{max u1 u2} +/-- An alias for `AddGrp.{max u v}`, to deal around unification issues. -/ @[nolint checkUnivs] -abbrev AddGroupCatMax.{u1, u2} := AddGroupCat.{max u1 u2} +abbrev AddGrpMax.{u1, u2} := AddGrp.{max u1 u2} -/-- An alias for `CommGroupCat.{max u v}`, to deal around unification issues. -/ -@[to_additive (attr := nolint checkUnivs) AddCommGroupCatMaxAux - "An alias for `AddCommGroupCat.{max u v}`, to deal around unification issues."] -abbrev CommGroupCatMax.{u1, u2} := CommGroupCat.{max u1 u2} -/-- An alias for `AddCommGroupCat.{max u v}`, to deal around unification issues. -/ +/-- An alias for `CommGrp.{max u v}`, to deal around unification issues. -/ +@[to_additive (attr := nolint checkUnivs) AddCommGrpMaxAux + "An alias for `AddCommGrp.{max u v}`, to deal around unification issues."] +abbrev CommGrpMax.{u1, u2} := CommGrp.{max u1 u2} +/-- An alias for `AddCommGrp.{max u v}`, to deal around unification issues. -/ @[nolint checkUnivs] -abbrev AddCommGroupCatMax.{u1, u2} := AddCommGroupCat.{max u1 u2} +abbrev AddCommGrpMax.{u1, u2} := AddCommGrp.{max u1 u2} /-! `@[simp]` lemmas for `MonoidHom.comp` and categorical identities. -/ -@[to_additive (attr := simp)] theorem MonoidHom.comp_id_groupCat - {G : GroupCat.{u}} {H : Type u} [Group H] (f : G →* H) : f.comp (𝟙 G) = f := - Category.id_comp (GroupCat.ofHom f) -@[to_additive (attr := simp)] theorem MonoidHom.id_groupCat_comp - {G : Type u} [Group G] {H : GroupCat.{u}} (f : G →* H) : MonoidHom.comp (𝟙 H) f = f := - Category.comp_id (GroupCat.ofHom f) - -@[to_additive (attr := simp)] theorem MonoidHom.comp_id_commGroupCat - {G : CommGroupCat.{u}} {H : Type u} [CommGroup H] (f : G →* H) : f.comp (𝟙 G) = f := - Category.id_comp (CommGroupCat.ofHom f) -@[to_additive (attr := simp)] theorem MonoidHom.id_commGroupCat_comp - {G : Type u} [CommGroup G] {H : CommGroupCat.{u}} (f : G →* H) : MonoidHom.comp (𝟙 H) f = f := - Category.comp_id (CommGroupCat.ofHom f) +@[to_additive (attr := simp)] theorem MonoidHom.comp_id_grp + {G : Grp.{u}} {H : Type u} [Group H] (f : G →* H) : f.comp (𝟙 G) = f := + Category.id_comp (Grp.ofHom f) +@[to_additive (attr := simp)] theorem MonoidHom.id_grp_comp + {G : Type u} [Group G] {H : Grp.{u}} (f : G →* H) : MonoidHom.comp (𝟙 H) f = f := + Category.comp_id (Grp.ofHom f) + +@[to_additive (attr := simp)] theorem MonoidHom.comp_id_commGrp + {G : CommGrp.{u}} {H : Type u} [CommGroup H] (f : G →* H) : f.comp (𝟙 G) = f := + Category.id_comp (CommGrp.ofHom f) +@[to_additive (attr := simp)] theorem MonoidHom.id_commGrp_comp + {G : Type u} [CommGroup G] {H : CommGrp.{u}} (f : G →* H) : MonoidHom.comp (𝟙 H) f = f := + Category.comp_id (CommGrp.ofHom f) diff --git a/Mathlib/Algebra/Category/GroupCat/Biproducts.lean b/Mathlib/Algebra/Category/Grp/Biproducts.lean similarity index 61% rename from Mathlib/Algebra/Category/GroupCat/Biproducts.lean rename to Mathlib/Algebra/Category/Grp/Biproducts.lean index 157b1ac831d9e..e4cb3e2bd6b35 100644 --- a/Mathlib/Algebra/Category/GroupCat/Biproducts.lean +++ b/Mathlib/Algebra/Category/Grp/Biproducts.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Algebra.Group.Pi.Lemmas -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.CategoryTheory.Preadditive.Biproducts -import Mathlib.Algebra.Category.GroupCat.Limits +import Mathlib.Algebra.Category.Grp.Limits import Mathlib.Tactic.CategoryTheory.Elementwise #align_import algebra.category.Group.biproducts from "leanprover-community/mathlib"@"234ddfeaa5572bc13716dd215c6444410a679a8e" @@ -22,25 +22,25 @@ open CategoryTheory.Limits universe w u -namespace AddCommGroupCat +namespace AddCommGrp set_option linter.uppercaseLean3 false -- `AddCommGroup` --- As `AddCommGroupCat` is preadditive, and has all limits, it automatically has biproducts. -instance : HasBinaryBiproducts AddCommGroupCat := +-- As `AddCommGrp` is preadditive, and has all limits, it automatically has biproducts. +instance : HasBinaryBiproducts AddCommGrp := HasBinaryBiproducts.of_hasBinaryProducts -instance : HasFiniteBiproducts AddCommGroupCat := +instance : HasFiniteBiproducts AddCommGrp := HasFiniteBiproducts.of_hasFiniteProducts -- We now construct explicit limit data, -- so we can compare the biproducts to the usual unbundled constructions. -/-- Construct limit data for a binary product in `AddCommGroupCat`, using -`AddCommGroupCat.of (G × H)`. +/-- Construct limit data for a binary product in `AddCommGrp`, using +`AddCommGrp.of (G × H)`. -/ @[simps cone_pt isLimit_lift] -def binaryProductLimitCone (G H : AddCommGroupCat.{u}) : Limits.LimitCone (pair G H) where +def binaryProductLimitCone (G H : AddCommGrp.{u}) : Limits.LimitCone (pair G H) where cone := - { pt := AddCommGroupCat.of (G × H) + { pt := AddCommGrp.of (G × H) π := { app := fun j => Discrete.casesOn j fun j => @@ -52,57 +52,57 @@ def binaryProductLimitCone (G H : AddCommGroupCat.{u}) : Limits.LimitCone (pair uniq := fun s m w => by simp_rw [← w ⟨WalkingPair.left⟩, ← w ⟨WalkingPair.right⟩] rfl } -#align AddCommGroup.binary_product_limit_cone AddCommGroupCat.binaryProductLimitCone +#align AddCommGroup.binary_product_limit_cone AddCommGrp.binaryProductLimitCone @[simp] -theorem binaryProductLimitCone_cone_π_app_left (G H : AddCommGroupCat.{u}) : +theorem binaryProductLimitCone_cone_π_app_left (G H : AddCommGrp.{u}) : (binaryProductLimitCone G H).cone.π.app ⟨WalkingPair.left⟩ = AddMonoidHom.fst G H := rfl -#align AddCommGroup.binary_product_limit_cone_cone_π_app_left AddCommGroupCat.binaryProductLimitCone_cone_π_app_left +#align AddCommGroup.binary_product_limit_cone_cone_π_app_left AddCommGrp.binaryProductLimitCone_cone_π_app_left @[simp] -theorem binaryProductLimitCone_cone_π_app_right (G H : AddCommGroupCat.{u}) : +theorem binaryProductLimitCone_cone_π_app_right (G H : AddCommGrp.{u}) : (binaryProductLimitCone G H).cone.π.app ⟨WalkingPair.right⟩ = AddMonoidHom.snd G H := rfl -#align AddCommGroup.binary_product_limit_cone_cone_π_app_right AddCommGroupCat.binaryProductLimitCone_cone_π_app_right +#align AddCommGroup.binary_product_limit_cone_cone_π_app_right AddCommGrp.binaryProductLimitCone_cone_π_app_right -/-- We verify that the biproduct in `AddCommGroupCat` is isomorphic to +/-- We verify that the biproduct in `AddCommGrp` is isomorphic to the cartesian product of the underlying types: -/ @[simps! hom_apply] -noncomputable def biprodIsoProd (G H : AddCommGroupCat.{u}) : - (G ⊞ H : AddCommGroupCat) ≅ AddCommGroupCat.of (G × H) := +noncomputable def biprodIsoProd (G H : AddCommGrp.{u}) : + (G ⊞ H : AddCommGrp) ≅ AddCommGrp.of (G × H) := IsLimit.conePointUniqueUpToIso (BinaryBiproduct.isLimit G H) (binaryProductLimitCone G H).isLimit -#align AddCommGroup.biprod_iso_prod AddCommGroupCat.biprodIsoProd +#align AddCommGroup.biprod_iso_prod AddCommGrp.biprodIsoProd -- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] AddCommGroupCat.biprodIsoProd_hom_apply +attribute [nolint simpNF] AddCommGrp.biprodIsoProd_hom_apply @[simp, elementwise] -theorem biprodIsoProd_inv_comp_fst (G H : AddCommGroupCat.{u}) : +theorem biprodIsoProd_inv_comp_fst (G H : AddCommGrp.{u}) : (biprodIsoProd G H).inv ≫ biprod.fst = AddMonoidHom.fst G H := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk WalkingPair.left) -#align AddCommGroup.biprod_iso_prod_inv_comp_fst AddCommGroupCat.biprodIsoProd_inv_comp_fst +#align AddCommGroup.biprod_iso_prod_inv_comp_fst AddCommGrp.biprodIsoProd_inv_comp_fst @[simp, elementwise] -theorem biprodIsoProd_inv_comp_snd (G H : AddCommGroupCat.{u}) : +theorem biprodIsoProd_inv_comp_snd (G H : AddCommGrp.{u}) : (biprodIsoProd G H).inv ≫ biprod.snd = AddMonoidHom.snd G H := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk WalkingPair.right) -#align AddCommGroup.biprod_iso_prod_inv_comp_snd AddCommGroupCat.biprodIsoProd_inv_comp_snd +#align AddCommGroup.biprod_iso_prod_inv_comp_snd AddCommGrp.biprodIsoProd_inv_comp_snd namespace HasLimit -variable {J : Type w} (f : J → AddCommGroupCat.{max w u}) +variable {J : Type w} (f : J → AddCommGrp.{max w u}) /-- The map from an arbitrary cone over an indexed family of abelian groups to the cartesian product of those groups. -/ --- This was marked `@[simps]` until we made `AddCommGroupCat.coe_of` a simp lemma, +-- This was marked `@[simps]` until we made `AddCommGrp.coe_of` a simp lemma, -- after which the simp normal form linter complains. -- The generated simp lemmas were not used in Mathlib. -- Possible solution: higher priority function coercions that remove the `of`? -- @[simps] -def lift (s : Fan f) : s.pt ⟶ AddCommGroupCat.of (∀ j, f j) where +def lift (s : Fan f) : s.pt ⟶ AddCommGrp.of (∀ j, f j) where toFun x j := s.π.app ⟨j⟩ x map_zero' := by simp only [Functor.const_obj_obj, map_zero] @@ -110,15 +110,15 @@ def lift (s : Fan f) : s.pt ⟶ AddCommGroupCat.of (∀ j, f j) where map_add' x y := by simp only [Functor.const_obj_obj, map_add] rfl -#align AddCommGroup.has_limit.lift AddCommGroupCat.HasLimit.lift +#align AddCommGroup.has_limit.lift AddCommGrp.HasLimit.lift -/-- Construct limit data for a product in `AddCommGroupCat`, using -`AddCommGroupCat.of (∀ j, F.obj j)`. +/-- Construct limit data for a product in `AddCommGrp`, using +`AddCommGrp.of (∀ j, F.obj j)`. -/ @[simps] def productLimitCone : Limits.LimitCone (Discrete.functor f) where cone := - { pt := AddCommGroupCat.of (∀ j, f j) + { pt := AddCommGrp.of (∀ j, f j) π := Discrete.natTrans fun j => Pi.evalAddMonoidHom (fun j => f j) j.as } isLimit := { lift := lift.{_, u} f @@ -127,7 +127,7 @@ def productLimitCone : Limits.LimitCone (Discrete.functor f) where ext x funext j exact congr_arg (fun g : s.pt ⟶ f j => (g : s.pt → f j) x) (w ⟨j⟩) } -#align AddCommGroup.has_limit.product_limit_cone AddCommGroupCat.HasLimit.productLimitCone +#align AddCommGroup.has_limit.product_limit_cone AddCommGrp.HasLimit.productLimitCone end HasLimit @@ -135,22 +135,22 @@ open HasLimit variable {J : Type} [Finite J] -/-- We verify that the biproduct we've just defined is isomorphic to the `AddCommGroupCat` structure +/-- We verify that the biproduct we've just defined is isomorphic to the `AddCommGrp` structure on the dependent function type. -/ @[simps! hom_apply] -noncomputable def biproductIsoPi (f : J → AddCommGroupCat.{u}) : - (⨁ f : AddCommGroupCat) ≅ AddCommGroupCat.of (∀ j, f j) := +noncomputable def biproductIsoPi (f : J → AddCommGrp.{u}) : + (⨁ f : AddCommGrp) ≅ AddCommGrp.of (∀ j, f j) := IsLimit.conePointUniqueUpToIso (biproduct.isLimit f) (productLimitCone f).isLimit -#align AddCommGroup.biproduct_iso_pi AddCommGroupCat.biproductIsoPi +#align AddCommGroup.biproduct_iso_pi AddCommGrp.biproductIsoPi -- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] AddCommGroupCat.biproductIsoPi_hom_apply +attribute [nolint simpNF] AddCommGrp.biproductIsoPi_hom_apply @[simp, elementwise] -theorem biproductIsoPi_inv_comp_π (f : J → AddCommGroupCat.{u}) (j : J) : +theorem biproductIsoPi_inv_comp_π (f : J → AddCommGrp.{u}) (j : J) : (biproductIsoPi f).inv ≫ biproduct.π f j = Pi.evalAddMonoidHom (fun j => f j) j := IsLimit.conePointUniqueUpToIso_inv_comp _ _ (Discrete.mk j) -#align AddCommGroup.biproduct_iso_pi_inv_comp_π AddCommGroupCat.biproductIsoPi_inv_comp_π +#align AddCommGroup.biproduct_iso_pi_inv_comp_π AddCommGrp.biproductIsoPi_inv_comp_π -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Colimits.lean b/Mathlib/Algebra/Category/Grp/Colimits.lean similarity index 77% rename from Mathlib/Algebra/Category/GroupCat/Colimits.lean rename to Mathlib/Algebra/Category/Grp/Colimits.lean index 5dd0ab37b2ba6..c1b4d321c1e41 100644 --- a/Mathlib/Algebra/Category/GroupCat/Colimits.lean +++ b/Mathlib/Algebra/Category/Grp/Colimits.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.GroupTheory.QuotientGroup import Mathlib.CategoryTheory.Limits.Shapes.Kernels import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits @@ -19,7 +19,7 @@ It is a very uniform approach, that conceivably could be synthesised directly by a tactic that analyses the shape of `AddCommGroup` and `MonoidHom`. TODO: -In fact, in `AddCommGroupCat` there is a much nicer model of colimits as quotients +In fact, in `AddCommGrp` there is a much nicer model of colimits as quotients of finitely supported functions, and we really should implement this as well (or instead). -/ @@ -33,14 +33,14 @@ open CategoryTheory Limits -- [ROBOT VOICE]: -- You should pretend for now that this file was automatically generated. -- It follows the same template as colimits in Mon. -namespace AddCommGroupCat +namespace AddCommGrp -variable {J : Type u} [Category.{v} J] (F : J ⥤ AddCommGroupCat.{max u v w}) +variable {J : Type u} [Category.{v} J] (F : J ⥤ AddCommGrp.{max u v w}) namespace Colimits /-! -We build the colimit of a diagram in `AddCommGroupCat` by constructing the +We build the colimit of a diagram in `AddCommGrp` by constructing the free group on the disjoint union of all the abelian groups in the diagram, then taking the quotient by the abelian group laws within each abelian group, and the identifications given by the morphisms in the diagram. @@ -56,7 +56,7 @@ inductive Prequotient | zero : Prequotient | neg : Prequotient → Prequotient | add : Prequotient → Prequotient → Prequotient -#align AddCommGroup.colimits.prequotient AddCommGroupCat.Colimits.Prequotient +#align AddCommGroup.colimits.prequotient AddCommGrp.Colimits.Prequotient instance : Inhabited (Prequotient.{w} F) := ⟨Prequotient.zero⟩ @@ -90,7 +90,7 @@ inductive Relation : Prequotient.{w} F → Prequotient.{w} F → Prop | add_left_neg : ∀ x, Relation (add (neg x) x) zero | add_comm : ∀ x y, Relation (add x y) (add y x) | add_assoc : ∀ x y z, Relation (add (add x y) z) (add x (add y z)) -#align AddCommGroup.colimits.relation AddCommGroupCat.Colimits.Relation +#align AddCommGroup.colimits.relation AddCommGrp.Colimits.Relation /-- The setoid corresponding to group expressions modulo abelian group relations and identifications. @@ -98,15 +98,15 @@ The setoid corresponding to group expressions modulo abelian group relations and def colimitSetoid : Setoid (Prequotient.{w} F) where r := Relation F iseqv := ⟨Relation.refl, fun r => Relation.symm _ _ r, fun r => Relation.trans _ _ _ r⟩ -#align AddCommGroup.colimits.colimit_setoid AddCommGroupCat.Colimits.colimitSetoid +#align AddCommGroup.colimits.colimit_setoid AddCommGrp.Colimits.colimitSetoid attribute [instance] colimitSetoid -/-- The underlying type of the colimit of a diagram in `AddCommGroupCat`. +/-- The underlying type of the colimit of a diagram in `AddCommGrp`. -/ def ColimitType : Type max u v w := Quotient (colimitSetoid.{w} F) -#align AddCommGroup.colimits.colimit_type AddCommGroupCat.Colimits.ColimitType +#align AddCommGroup.colimits.colimit_type AddCommGrp.Colimits.ColimitType instance : Zero (ColimitType.{w} F) where zero := Quotient.mk _ zero @@ -133,7 +133,7 @@ instance ColimitTypeInhabited : Inhabited (ColimitType.{w} F) := ⟨0⟩ @[simp] theorem quot_zero : Quot.mk Setoid.r zero = (0 : ColimitType.{w} F) := rfl -#align AddCommGroup.colimits.quot_zero AddCommGroupCat.Colimits.quot_zero +#align AddCommGroup.colimits.quot_zero AddCommGrp.Colimits.quot_zero @[simp] theorem quot_neg (x) : @@ -141,7 +141,7 @@ theorem quot_neg (x) : (by exact Quot.mk Setoid.r (neg x) : ColimitType.{w} F) = -(by exact Quot.mk Setoid.r x) := rfl -#align AddCommGroup.colimits.quot_neg AddCommGroupCat.Colimits.quot_neg +#align AddCommGroup.colimits.quot_neg AddCommGrp.Colimits.quot_neg @[simp] theorem quot_add (x y) : @@ -149,17 +149,17 @@ theorem quot_add (x y) : -- Porting note: force Lean to treat `ColimitType F` no as `Quot _` (by exact Quot.mk Setoid.r x) + (by exact Quot.mk Setoid.r y) := rfl -#align AddCommGroup.colimits.quot_add AddCommGroupCat.Colimits.quot_add +#align AddCommGroup.colimits.quot_add AddCommGrp.Colimits.quot_add /-- The bundled abelian group giving the colimit of a diagram. -/ -def colimit : AddCommGroupCat := - AddCommGroupCat.of (ColimitType.{w} F) -#align AddCommGroup.colimits.colimit AddCommGroupCat.Colimits.colimit +def colimit : AddCommGrp := + AddCommGrp.of (ColimitType.{w} F) +#align AddCommGroup.colimits.colimit AddCommGrp.Colimits.colimit /-- The function from a given abelian group in the diagram to the colimit abelian group. -/ def coconeFun (j : J) (x : F.obj j) : ColimitType.{w} F := Quot.mk _ (Prequotient.of j x) -#align AddCommGroup.colimits.cocone_fun AddCommGroupCat.Colimits.coconeFun +#align AddCommGroup.colimits.cocone_fun AddCommGrp.Colimits.coconeFun /-- The group homomorphism from a given abelian group in the diagram to the colimit abelian group. -/ @@ -167,7 +167,7 @@ def coconeMorphism (j : J) : F.obj j ⟶ colimit.{w} F where toFun := coconeFun F j map_zero' := by apply Quot.sound; apply Relation.zero map_add' := by intros; apply Quot.sound; apply Relation.add -#align AddCommGroup.colimits.cocone_morphism AddCommGroupCat.Colimits.coconeMorphism +#align AddCommGroup.colimits.cocone_morphism AddCommGrp.Colimits.coconeMorphism @[simp] theorem cocone_naturality {j j' : J} (f : j ⟶ j') : @@ -175,20 +175,20 @@ theorem cocone_naturality {j j' : J} (f : j ⟶ j') : ext apply Quot.sound apply Relation.map -#align AddCommGroup.colimits.cocone_naturality AddCommGroupCat.Colimits.cocone_naturality +#align AddCommGroup.colimits.cocone_naturality AddCommGrp.Colimits.cocone_naturality @[simp] theorem cocone_naturality_components (j j' : J) (f : j ⟶ j') (x : F.obj j) : (coconeMorphism.{w} F j') (F.map f x) = (coconeMorphism F j) x := by rw [← cocone_naturality F f] rfl -#align AddCommGroup.colimits.cocone_naturality_components AddCommGroupCat.Colimits.cocone_naturality_components +#align AddCommGroup.colimits.cocone_naturality_components AddCommGrp.Colimits.cocone_naturality_components /-- The cocone over the proposed colimit abelian group. -/ def colimitCocone : Cocone F where pt := colimit.{w} F ι := { app := coconeMorphism F } -#align AddCommGroup.colimits.colimit_cocone AddCommGroupCat.Colimits.colimitCocone +#align AddCommGroup.colimits.colimit_cocone AddCommGrp.Colimits.colimitCocone /-- The function from the free abelian group on the diagram to the cone point of any other cocone. -/ @@ -198,7 +198,7 @@ def descFunLift (s : Cocone F) : Prequotient.{w} F → s.pt | zero => 0 | neg x => -descFunLift s x | add x y => descFunLift s x + descFunLift s y -#align AddCommGroup.colimits.desc_fun_lift AddCommGroupCat.Colimits.descFunLift +#align AddCommGroup.colimits.desc_fun_lift AddCommGrp.Colimits.descFunLift /-- The function from the colimit abelian group to the cone point of any other cocone. -/ def descFun (s : Cocone F) : ColimitType.{w} F → s.pt := by @@ -222,7 +222,7 @@ def descFun (s : Cocone F) : ColimitType.{w} F → s.pt := by | add_left_neg => dsimp; rw [add_left_neg] | add_comm => dsimp; rw [add_comm] | add_assoc => dsimp; rw [add_assoc] -#align AddCommGroup.colimits.desc_fun AddCommGroupCat.Colimits.descFun +#align AddCommGroup.colimits.desc_fun AddCommGrp.Colimits.descFun /-- The group homomorphism from the colimit abelian group to the cone point of any other cocone. -/ def descMorphism (s : Cocone F) : colimit.{w} F ⟶ s.pt where @@ -230,7 +230,7 @@ def descMorphism (s : Cocone F) : colimit.{w} F ⟶ s.pt where map_zero' := rfl -- Porting note: in `mathlib3`, nothing needs to be done after `induction` map_add' x y := Quot.induction_on₂ x y fun _ _ => by dsimp; rw [← quot_add F]; rfl -#align AddCommGroup.colimits.desc_morphism AddCommGroupCat.Colimits.descMorphism +#align AddCommGroup.colimits.desc_morphism AddCommGrp.Colimits.descMorphism /-- Evidence that the proposed colimit is the colimit. -/ def colimitCoconeIsColimit : IsColimit (colimitCocone.{w} F) where @@ -248,7 +248,7 @@ def colimitCoconeIsColimit : IsColimit (colimitCocone.{w} F) where | add x y ihx ihy => simp only [quot_add] erw [m.map_add, (descMorphism F s).map_add, ihx, ihy] -#align AddCommGroup.colimits.colimit_cocone_is_colimit AddCommGroupCat.Colimits.colimitCoconeIsColimit +#align AddCommGroup.colimits.colimit_cocone_is_colimit AddCommGrp.Colimits.colimitCoconeIsColimit end Colimits @@ -256,41 +256,41 @@ lemma hasColimit : HasColimit F := ⟨_, Colimits.colimitCoconeIsColimit.{w} F variable (J) -lemma hasColimitsOfShape : HasColimitsOfShape J AddCommGroupCat.{max u v w} where +lemma hasColimitsOfShape : HasColimitsOfShape J AddCommGrp.{max u v w} where has_colimit F := hasColimit.{w} F -lemma hasColimitsOfSize : HasColimitsOfSize.{v, u} AddCommGroupCat.{max u v w} := +lemma hasColimitsOfSize : HasColimitsOfSize.{v, u} AddCommGrp.{max u v w} := ⟨fun _ => hasColimitsOfShape.{w} _⟩ -instance hasColimits : HasColimits AddCommGroupCat.{w} := hasColimitsOfSize.{w} -#align AddCommGroup.colimits.has_colimits_AddCommGroup AddCommGroupCat.hasColimits +instance hasColimits : HasColimits AddCommGrp.{w} := hasColimitsOfSize.{w} +#align AddCommGroup.colimits.has_colimits_AddCommGroup AddCommGrp.hasColimits -instance : HasColimitsOfSize.{v, v} (AddCommGroupCatMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{u, u} (AddCommGroupCatMax.{u, v}) := hasColimitsOfSize.{v} -instance : HasColimitsOfSize.{u, v} (AddCommGroupCatMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{v, u} (AddCommGroupCatMax.{u, v}) := hasColimitsOfSize.{u} -instance : HasColimitsOfSize.{0, 0} (AddCommGroupCat.{u}) := hasColimitsOfSize.{u, 0, 0} +instance : HasColimitsOfSize.{v, v} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} +instance : HasColimitsOfSize.{u, u} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{v} +instance : HasColimitsOfSize.{u, v} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} +instance : HasColimitsOfSize.{v, u} (AddCommGrpMax.{u, v}) := hasColimitsOfSize.{u} +instance : HasColimitsOfSize.{0, 0} (AddCommGrp.{u}) := hasColimitsOfSize.{u, 0, 0} -example : HasColimits AddCommGroupCatMax.{v, u} := +example : HasColimits AddCommGrpMax.{v, u} := inferInstance -example : HasColimits AddCommGroupCatMax.{u, v} := +example : HasColimits AddCommGrpMax.{u, v} := inferInstance -example : HasColimits AddCommGroupCat.{u} := +example : HasColimits AddCommGrp.{u} := inferInstance -end AddCommGroupCat +end AddCommGrp -namespace AddCommGroupCat +namespace AddCommGrp open QuotientAddGroup -/-- The categorical cokernel of a morphism in `AddCommGroupCat` +/-- The categorical cokernel of a morphism in `AddCommGrp` agrees with the usual group-theoretical quotient. -/ -noncomputable def cokernelIsoQuotient {G H : AddCommGroupCat.{u}} (f : G ⟶ H) : - cokernel f ≅ AddCommGroupCat.of (H ⧸ AddMonoidHom.range f) where +noncomputable def cokernelIsoQuotient {G H : AddCommGrp.{u}} (f : G ⟶ H) : + cokernel f ≅ AddCommGrp.of (H ⧸ AddMonoidHom.range f) where hom := cokernel.desc f (mk' _) <| by ext x apply Quotient.sound @@ -309,6 +309,6 @@ noncomputable def cokernelIsoQuotient {G H : AddCommGroupCat.{u}} (f : G ⟶ H) inv_hom_id := by ext x exact QuotientAddGroup.induction_on x <| cokernel.π_desc_apply f _ _ -#align AddCommGroup.cokernel_iso_quotient AddCommGroupCat.cokernelIsoQuotient +#align AddCommGroup.cokernel_iso_quotient AddCommGrp.cokernelIsoQuotient -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/EnoughInjectives.lean b/Mathlib/Algebra/Category/Grp/EnoughInjectives.lean similarity index 61% rename from Mathlib/Algebra/Category/GroupCat/EnoughInjectives.lean rename to Mathlib/Algebra/Category/Grp/EnoughInjectives.lean index a4ba112ad0f75..781920a5448fc 100644 --- a/Mathlib/Algebra/Category/GroupCat/EnoughInjectives.lean +++ b/Mathlib/Algebra/Category/Grp/EnoughInjectives.lean @@ -5,8 +5,8 @@ Authors: Jujian Zhang, Junyan Xu -/ import Mathlib.Algebra.Module.CharacterModule -import Mathlib.Algebra.Category.GroupCat.EquivalenceGroupAddGroup -import Mathlib.Algebra.Category.GroupCat.EpiMono +import Mathlib.Algebra.Category.Grp.EquivalenceGroupAddGroup +import Mathlib.Algebra.Category.Grp.EpiMono /-! @@ -17,32 +17,32 @@ injective presentation for `A`, hence category of abelian groups has enough inje ## Implementation notes -This file is split from `Mathlib.Algebra.GroupCat.Injective` is to prevent import loop. -This file's dependency imports `Mathlib.Algebra.GroupCat.Injective`. +This file is split from `Mathlib.Algebra.Grp.Injective` is to prevent import loop. +This file's dependency imports `Mathlib.Algebra.Grp.Injective`. -/ open CategoryTheory universe u -namespace AddCommGroupCat +namespace AddCommGrp open CharacterModule -instance enoughInjectives : EnoughInjectives AddCommGroupCat.{u} where +instance enoughInjectives : EnoughInjectives AddCommGrp.{u} where presentation A_ := Nonempty.intro { J := of <| (CharacterModule A_) → ULift.{u} (AddCircle (1 : ℚ)) injective := injective_of_divisible _ f := ⟨⟨fun a i ↦ ULift.up (i a), by aesop⟩, by aesop⟩ - mono := (AddCommGroupCat.mono_iff_injective _).mpr <| (injective_iff_map_eq_zero _).mpr + mono := (AddCommGrp.mono_iff_injective _).mpr <| (injective_iff_map_eq_zero _).mpr fun a h0 ↦ eq_zero_of_character_apply (congr_arg ULift.down <| congr_fun h0 ·) } -end AddCommGroupCat +end AddCommGrp -namespace CommGroupCat +namespace CommGrp -instance enoughInjectives : EnoughInjectives CommGroupCat.{u} := +instance enoughInjectives : EnoughInjectives CommGrp.{u} := EnoughInjectives.of_equivalence commGroupAddCommGroupEquivalence.functor -end CommGroupCat +end CommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/EpiMono.lean b/Mathlib/Algebra/Category/Grp/EpiMono.lean similarity index 70% rename from Mathlib/Algebra/Category/GroupCat/EpiMono.lean rename to Mathlib/Algebra/Category/Grp/EpiMono.lean index f5b70cf822869..aee84486d0f1d 100644 --- a/Mathlib/Algebra/Category/GroupCat/EpiMono.lean +++ b/Mathlib/Algebra/Category/Grp/EpiMono.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import Mathlib.Algebra.Category.GroupCat.EquivalenceGroupAddGroup +import Mathlib.Algebra.Category.Grp.EquivalenceGroupAddGroup import Mathlib.GroupTheory.QuotientGroup import Mathlib.CategoryTheory.ConcreteCategory.EpiMono import Mathlib.CategoryTheory.Limits.Constructions.EpiMono @@ -67,36 +67,36 @@ section open CategoryTheory -namespace GroupCat +namespace Grp set_option linter.uppercaseLean3 false -- Porting note: already have Group G but Lean can't use that @[to_additive] -instance (G : GroupCat) : Group G.α := +instance (G : Grp) : Group G.α := G.str -variable {A B : GroupCat.{u}} (f : A ⟶ B) +variable {A B : Grp.{u}} (f : A ⟶ B) @[to_additive] theorem ker_eq_bot_of_mono [Mono f] : f.ker = ⊥ := MonoidHom.ker_eq_bot_of_cancel fun u _ => - (@cancel_mono _ _ _ _ _ f _ (show GroupCat.of f.ker ⟶ A from u) _).1 -#align Group.ker_eq_bot_of_mono GroupCat.ker_eq_bot_of_mono -#align AddGroup.ker_eq_bot_of_mono AddGroupCat.ker_eq_bot_of_mono + (@cancel_mono _ _ _ _ _ f _ (show Grp.of f.ker ⟶ A from u) _).1 +#align Group.ker_eq_bot_of_mono Grp.ker_eq_bot_of_mono +#align AddGroup.ker_eq_bot_of_mono AddGrp.ker_eq_bot_of_mono @[to_additive] theorem mono_iff_ker_eq_bot : Mono f ↔ f.ker = ⊥ := ⟨fun _ => ker_eq_bot_of_mono f, fun h => ConcreteCategory.mono_of_injective _ <| (MonoidHom.ker_eq_bot_iff f).1 h⟩ -#align Group.mono_iff_ker_eq_bot GroupCat.mono_iff_ker_eq_bot -#align AddGroup.mono_iff_ker_eq_bot AddGroupCat.mono_iff_ker_eq_bot +#align Group.mono_iff_ker_eq_bot Grp.mono_iff_ker_eq_bot +#align AddGroup.mono_iff_ker_eq_bot AddGrp.mono_iff_ker_eq_bot @[to_additive] theorem mono_iff_injective : Mono f ↔ Function.Injective f := Iff.trans (mono_iff_ker_eq_bot f) <| MonoidHom.ker_eq_bot_iff f -#align Group.mono_iff_injective GroupCat.mono_iff_injective -#align AddGroup.mono_iff_injective AddGroupCat.mono_iff_injective +#align Group.mono_iff_injective Grp.mono_iff_injective +#align AddGroup.mono_iff_injective AddGrp.mono_iff_injective namespace SurjectiveOfEpiAuxs @@ -108,7 +108,7 @@ local notation "X" => Set.range (· • (f.range : Set B) : B → Set B) inductive XWithInfinity | fromCoset : Set.range (· • (f.range : Set B) : B → Set B) → XWithInfinity | infinity : XWithInfinity -#align Group.surjective_of_epi_auxs.X_with_infinity GroupCat.SurjectiveOfEpiAuxs.XWithInfinity +#align Group.surjective_of_epi_auxs.X_with_infinity Grp.SurjectiveOfEpiAuxs.XWithInfinity open XWithInfinity Equiv.Perm @@ -134,7 +134,7 @@ theorem mul_smul (b b' : B) (x : X') : (b * b') • x = b • b' • x := change fromCoset _ = fromCoset _ simp only [leftCoset_assoc] | ∞ => rfl -#align Group.surjective_of_epi_auxs.mul_smul GroupCat.SurjectiveOfEpiAuxs.mul_smul +#align Group.surjective_of_epi_auxs.mul_smul Grp.SurjectiveOfEpiAuxs.mul_smul theorem one_smul (x : X') : (1 : B) • x = x := match x with @@ -142,7 +142,7 @@ theorem one_smul (x : X') : (1 : B) • x = x := change fromCoset _ = fromCoset _ simp only [one_leftCoset, Subtype.ext_iff_val] | ∞ => rfl -#align Group.surjective_of_epi_auxs.one_smul GroupCat.SurjectiveOfEpiAuxs.one_smul +#align Group.surjective_of_epi_auxs.one_smul Grp.SurjectiveOfEpiAuxs.one_smul theorem fromCoset_eq_of_mem_range {b : B} (hb : b ∈ f.range) : fromCoset ⟨b • ↑f.range, b, rfl⟩ = fromCoset ⟨f.range, 1, one_leftCoset _⟩ := by @@ -152,7 +152,7 @@ theorem fromCoset_eq_of_mem_range {b : B} (hb : b ∈ f.range) : nth_rw 2 [show (f.range : Set B.α) = (1 : B) • f.range from (one_leftCoset _).symm] rw [leftCoset_eq_iff, mul_one] exact Subgroup.inv_mem _ hb -#align Group.surjective_of_epi_auxs.from_coset_eq_of_mem_range GroupCat.SurjectiveOfEpiAuxs.fromCoset_eq_of_mem_range +#align Group.surjective_of_epi_auxs.from_coset_eq_of_mem_range Grp.SurjectiveOfEpiAuxs.fromCoset_eq_of_mem_range example (G : Type) [Group G] (S : Subgroup G) : Set G := S @@ -166,7 +166,7 @@ theorem fromCoset_ne_of_nin_range {b : B} (hb : b ∉ f.range) : nth_rw 2 [show (f.range : Set B.α) = (1 : B) • f.range from (one_leftCoset _).symm] at r rw [leftCoset_eq_iff, mul_one] at r exact hb (inv_inv b ▸ Subgroup.inv_mem _ r) -#align Group.surjective_of_epi_auxs.from_coset_ne_of_nin_range GroupCat.SurjectiveOfEpiAuxs.fromCoset_ne_of_nin_range +#align Group.surjective_of_epi_auxs.from_coset_ne_of_nin_range Grp.SurjectiveOfEpiAuxs.fromCoset_ne_of_nin_range instance : DecidableEq X' := Classical.decEq _ @@ -175,31 +175,31 @@ instance : DecidableEq X' := -/ noncomputable def tau : SX' := Equiv.swap (fromCoset ⟨↑f.range, ⟨1, one_leftCoset _⟩⟩) ∞ -#align Group.surjective_of_epi_auxs.tau GroupCat.SurjectiveOfEpiAuxs.tau +#align Group.surjective_of_epi_auxs.tau Grp.SurjectiveOfEpiAuxs.tau local notation "τ" => tau f theorem τ_apply_infinity : τ ∞ = fromCoset ⟨f.range, 1, one_leftCoset _⟩ := Equiv.swap_apply_right _ _ -#align Group.surjective_of_epi_auxs.τ_apply_infinity GroupCat.SurjectiveOfEpiAuxs.τ_apply_infinity +#align Group.surjective_of_epi_auxs.τ_apply_infinity Grp.SurjectiveOfEpiAuxs.τ_apply_infinity theorem τ_apply_fromCoset : τ (fromCoset ⟨f.range, 1, one_leftCoset _⟩) = ∞ := Equiv.swap_apply_left _ _ -#align Group.surjective_of_epi_auxs.τ_apply_fromCoset GroupCat.SurjectiveOfEpiAuxs.τ_apply_fromCoset +#align Group.surjective_of_epi_auxs.τ_apply_fromCoset Grp.SurjectiveOfEpiAuxs.τ_apply_fromCoset theorem τ_apply_fromCoset' (x : B) (hx : x ∈ f.range) : τ (fromCoset ⟨x • ↑f.range, ⟨x, rfl⟩⟩) = ∞ := (fromCoset_eq_of_mem_range _ hx).symm ▸ τ_apply_fromCoset _ -#align Group.surjective_of_epi_auxs.τ_apply_fromCoset' GroupCat.SurjectiveOfEpiAuxs.τ_apply_fromCoset' +#align Group.surjective_of_epi_auxs.τ_apply_fromCoset' Grp.SurjectiveOfEpiAuxs.τ_apply_fromCoset' theorem τ_symm_apply_fromCoset : Equiv.symm τ (fromCoset ⟨f.range, 1, one_leftCoset _⟩) = ∞ := by rw [tau, Equiv.symm_swap, Equiv.swap_apply_left] -#align Group.surjective_of_epi_auxs.τ_symm_apply_fromCoset GroupCat.SurjectiveOfEpiAuxs.τ_symm_apply_fromCoset +#align Group.surjective_of_epi_auxs.τ_symm_apply_fromCoset Grp.SurjectiveOfEpiAuxs.τ_symm_apply_fromCoset theorem τ_symm_apply_infinity : Equiv.symm τ ∞ = fromCoset ⟨f.range, 1, one_leftCoset _⟩ := by rw [tau, Equiv.symm_swap, Equiv.swap_apply_right] -#align Group.surjective_of_epi_auxs.τ_symm_apply_infinity GroupCat.SurjectiveOfEpiAuxs.τ_symm_apply_infinity +#align Group.surjective_of_epi_auxs.τ_symm_apply_infinity Grp.SurjectiveOfEpiAuxs.τ_symm_apply_infinity /-- Let `g : B ⟶ S(X')` be defined as such that, for any `β : B`, `g(β)` is the function sending point at infinity to point at infinity and sending coset `y` to `β • y`. @@ -220,7 +220,7 @@ def g : B →* SX' where map_mul' b1 b2 := by ext simp [mul_smul] -#align Group.surjective_of_epi_auxs.G GroupCat.SurjectiveOfEpiAuxs.g +#align Group.surjective_of_epi_auxs.G Grp.SurjectiveOfEpiAuxs.g local notation "g" => g f @@ -235,7 +235,7 @@ def h : B →* SX' where map_mul' b1 b2 := by ext simp -#align Group.surjective_of_epi_auxs.H GroupCat.SurjectiveOfEpiAuxs.h +#align Group.surjective_of_epi_auxs.H Grp.SurjectiveOfEpiAuxs.h local notation "h" => h f @@ -250,29 +250,29 @@ The strategy is the following: assuming `epi f` theorem g_apply_fromCoset (x : B) (y : Set.range (· • (f.range : Set B) : B → Set B)) : g x (fromCoset y) = fromCoset ⟨x • ↑y, by obtain ⟨z, hz⟩ := y.2; exact ⟨x * z, by simp [← hz, smul_smul]⟩⟩ := rfl -#align Group.surjective_of_epi_auxs.g_apply_fromCoset GroupCat.SurjectiveOfEpiAuxs.g_apply_fromCoset +#align Group.surjective_of_epi_auxs.g_apply_fromCoset Grp.SurjectiveOfEpiAuxs.g_apply_fromCoset theorem g_apply_infinity (x : B) : (g x) ∞ = ∞ := rfl -#align Group.surjective_of_epi_auxs.g_apply_infinity GroupCat.SurjectiveOfEpiAuxs.g_apply_infinity +#align Group.surjective_of_epi_auxs.g_apply_infinity Grp.SurjectiveOfEpiAuxs.g_apply_infinity theorem h_apply_infinity (x : B) (hx : x ∈ f.range) : (h x) ∞ = ∞ := by change ((τ).symm.trans (g x)).trans τ _ = _ simp only [MonoidHom.coe_mk, Equiv.toFun_as_coe, Equiv.coe_trans, Function.comp_apply] rw [τ_symm_apply_infinity, g_apply_fromCoset] simpa only using τ_apply_fromCoset' f x hx -#align Group.surjective_of_epi_auxs.h_apply_infinity GroupCat.SurjectiveOfEpiAuxs.h_apply_infinity +#align Group.surjective_of_epi_auxs.h_apply_infinity Grp.SurjectiveOfEpiAuxs.h_apply_infinity theorem h_apply_fromCoset (x : B) : (h x) (fromCoset ⟨f.range, 1, one_leftCoset _⟩) = fromCoset ⟨f.range, 1, one_leftCoset _⟩ := by change ((τ).symm.trans (g x)).trans τ _ = _ simp [-MonoidHom.coe_range, τ_symm_apply_fromCoset, g_apply_infinity, τ_apply_infinity] -#align Group.surjective_of_epi_auxs.h_apply_fromCoset GroupCat.SurjectiveOfEpiAuxs.h_apply_fromCoset +#align Group.surjective_of_epi_auxs.h_apply_fromCoset Grp.SurjectiveOfEpiAuxs.h_apply_fromCoset theorem h_apply_fromCoset' (x : B) (b : B) (hb : b ∈ f.range) : h x (fromCoset ⟨b • f.range, b, rfl⟩) = fromCoset ⟨b • ↑f.range, b, rfl⟩ := (fromCoset_eq_of_mem_range _ hb).symm ▸ h_apply_fromCoset f x -#align Group.surjective_of_epi_auxs.h_apply_fromCoset' GroupCat.SurjectiveOfEpiAuxs.h_apply_fromCoset' +#align Group.surjective_of_epi_auxs.h_apply_fromCoset' Grp.SurjectiveOfEpiAuxs.h_apply_fromCoset' theorem h_apply_fromCoset_nin_range (x : B) (hx : x ∈ f.range) (b : B) (hb : b ∉ f.range) : h x (fromCoset ⟨b • f.range, b, rfl⟩) = fromCoset ⟨(x * b) • ↑f.range, x * b, rfl⟩ := by @@ -285,7 +285,7 @@ theorem h_apply_fromCoset_nin_range (x : B) (hx : x ∈ f.range) (b : B) (hb : b refine Equiv.swap_apply_of_ne_of_ne (fromCoset_ne_of_nin_range _ fun r => hb ?_) (by simp) convert Subgroup.mul_mem _ (Subgroup.inv_mem _ hx) r rw [← mul_assoc, mul_left_inv, one_mul] -#align Group.surjective_of_epi_auxs.h_apply_fromCoset_nin_range GroupCat.SurjectiveOfEpiAuxs.h_apply_fromCoset_nin_range +#align Group.surjective_of_epi_auxs.h_apply_fromCoset_nin_range Grp.SurjectiveOfEpiAuxs.h_apply_fromCoset_nin_range theorem agree : f.range = { x | h x = g x } := by refine Set.ext fun b => ⟨?_, fun hb : h b = g b => by_contradiction fun r => ?_⟩ @@ -308,16 +308,16 @@ theorem agree : f.range = { x | h x = g x } := by have eq2 : g b (fromCoset ⟨f.range, 1, one_leftCoset _⟩) = fromCoset ⟨b • ↑f.range, b, rfl⟩ := rfl exact (fromCoset_ne_of_nin_range _ r).symm (by rw [← eq1, ← eq2, DFunLike.congr_fun hb]) -#align Group.surjective_of_epi_auxs.agree GroupCat.SurjectiveOfEpiAuxs.agree +#align Group.surjective_of_epi_auxs.agree Grp.SurjectiveOfEpiAuxs.agree -theorem comp_eq : (f ≫ show B ⟶ GroupCat.of SX' from g) = f ≫ show B ⟶ GroupCat.of SX' from h := by +theorem comp_eq : (f ≫ show B ⟶ Grp.of SX' from g) = f ≫ show B ⟶ Grp.of SX' from h := by ext a change g (f a) = h (f a) have : f a ∈ { b | h b = g b } := by rw [← agree] use a rw [this] -#align Group.surjective_of_epi_auxs.comp_eq GroupCat.SurjectiveOfEpiAuxs.comp_eq +#align Group.surjective_of_epi_auxs.comp_eq Grp.SurjectiveOfEpiAuxs.comp_eq theorem g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := by intro r @@ -329,7 +329,7 @@ theorem g_ne_h (x : B) (hx : x ∉ f.range) : g ≠ h := by Equiv.coe_trans, Function.comp_apply] at r erw [Equiv.swap_apply_left, g_apply_infinity, Equiv.swap_apply_right] at r exact fromCoset_ne_of_nin_range _ hx r -#align Group.surjective_of_epi_auxs.g_ne_h GroupCat.SurjectiveOfEpiAuxs.g_ne_h +#align Group.surjective_of_epi_auxs.g_ne_h Grp.SurjectiveOfEpiAuxs.g_ne_h end SurjectiveOfEpiAuxs @@ -341,125 +341,125 @@ theorem surjective_of_epi [Epi f] : Function.Surjective f := by exact SurjectiveOfEpiAuxs.g_ne_h f b (fun ⟨c, hc⟩ => hb _ hc) ((cancel_epi f).1 (SurjectiveOfEpiAuxs.comp_eq f)) -#align Group.surjective_of_epi GroupCat.surjective_of_epi +#align Group.surjective_of_epi Grp.surjective_of_epi theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := ⟨fun _ => surjective_of_epi f, ConcreteCategory.epi_of_surjective f⟩ -#align Group.epi_iff_surjective GroupCat.epi_iff_surjective +#align Group.epi_iff_surjective Grp.epi_iff_surjective theorem epi_iff_range_eq_top : Epi f ↔ f.range = ⊤ := Iff.trans (epi_iff_surjective _) (Subgroup.eq_top_iff' f.range).symm -#align Group.epi_iff_range_eq_top GroupCat.epi_iff_range_eq_top +#align Group.epi_iff_range_eq_top Grp.epi_iff_range_eq_top -end GroupCat +end Grp -namespace AddGroupCat +namespace AddGrp set_option linter.uppercaseLean3 false -variable {A B : AddGroupCat.{u}} (f : A ⟶ B) +variable {A B : AddGrp.{u}} (f : A ⟶ B) theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := by have i1 : Epi f ↔ Epi (groupAddGroupEquivalence.inverse.map f) := by refine ⟨?_, groupAddGroupEquivalence.inverse.epi_of_epi_map⟩ intro e' apply groupAddGroupEquivalence.inverse.map_epi - rwa [GroupCat.epi_iff_surjective] at i1 -#align AddGroup.epi_iff_surjective AddGroupCat.epi_iff_surjective + rwa [Grp.epi_iff_surjective] at i1 +#align AddGroup.epi_iff_surjective AddGrp.epi_iff_surjective theorem epi_iff_range_eq_top : Epi f ↔ f.range = ⊤ := Iff.trans (epi_iff_surjective _) (AddSubgroup.eq_top_iff' f.range).symm -#align AddGroup.epi_iff_range_eq_top AddGroupCat.epi_iff_range_eq_top +#align AddGroup.epi_iff_range_eq_top AddGrp.epi_iff_range_eq_top -end AddGroupCat +end AddGrp -namespace GroupCat +namespace Grp set_option linter.uppercaseLean3 false -variable {A B : GroupCat.{u}} (f : A ⟶ B) +variable {A B : Grp.{u}} (f : A ⟶ B) -@[to_additive AddGroupCat.forget_groupCat_preserves_mono] -instance forget_groupCat_preserves_mono : (forget GroupCat).PreservesMonomorphisms where +@[to_additive AddGrp.forget_grp_preserves_mono] +instance forget_grp_preserves_mono : (forget Grp).PreservesMonomorphisms where preserves f e := by rwa [mono_iff_injective, ← CategoryTheory.mono_iff_injective] at e -#align Group.forget_Group_preserves_mono GroupCat.forget_groupCat_preserves_mono -#align AddGroup.forget_Group_preserves_mono AddGroupCat.forget_groupCat_preserves_mono +#align Group.forget_Group_preserves_mono Grp.forget_grp_preserves_mono +#align AddGroup.forget_Group_preserves_mono AddGrp.forget_grp_preserves_mono -@[to_additive AddGroupCat.forget_groupCat_preserves_epi] -instance forget_groupCat_preserves_epi : (forget GroupCat).PreservesEpimorphisms where +@[to_additive AddGrp.forget_grp_preserves_epi] +instance forget_grp_preserves_epi : (forget Grp).PreservesEpimorphisms where preserves f e := by rwa [epi_iff_surjective, ← CategoryTheory.epi_iff_surjective] at e -#align Group.forget_Group_preserves_epi GroupCat.forget_groupCat_preserves_epi -#align AddGroup.forget_Group_preserves_epi AddGroupCat.forget_groupCat_preserves_epi +#align Group.forget_Group_preserves_epi Grp.forget_grp_preserves_epi +#align AddGroup.forget_Group_preserves_epi AddGrp.forget_grp_preserves_epi -end GroupCat +end Grp -namespace CommGroupCat +namespace CommGrp set_option linter.uppercaseLean3 false -variable {A B : CommGroupCat.{u}} (f : A ⟶ B) +variable {A B : CommGrp.{u}} (f : A ⟶ B) -- Porting note: again to help with non-transparency -private instance (A : CommGroupCat) : CommGroup A.α := A.str -private instance (A : CommGroupCat) : Group A.α := A.str.toGroup +private instance (A : CommGrp) : CommGroup A.α := A.str +private instance (A : CommGrp) : Group A.α := A.str.toGroup @[to_additive] theorem ker_eq_bot_of_mono [Mono f] : f.ker = ⊥ := MonoidHom.ker_eq_bot_of_cancel fun u _ => - (@cancel_mono _ _ _ _ _ f _ (show CommGroupCat.of f.ker ⟶ A from u) _).1 -#align CommGroup.ker_eq_bot_of_mono CommGroupCat.ker_eq_bot_of_mono -#align AddCommGroup.ker_eq_bot_of_mono AddCommGroupCat.ker_eq_bot_of_mono + (@cancel_mono _ _ _ _ _ f _ (show CommGrp.of f.ker ⟶ A from u) _).1 +#align CommGroup.ker_eq_bot_of_mono CommGrp.ker_eq_bot_of_mono +#align AddCommGroup.ker_eq_bot_of_mono AddCommGrp.ker_eq_bot_of_mono @[to_additive] theorem mono_iff_ker_eq_bot : Mono f ↔ f.ker = ⊥ := ⟨fun _ => ker_eq_bot_of_mono f, fun h => ConcreteCategory.mono_of_injective _ <| (MonoidHom.ker_eq_bot_iff f).1 h⟩ -#align CommGroup.mono_iff_ker_eq_bot CommGroupCat.mono_iff_ker_eq_bot -#align AddCommGroup.mono_iff_ker_eq_bot AddCommGroupCat.mono_iff_ker_eq_bot +#align CommGroup.mono_iff_ker_eq_bot CommGrp.mono_iff_ker_eq_bot +#align AddCommGroup.mono_iff_ker_eq_bot AddCommGrp.mono_iff_ker_eq_bot @[to_additive] theorem mono_iff_injective : Mono f ↔ Function.Injective f := Iff.trans (mono_iff_ker_eq_bot f) <| MonoidHom.ker_eq_bot_iff f -#align CommGroup.mono_iff_injective CommGroupCat.mono_iff_injective -#align AddCommGroup.mono_iff_injective AddCommGroupCat.mono_iff_injective +#align CommGroup.mono_iff_injective CommGrp.mono_iff_injective +#align AddCommGroup.mono_iff_injective AddCommGrp.mono_iff_injective @[to_additive] theorem range_eq_top_of_epi [Epi f] : f.range = ⊤ := MonoidHom.range_eq_top_of_cancel fun u v h => (@cancel_epi _ _ _ _ _ f _ (show B ⟶ ⟨B ⧸ MonoidHom.range f, inferInstance⟩ from u) v).1 h -#align CommGroup.range_eq_top_of_epi CommGroupCat.range_eq_top_of_epi -#align AddCommGroup.range_eq_top_of_epi AddCommGroupCat.range_eq_top_of_epi +#align CommGroup.range_eq_top_of_epi CommGrp.range_eq_top_of_epi +#align AddCommGroup.range_eq_top_of_epi AddCommGrp.range_eq_top_of_epi -- Porting note: again lack of transparency @[to_additive] -instance (G : CommGroupCat) : CommGroup <| (forget CommGroupCat).obj G := +instance (G : CommGrp) : CommGroup <| (forget CommGrp).obj G := G.str @[to_additive] theorem epi_iff_range_eq_top : Epi f ↔ f.range = ⊤ := ⟨fun _ => range_eq_top_of_epi _, fun hf => ConcreteCategory.epi_of_surjective _ <| MonoidHom.range_top_iff_surjective.mp hf⟩ -#align CommGroup.epi_iff_range_eq_top CommGroupCat.epi_iff_range_eq_top -#align AddCommGroup.epi_iff_range_eq_top AddCommGroupCat.epi_iff_range_eq_top +#align CommGroup.epi_iff_range_eq_top CommGrp.epi_iff_range_eq_top +#align AddCommGroup.epi_iff_range_eq_top AddCommGrp.epi_iff_range_eq_top @[to_additive] theorem epi_iff_surjective : Epi f ↔ Function.Surjective f := by rw [epi_iff_range_eq_top, MonoidHom.range_top_iff_surjective] -#align CommGroup.epi_iff_surjective CommGroupCat.epi_iff_surjective -#align AddCommGroup.epi_iff_surjective AddCommGroupCat.epi_iff_surjective +#align CommGroup.epi_iff_surjective CommGrp.epi_iff_surjective +#align AddCommGroup.epi_iff_surjective AddCommGrp.epi_iff_surjective -@[to_additive AddCommGroupCat.forget_commGroupCat_preserves_mono] -instance forget_commGroupCat_preserves_mono : (forget CommGroupCat).PreservesMonomorphisms where +@[to_additive AddCommGrp.forget_commGrp_preserves_mono] +instance forget_commGrp_preserves_mono : (forget CommGrp).PreservesMonomorphisms where preserves f e := by rwa [mono_iff_injective, ← CategoryTheory.mono_iff_injective] at e -#align CommGroup.forget_CommGroup_preserves_mono CommGroupCat.forget_commGroupCat_preserves_mono -#align AddCommGroup.forget_CommGroup_preserves_mono AddCommGroupCat.forget_commGroupCat_preserves_mono +#align CommGroup.forget_CommGroup_preserves_mono CommGrp.forget_commGrp_preserves_mono +#align AddCommGroup.forget_CommGroup_preserves_mono AddCommGrp.forget_commGrp_preserves_mono -@[to_additive AddCommGroupCat.forget_commGroupCat_preserves_epi] -instance forget_commGroupCat_preserves_epi : (forget CommGroupCat).PreservesEpimorphisms where +@[to_additive AddCommGrp.forget_commGrp_preserves_epi] +instance forget_commGrp_preserves_epi : (forget CommGrp).PreservesEpimorphisms where preserves f e := by rwa [epi_iff_surjective, ← CategoryTheory.epi_iff_surjective] at e -#align CommGroup.forget_CommGroup_preserves_epi CommGroupCat.forget_commGroupCat_preserves_epi -#align AddCommGroup.forget_CommGroup_preserves_epi AddCommGroupCat.forget_commGroupCat_preserves_epi +#align CommGroup.forget_CommGroup_preserves_epi CommGrp.forget_commGrp_preserves_epi +#align AddCommGroup.forget_CommGroup_preserves_epi AddCommGrp.forget_commGrp_preserves_epi -end CommGroupCat +end CommGrp end diff --git a/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean new file mode 100644 index 0000000000000..2e674438b286e --- /dev/null +++ b/Mathlib/Algebra/Category/Grp/EquivalenceGroupAddGroup.lean @@ -0,0 +1,94 @@ +/- +Copyright (c) 2022 Jujian Zhang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jujian Zhang +-/ +import Mathlib.Algebra.Category.Grp.Basic +import Mathlib.Algebra.Group.Equiv.TypeTags + +#align_import algebra.category.Group.equivalence_Group_AddGroup from "leanprover-community/mathlib"@"47b51515e69f59bca5cf34ef456e6000fe205a69" + +/-! +# Equivalence between `Group` and `AddGroup` + +This file contains two equivalences: +* `groupAddGroupEquivalence` : the equivalence between `Grp` and `AddGrp` by sending + `X : Grp` to `Additive X` and `Y : AddGrp` to `Multiplicative Y`. +* `commGroupAddCommGroupEquivalence` : the equivalence between `CommGrp` and `AddCommGrp` + by sending `X : CommGrp` to `Additive X` and `Y : AddCommGrp` to `Multiplicative Y`. +-/ + +-- Porting note: everything is upper case +set_option linter.uppercaseLean3 false + +open CategoryTheory + +namespace Grp + +-- Porting note: Lean cannot find these now +private instance (X : Grp) : MulOneClass X.α := X.str.toMulOneClass +private instance (X : CommGrp) : MulOneClass X.α := X.str.toMulOneClass +private instance (X : AddGrp) : AddZeroClass X.α := X.str.toAddZeroClass +private instance (X : AddCommGrp) : AddZeroClass X.α := X.str.toAddZeroClass + +/-- The functor `Group ⥤ AddGroup` by sending `X ↦ additive X` and `f ↦ f`. +-/ +@[simps] +def toAddGrp : Grp ⥤ AddGrp where + obj X := AddGrp.of (Additive X) + map {X} {Y} := MonoidHom.toAdditive +#align Group.to_AddGroup Grp.toAddGrp + +end Grp + +namespace CommGrp + +/-- The functor `CommGroup ⥤ AddCommGroup` by sending `X ↦ additive X` and `f ↦ f`. +-/ +@[simps] +def toAddCommGrp : CommGrp ⥤ AddCommGrp where + obj X := AddCommGrp.of (Additive X) + map {X} {Y} := MonoidHom.toAdditive +#align CommGroup.to_AddCommGroup CommGrp.toAddCommGrp + +end CommGrp + +namespace AddGrp + +/-- The functor `AddGroup ⥤ Group` by sending `X ↦ multiplicative Y` and `f ↦ f`. +-/ +@[simps] +def toGrp : AddGrp ⥤ Grp where + obj X := Grp.of (Multiplicative X) + map {X} {Y} := AddMonoidHom.toMultiplicative +#align AddGroup.to_Group AddGrp.toGrp + +end AddGrp + +namespace AddCommGrp + +/-- The functor `AddCommGroup ⥤ CommGroup` by sending `X ↦ multiplicative Y` and `f ↦ f`. +-/ +@[simps] +def toCommGrp : AddCommGrp ⥤ CommGrp where + obj X := CommGrp.of (Multiplicative X) + map {X} {Y} := AddMonoidHom.toMultiplicative +#align AddCommGroup.to_CommGroup AddCommGrp.toCommGrp + +end AddCommGrp + +/-- The equivalence of categories between `Group` and `AddGroup` +-/ +def groupAddGroupEquivalence : Grp ≌ AddGrp := + CategoryTheory.Equivalence.mk Grp.toAddGrp AddGrp.toGrp + (NatIso.ofComponents fun X => MulEquiv.toGrpIso (MulEquiv.multiplicativeAdditive X)) + (NatIso.ofComponents fun X => AddEquiv.toAddGrpIso (AddEquiv.additiveMultiplicative X)) +#align Group_AddGroup_equivalence groupAddGroupEquivalence + +/-- The equivalence of categories between `CommGroup` and `AddCommGroup`. +-/ +def commGroupAddCommGroupEquivalence : CommGrp ≌ AddCommGrp := + CategoryTheory.Equivalence.mk CommGrp.toAddCommGrp AddCommGrp.toCommGrp + (NatIso.ofComponents fun X => MulEquiv.toCommGrpIso (MulEquiv.multiplicativeAdditive X)) + (NatIso.ofComponents fun X => AddEquiv.toAddCommGrpIso (AddEquiv.additiveMultiplicative X)) +#align CommGroup_AddCommGroup_equivalence commGroupAddCommGroupEquivalence diff --git a/Mathlib/Algebra/Category/GroupCat/FilteredColimits.lean b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean similarity index 52% rename from Mathlib/Algebra/Category/GroupCat/FilteredColimits.lean rename to Mathlib/Algebra/Category/Grp/FilteredColimits.lean index 1c36cdc002329..f1bb076f925de 100644 --- a/Mathlib/Algebra/Category/GroupCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Grp/FilteredColimits.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Justus Springer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Justus Springer -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.Algebra.Category.MonCat.FilteredColimits #align_import algebra.category.Group.filtered_colimits from "leanprover-community/mathlib"@"c43486ecf2a5a17479a32ce09e4818924145e90e" @@ -14,12 +14,12 @@ import Mathlib.Algebra.Category.MonCat.FilteredColimits Forgetful functors from algebraic categories usually don't preserve colimits. However, they tend to preserve _filtered_ colimits. -In this file, we start with a small filtered category `J` and a functor `F : J ⥤ GroupCat`. -We show that the colimit of `F ⋙ forget₂ GroupCat MonCat` (in `MonCat`) carries the structure of a +In this file, we start with a small filtered category `J` and a functor `F : J ⥤ Grp`. +We show that the colimit of `F ⋙ forget₂ Grp MonCat` (in `MonCat`) carries the structure of a group, -thereby showing that the forgetful functor `forget₂ GroupCat MonCat` preserves filtered colimits. -In particular, this implies that `forget GroupCat` preserves filtered colimits. -Similarly for `AddGroupCat`, `CommGroupCat` and `AddCommGroupCat`. +thereby showing that the forgetful functor `forget₂ Grp MonCat` preserves filtered colimits. +In particular, this implies that `forget Grp` preserves filtered colimits. +Similarly for `AddGrp`, `CommGrp` and `AddCommGrp`. -/ @@ -37,7 +37,7 @@ open CategoryTheory.Limits open CategoryTheory.IsFiltered renaming max → max' -- avoid name collision with `_root_.max`. -namespace GroupCat.FilteredColimits +namespace Grp.FilteredColimits section @@ -45,52 +45,52 @@ open MonCat.FilteredColimits (colimit_one_eq colimit_mul_mk_eq) -- Mathlib3 used parameters here, mainly so we could have the abbreviations `G` and `G.mk` below, -- without passing around `F` all the time. -variable {J : Type v} [SmallCategory J] [IsFiltered J] (F : J ⥤ GroupCat.{max v u}) +variable {J : Type v} [SmallCategory J] [IsFiltered J] (F : J ⥤ Grp.{max v u}) -/-- The colimit of `F ⋙ forget₂ GroupCat MonCat` in the category `MonCat`. +/-- The colimit of `F ⋙ forget₂ Grp MonCat` in the category `MonCat`. In the following, we will show that this has the structure of a group. -/ @[to_additive - "The colimit of `F ⋙ forget₂ AddGroupCat AddMonCat` in the category `AddMonCat`. + "The colimit of `F ⋙ forget₂ AddGrp AddMonCat` in the category `AddMonCat`. In the following, we will show that this has the structure of an additive group."] noncomputable abbrev G : MonCat := - MonCat.FilteredColimits.colimit.{v, u} (F ⋙ forget₂ GroupCat MonCat.{max v u}) -#align Group.filtered_colimits.G GroupCat.FilteredColimits.G -#align AddGroup.filtered_colimits.G AddGroupCat.FilteredColimits.G + MonCat.FilteredColimits.colimit.{v, u} (F ⋙ forget₂ Grp MonCat.{max v u}) +#align Group.filtered_colimits.G Grp.FilteredColimits.G +#align AddGroup.filtered_colimits.G AddGrp.FilteredColimits.G /-- The canonical projection into the colimit, as a quotient type. -/ @[to_additive "The canonical projection into the colimit, as a quotient type."] abbrev G.mk : (Σ j, F.obj j) → G.{v, u} F := - Quot.mk (Types.Quot.Rel (F ⋙ forget GroupCat.{max v u})) -#align Group.filtered_colimits.G.mk GroupCat.FilteredColimits.G.mk -#align AddGroup.filtered_colimits.G.mk AddGroupCat.FilteredColimits.G.mk + Quot.mk (Types.Quot.Rel (F ⋙ forget Grp.{max v u})) +#align Group.filtered_colimits.G.mk Grp.FilteredColimits.G.mk +#align AddGroup.filtered_colimits.G.mk AddGrp.FilteredColimits.G.mk @[to_additive] theorem G.mk_eq (x y : Σ j, F.obj j) (h : ∃ (k : J) (f : x.1 ⟶ k) (g : y.1 ⟶ k), F.map f x.2 = F.map g y.2) : G.mk.{v, u} F x = G.mk F y := - Quot.EqvGen_sound (Types.FilteredColimit.eqvGen_quot_rel_of_rel (F ⋙ forget GroupCat) x y h) -#align Group.filtered_colimits.G.mk_eq GroupCat.FilteredColimits.G.mk_eq -#align AddGroup.filtered_colimits.G.mk_eq AddGroupCat.FilteredColimits.G.mk_eq + Quot.EqvGen_sound (Types.FilteredColimit.eqvGen_quot_rel_of_rel (F ⋙ forget Grp) x y h) +#align Group.filtered_colimits.G.mk_eq Grp.FilteredColimits.G.mk_eq +#align AddGroup.filtered_colimits.G.mk_eq AddGrp.FilteredColimits.G.mk_eq /-- The "unlifted" version of taking inverses in the colimit. -/ @[to_additive "The \"unlifted\" version of negation in the colimit."] def colimitInvAux (x : Σ j, F.obj j) : G.{v, u} F := G.mk F ⟨x.1, x.2⁻¹⟩ -#align Group.filtered_colimits.colimit_inv_aux GroupCat.FilteredColimits.colimitInvAux -#align AddGroup.filtered_colimits.colimit_neg_aux AddGroupCat.FilteredColimits.colimitNegAux +#align Group.filtered_colimits.colimit_inv_aux Grp.FilteredColimits.colimitInvAux +#align AddGroup.filtered_colimits.colimit_neg_aux AddGrp.FilteredColimits.colimitNegAux @[to_additive] theorem colimitInvAux_eq_of_rel (x y : Σ j, F.obj j) - (h : Types.FilteredColimit.Rel (F ⋙ forget GroupCat) x y) : + (h : Types.FilteredColimit.Rel (F ⋙ forget Grp) x y) : colimitInvAux.{v, u} F x = colimitInvAux F y := by apply G.mk_eq obtain ⟨k, f, g, hfg⟩ := h use k, f, g rw [MonoidHom.map_inv, MonoidHom.map_inv, inv_inj] exact hfg -#align Group.filtered_colimits.colimit_inv_aux_eq_of_rel GroupCat.FilteredColimits.colimitInvAux_eq_of_rel -#align AddGroup.filtered_colimits.colimit_neg_aux_eq_of_rel AddGroupCat.FilteredColimits.colimitNegAux_eq_of_rel +#align Group.filtered_colimits.colimit_inv_aux_eq_of_rel Grp.FilteredColimits.colimitInvAux_eq_of_rel +#align AddGroup.filtered_colimits.colimit_neg_aux_eq_of_rel AddGrp.FilteredColimits.colimitNegAux_eq_of_rel /-- Taking inverses in the colimit. See also `colimitInvAux`. -/ @[to_additive "Negation in the colimit. See also `colimitNegAux`."] @@ -101,14 +101,14 @@ instance colimitInv : Inv (G.{v, u} F) where apply colimitInvAux_eq_of_rel apply Types.FilteredColimit.rel_of_quot_rel exact h -#align Group.filtered_colimits.colimit_has_inv GroupCat.FilteredColimits.colimitInv -#align AddGroup.filtered_colimits.colimit_has_neg AddGroupCat.FilteredColimits.colimitNeg +#align Group.filtered_colimits.colimit_has_inv Grp.FilteredColimits.colimitInv +#align AddGroup.filtered_colimits.colimit_has_neg AddGrp.FilteredColimits.colimitNeg @[to_additive (attr := simp)] theorem colimit_inv_mk_eq (x : Σ j, F.obj j) : (G.mk.{v, u} F x)⁻¹ = G.mk F ⟨x.1, x.2⁻¹⟩ := rfl -#align Group.filtered_colimits.colimit_inv_mk_eq GroupCat.FilteredColimits.colimit_inv_mk_eq -#align AddGroup.filtered_colimits.colimit_neg_mk_eq AddGroupCat.FilteredColimits.colimit_neg_mk_eq +#align Group.filtered_colimits.colimit_inv_mk_eq Grp.FilteredColimits.colimit_inv_mk_eq +#align AddGroup.filtered_colimits.colimit_neg_mk_eq AddGrp.FilteredColimits.colimit_neg_mk_eq @[to_additive] noncomputable instance colimitGroup : Group (G.{v, u} F) := @@ -117,148 +117,148 @@ noncomputable instance colimitGroup : Group (G.{v, u} F) := refine Quot.inductionOn x ?_; clear x; intro x cases' x with j x erw [colimit_inv_mk_eq, - colimit_mul_mk_eq (F ⋙ forget₂ GroupCat MonCat.{max v u}) ⟨j, _⟩ ⟨j, _⟩ j (𝟙 j) (𝟙 j), - colimit_one_eq (F ⋙ forget₂ GroupCat MonCat.{max v u}) j] + colimit_mul_mk_eq (F ⋙ forget₂ Grp MonCat.{max v u}) ⟨j, _⟩ ⟨j, _⟩ j (𝟙 j) (𝟙 j), + colimit_one_eq (F ⋙ forget₂ Grp MonCat.{max v u}) j] dsimp erw [CategoryTheory.Functor.map_id, mul_left_inv] } -#align Group.filtered_colimits.colimit_group GroupCat.FilteredColimits.colimitGroup -#align AddGroup.filtered_colimits.colimit_add_group AddGroupCat.FilteredColimits.colimitAddGroup +#align Group.filtered_colimits.colimit_group Grp.FilteredColimits.colimitGroup +#align AddGroup.filtered_colimits.colimit_add_group AddGrp.FilteredColimits.colimitAddGroup /-- The bundled group giving the filtered colimit of a diagram. -/ @[to_additive "The bundled additive group giving the filtered colimit of a diagram."] -noncomputable def colimit : GroupCat.{max v u} := - GroupCat.of (G.{v, u} F) -#align Group.filtered_colimits.colimit GroupCat.FilteredColimits.colimit -#align AddGroup.filtered_colimits.colimit AddGroupCat.FilteredColimits.colimit +noncomputable def colimit : Grp.{max v u} := + Grp.of (G.{v, u} F) +#align Group.filtered_colimits.colimit Grp.FilteredColimits.colimit +#align AddGroup.filtered_colimits.colimit AddGrp.FilteredColimits.colimit /-- The cocone over the proposed colimit group. -/ @[to_additive "The cocone over the proposed colimit additive group."] noncomputable def colimitCocone : Cocone F where pt := colimit.{v, u} F - ι := { (MonCat.FilteredColimits.colimitCocone (F ⋙ forget₂ GroupCat MonCat.{max v u})).ι with } -#align Group.filtered_colimits.colimit_cocone GroupCat.FilteredColimits.colimitCocone -#align AddGroup.filtered_colimits.colimit_cocone AddGroupCat.FilteredColimits.colimitCocone + ι := { (MonCat.FilteredColimits.colimitCocone (F ⋙ forget₂ Grp MonCat.{max v u})).ι with } +#align Group.filtered_colimits.colimit_cocone Grp.FilteredColimits.colimitCocone +#align AddGroup.filtered_colimits.colimit_cocone AddGrp.FilteredColimits.colimitCocone -/-- The proposed colimit cocone is a colimit in `GroupCat`. -/ +/-- The proposed colimit cocone is a colimit in `Grp`. -/ @[to_additive "The proposed colimit cocone is a colimit in `AddGroup`."] def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where desc t := - MonCat.FilteredColimits.colimitDesc.{v, u} (F ⋙ forget₂ GroupCat MonCat.{max v u}) - ((forget₂ GroupCat MonCat).mapCocone t) + MonCat.FilteredColimits.colimitDesc.{v, u} (F ⋙ forget₂ Grp MonCat.{max v u}) + ((forget₂ Grp MonCat).mapCocone t) fac t j := DFunLike.coe_injective <| - (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget GroupCat)).fac - ((forget GroupCat).mapCocone t) j + (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget Grp)).fac + ((forget Grp).mapCocone t) j uniq t _ h := DFunLike.coe_injective' <| - (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget GroupCat)).uniq - ((forget GroupCat).mapCocone t) _ + (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget Grp)).uniq + ((forget Grp).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x -#align Group.filtered_colimits.colimit_cocone_is_colimit GroupCat.FilteredColimits.colimitCoconeIsColimit -#align AddGroup.filtered_colimits.colimit_cocone_is_colimit AddGroupCat.FilteredColimits.colimitCoconeIsColimit +#align Group.filtered_colimits.colimit_cocone_is_colimit Grp.FilteredColimits.colimitCoconeIsColimit +#align AddGroup.filtered_colimits.colimit_cocone_is_colimit AddGrp.FilteredColimits.colimitCoconeIsColimit @[to_additive forget₂AddMonPreservesFilteredColimits] noncomputable instance forget₂MonPreservesFilteredColimits : - PreservesFilteredColimits.{u} (forget₂ GroupCat.{u} MonCat.{u}) where + PreservesFilteredColimits.{u} (forget₂ Grp.{u} MonCat.{u}) where preserves_filtered_colimits x hx1 _ := letI : Category.{u, u} x := hx1 ⟨fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) (MonCat.FilteredColimits.colimitCoconeIsColimit.{u, u} _)⟩ -#align Group.filtered_colimits.forget₂_Mon_preserves_filtered_colimits GroupCat.FilteredColimits.forget₂MonPreservesFilteredColimits -#align AddGroup.filtered_colimits.forget₂_AddMon_preserves_filtered_colimits AddGroupCat.FilteredColimits.forget₂AddMonPreservesFilteredColimits +#align Group.filtered_colimits.forget₂_Mon_preserves_filtered_colimits Grp.FilteredColimits.forget₂MonPreservesFilteredColimits +#align AddGroup.filtered_colimits.forget₂_AddMon_preserves_filtered_colimits AddGrp.FilteredColimits.forget₂AddMonPreservesFilteredColimits @[to_additive] noncomputable instance forgetPreservesFilteredColimits : - PreservesFilteredColimits (forget GroupCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ GroupCat MonCat) (forget MonCat.{u}) -#align Group.filtered_colimits.forget_preserves_filtered_colimits GroupCat.FilteredColimits.forgetPreservesFilteredColimits -#align AddGroup.filtered_colimits.forget_preserves_filtered_colimits AddGroupCat.FilteredColimits.forgetPreservesFilteredColimits + PreservesFilteredColimits (forget Grp.{u}) := + Limits.compPreservesFilteredColimits (forget₂ Grp MonCat) (forget MonCat.{u}) +#align Group.filtered_colimits.forget_preserves_filtered_colimits Grp.FilteredColimits.forgetPreservesFilteredColimits +#align AddGroup.filtered_colimits.forget_preserves_filtered_colimits AddGrp.FilteredColimits.forgetPreservesFilteredColimits end -end GroupCat.FilteredColimits +end Grp.FilteredColimits -namespace CommGroupCat.FilteredColimits +namespace CommGrp.FilteredColimits section -- We use parameters here, mainly so we can have the abbreviation `G` below, without -- passing around `F` all the time. -variable {J : Type v} [SmallCategory J] [IsFiltered J] (F : J ⥤ CommGroupCat.{max v u}) +variable {J : Type v} [SmallCategory J] [IsFiltered J] (F : J ⥤ CommGrp.{max v u}) -/-- The colimit of `F ⋙ forget₂ CommGroupCat GroupCat` in the category `GroupCat`. +/-- The colimit of `F ⋙ forget₂ CommGrp Grp` in the category `Grp`. In the following, we will show that this has the structure of a _commutative_ group. -/ @[to_additive - "The colimit of `F ⋙ forget₂ AddCommGroupCat AddGroupCat` in the category `AddGroupCat`. + "The colimit of `F ⋙ forget₂ AddCommGrp AddGrp` in the category `AddGrp`. In the following, we will show that this has the structure of a _commutative_ additive group."] -noncomputable abbrev G : GroupCat.{max v u} := - GroupCat.FilteredColimits.colimit.{v, u} (F ⋙ forget₂ CommGroupCat.{max v u} GroupCat.{max v u}) -#align CommGroup.filtered_colimits.G CommGroupCat.FilteredColimits.G -#align AddCommGroup.filtered_colimits.G AddCommGroupCat.FilteredColimits.G +noncomputable abbrev G : Grp.{max v u} := + Grp.FilteredColimits.colimit.{v, u} (F ⋙ forget₂ CommGrp.{max v u} Grp.{max v u}) +#align CommGroup.filtered_colimits.G CommGrp.FilteredColimits.G +#align AddCommGroup.filtered_colimits.G AddCommGrp.FilteredColimits.G @[to_additive] noncomputable instance colimitCommGroup : CommGroup.{max v u} (G.{v, u} F) := { (G F).str, CommMonCat.FilteredColimits.colimitCommMonoid - (F ⋙ forget₂ CommGroupCat CommMonCat.{max v u}) with } -#align CommGroup.filtered_colimits.colimit_comm_group CommGroupCat.FilteredColimits.colimitCommGroup -#align AddCommGroup.filtered_colimits.colimit_add_comm_group AddCommGroupCat.FilteredColimits.colimitAddCommGroup + (F ⋙ forget₂ CommGrp CommMonCat.{max v u}) with } +#align CommGroup.filtered_colimits.colimit_comm_group CommGrp.FilteredColimits.colimitCommGroup +#align AddCommGroup.filtered_colimits.colimit_add_comm_group AddCommGrp.FilteredColimits.colimitAddCommGroup /-- The bundled commutative group giving the filtered colimit of a diagram. -/ @[to_additive "The bundled additive commutative group giving the filtered colimit of a diagram."] -noncomputable def colimit : CommGroupCat := - CommGroupCat.of (G.{v, u} F) -#align CommGroup.filtered_colimits.colimit CommGroupCat.FilteredColimits.colimit -#align AddCommGroup.filtered_colimits.colimit AddCommGroupCat.FilteredColimits.colimit +noncomputable def colimit : CommGrp := + CommGrp.of (G.{v, u} F) +#align CommGroup.filtered_colimits.colimit CommGrp.FilteredColimits.colimit +#align AddCommGroup.filtered_colimits.colimit AddCommGrp.FilteredColimits.colimit /-- The cocone over the proposed colimit commutative group. -/ @[to_additive "The cocone over the proposed colimit additive commutative group."] noncomputable def colimitCocone : Cocone F where pt := colimit.{v, u} F ι := - { (GroupCat.FilteredColimits.colimitCocone - (F ⋙ forget₂ CommGroupCat GroupCat.{max v u})).ι with } -#align CommGroup.filtered_colimits.colimit_cocone CommGroupCat.FilteredColimits.colimitCocone -#align AddCommGroup.filtered_colimits.colimit_cocone AddCommGroupCat.FilteredColimits.colimitCocone + { (Grp.FilteredColimits.colimitCocone + (F ⋙ forget₂ CommGrp Grp.{max v u})).ι with } +#align CommGroup.filtered_colimits.colimit_cocone CommGrp.FilteredColimits.colimitCocone +#align AddCommGroup.filtered_colimits.colimit_cocone AddCommGrp.FilteredColimits.colimitCocone -/-- The proposed colimit cocone is a colimit in `CommGroupCat`. -/ +/-- The proposed colimit cocone is a colimit in `CommGrp`. -/ @[to_additive "The proposed colimit cocone is a colimit in `AddCommGroup`."] def colimitCoconeIsColimit : IsColimit (colimitCocone.{v, u} F) where desc t := - (GroupCat.FilteredColimits.colimitCoconeIsColimit.{v, u} - (F ⋙ forget₂ CommGroupCat GroupCat.{max v u})).desc - ((forget₂ CommGroupCat GroupCat.{max v u}).mapCocone t) + (Grp.FilteredColimits.colimitCoconeIsColimit.{v, u} + (F ⋙ forget₂ CommGrp Grp.{max v u})).desc + ((forget₂ CommGrp Grp.{max v u}).mapCocone t) fac t j := DFunLike.coe_injective <| - (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget CommGroupCat)).fac - ((forget CommGroupCat).mapCocone t) j + (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget CommGrp)).fac + ((forget CommGrp).mapCocone t) j uniq t _ h := DFunLike.coe_injective <| - (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget CommGroupCat)).uniq - ((forget CommGroupCat).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x -#align CommGroup.filtered_colimits.colimit_cocone_is_colimit CommGroupCat.FilteredColimits.colimitCoconeIsColimit -#align AddCommGroup.filtered_colimits.colimit_cocone_is_colimit AddCommGroupCat.FilteredColimits.colimitCoconeIsColimit + (Types.TypeMax.colimitCoconeIsColimit.{v, u} (F ⋙ forget CommGrp)).uniq + ((forget CommGrp).mapCocone t) _ fun j => funext fun x => DFunLike.congr_fun (h j) x +#align CommGroup.filtered_colimits.colimit_cocone_is_colimit CommGrp.FilteredColimits.colimitCoconeIsColimit +#align AddCommGroup.filtered_colimits.colimit_cocone_is_colimit AddCommGrp.FilteredColimits.colimitCoconeIsColimit @[to_additive] noncomputable instance forget₂GroupPreservesFilteredColimits : - PreservesFilteredColimits (forget₂ CommGroupCat GroupCat.{u}) where + PreservesFilteredColimits (forget₂ CommGrp Grp.{u}) where preserves_filtered_colimits J hJ1 _ := letI : Category J := hJ1 { preservesColimit := fun {F} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit.{u, u} F) - (GroupCat.FilteredColimits.colimitCoconeIsColimit.{u, u} - (F ⋙ forget₂ CommGroupCat GroupCat.{u})) } -#align CommGroup.filtered_colimits.forget₂_Group_preserves_filtered_colimits CommGroupCat.FilteredColimits.forget₂GroupPreservesFilteredColimits -#align AddCommGroup.filtered_colimits.forget₂_AddGroup_preserves_filtered_colimits AddCommGroupCat.FilteredColimits.forget₂AddGroupPreservesFilteredColimits + (Grp.FilteredColimits.colimitCoconeIsColimit.{u, u} + (F ⋙ forget₂ CommGrp Grp.{u})) } +#align CommGroup.filtered_colimits.forget₂_Group_preserves_filtered_colimits CommGrp.FilteredColimits.forget₂GroupPreservesFilteredColimits +#align AddCommGroup.filtered_colimits.forget₂_AddGroup_preserves_filtered_colimits AddCommGrp.FilteredColimits.forget₂AddGroupPreservesFilteredColimits @[to_additive] noncomputable instance forgetPreservesFilteredColimits : - PreservesFilteredColimits (forget CommGroupCat.{u}) := - Limits.compPreservesFilteredColimits (forget₂ CommGroupCat GroupCat) (forget GroupCat.{u}) -#align CommGroup.filtered_colimits.forget_preserves_filtered_colimits CommGroupCat.FilteredColimits.forgetPreservesFilteredColimits -#align AddCommGroup.filtered_colimits.forget_preserves_filtered_colimits AddCommGroupCat.FilteredColimits.forgetPreservesFilteredColimits + PreservesFilteredColimits (forget CommGrp.{u}) := + Limits.compPreservesFilteredColimits (forget₂ CommGrp Grp) (forget Grp.{u}) +#align CommGroup.filtered_colimits.forget_preserves_filtered_colimits CommGrp.FilteredColimits.forgetPreservesFilteredColimits +#align AddCommGroup.filtered_colimits.forget_preserves_filtered_colimits AddCommGrp.FilteredColimits.forgetPreservesFilteredColimits end -end CommGroupCat.FilteredColimits +end CommGrp.FilteredColimits diff --git a/Mathlib/Algebra/Category/GroupCat/ForgetCorepresentable.lean b/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean similarity index 65% rename from Mathlib/Algebra/Category/GroupCat/ForgetCorepresentable.lean rename to Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean index fd14041e4b112..7f0cd7d6755d1 100644 --- a/Mathlib/Algebra/Category/GroupCat/ForgetCorepresentable.lean +++ b/Mathlib/Algebra/Category/Grp/ForgetCorepresentable.lean @@ -3,16 +3,16 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.Algebra.Group.ULift import Mathlib.CategoryTheory.Yoneda /-! # The forget functor is corepresentable -It is shown that the forget functor `AddCommGroupCat.{u} ⥤ Type u` is corepresentable -by `ULift ℤ`. Similar results are obtained for the variants `CommGroupCat`, `AddGroupCat` -and `GroupCat`. +It is shown that the forget functor `AddCommGrp.{u} ⥤ Type u` is corepresentable +by `ULift ℤ`. Similar results are obtained for the variants `CommGrp`, `AddGrp` +and `Grp`. -/ @@ -75,38 +75,38 @@ def fromULiftIntEquiv (α : Type u) [AddGroup α] : (ULift.{u} ℤ →+ α) ≃ end AddMonoidHom -/-- The forget functor `GroupCat.{u} ⥤ Type u` is corepresentable. -/ -def GroupCat.coyonedaObjIsoForget : - coyoneda.obj (op (of (ULift.{u} (Multiplicative ℤ)))) ≅ forget GroupCat.{u} := +/-- The forget functor `Grp.{u} ⥤ Type u` is corepresentable. -/ +def Grp.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} (Multiplicative ℤ)))) ≅ forget Grp.{u} := (NatIso.ofComponents (fun M => (MonoidHom.fromULiftMultiplicativeIntEquiv M.α).toIso)) -/-- The forget functor `CommGroupCat.{u} ⥤ Type u` is corepresentable. -/ -def CommGroupCat.coyonedaObjIsoForget : - coyoneda.obj (op (of (ULift.{u} (Multiplicative ℤ)))) ≅ forget CommGroupCat.{u} := +/-- The forget functor `CommGrp.{u} ⥤ Type u` is corepresentable. -/ +def CommGrp.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} (Multiplicative ℤ)))) ≅ forget CommGrp.{u} := (NatIso.ofComponents (fun M => (MonoidHom.fromULiftMultiplicativeIntEquiv M.α).toIso)) -/-- The forget functor `AddGroupCat.{u} ⥤ Type u` is corepresentable. -/ -def AddGroupCat.coyonedaObjIsoForget : - coyoneda.obj (op (of (ULift.{u} ℤ))) ≅ forget AddGroupCat.{u} := +/-- The forget functor `AddGrp.{u} ⥤ Type u` is corepresentable. -/ +def AddGrp.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} ℤ))) ≅ forget AddGrp.{u} := (NatIso.ofComponents (fun M => (AddMonoidHom.fromULiftIntEquiv M.α).toIso)) -/-- The forget functor `AddCommGroupCat.{u} ⥤ Type u` is corepresentable. -/ -def AddCommGroupCat.coyonedaObjIsoForget : - coyoneda.obj (op (of (ULift.{u} ℤ))) ≅ forget AddCommGroupCat.{u} := +/-- The forget functor `AddCommGrp.{u} ⥤ Type u` is corepresentable. -/ +def AddCommGrp.coyonedaObjIsoForget : + coyoneda.obj (op (of (ULift.{u} ℤ))) ≅ forget AddCommGrp.{u} := (NatIso.ofComponents (fun M => (AddMonoidHom.fromULiftIntEquiv M.α).toIso)) -instance GroupCat.forget_corepresentable : - (forget GroupCat.{u}).Corepresentable where - has_corepresentation := ⟨_, ⟨GroupCat.coyonedaObjIsoForget⟩⟩ +instance Grp.forget_corepresentable : + (forget Grp.{u}).Corepresentable where + has_corepresentation := ⟨_, ⟨Grp.coyonedaObjIsoForget⟩⟩ -instance CommGroupCat.forget_corepresentable : - (forget CommGroupCat.{u}).Corepresentable where - has_corepresentation := ⟨_, ⟨CommGroupCat.coyonedaObjIsoForget⟩⟩ +instance CommGrp.forget_corepresentable : + (forget CommGrp.{u}).Corepresentable where + has_corepresentation := ⟨_, ⟨CommGrp.coyonedaObjIsoForget⟩⟩ -instance AddGroupCat.forget_corepresentable : - (forget AddGroupCat.{u}).Corepresentable where - has_corepresentation := ⟨_, ⟨AddGroupCat.coyonedaObjIsoForget⟩⟩ +instance AddGrp.forget_corepresentable : + (forget AddGrp.{u}).Corepresentable where + has_corepresentation := ⟨_, ⟨AddGrp.coyonedaObjIsoForget⟩⟩ -instance AddCommGroupCat.forget_corepresentable : - (forget AddCommGroupCat.{u}).Corepresentable where - has_corepresentation := ⟨_, ⟨AddCommGroupCat.coyonedaObjIsoForget⟩⟩ +instance AddCommGrp.forget_corepresentable : + (forget AddCommGrp.{u}).Corepresentable where + has_corepresentation := ⟨_, ⟨AddCommGrp.coyonedaObjIsoForget⟩⟩ diff --git a/Mathlib/Algebra/Category/GroupCat/Images.lean b/Mathlib/Algebra/Category/Grp/Images.lean similarity index 65% rename from Mathlib/Algebra/Category/GroupCat/Images.lean rename to Mathlib/Algebra/Category/Grp/Images.lean index 175f7b95fab08..6c77845b7fd6d 100644 --- a/Mathlib/Algebra/Category/GroupCat/Images.lean +++ b/Mathlib/Algebra/Category/Grp/Images.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Algebra.Category.GroupCat.Abelian +import Mathlib.Algebra.Category.Grp.Abelian import Mathlib.CategoryTheory.Limits.Shapes.Images #align_import algebra.category.Group.images from "leanprover-community/mathlib"@"70fd9563a21e7b963887c9360bd29b2393e6225a" @@ -12,7 +12,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Images # The category of commutative additive groups has images. Note that we don't need to register any of the constructions here as instances, because we get them -from the fact that `AddCommGroupCat` is an abelian category. +from the fact that `AddCommGrp` is an abelian category. -/ @@ -22,28 +22,28 @@ open CategoryTheory.Limits universe u -namespace AddCommGroupCat +namespace AddCommGrp set_option linter.uppercaseLean3 false -- Note that because `injective_of_mono` is currently only proved in `Type 0`, -- we restrict to the lowest universe here for now. -variable {G H : AddCommGroupCat.{0}} (f : G ⟶ H) +variable {G H : AddCommGrp.{0}} (f : G ⟶ H) attribute [local ext] Subtype.ext_val section --- implementation details of `IsImage` for `AddCommGroupCat`; use the API, not these -/-- the image of a morphism in `AddCommGroupCat` is just the bundling of `AddMonoidHom.range f` -/ -def image : AddCommGroupCat := - AddCommGroupCat.of (AddMonoidHom.range f) -#align AddCommGroup.image AddCommGroupCat.image +-- implementation details of `IsImage` for `AddCommGrp`; use the API, not these +/-- the image of a morphism in `AddCommGrp` is just the bundling of `AddMonoidHom.range f` -/ +def image : AddCommGrp := + AddCommGrp.of (AddMonoidHom.range f) +#align AddCommGroup.image AddCommGrp.image /-- the inclusion of `image f` into the target -/ def image.ι : image f ⟶ H := f.range.subtype -#align AddCommGroup.image.ι AddCommGroupCat.image.ι +#align AddCommGroup.image.ι AddCommGrp.image.ι instance : Mono (image.ι f) := ConcreteCategory.mono_of_injective (image.ι f) Subtype.val_injective @@ -51,12 +51,12 @@ instance : Mono (image.ι f) := /-- the corestriction map to the image -/ def factorThruImage : G ⟶ image f := f.rangeRestrict -#align AddCommGroup.factor_thru_image AddCommGroupCat.factorThruImage +#align AddCommGroup.factor_thru_image AddCommGrp.factorThruImage theorem image.fac : factorThruImage f ≫ image.ι f = f := by ext rfl -#align AddCommGroup.image.fac AddCommGroupCat.image.fac +#align AddCommGroup.image.fac AddCommGrp.image.fac attribute [local simp] image.fac @@ -82,37 +82,37 @@ noncomputable def image.lift (F' : MonoFactorisation f) : image f ⟶ F'.I where rw [(Classical.indefiniteDescription (fun z => f z = _) _).2] rw [(Classical.indefiniteDescription (fun z => f z = _) _).2] rfl -#align AddCommGroup.image.lift AddCommGroupCat.image.lift +#align AddCommGroup.image.lift AddCommGrp.image.lift theorem image.lift_fac (F' : MonoFactorisation f) : image.lift F' ≫ F'.m = image.ι f := by ext x change (F'.e ≫ F'.m) _ = _ rw [F'.fac, (Classical.indefiniteDescription _ x.2).2] rfl -#align AddCommGroup.image.lift_fac AddCommGroupCat.image.lift_fac +#align AddCommGroup.image.lift_fac AddCommGrp.image.lift_fac end -/-- the factorisation of any morphism in `AddCommGroupCat` through a mono. -/ +/-- the factorisation of any morphism in `AddCommGrp` through a mono. -/ def monoFactorisation : MonoFactorisation f where I := image f m := image.ι f e := factorThruImage f -#align AddCommGroup.mono_factorisation AddCommGroupCat.monoFactorisation +#align AddCommGroup.mono_factorisation AddCommGrp.monoFactorisation -/-- the factorisation of any morphism in `AddCommGroupCat` through a mono has +/-- the factorisation of any morphism in `AddCommGrp` through a mono has the universal property of the image. -/ noncomputable def isImage : IsImage (monoFactorisation f) where lift := image.lift lift_fac := image.lift_fac -#align AddCommGroup.is_image AddCommGroupCat.isImage +#align AddCommGroup.is_image AddCommGrp.isImage -/-- The categorical image of a morphism in `AddCommGroupCat` +/-- The categorical image of a morphism in `AddCommGrp` agrees with the usual group-theoretical range. -/ -noncomputable def imageIsoRange {G H : AddCommGroupCat.{0}} (f : G ⟶ H) : - Limits.image f ≅ AddCommGroupCat.of f.range := +noncomputable def imageIsoRange {G H : AddCommGrp.{0}} (f : G ⟶ H) : + Limits.image f ≅ AddCommGrp.of f.range := IsImage.isoExt (Image.isImage f) (isImage f) -#align AddCommGroup.image_iso_range AddCommGroupCat.imageIsoRange +#align AddCommGroup.image_iso_range AddCommGrp.imageIsoRange -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Injective.lean b/Mathlib/Algebra/Category/Grp/Injective.lean similarity index 76% rename from Mathlib/Algebra/Category/GroupCat/Injective.lean rename to Mathlib/Algebra/Category/Grp/Injective.lean index 0e381e2c762a3..95dc807ecd16f 100644 --- a/Mathlib/Algebra/Category/GroupCat/Injective.lean +++ b/Mathlib/Algebra/Category/Grp/Injective.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ -import Mathlib.Algebra.Category.GroupCat.ZModuleEquivalence +import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.Module.Injective import Mathlib.RingTheory.PrincipalIdealDomain import Mathlib.Topology.Instances.AddCircle @@ -19,10 +19,10 @@ groups and that the category of abelian groups has enough injective objects. ## Main results -- `AddCommGroupCat.injective_of_divisible` : a divisible group is also an injective object. -- `AddCommGroupCat.enoughInjectives` : the category of abelian groups (written additively) has +- `AddCommGrp.injective_of_divisible` : a divisible group is also an injective object. +- `AddCommGrp.enoughInjectives` : the category of abelian groups (written additively) has enough injectives. -- `CommGroupCat.enoughInjectives` : the category of abelian group (written multiplicatively) has +- `CommGrp.enoughInjectives` : the category of abelian group (written multiplicatively) has enough injectives. ## Implementation notes @@ -57,21 +57,21 @@ theorem Module.Baer.of_divisible [DivisibleBy A ℤ] : Module.Baer ℤ A := fun rw [map_zsmul, LinearMap.toSpanSingleton_apply, DivisibleBy.div_cancel gₘ h0, ← map_zsmul g, SetLike.mk_smul_mk] -namespace AddCommGroupCat +namespace AddCommGrp theorem injective_as_module_iff : Injective (⟨A⟩ : ModuleCat ℤ) ↔ - Injective (⟨A,inferInstance⟩ : AddCommGroupCat) := - ((forget₂ (ModuleCat ℤ) AddCommGroupCat).asEquivalence.map_injective_iff ⟨A⟩).symm + Injective (⟨A,inferInstance⟩ : AddCommGrp) := + ((forget₂ (ModuleCat ℤ) AddCommGrp).asEquivalence.map_injective_iff ⟨A⟩).symm #noalign AddCommGroup.injective_of_injective_as_module #noalign AddCommGroup.injective_as_module_of_injective_as_Ab instance injective_of_divisible [DivisibleBy A ℤ] : - Injective (⟨A,inferInstance⟩ : AddCommGroupCat) := + Injective (⟨A,inferInstance⟩ : AddCommGrp) := (injective_as_module_iff A).mp <| Module.injective_object_of_injective_module (inj := (Module.Baer.of_divisible A).injective) -#align AddCommGroup.injective_of_divisible AddCommGroupCat.injective_of_divisible +#align AddCommGroup.injective_of_divisible AddCommGrp.injective_of_divisible instance injective_ratCircle : Injective <| of <| ULift.{u} <| AddCircle (1 : ℚ) := injective_of_divisible _ -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Kernels.lean b/Mathlib/Algebra/Category/Grp/Kernels.lean similarity index 89% rename from Mathlib/Algebra/Category/GroupCat/Kernels.lean rename to Mathlib/Algebra/Category/Grp/Kernels.lean index 425ccccaa958d..818c11a8823e6 100644 --- a/Mathlib/Algebra/Category/GroupCat/Kernels.lean +++ b/Mathlib/Algebra/Category/Grp/Kernels.lean @@ -3,21 +3,21 @@ Copyright (c) 2023 Moritz Firsching. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: David Kurniadi Angdinata, Moritz Firsching, Nikolas Kuhn -/ -import Mathlib.Algebra.Category.GroupCat.EpiMono -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.EpiMono +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.CategoryTheory.Limits.Shapes.Kernels /-! # The concrete (co)kernels in the category of abelian groups are categorical (co)kernels. -/ -namespace AddCommGroupCat +namespace AddCommGrp open AddMonoidHom CategoryTheory Limits QuotientAddGroup universe u -variable {G H : AddCommGroupCat.{u}} (f : G ⟶ H) +variable {G H : AddCommGrp.{u}} (f : G ⟶ H) /-- The kernel cone induced by the concrete kernel. -/ def kernelCone : KernelFork f := @@ -44,4 +44,4 @@ def cokernelIsColimit : IsColimit <| cokernelCocone f := (fun _ _ h => have : Epi (cokernelCocone f).π := (epi_iff_surjective _).mpr <| mk'_surjective _ (cancel_epi _).mp <| by simpa only [parallelPair_obj_one] using h) -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Limits.lean b/Mathlib/Algebra/Category/Grp/Limits.lean similarity index 54% rename from Mathlib/Algebra/Category/GroupCat/Limits.lean rename to Mathlib/Algebra/Category/Grp/Limits.lean index 9f0707ed84966..5683aed89ea9a 100644 --- a/Mathlib/Algebra/Category/GroupCat/Limits.lean +++ b/Mathlib/Algebra/Category/Grp/Limits.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Algebra.Category.MonCat.Limits -import Mathlib.Algebra.Category.GroupCat.ForgetCorepresentable -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.ForgetCorepresentable +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.Algebra.Group.Subgroup.Basic import Mathlib.CategoryTheory.Comma.Over import Mathlib.CategoryTheory.Limits.ConcreteCategory @@ -30,104 +30,104 @@ noncomputable section variable {J : Type v} [Category.{w} J] -namespace GroupCat +namespace Grp -variable (F : J ⥤ GroupCat.{u}) +variable (F : J ⥤ Grp.{u}) @[to_additive] -instance groupObj (j) : Group ((F ⋙ forget GroupCat).obj j) := +instance groupObj (j) : Group ((F ⋙ forget Grp).obj j) := inferInstanceAs <| Group (F.obj j) set_option linter.uppercaseLean3 false in -#align Group.group_obj GroupCat.groupObj +#align Group.group_obj Grp.groupObj set_option linter.uppercaseLean3 false in -#align AddGroup.add_group_obj AddGroupCat.addGroupObj +#align AddGroup.add_group_obj AddGrp.addGroupObj -/-- The flat sections of a functor into `GroupCat` form a subgroup of all sections. +/-- The flat sections of a functor into `Grp` form a subgroup of all sections. -/ @[to_additive - "The flat sections of a functor into `AddGroupCat` form an additive subgroup of all sections."] + "The flat sections of a functor into `AddGrp` form an additive subgroup of all sections."] def sectionsSubgroup : Subgroup (∀ j, F.obj j) := - { MonCat.sectionsSubmonoid (F ⋙ forget₂ GroupCat MonCat) with - carrier := (F ⋙ forget GroupCat).sections + { MonCat.sectionsSubmonoid (F ⋙ forget₂ Grp MonCat) with + carrier := (F ⋙ forget Grp).sections inv_mem' := fun {a} ah j j' f => by simp only [Functor.comp_map, Pi.inv_apply, MonoidHom.map_inv, inv_inj] dsimp [Functor.sections] at ah ⊢ rw [(F.map f).map_inv (a j), ah f] } set_option linter.uppercaseLean3 false in -#align Group.sections_subgroup GroupCat.sectionsSubgroup +#align Group.sections_subgroup Grp.sectionsSubgroup set_option linter.uppercaseLean3 false in -#align AddGroup.sections_add_subgroup AddGroupCat.sectionsAddSubgroup +#align AddGroup.sections_add_subgroup AddGrp.sectionsAddSubgroup @[to_additive] -instance sectionsGroup : Group (F ⋙ forget GroupCat.{u}).sections := +instance sectionsGroup : Group (F ⋙ forget Grp.{u}).sections := (sectionsSubgroup F).toGroup /-- The projection from `Functor.sections` to a factor as a `MonoidHom`. -/ @[to_additive "The projection from `Functor.sections` to a factor as an `AddMonoidHom`."] -def sectionsπMonoidHom (j : J) : (F ⋙ forget GroupCat.{u}).sections →* F.obj j where +def sectionsπMonoidHom (j : J) : (F ⋙ forget Grp.{u}).sections →* F.obj j where toFun x := x.val j map_one' := rfl map_mul' _ _ := rfl section -variable [Small.{u} (Functor.sections (F ⋙ forget GroupCat))] +variable [Small.{u} (Functor.sections (F ⋙ forget Grp))] @[to_additive] noncomputable instance limitGroup : - Group (Types.Small.limitCone.{v, u} (F ⋙ forget GroupCat.{u})).pt := - inferInstanceAs <| Group (Shrink (F ⋙ forget GroupCat.{u}).sections) + Group (Types.Small.limitCone.{v, u} (F ⋙ forget Grp.{u})).pt := + inferInstanceAs <| Group (Shrink (F ⋙ forget Grp.{u}).sections) set_option linter.uppercaseLean3 false in -#align Group.limit_group GroupCat.limitGroup +#align Group.limit_group Grp.limitGroup set_option linter.uppercaseLean3 false in -#align AddGroup.limit_add_group AddGroupCat.limitAddGroup +#align AddGroup.limit_add_group AddGrp.limitAddGroup @[to_additive] -instance : Small.{u} (Functor.sections ((F ⋙ forget₂ GroupCat MonCat) ⋙ forget MonCat)) := - inferInstanceAs <| Small.{u} (Functor.sections (F ⋙ forget GroupCat)) +instance : Small.{u} (Functor.sections ((F ⋙ forget₂ Grp MonCat) ⋙ forget MonCat)) := + inferInstanceAs <| Small.{u} (Functor.sections (F ⋙ forget Grp)) -/-- We show that the forgetful functor `GroupCat ⥤ MonCat` creates limits. +/-- We show that the forgetful functor `Grp ⥤ MonCat` creates limits. All we need to do is notice that the limit point has a `Group` instance available, and then reuse the existing limit. -/ -@[to_additive "We show that the forgetful functor `AddGroupCat ⥤ AddMonCat` creates limits. +@[to_additive "We show that the forgetful functor `AddGrp ⥤ AddMonCat` creates limits. All we need to do is notice that the limit point has an `AddGroup` instance available, and then reuse the existing limit."] noncomputable instance Forget₂.createsLimit : - CreatesLimit F (forget₂ GroupCat.{u} MonCat.{u}) := + CreatesLimit F (forget₂ Grp.{u} MonCat.{u}) := -- Porting note: need to add `forget₂ GrpCat MonCat` reflects isomorphism - letI : (forget₂ GroupCat.{u} MonCat.{u}).ReflectsIsomorphisms := + letI : (forget₂ Grp.{u} MonCat.{u}).ReflectsIsomorphisms := CategoryTheory.reflectsIsomorphisms_forget₂ _ _ - createsLimitOfReflectsIso (K := F) (F := (forget₂ GroupCat.{u} MonCat.{u})) + createsLimitOfReflectsIso (K := F) (F := (forget₂ Grp.{u} MonCat.{u})) fun c' t => { liftedCone := - { pt := GroupCat.of (Types.Small.limitCone (F ⋙ forget GroupCat)).pt + { pt := Grp.of (Types.Small.limitCone (F ⋙ forget Grp)).pt π := - { app := MonCat.limitπMonoidHom (F ⋙ forget₂ GroupCat MonCat) + { app := MonCat.limitπMonoidHom (F ⋙ forget₂ Grp MonCat) naturality := (MonCat.HasLimits.limitCone - (F ⋙ forget₂ GroupCat MonCat.{u})).π.naturality } } + (F ⋙ forget₂ Grp MonCat.{u})).π.naturality } } validLift := by apply IsLimit.uniqueUpToIso (MonCat.HasLimits.limitConeIsLimit.{v, u} _) t makesLimit := - IsLimit.ofFaithful (forget₂ GroupCat MonCat.{u}) (MonCat.HasLimits.limitConeIsLimit _) + IsLimit.ofFaithful (forget₂ Grp MonCat.{u}) (MonCat.HasLimits.limitConeIsLimit _) (fun s => _) fun s => rfl } set_option linter.uppercaseLean3 false in -#align Group.forget₂.creates_limit GroupCat.Forget₂.createsLimit +#align Group.forget₂.creates_limit Grp.Forget₂.createsLimit set_option linter.uppercaseLean3 false in -#align AddGroup.forget₂.creates_limit AddGroupCat.Forget₂.createsLimit +#align AddGroup.forget₂.creates_limit AddGrp.Forget₂.createsLimit -/-- A choice of limit cone for a functor into `GroupCat`. +/-- A choice of limit cone for a functor into `Grp`. (Generally, you'll just want to use `limit F`.) -/ -@[to_additive "A choice of limit cone for a functor into `GroupCat`. +@[to_additive "A choice of limit cone for a functor into `Grp`. (Generally, you'll just want to use `limit F`.)"] noncomputable def limitCone : Cone F := - liftLimit (limit.isLimit (F ⋙ forget₂ GroupCat.{u} MonCat.{u})) + liftLimit (limit.isLimit (F ⋙ forget₂ Grp.{u} MonCat.{u})) set_option linter.uppercaseLean3 false in -#align Group.limit_cone GroupCat.limitCone +#align Group.limit_cone Grp.limitCone set_option linter.uppercaseLean3 false in -#align AddGroup.limit_cone AddGroupCat.limitCone +#align AddGroup.limit_cone AddGrp.limitCone /-- The chosen cone is a limit cone. (Generally, you'll just want to use `limit.cone F`.) @@ -137,12 +137,12 @@ set_option linter.uppercaseLean3 false in noncomputable def limitConeIsLimit : IsLimit (limitCone F) := liftedLimitIsLimit _ set_option linter.uppercaseLean3 false in -#align Group.limit_cone_is_limit GroupCat.limitConeIsLimit +#align Group.limit_cone_is_limit Grp.limitConeIsLimit set_option linter.uppercaseLean3 false in -#align AddGroup.limit_cone_is_limit AddGroupCat.limitConeIsLimit +#align AddGroup.limit_cone_is_limit AddGrp.limitConeIsLimit -/-- If `(F ⋙ forget GroupCat).sections` is `u`-small, `F` has a limit. -/ -@[to_additive "If `(F ⋙ forget AddGroupCat).sections` is `u`-small, `F` has a limit."] +/-- If `(F ⋙ forget Grp).sections` is `u`-small, `F` has a limit. -/ +@[to_additive "If `(F ⋙ forget AddGrp).sections` is `u`-small, `F` has a limit."] instance hasLimit : HasLimit F := HasLimit.mk { cone := limitCone F @@ -151,72 +151,72 @@ instance hasLimit : HasLimit F := end -/-- A functor `F : J ⥤ GroupCat.{u}` has a limit iff `(F ⋙ forget GroupCat).sections` is +/-- A functor `F : J ⥤ Grp.{u}` has a limit iff `(F ⋙ forget Grp).sections` is `u`-small. -/ -@[to_additive "A functor `F : J ⥤ AddGroupCat.{u}` has a limit iff -`(F ⋙ forget AddGroupCat).sections` is `u`-small."] +@[to_additive "A functor `F : J ⥤ AddGrp.{u}` has a limit iff +`(F ⋙ forget AddGrp).sections` is `u`-small."] lemma hasLimit_iff_small_sections : - HasLimit F ↔ Small.{u} (F ⋙ forget GroupCat).sections := by + HasLimit F ↔ Small.{u} (F ⋙ forget Grp).sections := by constructor · apply Concrete.small_sections_of_hasLimit · intro infer_instance -/-- If `J` is `u`-small, `GroupCat.{u}` has limits of shape `J`. -/ -@[to_additive "If `J` is `u`-small, `AddGroupCat.{u}` has limits of shape `J`."] -instance hasLimitsOfShape [Small.{u} J] : HasLimitsOfShape J GroupCat.{u} where +/-- If `J` is `u`-small, `Grp.{u}` has limits of shape `J`. -/ +@[to_additive "If `J` is `u`-small, `AddGrp.{u}` has limits of shape `J`."] +instance hasLimitsOfShape [Small.{u} J] : HasLimitsOfShape J Grp.{u} where has_limit _ := inferInstance /-- The category of groups has all limits. -/ @[to_additive "The category of additive groups has all limits.", to_additive_relevant_arg 2] -instance hasLimitsOfSize [UnivLE.{v, u}] : HasLimitsOfSize.{w, v} GroupCat.{u} where +instance hasLimitsOfSize [UnivLE.{v, u}] : HasLimitsOfSize.{w, v} Grp.{u} where has_limits_of_shape J _ := { } set_option linter.uppercaseLean3 false in -#align Group.has_limits_of_size GroupCat.hasLimitsOfSize +#align Group.has_limits_of_size Grp.hasLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddGroup.has_limits_of_size AddGroupCat.hasLimitsOfSize +#align AddGroup.has_limits_of_size AddGrp.hasLimitsOfSize @[to_additive] -instance hasLimits : HasLimits GroupCat.{u} := - GroupCat.hasLimitsOfSize.{u, u} +instance hasLimits : HasLimits Grp.{u} := + Grp.hasLimitsOfSize.{u, u} set_option linter.uppercaseLean3 false in -#align Group.has_limits GroupCat.hasLimits +#align Group.has_limits Grp.hasLimits set_option linter.uppercaseLean3 false in -#align AddGroup.has_limits AddGroupCat.hasLimits +#align AddGroup.has_limits AddGrp.hasLimits /-- The forgetful functor from groups to monoids preserves all limits. This means the underlying monoid of a limit can be computed as a limit in the category of monoids. -/ -@[to_additive AddGroupCat.forget₂AddMonPreservesLimitsOfSize +@[to_additive AddGrp.forget₂AddMonPreservesLimitsOfSize "The forgetful functor from additive groups to additive monoids preserves all limits. This means the underlying additive monoid of a limit can be computed as a limit in the category of additive monoids.", to_additive_relevant_arg 2] noncomputable instance forget₂MonPreservesLimitsOfSize [UnivLE.{v, u}] : - PreservesLimitsOfSize.{w, v} (forget₂ GroupCat.{u} MonCat.{u}) where + PreservesLimitsOfSize.{w, v} (forget₂ Grp.{u} MonCat.{u}) where preservesLimitsOfShape {J _} := { } set_option linter.uppercaseLean3 false in -#align Group.forget₂_Mon_preserves_limits_of_size GroupCat.forget₂MonPreservesLimitsOfSize +#align Group.forget₂_Mon_preserves_limits_of_size Grp.forget₂MonPreservesLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddGroup.forget₂_AddMon_preserves_limits AddGroupCat.forget₂AddMonPreservesLimitsOfSize +#align AddGroup.forget₂_AddMon_preserves_limits AddGrp.forget₂AddMonPreservesLimitsOfSize @[to_additive] noncomputable instance forget₂MonPreservesLimits : - PreservesLimits (forget₂ GroupCat.{u} MonCat.{u}) := - GroupCat.forget₂MonPreservesLimitsOfSize.{u, u} + PreservesLimits (forget₂ Grp.{u} MonCat.{u}) := + Grp.forget₂MonPreservesLimitsOfSize.{u, u} set_option linter.uppercaseLean3 false in -#align Group.forget₂_Mon_preserves_limits GroupCat.forget₂MonPreservesLimits +#align Group.forget₂_Mon_preserves_limits Grp.forget₂MonPreservesLimits set_option linter.uppercaseLean3 false in -#align AddGroup.forget₂_Mon_preserves_limits AddGroupCat.forget₂MonPreservesLimits +#align AddGroup.forget₂_Mon_preserves_limits AddGrp.forget₂MonPreservesLimits -/-- If `J` is `u`-small, the forgetful functor from `GroupCat.{u}` preserves limits of shape `J`. -/ -@[to_additive "If `J` is `u`-small, the forgetful functor from `AddGroupCat.{u}`\n +/-- If `J` is `u`-small, the forgetful functor from `Grp.{u}` preserves limits of shape `J`. -/ +@[to_additive "If `J` is `u`-small, the forgetful functor from `AddGrp.{u}`\n preserves limits of shape `J`."] noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : - PreservesLimitsOfShape J (forget GroupCat.{u}) where + PreservesLimitsOfShape J (forget Grp.{u}) where preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) @@ -229,103 +229,103 @@ This means the underlying type of a limit can be computed as a limit in the cate This means the underlying type of a limit can be computed as a limit in the category of types.", to_additive_relevant_arg 2] noncomputable instance forgetPreservesLimitsOfSize : - PreservesLimitsOfSize.{w, v} (forget GroupCat.{u}) := inferInstance + PreservesLimitsOfSize.{w, v} (forget Grp.{u}) := inferInstance set_option linter.uppercaseLean3 false in -#align Group.forget_preserves_limits_of_size GroupCat.forgetPreservesLimitsOfSize +#align Group.forget_preserves_limits_of_size Grp.forgetPreservesLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddGroup.forget_preserves_limits_of_size AddGroupCat.forgetPreservesLimitsOfSize +#align AddGroup.forget_preserves_limits_of_size AddGrp.forgetPreservesLimitsOfSize @[to_additive] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget GroupCat.{u}) := - GroupCat.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance forgetPreservesLimits : PreservesLimits (forget Grp.{u}) := + Grp.forgetPreservesLimitsOfSize.{u, u} set_option linter.uppercaseLean3 false in -#align Group.forget_preserves_limits GroupCat.forgetPreservesLimits +#align Group.forget_preserves_limits Grp.forgetPreservesLimits set_option linter.uppercaseLean3 false in -#align AddGroup.forget_preserves_limits AddGroupCat.forgetPreservesLimits +#align AddGroup.forget_preserves_limits AddGrp.forgetPreservesLimits -end GroupCat +end Grp -namespace CommGroupCat +namespace CommGrp -variable (F : J ⥤ CommGroupCat.{u}) +variable (F : J ⥤ CommGrp.{u}) @[to_additive] -instance commGroupObj (j) : CommGroup ((F ⋙ forget CommGroupCat).obj j) := +instance commGroupObj (j) : CommGroup ((F ⋙ forget CommGrp).obj j) := inferInstanceAs <| CommGroup (F.obj j) set_option linter.uppercaseLean3 false in -#align CommGroup.comm_group_obj CommGroupCat.commGroupObj +#align CommGroup.comm_group_obj CommGrp.commGroupObj set_option linter.uppercaseLean3 false in -#align AddCommGroup.add_comm_group_obj AddCommGroupCat.addCommGroupObj +#align AddCommGroup.add_comm_group_obj AddCommGrp.addCommGroupObj @[to_additive] noncomputable instance limitCommGroup - [Small.{u} (Functor.sections (F ⋙ forget CommGroupCat))] : - CommGroup (Types.Small.limitCone.{v, u} (F ⋙ forget CommGroupCat.{u})).pt := - letI : CommGroup (F ⋙ forget CommGroupCat.{u}).sections := + [Small.{u} (Functor.sections (F ⋙ forget CommGrp))] : + CommGroup (Types.Small.limitCone.{v, u} (F ⋙ forget CommGrp.{u})).pt := + letI : CommGroup (F ⋙ forget CommGrp.{u}).sections := @Subgroup.toCommGroup (∀ j, F.obj j) _ - (GroupCat.sectionsSubgroup (F ⋙ forget₂ CommGroupCat.{u} GroupCat.{u})) - inferInstanceAs <| CommGroup (Shrink (F ⋙ forget CommGroupCat.{u}).sections) + (Grp.sectionsSubgroup (F ⋙ forget₂ CommGrp.{u} Grp.{u})) + inferInstanceAs <| CommGroup (Shrink (F ⋙ forget CommGrp.{u}).sections) set_option linter.uppercaseLean3 false in -#align CommGroup.limit_comm_group CommGroupCat.limitCommGroup +#align CommGroup.limit_comm_group CommGrp.limitCommGroup set_option linter.uppercaseLean3 false in -#align AddCommGroup.limit_add_comm_group AddCommGroupCat.limitAddCommGroup +#align AddCommGroup.limit_add_comm_group AddCommGrp.limitAddCommGroup @[to_additive] -instance : (forget₂ CommGroupCat.{u} GroupCat.{u}).ReflectsIsomorphisms := +instance : (forget₂ CommGrp.{u} Grp.{u}).ReflectsIsomorphisms := reflectsIsomorphisms_forget₂ _ _ -/-- We show that the forgetful functor `CommGroupCat ⥤ GroupCat` creates limits. +/-- We show that the forgetful functor `CommGrp ⥤ Grp` creates limits. All we need to do is notice that the limit point has a `CommGroup` instance available, and then reuse the existing limit. -/ -@[to_additive "We show that the forgetful functor `AddCommGroupCat ⥤ AddGroupCat` creates limits. +@[to_additive "We show that the forgetful functor `AddCommGrp ⥤ AddGrp` creates limits. All we need to do is notice that the limit point has an `AddCommGroup` instance available, and then reuse the existing limit."] noncomputable instance Forget₂.createsLimit : - CreatesLimit F (forget₂ CommGroupCat GroupCat.{u}) := + CreatesLimit F (forget₂ CommGrp Grp.{u}) := createsLimitOfReflectsIso (fun c hc => by have : HasLimit _ := ⟨_, hc⟩ - have : Small.{u} (F ⋙ forget CommGroupCat).sections := - Concrete.small_sections_of_hasLimit (F ⋙ forget₂ CommGroupCat GroupCat) - have : Small.{u} ((F ⋙ forget₂ CommGroupCat GroupCat ⋙ forget₂ GroupCat MonCat) ⋙ + have : Small.{u} (F ⋙ forget CommGrp).sections := + Concrete.small_sections_of_hasLimit (F ⋙ forget₂ CommGrp Grp) + have : Small.{u} ((F ⋙ forget₂ CommGrp Grp ⋙ forget₂ Grp MonCat) ⋙ forget MonCat).sections := this - have : Small.{u} ((F ⋙ forget₂ CommGroupCat GroupCat) ⋙ forget GroupCat).sections := this + have : Small.{u} ((F ⋙ forget₂ CommGrp Grp) ⋙ forget Grp).sections := this exact { liftedCone := - { pt := CommGroupCat.of (Types.Small.limitCone.{v, u} (F ⋙ forget CommGroupCat)).pt + { pt := CommGrp.of (Types.Small.limitCone.{v, u} (F ⋙ forget CommGrp)).pt π := { app := MonCat.limitπMonoidHom - (F ⋙ forget₂ CommGroupCat GroupCat.{u} ⋙ forget₂ GroupCat MonCat.{u}) + (F ⋙ forget₂ CommGrp Grp.{u} ⋙ forget₂ Grp MonCat.{u}) naturality := (MonCat.HasLimits.limitCone _).π.naturality } } - validLift := by apply IsLimit.uniqueUpToIso (GroupCat.limitConeIsLimit _) hc + validLift := by apply IsLimit.uniqueUpToIso (Grp.limitConeIsLimit _) hc makesLimit := - IsLimit.ofFaithful (forget₂ _ GroupCat.{u} ⋙ forget₂ _ MonCat.{u}) + IsLimit.ofFaithful (forget₂ _ Grp.{u} ⋙ forget₂ _ MonCat.{u}) (by apply MonCat.HasLimits.limitConeIsLimit _) (fun s => _) fun s => rfl }) set_option linter.uppercaseLean3 false in -#align CommGroup.forget₂.creates_limit CommGroupCat.Forget₂.createsLimit +#align CommGroup.forget₂.creates_limit CommGrp.Forget₂.createsLimit set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget₂.creates_limit AddCommGroupCat.Forget₂.createsLimit +#align AddCommGroup.forget₂.creates_limit AddCommGrp.Forget₂.createsLimit section -variable [Small.{u} (Functor.sections (F ⋙ forget CommGroupCat))] +variable [Small.{u} (Functor.sections (F ⋙ forget CommGrp))] -/-- A choice of limit cone for a functor into `CommGroupCat`. +/-- A choice of limit cone for a functor into `CommGrp`. (Generally, you'll just want to use `limit F`.) -/ @[to_additive - "A choice of limit cone for a functor into `AddCommGroupCat`. + "A choice of limit cone for a functor into `AddCommGrp`. (Generally, you'll just want to use `limit F`.)"] noncomputable def limitCone : Cone F := - letI : Small.{u} (Functor.sections ((F ⋙ forget₂ CommGroupCat GroupCat) ⋙ forget GroupCat)) := - inferInstanceAs <| Small (Functor.sections (F ⋙ forget CommGroupCat)) - liftLimit (limit.isLimit (F ⋙ forget₂ CommGroupCat.{u} GroupCat.{u})) + letI : Small.{u} (Functor.sections ((F ⋙ forget₂ CommGrp Grp) ⋙ forget Grp)) := + inferInstanceAs <| Small (Functor.sections (F ⋙ forget CommGrp)) + liftLimit (limit.isLimit (F ⋙ forget₂ CommGrp.{u} Grp.{u})) set_option linter.uppercaseLean3 false in -#align CommGroup.limit_cone CommGroupCat.limitCone +#align CommGroup.limit_cone CommGrp.limitCone set_option linter.uppercaseLean3 false in -#align AddCommGroup.limit_cone AddCommGroupCat.limitCone +#align AddCommGroup.limit_cone AddCommGrp.limitCone /-- The chosen cone is a limit cone. (Generally, you'll just want to use `limit.cone F`.) @@ -336,12 +336,12 @@ set_option linter.uppercaseLean3 false in noncomputable def limitConeIsLimit : IsLimit (limitCone.{v, u} F) := liftedLimitIsLimit _ set_option linter.uppercaseLean3 false in -#align CommGroup.limit_cone_is_limit CommGroupCat.limitConeIsLimit +#align CommGroup.limit_cone_is_limit CommGrp.limitConeIsLimit set_option linter.uppercaseLean3 false in -#align AddCommGroup.limit_cone_is_limit AddCommGroupCat.limitConeIsLimit +#align AddCommGroup.limit_cone_is_limit AddCommGrp.limitConeIsLimit -/-- If `(F ⋙ forget CommGroupCat).sections` is `u`-small, `F` has a limit. -/ -@[to_additive "If `(F ⋙ forget AddCommGroupCat).sections` is `u`-small, `F` has a limit."] +/-- If `(F ⋙ forget CommGrp).sections` is `u`-small, `F` has a limit. -/ +@[to_additive "If `(F ⋙ forget AddCommGrp).sections` is `u`-small, `F` has a limit."] instance hasLimit : HasLimit F := HasLimit.mk { cone := limitCone F @@ -350,54 +350,54 @@ instance hasLimit : HasLimit F := end -/-- A functor `F : J ⥤ CommGroupCat.{u}` has a limit iff `(F ⋙ forget CommGroupCat).sections` is +/-- A functor `F : J ⥤ CommGrp.{u}` has a limit iff `(F ⋙ forget CommGrp).sections` is `u`-small. -/ -@[to_additive "A functor `F : J ⥤ AddCommGroupCat.{u}` has a limit iff -`(F ⋙ forget AddCommGroupCat).sections` is `u`-small."] +@[to_additive "A functor `F : J ⥤ AddCommGrp.{u}` has a limit iff +`(F ⋙ forget AddCommGrp).sections` is `u`-small."] lemma hasLimit_iff_small_sections : - HasLimit F ↔ Small.{u} (F ⋙ forget CommGroupCat).sections := by + HasLimit F ↔ Small.{u} (F ⋙ forget CommGrp).sections := by constructor · apply Concrete.small_sections_of_hasLimit · intro infer_instance -/-- If `J` is `u`-small, `CommGroupCat.{u}` has limits of shape `J`. -/ -@[to_additive "If `J` is `u`-small, `AddCommGroupCat.{u}` has limits of shape `J`."] -instance hasLimitsOfShape [Small.{u} J] : HasLimitsOfShape J CommGroupCat.{u} where +/-- If `J` is `u`-small, `CommGrp.{u}` has limits of shape `J`. -/ +@[to_additive "If `J` is `u`-small, `AddCommGrp.{u}` has limits of shape `J`."] +instance hasLimitsOfShape [Small.{u} J] : HasLimitsOfShape J CommGrp.{u} where has_limit _ := inferInstance /-- The category of commutative groups has all limits. -/ @[to_additive "The category of additive commutative groups has all limits.", to_additive_relevant_arg 2] -instance hasLimitsOfSize [UnivLE.{v, u}] : HasLimitsOfSize.{w, v} CommGroupCat.{u} +instance hasLimitsOfSize [UnivLE.{v, u}] : HasLimitsOfSize.{w, v} CommGrp.{u} where has_limits_of_shape _ _ := { } set_option linter.uppercaseLean3 false in -#align CommGroup.has_limits_of_size CommGroupCat.hasLimitsOfSize +#align CommGroup.has_limits_of_size CommGrp.hasLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddCommGroup.has_limits_of_size AddCommGroupCat.hasLimitsOfSize +#align AddCommGroup.has_limits_of_size AddCommGrp.hasLimitsOfSize @[to_additive] -instance hasLimits : HasLimits CommGroupCat.{u} := - CommGroupCat.hasLimitsOfSize.{u, u} +instance hasLimits : HasLimits CommGrp.{u} := + CommGrp.hasLimitsOfSize.{u, u} set_option linter.uppercaseLean3 false in -#align CommGroup.has_limits CommGroupCat.hasLimits +#align CommGroup.has_limits CommGrp.hasLimits set_option linter.uppercaseLean3 false in -#align AddCommGroup.has_limits AddCommGroupCat.hasLimits +#align AddCommGroup.has_limits AddCommGrp.hasLimits @[to_additive] noncomputable instance forget₂GroupPreservesLimit : - PreservesLimit F (forget₂ CommGroupCat.{u} GroupCat.{u}) where + PreservesLimit F (forget₂ CommGrp.{u} Grp.{u}) where preserves {c} hc := by - have : HasLimit (F ⋙ forget₂ CommGroupCat GroupCat) := by - rw [GroupCat.hasLimit_iff_small_sections] - change Small.{u} (F ⋙ forget CommGroupCat).sections - rw [← CommGroupCat.hasLimit_iff_small_sections] + have : HasLimit (F ⋙ forget₂ CommGrp Grp) := by + rw [Grp.hasLimit_iff_small_sections] + change Small.{u} (F ⋙ forget CommGrp).sections + rw [← CommGrp.hasLimit_iff_small_sections] exact ⟨_, hc⟩ exact isLimitOfPreserves _ hc @[to_additive] noncomputable instance forget₂GroupPreservesLimitsOfShape : - PreservesLimitsOfShape J (forget₂ CommGroupCat.{u} GroupCat.{u}) where + PreservesLimitsOfShape J (forget₂ CommGrp.{u} Grp.{u}) where /-- The forgetful functor from commutative groups to groups preserves all limits. (That is, the underlying group could have been computed instead as limits in the category @@ -409,43 +409,43 @@ of groups.) of additive groups.)", to_additive_relevant_arg 2] noncomputable instance forget₂GroupPreservesLimitsOfSize : - PreservesLimitsOfSize.{w, v} (forget₂ CommGroupCat.{u} GroupCat.{u}) where + PreservesLimitsOfSize.{w, v} (forget₂ CommGrp.{u} Grp.{u}) where set_option linter.uppercaseLean3 false in -#align CommGroup.forget₂_Group_preserves_limits_of_size CommGroupCat.forget₂GroupPreservesLimitsOfSize +#align CommGroup.forget₂_Group_preserves_limits_of_size CommGrp.forget₂GroupPreservesLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget₂_AddGroup_preserves_limits AddCommGroupCat.forget₂AddGroupPreservesLimitsOfSize +#align AddCommGroup.forget₂_AddGroup_preserves_limits AddCommGrp.forget₂AddGroupPreservesLimitsOfSize @[to_additive] noncomputable instance forget₂GroupPreservesLimits : - PreservesLimits (forget₂ CommGroupCat GroupCat.{u}) := - CommGroupCat.forget₂GroupPreservesLimitsOfSize.{u, u} + PreservesLimits (forget₂ CommGrp Grp.{u}) := + CommGrp.forget₂GroupPreservesLimitsOfSize.{u, u} set_option linter.uppercaseLean3 false in -#align CommGroup.forget₂_Group_preserves_limits CommGroupCat.forget₂GroupPreservesLimits +#align CommGroup.forget₂_Group_preserves_limits CommGrp.forget₂GroupPreservesLimits set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget₂_Group_preserves_limits AddCommGroupCat.forget₂AddGroupPreservesLimits +#align AddCommGroup.forget₂_Group_preserves_limits AddCommGrp.forget₂AddGroupPreservesLimits /-- An auxiliary declaration to speed up typechecking. -/ -@[to_additive AddCommGroupCat.forget₂AddCommMonPreservesLimitsAux +@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsAux "An auxiliary declaration to speed up typechecking."] noncomputable def forget₂CommMonPreservesLimitsAux - [Small.{u} (F ⋙ forget CommGroupCat).sections] : - IsLimit ((forget₂ CommGroupCat.{u} CommMonCat.{u}).mapCone (limitCone.{v, u} F)) := + [Small.{u} (F ⋙ forget CommGrp).sections] : + IsLimit ((forget₂ CommGrp.{u} CommMonCat.{u}).mapCone (limitCone.{v, u} F)) := letI : Small.{u} (Functor.sections ((F ⋙ forget₂ _ CommMonCat) ⋙ forget CommMonCat)) := - inferInstanceAs <| Small (Functor.sections (F ⋙ forget CommGroupCat)) - CommMonCat.limitConeIsLimit.{v, u} (F ⋙ forget₂ CommGroupCat.{u} CommMonCat.{u}) + inferInstanceAs <| Small (Functor.sections (F ⋙ forget CommGrp)) + CommMonCat.limitConeIsLimit.{v, u} (F ⋙ forget₂ CommGrp.{u} CommMonCat.{u}) set_option linter.uppercaseLean3 false in -#align CommGroup.forget₂_CommMon_preserves_limits_aux CommGroupCat.forget₂CommMonPreservesLimitsAux +#align CommGroup.forget₂_CommMon_preserves_limits_aux CommGrp.forget₂CommMonPreservesLimitsAux set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget₂_AddCommMon_preserves_limits_aux AddCommGroupCat.forget₂AddCommMonPreservesLimitsAux +#align AddCommGroup.forget₂_AddCommMon_preserves_limits_aux AddCommGrp.forget₂AddCommMonPreservesLimitsAux -/-- If `J` is `u`-small, the forgetful functor from `CommGroupCat.{u}` to `CommMonCat.{u}` +/-- If `J` is `u`-small, the forgetful functor from `CommGrp.{u}` to `CommMonCat.{u}` preserves limits of shape `J`. -/ -@[to_additive AddCommGroupCat.forget₂AddCommMonPreservesLimitsOfShape - "If `J` is `u`-small, the forgetful functor from `AddCommGroupCat.{u}` +@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsOfShape + "If `J` is `u`-small, the forgetful functor from `AddCommGrp.{u}` to `AddCommMonCat.{u}` preserves limits of shape `J`."] noncomputable instance forget₂CommMonPreservesLimitsOfShape [Small.{u} J] : - PreservesLimitsOfShape J (forget₂ CommGroupCat.{u} CommMonCat.{u}) where + PreservesLimitsOfShape J (forget₂ CommGrp.{u} CommMonCat.{u}) where preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) (forget₂CommMonPreservesLimitsAux.{v, u} F) @@ -453,24 +453,24 @@ noncomputable instance forget₂CommMonPreservesLimitsOfShape [Small.{u} J] : (That is, the underlying commutative monoids could have been computed instead as limits in the category of commutative monoids.) -/ -@[to_additive AddCommGroupCat.forget₂AddCommMonPreservesLimitsOfSize +@[to_additive AddCommGrp.forget₂AddCommMonPreservesLimitsOfSize "The forgetful functor from additive commutative groups to additive commutative monoids preserves all limits. (That is, the underlying additive commutative monoids could have been computed instead as limits in the category of additive commutative monoids.)"] noncomputable instance forget₂CommMonPreservesLimitsOfSize [UnivLE.{v, u}] : - PreservesLimitsOfSize.{w, v} (forget₂ CommGroupCat CommMonCat.{u}) where + PreservesLimitsOfSize.{w, v} (forget₂ CommGrp CommMonCat.{u}) where preservesLimitsOfShape := { } set_option linter.uppercaseLean3 false in -#align CommGroup.forget₂_CommMon_preserves_limits_of_size CommGroupCat.forget₂CommMonPreservesLimitsOfSize +#align CommGroup.forget₂_CommMon_preserves_limits_of_size CommGrp.forget₂CommMonPreservesLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget₂_AddCommMon_preserves_limits AddCommGroupCat.forget₂AddCommMonPreservesLimitsOfSize +#align AddCommGroup.forget₂_AddCommMon_preserves_limits AddCommGrp.forget₂AddCommMonPreservesLimitsOfSize -/-- If `J` is `u`-small, the forgetful functor from `CommGroupCat.{u}` preserves limits of +/-- If `J` is `u`-small, the forgetful functor from `CommGrp.{u}` preserves limits of shape `J`. -/ -@[to_additive "If `J` is `u`-small, the forgetful functor from `AddCommGroupCat.{u}`\n +@[to_additive "If `J` is `u`-small, the forgetful functor from `AddCommGrp.{u}`\n preserves limits of shape `J`."] noncomputable instance forgetPreservesLimitsOfShape [Small.{u} J] : - PreservesLimitsOfShape J (forget CommGroupCat.{u}) where + PreservesLimitsOfShape J (forget CommGrp.{u}) where preservesLimit {F} := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) (Types.Small.limitConeIsLimit (F ⋙ forget _)) @@ -482,32 +482,32 @@ underlying types could have been computed instead as limits in the category of t (That is, the underlying types could have been computed instead as limits in the category of types.)"] noncomputable instance forgetPreservesLimitsOfSize : - PreservesLimitsOfSize.{w, v} (forget CommGroupCat.{u}) := inferInstance + PreservesLimitsOfSize.{w, v} (forget CommGrp.{u}) := inferInstance set_option linter.uppercaseLean3 false in -#align CommGroup.forget_preserves_limits_of_size CommGroupCat.forgetPreservesLimitsOfSize +#align CommGroup.forget_preserves_limits_of_size CommGrp.forgetPreservesLimitsOfSize set_option linter.uppercaseLean3 false in -#align AddCommGroup.forget_preserves_limits AddCommGroupCat.forgetPreservesLimitsOfSize +#align AddCommGroup.forget_preserves_limits AddCommGrp.forgetPreservesLimitsOfSize -noncomputable instance _root_.AddCommGroupCat.forgetPreservesLimits : - PreservesLimits (forget AddCommGroupCat.{u}) := - AddCommGroupCat.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance _root_.AddCommGrp.forgetPreservesLimits : + PreservesLimits (forget AddCommGrp.{u}) := + AddCommGrp.forgetPreservesLimitsOfSize.{u, u} @[to_additive existing] -noncomputable instance forgetPreservesLimits : PreservesLimits (forget CommGroupCat.{u}) := - CommGroupCat.forgetPreservesLimitsOfSize.{u, u} +noncomputable instance forgetPreservesLimits : PreservesLimits (forget CommGrp.{u}) := + CommGrp.forgetPreservesLimitsOfSize.{u, u} -- Verify we can form limits indexed over smaller categories. -example (f : ℕ → AddCommGroupCat) : HasProduct f := by infer_instance +example (f : ℕ → AddCommGrp) : HasProduct f := by infer_instance -end CommGroupCat +end CommGrp -namespace AddCommGroupCat +namespace AddCommGrp -/-- The categorical kernel of a morphism in `AddCommGroupCat` +/-- The categorical kernel of a morphism in `AddCommGrp` agrees with the usual group-theoretical kernel. -/ -def kernelIsoKer {G H : AddCommGroupCat.{u}} (f : G ⟶ H) : - kernel f ≅ AddCommGroupCat.of f.ker where +def kernelIsoKer {G H : AddCommGrp.{u}} (f : G ⟶ H) : + kernel f ≅ AddCommGrp.of f.ker where hom := { toFun := fun g => ⟨kernel.ι f g, DFunLike.congr_fun (kernel.condition f) g⟩ map_zero' := by @@ -531,28 +531,28 @@ def kernelIsoKer {G H : AddCommGroupCat.{u}} (f : G ⟶ H) : dsimp apply DFunLike.congr_fun (kernel.lift_ι f _ _) inv_hom_id := by - apply AddCommGroupCat.ext + apply AddCommGrp.ext simp only [AddMonoidHom.coe_mk, coe_id, coe_comp] rintro ⟨x, mem⟩ refine Subtype.ext ?_ simp only [ZeroHom.coe_mk, Function.comp_apply, id_eq] apply DFunLike.congr_fun (kernel.lift_ι f _ _) set_option linter.uppercaseLean3 false in -#align AddCommGroup.kernel_iso_ker AddCommGroupCat.kernelIsoKer +#align AddCommGroup.kernel_iso_ker AddCommGrp.kernelIsoKer @[simp] -theorem kernelIsoKer_hom_comp_subtype {G H : AddCommGroupCat.{u}} (f : G ⟶ H) : +theorem kernelIsoKer_hom_comp_subtype {G H : AddCommGrp.{u}} (f : G ⟶ H) : (kernelIsoKer f).hom ≫ AddSubgroup.subtype f.ker = kernel.ι f := by ext; rfl set_option linter.uppercaseLean3 false in -#align AddCommGroup.kernel_iso_ker_hom_comp_subtype AddCommGroupCat.kernelIsoKer_hom_comp_subtype +#align AddCommGroup.kernel_iso_ker_hom_comp_subtype AddCommGrp.kernelIsoKer_hom_comp_subtype @[simp] -theorem kernelIsoKer_inv_comp_ι {G H : AddCommGroupCat.{u}} (f : G ⟶ H) : +theorem kernelIsoKer_inv_comp_ι {G H : AddCommGrp.{u}} (f : G ⟶ H) : (kernelIsoKer f).inv ≫ kernel.ι f = AddSubgroup.subtype f.ker := by ext simp [kernelIsoKer] set_option linter.uppercaseLean3 false in -#align AddCommGroup.kernel_iso_ker_inv_comp_ι AddCommGroupCat.kernelIsoKer_inv_comp_ι +#align AddCommGroup.kernel_iso_ker_inv_comp_ι AddCommGrp.kernelIsoKer_inv_comp_ι -- Porting note: explicitly add what to be synthesized under `simps!`, because other lemmas -- automatically generated is not in normal form @@ -560,14 +560,14 @@ set_option linter.uppercaseLean3 false in agrees with the `AddSubgroup.subtype` map. -/ @[simps! hom_left_apply_coe inv_left_apply] -def kernelIsoKerOver {G H : AddCommGroupCat.{u}} (f : G ⟶ H) : - Over.mk (kernel.ι f) ≅ @Over.mk _ _ G (AddCommGroupCat.of f.ker) (AddSubgroup.subtype f.ker) := +def kernelIsoKerOver {G H : AddCommGrp.{u}} (f : G ⟶ H) : + Over.mk (kernel.ι f) ≅ @Over.mk _ _ G (AddCommGrp.of f.ker) (AddSubgroup.subtype f.ker) := Over.isoMk (kernelIsoKer f) set_option linter.uppercaseLean3 false in -#align AddCommGroup.kernel_iso_ker_over AddCommGroupCat.kernelIsoKerOver +#align AddCommGroup.kernel_iso_ker_over AddCommGrp.kernelIsoKerOver -- These lemmas have always been bad (#7657), but lean4#2644 made `simp` start noticing -attribute [nolint simpNF] AddCommGroupCat.kernelIsoKerOver_inv_left_apply - AddCommGroupCat.kernelIsoKerOver_hom_left_apply_coe +attribute [nolint simpNF] AddCommGrp.kernelIsoKerOver_inv_left_apply + AddCommGrp.kernelIsoKerOver_hom_left_apply_coe -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Preadditive.lean b/Mathlib/Algebra/Category/Grp/Preadditive.lean similarity index 70% rename from Mathlib/Algebra/Category/GroupCat/Preadditive.lean rename to Mathlib/Algebra/Category/Grp/Preadditive.lean index 5f6442c9cf077..ae9d60a475c0f 100644 --- a/Mathlib/Algebra/Category/GroupCat/Preadditive.lean +++ b/Mathlib/Algebra/Category/Grp/Preadditive.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.CategoryTheory.Preadditive.Basic #align_import algebra.category.Group.preadditive from "leanprover-community/mathlib"@"829895f162a1f29d0133f4b3538f4cd1fb5bffd3" @@ -16,26 +16,26 @@ open CategoryTheory universe u -namespace AddCommGroupCat +namespace AddCommGrp -- porting note (#10670): this instance was not necessary in mathlib -instance (P Q : AddCommGroupCat) : AddCommGroup (P ⟶ Q) := +instance (P Q : AddCommGrp) : AddCommGroup (P ⟶ Q) := (inferInstance : AddCommGroup (AddMonoidHom P Q)) -- porting note (#10688): this lemma was not necessary in mathlib @[simp] -lemma hom_add_apply {P Q : AddCommGroupCat} (f g : P ⟶ Q) (x : P) : (f + g) x = f x + g x := rfl +lemma hom_add_apply {P Q : AddCommGrp} (f g : P ⟶ Q) (x : P) : (f + g) x = f x + g x := rfl section -- Porting note: the simp attribute was locally deactivated here, --- otherwise Lean would try to infer `Preadditive AddCommGroupCat` +-- otherwise Lean would try to infer `Preadditive AddCommGrp` -- in order to prove the axioms `add_comp` and `comp_add` in the -- next instance declaration attribute [-simp] Preadditive.add_comp Preadditive.comp_add -instance : Preadditive AddCommGroupCat where +instance : Preadditive AddCommGrp where end -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/Subobject.lean b/Mathlib/Algebra/Category/Grp/Subobject.lean similarity index 56% rename from Mathlib/Algebra/Category/GroupCat/Subobject.lean rename to Mathlib/Algebra/Category/Grp/Subobject.lean index 54aa4f40c9f8c..5db28236f0503 100644 --- a/Mathlib/Algebra/Category/GroupCat/Subobject.lean +++ b/Mathlib/Algebra/Category/Grp/Subobject.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ -import Mathlib.Algebra.Category.GroupCat.ZModuleEquivalence +import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Algebra.Category.ModuleCat.Subobject #align_import algebra.category.Group.subobject from "leanprover-community/mathlib"@"29b834dfbba4ed1b7950628bbf5e69ab98c15b4c" @@ -17,11 +17,11 @@ open CategoryTheory universe u -namespace AddCommGroupCat +namespace AddCommGrp -instance wellPowered_addCommGroupCat : WellPowered AddCommGroupCat.{u} := - wellPowered_of_equiv (forget₂ (ModuleCat.{u} ℤ) AddCommGroupCat.{u}).asEquivalence +instance wellPowered_addCommGrp : WellPowered AddCommGrp.{u} := + wellPowered_of_equiv (forget₂ (ModuleCat.{u} ℤ) AddCommGrp.{u}).asEquivalence set_option linter.uppercaseLean3 false in -#align AddCommGroup.well_powered_AddCommGroup AddCommGroupCat.wellPowered_addCommGroupCat +#align AddCommGroup.well_powered_AddCommGroup AddCommGrp.wellPowered_addCommGrp -end AddCommGroupCat +end AddCommGrp diff --git a/Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean similarity index 82% rename from Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean rename to Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean index 94c16c839d277..506b9d3c38b2b 100644 --- a/Mathlib/Algebra/Category/GroupCat/ZModuleEquivalence.lean +++ b/Mathlib/Algebra/Category/Grp/ZModuleEquivalence.lean @@ -25,8 +25,8 @@ universe u namespace ModuleCat -/-- The forgetful functor from `ℤ` modules to `AddCommGroupCat` is full. -/ -instance forget₂_addCommGroup_full : (forget₂ (ModuleCat ℤ) AddCommGroupCat.{u}).Full where +/-- The forgetful functor from `ℤ` modules to `AddCommGrp` is full. -/ +instance forget₂_addCommGroup_full : (forget₂ (ModuleCat ℤ) AddCommGrp.{u}).Full where map_surjective {A B} -- `AddMonoidHom.toIntLinearMap` doesn't work here because `A` and `B` are not -- definitionally equal to the canonical `AddCommGroup.intModule` module @@ -40,17 +40,17 @@ instance forget₂_addCommGroup_full : (forget₂ (ModuleCat ℤ) AddCommGroupCa set_option linter.uppercaseLean3 false in #align Module.forget₂_AddCommGroup_full ModuleCat.forget₂_addCommGroup_full -/-- The forgetful functor from `ℤ` modules to `AddCommGroupCat` is essentially surjective. -/ -instance forget₂_addCommGroupCat_essSurj : (forget₂ (ModuleCat ℤ) AddCommGroupCat.{u}).EssSurj where +/-- The forgetful functor from `ℤ` modules to `AddCommGrp` is essentially surjective. -/ +instance forget₂_addCommGrp_essSurj : (forget₂ (ModuleCat ℤ) AddCommGrp.{u}).EssSurj where mem_essImage A := ⟨ModuleCat.of ℤ A, ⟨{ hom := 𝟙 A inv := 𝟙 A }⟩⟩ set_option linter.uppercaseLean3 false in -#align Module.forget₂_AddCommGroup_ess_surj ModuleCat.forget₂_addCommGroupCat_essSurj +#align Module.forget₂_AddCommGroup_ess_surj ModuleCat.forget₂_addCommGrp_essSurj noncomputable instance forget₂AddCommGroupIsEquivalence : - (forget₂ (ModuleCat ℤ) AddCommGroupCat.{u}).IsEquivalence where + (forget₂ (ModuleCat ℤ) AddCommGrp.{u}).IsEquivalence where set_option linter.uppercaseLean3 false in #align Module.forget₂_AddCommGroup_is_equivalence ModuleCat.forget₂AddCommGroupIsEquivalence diff --git a/Mathlib/Algebra/Category/GroupCat/Zero.lean b/Mathlib/Algebra/Category/Grp/Zero.lean similarity index 64% rename from Mathlib/Algebra/Category/GroupCat/Zero.lean rename to Mathlib/Algebra/Category/Grp/Zero.lean index a0d06d045ba25..a1e1c65ace197 100644 --- a/Mathlib/Algebra/Category/GroupCat/Zero.lean +++ b/Mathlib/Algebra/Category/Grp/Zero.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.CategoryTheory.Limits.Shapes.ZeroObjects #align_import algebra.category.Group.zero from "leanprover-community/mathlib"@"70fd9563a21e7b963887c9360bd29b2393e6225a" @@ -22,10 +22,10 @@ open CategoryTheory.Limits universe u -namespace GroupCat +namespace Grp @[to_additive] -theorem isZero_of_subsingleton (G : GroupCat) [Subsingleton G] : IsZero G := by +theorem isZero_of_subsingleton (G : Grp) [Subsingleton G] : IsZero G := by refine ⟨fun X => ⟨⟨⟨1⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨1⟩, fun f => ?_⟩⟩⟩ · ext x have : x = 1 := Subsingleton.elim _ _ @@ -33,20 +33,20 @@ theorem isZero_of_subsingleton (G : GroupCat) [Subsingleton G] : IsZero G := by · ext apply Subsingleton.elim set_option linter.uppercaseLean3 false in -#align Group.is_zero_of_subsingleton GroupCat.isZero_of_subsingleton +#align Group.is_zero_of_subsingleton Grp.isZero_of_subsingleton set_option linter.uppercaseLean3 false in -#align AddGroup.is_zero_of_subsingleton AddGroupCat.isZero_of_subsingleton +#align AddGroup.is_zero_of_subsingleton AddGrp.isZero_of_subsingleton -@[to_additive AddGroupCat.hasZeroObject] -instance : HasZeroObject GroupCat := +@[to_additive AddGrp.hasZeroObject] +instance : HasZeroObject Grp := ⟨⟨of PUnit, isZero_of_subsingleton _⟩⟩ -end GroupCat +end Grp -namespace CommGroupCat +namespace CommGrp @[to_additive] -theorem isZero_of_subsingleton (G : CommGroupCat) [Subsingleton G] : IsZero G := by +theorem isZero_of_subsingleton (G : CommGrp) [Subsingleton G] : IsZero G := by refine ⟨fun X => ⟨⟨⟨1⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨1⟩, fun f => ?_⟩⟩⟩ · ext x have : x = 1 := Subsingleton.elim _ _ @@ -54,12 +54,12 @@ theorem isZero_of_subsingleton (G : CommGroupCat) [Subsingleton G] : IsZero G := · ext apply Subsingleton.elim set_option linter.uppercaseLean3 false in -#align CommGroup.is_zero_of_subsingleton CommGroupCat.isZero_of_subsingleton +#align CommGroup.is_zero_of_subsingleton CommGrp.isZero_of_subsingleton set_option linter.uppercaseLean3 false in -#align AddCommGroup.is_zero_of_subsingleton AddCommGroupCat.isZero_of_subsingleton +#align AddCommGroup.is_zero_of_subsingleton AddCommGrp.isZero_of_subsingleton -@[to_additive AddCommGroupCat.hasZeroObject] -instance : HasZeroObject CommGroupCat := +@[to_additive AddCommGrp.hasZeroObject] +instance : HasZeroObject CommGrp := ⟨⟨of PUnit, isZero_of_subsingleton _⟩⟩ -end CommGroupCat +end CommGrp diff --git a/Mathlib/Algebra/Category/GroupWithZeroCat.lean b/Mathlib/Algebra/Category/GrpWithZero.lean similarity index 56% rename from Mathlib/Algebra/Category/GroupWithZeroCat.lean rename to Mathlib/Algebra/Category/GrpWithZero.lean index 5dd70eacaa7e9..cad8aac66bd74 100644 --- a/Mathlib/Algebra/Category/GroupWithZeroCat.lean +++ b/Mathlib/Algebra/Category/GrpWithZero.lean @@ -12,7 +12,7 @@ import Mathlib.CategoryTheory.Category.Bipointed /-! # The category of groups with zero -This file defines `GroupWithZeroCat`, the category of groups with zero. +This file defines `GrpWithZero`, the category of groups with zero. -/ universe u @@ -20,29 +20,29 @@ universe u open CategoryTheory Order /-- The category of groups with zero. -/ -def GroupWithZeroCat := +def GrpWithZero := Bundled GroupWithZero set_option linter.uppercaseLean3 false in -#align GroupWithZero GroupWithZeroCat +#align GroupWithZero GrpWithZero -namespace GroupWithZeroCat +namespace GrpWithZero -instance : CoeSort GroupWithZeroCat Type* := +instance : CoeSort GrpWithZero Type* := Bundled.coeSort -instance (X : GroupWithZeroCat) : GroupWithZero X := +instance (X : GrpWithZero) : GroupWithZero X := X.str -/-- Construct a bundled `GroupWithZeroCat` from a `GroupWithZero`. -/ -def of (α : Type*) [GroupWithZero α] : GroupWithZeroCat := +/-- Construct a bundled `GrpWithZero` from a `GroupWithZero`. -/ +def of (α : Type*) [GroupWithZero α] : GrpWithZero := Bundled.of α set_option linter.uppercaseLean3 false in -#align GroupWithZero.of GroupWithZeroCat.of +#align GroupWithZero.of GrpWithZero.of -instance : Inhabited GroupWithZeroCat := +instance : Inhabited GrpWithZero := ⟨of (WithZero PUnit)⟩ -instance : LargeCategory.{u} GroupWithZeroCat where +instance : LargeCategory.{u} GrpWithZero where Hom X Y := MonoidWithZeroHom X Y id X := MonoidWithZeroHom.id X comp f g := g.comp f @@ -50,7 +50,7 @@ instance : LargeCategory.{u} GroupWithZeroCat where comp_id := MonoidWithZeroHom.id_comp assoc _ _ _ := MonoidWithZeroHom.comp_assoc _ _ _ -instance {M N : GroupWithZeroCat} : FunLike (M ⟶ N) M N := +instance {M N : GrpWithZero} : FunLike (M ⟶ N) M N := ⟨fun f => f.toFun, fun f g h => by cases f cases g @@ -58,36 +58,36 @@ instance {M N : GroupWithZeroCat} : FunLike (M ⟶ N) M N := apply DFunLike.coe_injective' exact h⟩ -lemma coe_id {X : GroupWithZeroCat} : (𝟙 X : X → X) = id := rfl +lemma coe_id {X : GrpWithZero} : (𝟙 X : X → X) = id := rfl -lemma coe_comp {X Y Z : GroupWithZeroCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl +lemma coe_comp {X Y Z : GrpWithZero} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl -instance groupWithZeroConcreteCategory : ConcreteCategory GroupWithZeroCat where +instance groupWithZeroConcreteCategory : ConcreteCategory GrpWithZero where forget := { obj := fun G => G map := fun f => f.toFun } forget_faithful := ⟨fun h => DFunLike.coe_injective h⟩ -@[simp] lemma forget_map {X Y : GroupWithZeroCat} (f : X ⟶ Y) : - (forget GroupWithZeroCat).map f = f := rfl +@[simp] lemma forget_map {X Y : GrpWithZero} (f : X ⟶ Y) : + (forget GrpWithZero).map f = f := rfl -instance hasForgetToBipointed : HasForget₂ GroupWithZeroCat Bipointed where +instance hasForgetToBipointed : HasForget₂ GrpWithZero Bipointed where forget₂ := { obj := fun X => ⟨X, 0, 1⟩ map := fun f => ⟨f, f.map_zero', f.map_one'⟩ } set_option linter.uppercaseLean3 false in -#align GroupWithZero.has_forget_to_Bipointed GroupWithZeroCat.hasForgetToBipointed +#align GroupWithZero.has_forget_to_Bipointed GrpWithZero.hasForgetToBipointed -instance hasForgetToMon : HasForget₂ GroupWithZeroCat MonCat where +instance hasForgetToMon : HasForget₂ GrpWithZero MonCat where forget₂ := { obj := fun X => ⟨ X , _ ⟩ map := fun f => f.toMonoidHom } set_option linter.uppercaseLean3 false in -#align GroupWithZero.has_forget_to_Mon GroupWithZeroCat.hasForgetToMon +#align GroupWithZero.has_forget_to_Mon GrpWithZero.hasForgetToMon /-- Constructs an isomorphism of groups with zero from a group isomorphism between them. -/ @[simps] -def Iso.mk {α β : GroupWithZeroCat.{u}} (e : α ≃* β) : α ≅ β where +def Iso.mk {α β : GrpWithZero.{u}} (e : α ≃* β) : α ≅ β where hom := (e : α →*₀ β) inv := (e.symm : β →*₀ α) hom_inv_id := by @@ -97,6 +97,6 @@ def Iso.mk {α β : GroupWithZeroCat.{u}} (e : α ≃* β) : α ≅ β where ext exact e.apply_symm_apply _ set_option linter.uppercaseLean3 false in -#align GroupWithZero.iso.mk GroupWithZeroCat.Iso.mk +#align GroupWithZero.iso.mk GrpWithZero.Iso.mk -end GroupWithZeroCat +end GrpWithZero diff --git a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean index dd6f7fc97663f..065582e7c0d7e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Abelian.lean @@ -97,7 +97,7 @@ set_option linter.uppercaseLean3 false in #align Module.forget_reflects_limits_of_size ModuleCat.forgetReflectsLimitsOfSize instance forget₂ReflectsLimitsOfSize : - ReflectsLimitsOfSize.{v, v} (forget₂ (ModuleCatMax.{v, w} R) AddCommGroupCat.{max v w}) := + ReflectsLimitsOfSize.{v, v} (forget₂ (ModuleCatMax.{v, w} R) AddCommGrp.{max v w}) := reflectsLimitsOfReflectsIsomorphisms set_option linter.uppercaseLean3 false in #align Module.forget₂_reflects_limits_of_size ModuleCat.forget₂ReflectsLimitsOfSize @@ -107,7 +107,7 @@ instance forgetReflectsLimits : ReflectsLimits (forget (ModuleCat.{v} R)) := set_option linter.uppercaseLean3 false in #align Module.forget_reflects_limits ModuleCat.forgetReflectsLimits -instance forget₂ReflectsLimits : ReflectsLimits (forget₂ (ModuleCat.{v} R) AddCommGroupCat.{v}) := +instance forget₂ReflectsLimits : ReflectsLimits (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) := ModuleCat.forget₂ReflectsLimitsOfSize.{v, v} set_option linter.uppercaseLean3 false in #align Module.forget₂_reflects_limits ModuleCat.forget₂ReflectsLimits diff --git a/Mathlib/Algebra/Category/ModuleCat/Basic.lean b/Mathlib/Algebra/Category/ModuleCat/Basic.lean index c091c263f021f..7bf36d2ddfcdd 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Basic.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Robert A. Spencer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Robert A. Spencer, Markus Himmel -/ -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.CategoryTheory.Conj import Mathlib.CategoryTheory.Linear.Basic import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor @@ -126,10 +126,10 @@ instance {M : ModuleCat.{v} R} : Module R ((forget (ModuleCat R)).obj M) := lemma ext {M N : ModuleCat.{v} R} {f₁ f₂ : M ⟶ N} (h : ∀ (x : M), f₁ x = f₂ x) : f₁ = f₂ := DFunLike.ext _ _ h -instance hasForgetToAddCommGroup : HasForget₂ (ModuleCat R) AddCommGroupCat where +instance hasForgetToAddCommGroup : HasForget₂ (ModuleCat R) AddCommGrp where forget₂ := - { obj := fun M => AddCommGroupCat.of M - map := fun f => AddCommGroupCat.ofHom f.toAddMonoidHom } + { obj := fun M => AddCommGrp.of M + map := fun f => AddCommGrp.ofHom f.toAddMonoidHom } #align Module.has_forget_to_AddCommGroup ModuleCat.hasForgetToAddCommGroup /-- The object in the category of R-modules associated to an R-module -/ @@ -139,23 +139,23 @@ def of (X : Type v) [AddCommGroup X] [Module R X] : ModuleCat R := @[simp] theorem forget₂_obj (X : ModuleCat R) : - (forget₂ (ModuleCat R) AddCommGroupCat).obj X = AddCommGroupCat.of X := + (forget₂ (ModuleCat R) AddCommGrp).obj X = AddCommGrp.of X := rfl #align Module.forget₂_obj ModuleCat.forget₂_obj -- Porting note: the simpNF linter correctly doesn't like this. -- I'm not sure what this is for, actually. -- If it is really needed, better might be a simp lemma that says --- `AddCommGroupCat.of (ModuleCat.of R X) = AddCommGroupCat.of X`. +-- `AddCommGrp.of (ModuleCat.of R X) = AddCommGrp.of X`. -- @[simp 900] theorem forget₂_obj_moduleCat_of (X : Type v) [AddCommGroup X] [Module R X] : - (forget₂ (ModuleCat R) AddCommGroupCat).obj (of R X) = AddCommGroupCat.of X := + (forget₂ (ModuleCat R) AddCommGrp).obj (of R X) = AddCommGrp.of X := rfl #align Module.forget₂_obj_Module_of ModuleCat.forget₂_obj_moduleCat_of @[simp] theorem forget₂_map (X Y : ModuleCat R) (f : X ⟶ Y) : - (forget₂ (ModuleCat R) AddCommGroupCat).map f = LinearMap.toAddMonoidHom f := + (forget₂ (ModuleCat R) AddCommGrp).map f = LinearMap.toAddMonoidHom f := rfl #align Module.forget₂_map ModuleCat.forget₂_map @@ -322,9 +322,9 @@ instance : Preadditive (ModuleCat.{v} R) where erw [map_add] rfl -instance forget₂_addCommGroupCat_additive : - (forget₂ (ModuleCat.{v} R) AddCommGroupCat).Additive where -#align Module.forget₂_AddCommGroup_additive ModuleCat.forget₂_addCommGroupCat_additive +instance forget₂_addCommGrp_additive : + (forget₂ (ModuleCat.{v} R) AddCommGrp).Additive where +#align Module.forget₂_AddCommGroup_additive ModuleCat.forget₂_addCommGrp_additive section @@ -357,7 +357,7 @@ variable (M N : ModuleCat.{v} R) /-- The scalar multiplication on an object of `ModuleCat R` considered as a morphism of rings from `R` to the endomorphisms of the underlying abelian group. -/ -def smul : R →+* End ((forget₂ (ModuleCat R) AddCommGroupCat).obj M) where +def smul : R →+* End ((forget₂ (ModuleCat R) AddCommGrp).obj M) where toFun r := { toFun := fun (m : M) => r • m map_zero' := by dsimp; rw [smul_zero] @@ -368,17 +368,17 @@ def smul : R →+* End ((forget₂ (ModuleCat R) AddCommGroupCat).obj M) where map_add' r s := AddMonoidHom.ext (fun (x : M) => add_smul r s x) lemma smul_naturality {M N : ModuleCat.{v} R} (f : M ⟶ N) (r : R) : - (forget₂ (ModuleCat R) AddCommGroupCat).map f ≫ N.smul r = - M.smul r ≫ (forget₂ (ModuleCat R) AddCommGroupCat).map f := by + (forget₂ (ModuleCat R) AddCommGrp).map f ≫ N.smul r = + M.smul r ≫ (forget₂ (ModuleCat R) AddCommGrp).map f := by ext x exact (f.map_smul r x).symm variable (R) /-- The scalar multiplication on `ModuleCat R` considered as a morphism of rings -to the endomorphisms of the forgetful functor to `AddCommGroupCat)`. -/ +to the endomorphisms of the forgetful functor to `AddCommGrp)`. -/ @[simps] -def smulNatTrans : R →+* End (forget₂ (ModuleCat R) AddCommGroupCat) where +def smulNatTrans : R →+* End (forget₂ (ModuleCat R) AddCommGrp) where toFun r := { app := fun M => M.smul r naturality := fun _ _ _ => smul_naturality _ r } @@ -389,14 +389,14 @@ def smulNatTrans : R →+* End (forget₂ (ModuleCat R) AddCommGroupCat) where variable {R} -/-- Given `A : AddCommGroupCat` and a ring morphism `R →+* End A`, this is a type synonym +/-- Given `A : AddCommGrp` and a ring morphism `R →+* End A`, this is a type synonym for `A`, on which we shall define a structure of `R`-module. -/ @[nolint unusedArguments] -def mkOfSMul' {A : AddCommGroupCat} (_ : R →+* End A) := A +def mkOfSMul' {A : AddCommGrp} (_ : R →+* End A) := A section -variable {A : AddCommGroupCat} (φ : R →+* End A) +variable {A : AddCommGrp} (φ : R →+* End A) instance : AddCommGroup (mkOfSMul' φ) := by dsimp only [mkOfSMul'] @@ -416,7 +416,7 @@ instance : Module R (mkOfSMul' φ) where add_smul _ _ _ := by simp; rfl zero_smul := by simp -/-- Given `A : AddCommGroupCat` and a ring morphism `R →+* End A`, this is an object in +/-- Given `A : AddCommGrp` and a ring morphism `R →+* End A`, this is an object in `ModuleCat R`, whose underlying abelian group is `A` and whose scalar multiplication is given by `R`. -/ abbrev mkOfSMul := ModuleCat.of R (mkOfSMul' φ) @@ -430,12 +430,12 @@ end section variable {M N} - (φ : (forget₂ (ModuleCat R) AddCommGroupCat).obj M ⟶ - (forget₂ (ModuleCat R) AddCommGroupCat).obj N) + (φ : (forget₂ (ModuleCat R) AddCommGrp).obj M ⟶ + (forget₂ (ModuleCat R) AddCommGrp).obj N) (hφ : ∀ (r : R), φ ≫ N.smul r = M.smul r ≫ φ) /-- Constructor for morphisms in `ModuleCat R` which takes as inputs -a morphism between the underlying objects in `AddCommGroupCat` and the compatibility +a morphism between the underlying objects in `AddCommGrp` and the compatibility with the scalar multiplication. -/ @[simps] def homMk : M ⟶ N where @@ -444,7 +444,7 @@ def homMk : M ⟶ N where map_smul' r x := (congr_hom (hφ r) x).symm lemma forget₂_map_homMk : - (forget₂ (ModuleCat R) AddCommGroupCat).map (homMk φ hφ) = φ := rfl + (forget₂ (ModuleCat R) AddCommGrp).map (homMk φ hφ) = φ := rfl end @@ -454,10 +454,10 @@ instance : (forget (ModuleCat.{v} R)).ReflectsIsomorphisms where (asIso ((forget (ModuleCat R)).map f)).toEquiv.invFun (Equiv.left_inv _) (Equiv.right_inv _)).toModuleIso).hom) -instance : (forget₂ (ModuleCat.{v} R) AddCommGroupCat.{v}).ReflectsIsomorphisms where +instance : (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}).ReflectsIsomorphisms where reflects f _ := by have : IsIso ((forget _).map f) := by - change IsIso ((forget _).map ((forget₂ _ AddCommGroupCat).map f)) + change IsIso ((forget _).map ((forget₂ _ AddCommGrp).map f)) infer_instance apply isIso_of_reflects_iso _ (forget _) diff --git a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean index e98c33b9188dd..6b7ddec2e4d52 100644 --- a/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean +++ b/Mathlib/Algebra/Category/ModuleCat/ChangeOfRings.lean @@ -871,18 +871,18 @@ noncomputable instance preservesLimitRestrictScalars PreservesLimit F (restrictScalars f) := ⟨fun {c} hc => by have : Small.{v} ((F ⋙ restrictScalars f) ⋙ forget _).sections := by assumption - have hc' := isLimitOfPreserves (forget₂ _ AddCommGroupCat) hc - exact isLimitOfReflects (forget₂ _ AddCommGroupCat) hc'⟩ + have hc' := isLimitOfPreserves (forget₂ _ AddCommGrp) hc + exact isLimitOfReflects (forget₂ _ AddCommGrp) hc'⟩ instance preservesColimitRestrictScalars {R S : Type*} [Ring R] [Ring S] (f : R →+* S) {J : Type*} [Category J] (F : J ⥤ ModuleCat.{v} S) - [HasColimit (F ⋙ forget₂ _ AddCommGroupCat)] : + [HasColimit (F ⋙ forget₂ _ AddCommGrp)] : PreservesColimit F (ModuleCat.restrictScalars.{v} f) := by - have : HasColimit ((F ⋙ restrictScalars f) ⋙ forget₂ (ModuleCat R) AddCommGroupCat) := - inferInstanceAs (HasColimit (F ⋙ forget₂ _ AddCommGroupCat)) + have : HasColimit ((F ⋙ restrictScalars f) ⋙ forget₂ (ModuleCat R) AddCommGrp) := + inferInstanceAs (HasColimit (F ⋙ forget₂ _ AddCommGrp)) apply preservesColimitOfPreservesColimitCocone (HasColimit.isColimitColimitCocone F) - apply isColimitOfReflects (forget₂ _ AddCommGroupCat) - apply isColimitOfPreserves (forget₂ (ModuleCat.{v} S) AddCommGroupCat.{v}) + apply isColimitOfReflects (forget₂ _ AddCommGrp) + apply isColimitOfPreserves (forget₂ (ModuleCat.{v} S) AddCommGrp.{v}) exact HasColimit.isColimitColimitCocone F diff --git a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean index 1c452e36296f8..48c6a57ab0541 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Colimits.lean @@ -5,16 +5,16 @@ Authors: Scott Morrison, Joël Riou -/ import Mathlib.Algebra.Category.ModuleCat.Basic import Mathlib.CategoryTheory.ConcreteCategory.Elementwise -import Mathlib.Algebra.Category.GroupCat.Colimits +import Mathlib.Algebra.Category.Grp.Colimits #align_import algebra.category.Module.colimits from "leanprover-community/mathlib"@"5a684ce82399d820475609907c6ef8dba5b1b97c" /-! # The category of R-modules has all colimits. -From the existence of colimits in `AddCommGroupCat`, we deduce the existence of colimits +From the existence of colimits in `AddCommGrp`, we deduce the existence of colimits in `ModuleCat R`. This way, we get for free that the functor -`forget₂ (ModuleCat R) AddCommGroupCat` commutes with colimits. +`forget₂ (ModuleCat R) AddCommGrp` commutes with colimits. Note that finite colimits can already be obtained from the instance `Abelian (Module R)`. @@ -35,13 +35,13 @@ variable {J : Type u} [Category.{v} J] (F : J ⥤ ModuleCat.{w'} R) namespace HasColimit -variable [HasColimit (F ⋙ forget₂ _ AddCommGroupCat)] +variable [HasColimit (F ⋙ forget₂ _ AddCommGrp)] /-- The induced scalar multiplication on -`colimit (F ⋙ forget₂ _ AddCommGroupCat)`. -/ +`colimit (F ⋙ forget₂ _ AddCommGrp)`. -/ @[simps] noncomputable def coconePointSMul : - R →+* End (colimit (F ⋙ forget₂ _ AddCommGroupCat)) where + R →+* End (colimit (F ⋙ forget₂ _ AddCommGrp)) where toFun r := colimMap { app := fun j => (F.obj j).smul r naturality := fun X Y f => smul_naturality _ _ } @@ -54,26 +54,26 @@ noncomputable def coconePointSMul : map_mul' r s := colimit.hom_ext (fun j => by simp) /-- The cocone for `F` constructed from the colimit of -`(F ⋙ forget₂ (ModuleCat R) AddCommGroupCat)`. -/ +`(F ⋙ forget₂ (ModuleCat R) AddCommGrp)`. -/ @[simps] noncomputable def colimitCocone : Cocone F where pt := mkOfSMul (coconePointSMul F) ι := - { app := fun j => homMk (colimit.ι (F ⋙ forget₂ _ AddCommGroupCat) j) (fun r => by + { app := fun j => homMk (colimit.ι (F ⋙ forget₂ _ AddCommGrp) j) (fun r => by dsimp -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [mkOfSMul_smul] simp) naturality := fun i j f => by - apply (forget₂ _ AddCommGroupCat).map_injective + apply (forget₂ _ AddCommGrp).map_injective simp only [Functor.map_comp, forget₂_map_homMk] dsimp - erw [colimit.w (F ⋙ forget₂ _ AddCommGroupCat), comp_id] } + erw [colimit.w (F ⋙ forget₂ _ AddCommGrp), comp_id] } /-- The cocone for `F` constructed from the colimit of -`(F ⋙ forget₂ (ModuleCat R) AddCommGroupCat)` is a colimit cocone. -/ +`(F ⋙ forget₂ (ModuleCat R) AddCommGrp)` is a colimit cocone. -/ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) where - desc s := homMk (colimit.desc _ ((forget₂ _ AddCommGroupCat).mapCocone s)) (fun r => by + desc s := homMk (colimit.desc _ ((forget₂ _ AddCommGrp).mapCocone s)) (fun r => by apply colimit.hom_ext intro j dsimp @@ -85,50 +85,50 @@ noncomputable def isColimitColimitCocone : IsColimit (colimitCocone F) where Functor.mapCocone_pt, Functor.mapCocone_ι_app, forget₂_map] exact smul_naturality (s.ι.app j) r) fac s j := by - apply (forget₂ _ AddCommGroupCat).map_injective - exact colimit.ι_desc ((forget₂ _ AddCommGroupCat).mapCocone s) j + apply (forget₂ _ AddCommGrp).map_injective + exact colimit.ι_desc ((forget₂ _ AddCommGrp).mapCocone s) j uniq s m hm := by - apply (forget₂ _ AddCommGroupCat).map_injective + apply (forget₂ _ AddCommGrp).map_injective apply colimit.hom_ext intro j - erw [colimit.ι_desc ((forget₂ _ AddCommGroupCat).mapCocone s) j] + erw [colimit.ι_desc ((forget₂ _ AddCommGrp).mapCocone s) j] dsimp rw [← hm] rfl instance : HasColimit F := ⟨_, isColimitColimitCocone F⟩ -noncomputable instance : PreservesColimit F (forget₂ _ AddCommGroupCat) := +noncomputable instance : PreservesColimit F (forget₂ _ AddCommGrp) := preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) noncomputable instance reflectsColimit : - ReflectsColimit F (forget₂ (ModuleCat.{w'} R) AddCommGroupCat) := + ReflectsColimit F (forget₂ (ModuleCat.{w'} R) AddCommGrp) := reflectsColimitOfReflectsIsomorphisms _ _ end HasColimit variable (J R) -instance hasColimitsOfShape [HasColimitsOfShape J AddCommGroupCat.{w'}] : +instance hasColimitsOfShape [HasColimitsOfShape J AddCommGrp.{w'}] : HasColimitsOfShape J (ModuleCat.{w'} R) where -noncomputable instance reflectsColimitsOfShape [HasColimitsOfShape J AddCommGroupCat.{w'}] : - ReflectsColimitsOfShape J (forget₂ (ModuleCat.{w'} R) AddCommGroupCat) where +noncomputable instance reflectsColimitsOfShape [HasColimitsOfShape J AddCommGrp.{w'}] : + ReflectsColimitsOfShape J (forget₂ (ModuleCat.{w'} R) AddCommGrp) where -instance hasColimitsOfSize [HasColimitsOfSize.{v, u} AddCommGroupCat.{w'}] : +instance hasColimitsOfSize [HasColimitsOfSize.{v, u} AddCommGrp.{w'}] : HasColimitsOfSize.{v, u} (ModuleCat.{w'} R) where noncomputable instance forget₂PreservesColimitsOfShape - [HasColimitsOfShape J AddCommGroupCat.{w'}] : - PreservesColimitsOfShape J (forget₂ (ModuleCat.{w'} R) AddCommGroupCat) where + [HasColimitsOfShape J AddCommGrp.{w'}] : + PreservesColimitsOfShape J (forget₂ (ModuleCat.{w'} R) AddCommGrp) where noncomputable instance forget₂PreservesColimitsOfSize - [HasColimitsOfSize.{u, v} AddCommGroupCat.{w'}] : - PreservesColimitsOfSize.{u, v} (forget₂ (ModuleCat.{w'} R) AddCommGroupCat) where + [HasColimitsOfSize.{u, v} AddCommGrp.{w'}] : + PreservesColimitsOfSize.{u, v} (forget₂ (ModuleCat.{w'} R) AddCommGrp) where noncomputable instance - [HasColimitsOfSize.{u, v} AddCommGroupCatMax.{w, w'}] : - PreservesColimitsOfSize.{u, v} (forget₂ (ModuleCatMax.{w, w'} R) AddCommGroupCat) where + [HasColimitsOfSize.{u, v} AddCommGrpMax.{w, w'}] : + PreservesColimitsOfSize.{u, v} (forget₂ (ModuleCatMax.{w, w'} R) AddCommGrp) where instance : HasFiniteColimits (ModuleCat.{w'} R) := inferInstance @@ -149,6 +149,6 @@ example (R : Type u) [Ring R] : HasCoequalizers (ModuleCat.{u} R) := by instance : HasCoequalizers (ModuleCat.{v} R) where noncomputable example (R : Type u) [Ring R] : - PreservesColimits (forget₂ (ModuleCat.{u} R) AddCommGroupCat) := inferInstance + PreservesColimits (forget₂ (ModuleCat.{u} R) AddCommGrp) := inferInstance end ModuleCat diff --git a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean index b8df1d3336a85..1c60ce8a61055 100644 --- a/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/FilteredColimits.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Justus Springer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Justus Springer -/ -import Mathlib.Algebra.Category.GroupCat.FilteredColimits +import Mathlib.Algebra.Category.Grp.FilteredColimits import Mathlib.Algebra.Category.ModuleCat.Basic #align_import algebra.category.Module.filtered_colimits from "leanprover-community/mathlib"@"806bbb0132ba63b93d5edbe4789ea226f8329979" @@ -15,9 +15,9 @@ Forgetful functors from algebraic categories usually don't preserve colimits. Ho to preserve _filtered_ colimits. In this file, we start with a ring `R`, a small filtered category `J` and a functor -`F : J ⥤ ModuleCat R`. We show that the colimit of `F ⋙ forget₂ (ModuleCat R) AddCommGroupCat` -(in `AddCommGroupCat`) carries the structure of an `R`-module, thereby showing that the forgetful -functor `forget₂ (ModuleCat R) AddCommGroupCat` preserves filtered colimits. In particular, this +`F : J ⥤ ModuleCat R`. We show that the colimit of `F ⋙ forget₂ (ModuleCat R) AddCommGrp` +(in `AddCommGrp`) carries the structure of an `R`-module, thereby showing that the forgetful +functor `forget₂ (ModuleCat R) AddCommGrp` preserves filtered colimits. In particular, this implies that `forget (ModuleCat R)` preserves filtered colimits. -/ @@ -42,12 +42,12 @@ section variable {R : Type u} [Ring R] {J : Type v} [SmallCategory J] [IsFiltered J] variable (F : J ⥤ ModuleCatMax.{v, u, u} R) -/-- The colimit of `F ⋙ forget₂ (ModuleCat R) AddCommGroupCat` in the category `AddCommGroupCat`. +/-- The colimit of `F ⋙ forget₂ (ModuleCat R) AddCommGrp` in the category `AddCommGrp`. In the following, we will show that this has the structure of an `R`-module. -/ -abbrev M : AddCommGroupCat := - AddCommGroupCat.FilteredColimits.colimit.{v, u} - (F ⋙ forget₂ (ModuleCat R) AddCommGroupCat.{max v u}) +abbrev M : AddCommGrp := + AddCommGrp.FilteredColimits.colimit.{v, u} + (F ⋙ forget₂ (ModuleCat R) AddCommGrp.{max v u}) set_option linter.uppercaseLean3 false in #align Module.filtered_colimits.M ModuleCat.FilteredColimits.M @@ -155,8 +155,8 @@ set_option linter.uppercaseLean3 false in /-- The linear map from a given `R`-module in the diagram to the colimit module. -/ def coconeMorphism (j : J) : F.obj j ⟶ colimit F := - { (AddCommGroupCat.FilteredColimits.colimitCocone - (F ⋙ forget₂ (ModuleCat R) AddCommGroupCat.{max v u})).ι.app j with + { (AddCommGrp.FilteredColimits.colimitCocone + (F ⋙ forget₂ (ModuleCat R) AddCommGrp.{max v u})).ι.app j with map_smul' := fun r x => by erw [colimit_smul_mk_eq F r ⟨j, x⟩]; rfl } set_option linter.uppercaseLean3 false in #align Module.filtered_colimits.cocone_morphism ModuleCat.FilteredColimits.coconeMorphism @@ -177,9 +177,9 @@ We already know that this is a morphism between additive groups. The only thing it is a linear map, i.e. preserves scalar multiplication. -/ def colimitDesc (t : Cocone F) : colimit F ⟶ t.pt := - { (AddCommGroupCat.FilteredColimits.colimitCoconeIsColimit - (F ⋙ forget₂ (ModuleCatMax.{v, u} R) AddCommGroupCat.{max v u})).desc - ((forget₂ (ModuleCat R) AddCommGroupCat.{max v u}).mapCocone t) with + { (AddCommGrp.FilteredColimits.colimitCoconeIsColimit + (F ⋙ forget₂ (ModuleCatMax.{v, u} R) AddCommGrp.{max v u})).desc + ((forget₂ (ModuleCat R) AddCommGrp.{max v u}).mapCocone t) with map_smul' := fun r x => by refine Quot.inductionOn x ?_; clear x; intro x; cases' x with j x erw [colimit_smul_mk_eq] @@ -202,20 +202,20 @@ set_option linter.uppercaseLean3 false in #align Module.filtered_colimits.colimit_cocone_is_colimit ModuleCat.FilteredColimits.colimitCoconeIsColimit instance forget₂AddCommGroupPreservesFilteredColimits : - PreservesFilteredColimits (forget₂ (ModuleCat.{u} R) AddCommGroupCat.{u}) where + PreservesFilteredColimits (forget₂ (ModuleCat.{u} R) AddCommGrp.{u}) where preserves_filtered_colimits J _ _ := { -- Porting note: without the curly braces for `F` -- here we get a confusing error message about universes. preservesColimit := fun {F : J ⥤ ModuleCat.{u} R} => preservesColimitOfPreservesColimitCocone (colimitCoconeIsColimit F) - (AddCommGroupCat.FilteredColimits.colimitCoconeIsColimit - (F ⋙ forget₂ (ModuleCat.{u} R) AddCommGroupCat.{u})) } + (AddCommGrp.FilteredColimits.colimitCoconeIsColimit + (F ⋙ forget₂ (ModuleCat.{u} R) AddCommGrp.{u})) } set_option linter.uppercaseLean3 false in #align Module.filtered_colimits.forget₂_AddCommGroup_preserves_filtered_colimits ModuleCat.FilteredColimits.forget₂AddCommGroupPreservesFilteredColimits instance forgetPreservesFilteredColimits : PreservesFilteredColimits (forget (ModuleCat.{u} R)) := - Limits.compPreservesFilteredColimits (forget₂ (ModuleCat R) AddCommGroupCat) - (forget AddCommGroupCat) + Limits.compPreservesFilteredColimits (forget₂ (ModuleCat R) AddCommGrp) + (forget AddCommGrp) set_option linter.uppercaseLean3 false in #align Module.filtered_colimits.forget_preserves_filtered_colimits ModuleCat.FilteredColimits.forgetPreservesFilteredColimits diff --git a/Mathlib/Algebra/Category/ModuleCat/Injective.lean b/Mathlib/Algebra/Category/ModuleCat/Injective.lean index 6b55d3dc07698..ab64c32a4c6ea 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Injective.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Injective.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang -/ import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings -import Mathlib.Algebra.Category.GroupCat.EnoughInjectives -import Mathlib.Algebra.Category.GroupCat.ZModuleEquivalence +import Mathlib.Algebra.Category.Grp.EnoughInjectives +import Mathlib.Algebra.Category.Grp.ZModuleEquivalence import Mathlib.Logic.Equiv.TransferInstance /-! @@ -26,7 +26,7 @@ universe v u variable (R : Type u) [Ring R] instance : EnoughInjectives (ModuleCat.{v} ℤ) := - EnoughInjectives.of_equivalence (forget₂ (ModuleCat ℤ) AddCommGroupCat) + EnoughInjectives.of_equivalence (forget₂ (ModuleCat ℤ) AddCommGrp) lemma ModuleCat.enoughInjectives : EnoughInjectives (ModuleCat.{max v u} R) := EnoughInjectives.of_adjunction (ModuleCat.restrictCoextendScalarsAdj.{max v u} (algebraMap ℤ R)) diff --git a/Mathlib/Algebra/Category/ModuleCat/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Limits.lean index e740d61ec51bd..49d8bad6e20ec 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Limits.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import Mathlib.Algebra.Category.ModuleCat.Basic -import Mathlib.Algebra.Category.GroupCat.Limits +import Mathlib.Algebra.Category.Grp.Limits import Mathlib.Algebra.DirectLimit #align_import algebra.category.Module.limits from "leanprover-community/mathlib"@"c43486ecf2a5a17479a32ce09e4818924145e90e" @@ -45,9 +45,9 @@ instance moduleObj (j) : /-- The flat sections of a functor into `ModuleCat R` form a submodule of all sections. -/ def sectionsSubmodule : Submodule R (∀ j, F.obj j) := - { AddGroupCat.sectionsAddSubgroup.{v, w} - (F ⋙ forget₂ (ModuleCat R) AddCommGroupCat.{w} ⋙ - forget₂ AddCommGroupCat AddGroupCat.{w}) with + { AddGrp.sectionsAddSubgroup.{v, w} + (F ⋙ forget₂ (ModuleCat R) AddCommGrp.{w} ⋙ + forget₂ AddCommGrp AddGrp.{w}) with carrier := (F ⋙ forget (ModuleCat R)).sections smul_mem' := fun r s sh j j' f => by simp only [forget_map, Functor.comp_map, Pi.smul_apply, map_smul] @@ -162,33 +162,33 @@ instance (priority := high) hasLimits' : HasLimits (ModuleCat.{u} R) := /-- An auxiliary declaration to speed up typechecking. -/ def forget₂AddCommGroupPreservesLimitsAux : - IsLimit ((forget₂ (ModuleCat R) AddCommGroupCat).mapCone (limitCone F)) := - letI : Small.{w} (Functor.sections ((F ⋙ forget₂ _ AddCommGroupCat) ⋙ forget _)) := + IsLimit ((forget₂ (ModuleCat R) AddCommGrp).mapCone (limitCone F)) := + letI : Small.{w} (Functor.sections ((F ⋙ forget₂ _ AddCommGrp) ⋙ forget _)) := inferInstanceAs <| Small.{w} (Functor.sections (F ⋙ forget (ModuleCat R))) - AddCommGroupCat.limitConeIsLimit - (F ⋙ forget₂ (ModuleCat.{w} R) _ : J ⥤ AddCommGroupCat.{w}) + AddCommGrp.limitConeIsLimit + (F ⋙ forget₂ (ModuleCat.{w} R) _ : J ⥤ AddCommGrp.{w}) #align Module.forget₂_AddCommGroup_preserves_limits_aux ModuleCat.forget₂AddCommGroupPreservesLimitsAux /-- The forgetful functor from R-modules to abelian groups preserves all limits. -/ instance forget₂AddCommGroupPreservesLimit : - PreservesLimit F (forget₂ (ModuleCat R) AddCommGroupCat) := + PreservesLimit F (forget₂ (ModuleCat R) AddCommGrp) := preservesLimitOfPreservesLimitCone (limitConeIsLimit F) (forget₂AddCommGroupPreservesLimitsAux F) instance forget₂AddCommGroupReflectsLimit : - ReflectsLimit F (forget₂ (ModuleCat R) AddCommGroupCat) := + ReflectsLimit F (forget₂ (ModuleCat R) AddCommGrp) := reflectsLimitOfReflectsIsomorphisms _ _ /-- The forgetful functor from R-modules to abelian groups preserves all limits. -/ instance forget₂AddCommGroupPreservesLimitsOfSize [UnivLE.{v, w}] : PreservesLimitsOfSize.{t, v} - (forget₂ (ModuleCat.{w} R) AddCommGroupCat.{w}) where + (forget₂ (ModuleCat.{w} R) AddCommGrp.{w}) where preservesLimitsOfShape := { preservesLimit := inferInstance } #align Module.forget₂_AddCommGroup_preserves_limits_of_size ModuleCat.forget₂AddCommGroupPreservesLimitsOfSize instance forget₂AddCommGroupPreservesLimits : - PreservesLimits (forget₂ (ModuleCat R) AddCommGroupCat.{w}) := + PreservesLimits (forget₂ (ModuleCat R) AddCommGrp.{w}) := ModuleCat.forget₂AddCommGroupPreservesLimitsOfSize.{w, w} #align Module.forget₂_AddCommGroup_preserves_limits ModuleCat.forget₂AddCommGroupPreservesLimits diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean index cbe7ea536e274..bd4c9dcf6327b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf.lean @@ -39,7 +39,7 @@ variable {C : Type u₁} [Category.{v₁} C] described as a presheaf of abelian groups, and the extra data of the action at each object, and a condition relating functoriality and scalar multiplication. -/ structure PresheafOfModules (R : Cᵒᵖ ⥤ RingCat.{u}) where - presheaf : Cᵒᵖ ⥤ AddCommGroupCat.{v} + presheaf : Cᵒᵖ ⥤ AddCommGrp.{v} module : ∀ X : Cᵒᵖ, Module (R.obj X) (presheaf.obj X) := by infer_instance map_smul : ∀ {X Y : Cᵒᵖ} (f : X ⟶ Y) (r : R.obj X) (x : presheaf.obj X), presheaf.map f (r • x) = R.map f r • presheaf.map f x := by aesop_cat @@ -153,7 +153,7 @@ variable {P Q} instance : Add (P ⟶ Q) := ⟨fun f g => mk (f.hom + g.hom) (by intros - simp only [NatTrans.app_add, AddCommGroupCat.hom_add_apply, map_smul, smul_add])⟩ + simp only [NatTrans.app_add, AddCommGrp.hom_add_apply, map_smul, smul_add])⟩ @[simp] lemma add_app (f g : P ⟶ Q) (X : Cᵒᵖ) : (f + g).app X = f.app X + g.app X := rfl @@ -201,7 +201,7 @@ variable (R) to presheaves of abelian groups. -/ @[simps obj] -def toPresheaf : PresheafOfModules.{v} R ⥤ (Cᵒᵖ ⥤ AddCommGroupCat.{v}) where +def toPresheaf : PresheafOfModules.{v} R ⥤ (Cᵒᵖ ⥤ AddCommGrp.{v}) where obj P := P.presheaf map f := f.hom @@ -349,9 +349,9 @@ variable (M : CorePresheafOfModules R) /-- The presheaf of abelian groups attached to a `CorePresheafOfModules R`. -/ @[simps] -def presheaf : Cᵒᵖ ⥤ AddCommGroupCat.{v} where - obj X := AddCommGroupCat.of (M.obj X) - map f := AddCommGroupCat.ofHom (M.map f).toAddMonoidHom +def presheaf : Cᵒᵖ ⥤ AddCommGrp.{v} where + obj X := AddCommGrp.of (M.obj X) + map f := AddCommGrp.ofHom (M.map f).toAddMonoidHom instance (X : Cᵒᵖ) : Module (R.obj X) (M.presheaf.obj X) := M.module X diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean index fefcc51f4fd18..ac5a7ef32dde8 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Colimits.lean @@ -116,13 +116,13 @@ noncomputable instance evaluationPreservesColimit (X : Cᵒᵖ) : preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (colimit.isColimit _) variable [∀ X, PreservesColimit F - (evaluation R X ⋙ forget₂ (ModuleCat (R.obj X)) AddCommGroupCat)] + (evaluation R X ⋙ forget₂ (ModuleCat (R.obj X)) AddCommGrp)] noncomputable instance toPresheafPreservesColimit : PreservesColimit F (toPresheaf R) := preservesColimitOfPreservesColimitCocone (isColimitColimitCocone F) (Limits.evaluationJointlyReflectsColimits _ - (fun X => isColimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGroupCat) + (fun X => isColimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGrp) (isColimitColimitCocone F))) end Colimits @@ -131,7 +131,7 @@ variable (R J) section HasColimitsOfShape -variable [HasColimitsOfShape J AddCommGroupCat.{v}] +variable [HasColimitsOfShape J AddCommGrp.{v}] instance hasColimitsOfShape : HasColimitsOfShape J (PresheafOfModules.{v} R) where diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean index 5841f5dda1606..4fd891084e89b 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Limits.lean @@ -128,7 +128,7 @@ noncomputable instance toPresheafPreservesLimit : PreservesLimit F (toPresheaf R) := preservesLimitOfPreservesLimitCone (isLimitLimitCone F) (Limits.evaluationJointlyReflectsLimits _ - (fun X => isLimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGroupCat) + (fun X => isLimitOfPreserves (evaluation R X ⋙ forget₂ _ AddCommGrp) (isLimitLimitCone F))) end Limits diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean index ff9e7410500f4..8ec02b10ada53 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafification.lean @@ -29,13 +29,13 @@ open CategoryTheory Category Limits variable {C : Type u'} [Category.{v'} C] {J : GrothendieckTopology C} {R₀ : Cᵒᵖ ⥤ RingCat.{u}} {R : Sheaf J RingCat.{u}} (α : R₀ ⟶ R.val) [Presheaf.IsLocallyInjective J α] [Presheaf.IsLocallySurjective J α] - [J.WEqualsLocallyBijective AddCommGroupCat.{v}] + [J.WEqualsLocallyBijective AddCommGrp.{v}] namespace PresheafOfModules section -variable [HasWeakSheafify J AddCommGroupCat.{v}] +variable [HasWeakSheafify J AddCommGrp.{v}] /-- Given a locally bijective morphism `α : R₀ ⟶ R.val` where `R₀` is a presheaf of rings and `R` a sheaf of rings (i.e. `R` identifies to the sheafification of `R₀`), this is @@ -43,7 +43,7 @@ the associated sheaf of modules functor `PresheafOfModules.{v} R₀ ⥤ SheafOfM @[simps! (config := .lemmasOnly) map] noncomputable def sheafification : PresheafOfModules.{v} R₀ ⥤ SheafOfModules.{v} R where obj M₀ := sheafify α (CategoryTheory.toSheafify J M₀.presheaf) - map f := sheafifyMap _ _ _ f ((presheafToSheaf J AddCommGroupCat).map f.hom) (by simp) + map f := sheafifyMap _ _ _ f ((presheafToSheaf J AddCommGrp).map f.hom) (by simp) map_id M₀ := by ext1 apply (toPresheaf _).map_injective @@ -57,14 +57,14 @@ noncomputable def sheafification : PresheafOfModules.{v} R₀ ⥤ SheafOfModules forgets the module structures. -/ noncomputable def sheafificationCompToSheaf : sheafification.{v} α ⋙ SheafOfModules.toSheaf _ ≅ - toPresheaf _ ⋙ presheafToSheaf J AddCommGroupCat := + toPresheaf _ ⋙ presheafToSheaf J AddCommGrp := Iso.refl _ /-- The sheafification of presheaves of modules commutes with the functor which forgets the module structures. -/ noncomputable def sheafificationCompForgetCompToPresheaf : sheafification.{v} α ⋙ SheafOfModules.forget _ ⋙ toPresheaf _ ≅ - toPresheaf _ ⋙ presheafToSheaf J AddCommGroupCat ⋙ sheafToPresheaf J AddCommGroupCat := + toPresheaf _ ⋙ presheafToSheaf J AddCommGrp ⋙ sheafToPresheaf J AddCommGrp := Iso.refl _ /-- The bijection between types of morphisms which is part of the adjunction @@ -85,7 +85,7 @@ lemma sheafificationHomEquiv_hom {P : PresheafOfModules.{v} R₀} {F : SheafOfModules.{v} R} (f : (sheafification α).obj P ⟶ F) : (sheafificationHomEquiv α f).hom = - (sheafificationAdjunction J AddCommGroupCat).homEquiv P.presheaf + (sheafificationAdjunction J AddCommGrp).homEquiv P.presheaf ((SheafOfModules.toSheaf _).obj F) ((SheafOfModules.toSheaf _).map f) := by rw [sheafificationHomEquiv_hom', Adjunction.homEquiv_unit] dsimp @@ -94,10 +94,10 @@ lemma toSheaf_map_sheafificationHomEquiv_symm {P : PresheafOfModules.{v} R₀} {F : SheafOfModules.{v} R} (g : P ⟶ (restrictScalars α).obj ((SheafOfModules.forget _).obj F)) : (SheafOfModules.toSheaf _).map ((sheafificationHomEquiv α).symm g) = - (((sheafificationAdjunction J AddCommGroupCat).homEquiv + (((sheafificationAdjunction J AddCommGrp).homEquiv P.presheaf ((SheafOfModules.toSheaf R).obj F)).symm g.hom) := by obtain ⟨f, rfl⟩ := (sheafificationHomEquiv α).surjective g - apply ((sheafificationAdjunction J AddCommGroupCat).homEquiv _ _).injective + apply ((sheafificationAdjunction J AddCommGrp).homEquiv _ _).injective rw [Equiv.apply_symm_apply, Adjunction.homEquiv_unit, Equiv.symm_apply_apply] rfl @@ -134,11 +134,11 @@ end section -variable [HasSheafify J AddCommGroupCat.{v}] +variable [HasSheafify J AddCommGrp.{v}] noncomputable instance : PreservesFiniteLimits (sheafification.{v} α ⋙ SheafOfModules.toSheaf.{v} R) := - compPreservesFiniteLimits (toPresheaf.{v} R₀) (presheafToSheaf J AddCommGroupCat) + compPreservesFiniteLimits (toPresheaf.{v} R₀) (presheafToSheaf J AddCommGrp) instance : (SheafOfModules.toSheaf.{v} R ⋙ sheafToPresheaf _ _).ReflectsIsomorphisms := inferInstanceAs (SheafOfModules.forget.{v} R ⋙ toPresheaf _).ReflectsIsomorphisms diff --git a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean index b1c675fb578d7..07f18f8c8fc8e 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Presheaf/Sheafify.lean @@ -49,7 +49,7 @@ end smul section variable {R₀ R : Cᵒᵖ ⥤ RingCat.{u}} (α : R₀ ⟶ R) [Presheaf.IsLocallyInjective J α] - {M₀ : PresheafOfModules.{v} R₀} {A : Cᵒᵖ ⥤ AddCommGroupCat.{v}} (φ : M₀.presheaf ⟶ A) + {M₀ : PresheafOfModules.{v} R₀} {A : Cᵒᵖ ⥤ AddCommGrp.{v}} (φ : M₀.presheaf ⟶ A) [Presheaf.IsLocallyInjective J φ] (hA : Presheaf.IsSeparated J A) {X : C} (r : R.obj (Opposite.op X)) (m : A.obj (Opposite.op X)) {P : Presieve X} (r₀ : FamilyOfElements (R₀ ⋙ forget _) P) (m₀ : FamilyOfElements (M₀.presheaf ⋙ forget _) P) @@ -80,7 +80,7 @@ lemma isCompatible_map_smul_aux {Y Z : C} (f : Y ⟶ X) (g : Z ⟶ Y) rw [← PresheafOfModules.Sheafify.app_eq_of_isLocallyInjective α φ hA (R₀.map g.op r₀) r₀' (M₀.presheaf.map g.op m₀) m₀', M₀.map_smul] · rw [hr₀', R.map_comp, comp_apply, ← hr₀, NatTrans.naturality_apply] - · rw [hm₀', A.map_comp, AddCommGroupCat.coe_comp, Function.comp_apply, ← hm₀] + · rw [hm₀', A.map_comp, AddCommGrp.coe_comp, Function.comp_apply, ← hm₀] erw [NatTrans.naturality_apply] rfl @@ -123,7 +123,7 @@ variable {R₀ : Cᵒᵖ ⥤ RingCat.{u}} {R : Sheaf J RingCat.{u}} (α : R₀ namespace PresheafOfModules -variable {M₀ : PresheafOfModules.{v} R₀} {A : Sheaf J AddCommGroupCat.{v}} +variable {M₀ : PresheafOfModules.{v} R₀} {A : Sheaf J AddCommGrp.{v}} (φ : M₀.presheaf ⟶ A.val) [Presheaf.IsLocallyInjective J φ] [Presheaf.IsLocallySurjective J φ] @@ -326,7 +326,7 @@ instance : Presheaf.IsLocallyInjective J (toSheafify α φ).hom := by instance : Presheaf.IsLocallySurjective J (toSheafify α φ).hom := by dsimp; infer_instance -variable [J.WEqualsLocallyBijective AddCommGroupCat.{v}] +variable [J.WEqualsLocallyBijective AddCommGrp.{v}] /-- The bijection `((sheafify α φ).val ⟶ F) ≃ (M₀ ⟶ (restrictScalars α).obj F)` which is part of the universal property of the sheafification of the presheaf of modules `M₀`, @@ -355,7 +355,7 @@ noncomputable def sheafifyHomEquiv {F : SheafOfModules.{v} R} : section -variable {M₀' : PresheafOfModules.{v} R₀} {A' : Sheaf J AddCommGroupCat.{v}} +variable {M₀' : PresheafOfModules.{v} R₀} {A' : Sheaf J AddCommGrp.{v}} (φ' : M₀'.presheaf ⟶ A'.val) [Presheaf.IsLocallyInjective J φ'] [Presheaf.IsLocallySurjective J φ'] (τ₀ : M₀ ⟶ M₀') (τ : A ⟶ A') diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean index 6407817e043ff..9916606e3b760 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf.lean @@ -86,17 +86,17 @@ instance : (forget.{v} R).ReflectsIsomorphisms := (fullyFaithfulForget R).reflec def evaluation (X : Cᵒᵖ) : SheafOfModules.{v} R ⥤ ModuleCat.{v} (R.val.obj X) := forget _ ⋙ PresheafOfModules.evaluation _ X -/-- The forget functor `SheafOfModules R ⥤ Sheaf J AddCommGroupCat`. -/ +/-- The forget functor `SheafOfModules R ⥤ Sheaf J AddCommGrp`. -/ @[simps] -def toSheaf : SheafOfModules.{v} R ⥤ Sheaf J AddCommGroupCat.{v} where +def toSheaf : SheafOfModules.{v} R ⥤ Sheaf J AddCommGrp.{v} where obj M := ⟨_, M.isSheaf⟩ map f := { val := f.val.hom } /-- The canonical isomorphism between -`SheafOfModules.toSheaf R ⋙ sheafToPresheaf J AddCommGroupCat.{v}` +`SheafOfModules.toSheaf R ⋙ sheafToPresheaf J AddCommGrp.{v}` and `SheafOfModules.forget R ⋙ PresheafOfModules.toPresheaf R.val`. -/ def toSheafCompSheafToPresheafIso : - toSheaf R ⋙ sheafToPresheaf J AddCommGroupCat.{v} ≅ + toSheaf R ⋙ sheafToPresheaf J AddCommGrp.{v} ≅ forget R ⋙ PresheafOfModules.toPresheaf R.val := Iso.refl _ instance : (toSheaf.{v} R).Faithful := @@ -121,13 +121,13 @@ instance : (toSheaf R).Additive where /-- The type of sections of a sheaf of modules. -/ abbrev sections (M : SheafOfModules.{v} R) : Type _ := M.val.sections -variable [J.HasSheafCompose (forget₂ RingCat.{u} AddCommGroupCat.{u})] +variable [J.HasSheafCompose (forget₂ RingCat.{u} AddCommGrp.{u})] /-- The obvious free sheaf of modules of rank `1`. -/ @[simps] def unit : SheafOfModules R where val := PresheafOfModules.unit R.val - isSheaf := ((sheafCompose J (forget₂ RingCat.{u} AddCommGroupCat.{u})).obj R).cond + isSheaf := ((sheafCompose J (forget₂ RingCat.{u} AddCommGrp.{u})).obj R).cond variable {R} @@ -147,7 +147,7 @@ namespace PresheafOfModules variable {R : Cᵒᵖ ⥤ RingCat.{u}} {M₁ M₂ : PresheafOfModules.{v} R} (f : M₁ ⟶ M₂) {N : PresheafOfModules.{v} R} (hN : Presheaf.IsSheaf J N.presheaf) - [J.WEqualsLocallyBijective AddCommGroupCat.{v}] + [J.WEqualsLocallyBijective AddCommGrp.{v}] [Presheaf.IsLocallySurjective J f.hom] [Presheaf.IsLocallyInjective J f.hom] diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Abelian.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Abelian.lean index 72a67e3bd73d6..a69cc55e87222 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Abelian.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Abelian.lean @@ -12,12 +12,12 @@ import Mathlib.CategoryTheory.Abelian.Transfer In this file, it is shown that the category of sheaves of modules over a sheaf of rings `R` is an abelian category. More precisely, if `J` is a Grothendieck topology on a category `C` and `R : Sheaf J RingCat.{u}`, -then `SheafOfModules.{v} R` is abelian if the conditions `HasSheafify J AddCommGroupCat.{v}]` -and `J.WEqualsLocallyBijective AddCommGroupCat.{v}` are satisfied. +then `SheafOfModules.{v} R` is abelian if the conditions `HasSheafify J AddCommGrp.{v}]` +and `J.WEqualsLocallyBijective AddCommGrp.{v}` are satisfied. In particular, if `u = v` and `C : Type u` is a small category, then `SheafOfModules.{u} R` is abelian: this instance shall be -found automatically if this file and `Algebra.Category.GroupCat.FilteredColimits` +found automatically if this file and `Algebra.Category.Grp.FilteredColimits` are imported. -/ @@ -30,8 +30,8 @@ variable {C : Type u'} [Category.{v'} C] {J : GrothendieckTopology C} namespace SheafOfModules -variable (R : Sheaf J RingCat.{u}) [HasSheafify J AddCommGroupCat.{v}] - [J.WEqualsLocallyBijective AddCommGroupCat.{v}] +variable (R : Sheaf J RingCat.{u}) [HasSheafify J AddCommGrp.{v}] + [J.WEqualsLocallyBijective AddCommGrp.{v}] noncomputable instance : Abelian (SheafOfModules.{v} R) := by let adj := PresheafOfModules.sheafificationAdjunction (𝟙 R.val) diff --git a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean index 2ab4f3bbfb88c..bace928c718ea 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Sheaf/Limits.lean @@ -32,10 +32,10 @@ variable {R : Cᵒᵖ ⥤ RingCat.{u}} [∀ X, Small.{v} ((F ⋙ evaluation R X) ⋙ forget _).sections] {c : Cone F} (hc : IsLimit c) (hF : ∀ j, Presheaf.IsSheaf J (F.obj j).presheaf) - [HasLimitsOfShape D AddCommGroupCat.{v}] + [HasLimitsOfShape D AddCommGrp.{v}] lemma isSheaf_of_isLimit : Presheaf.IsSheaf J (c.pt.presheaf) := by - let G : D ⥤ Sheaf J AddCommGroupCat.{v} := + let G : D ⥤ Sheaf J AddCommGrp.{v} := { obj := fun j => ⟨(F.obj j).presheaf, hF j⟩ map := fun φ => ⟨(PresheafOfModules.toPresheaf R).map (F.map φ)⟩ } exact Sheaf.isSheaf_of_isLimit G _ (isLimitOfPreserves (toPresheaf R) hc) @@ -50,7 +50,7 @@ section Limits variable (F : D ⥤ SheafOfModules.{v} R) [∀ X, Small.{v} ((F ⋙ evaluation R X) ⋙ CategoryTheory.forget _).sections] - [HasLimitsOfShape D AddCommGroupCat.{v}] + [HasLimitsOfShape D AddCommGrp.{v}] instance (X : Cᵒᵖ) : Small.{v} (((F ⋙ forget _) ⋙ PresheafOfModules.evaluation _ X) ⋙ CategoryTheory.forget _).sections := by diff --git a/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean b/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean index 7b005e8e11076..abff69661e7bf 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Tannaka.lean @@ -25,10 +25,10 @@ A ring `R` is equivalent to the endomorphisms of the additive forgetful functor `Module R ⥤ AddCommGroup`. -/ def ringEquivEndForget₂ (R : Type u) [Ring R] : - R ≃+* End (AdditiveFunctor.of (forget₂ (ModuleCat.{u} R) AddCommGroupCat.{u})) where + R ≃+* End (AdditiveFunctor.of (forget₂ (ModuleCat.{u} R) AddCommGrp.{u})) where toFun r := { app := fun M => - @AddCommGroupCat.ofHom M.carrier M.carrier _ _ (DistribMulAction.toAddMonoidHom M r) + @AddCommGrp.ofHom M.carrier M.carrier _ _ (DistribMulAction.toAddMonoidHom M r) naturality := fun M N f => by ext exact (f.map_smul _ _).symm } @@ -48,14 +48,14 @@ def ringEquivEndForget₂ (R : Type u) [Ring R] : apply NatTrans.ext ext dsimp - simp only [AddCommGroupCat.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, add_smul] + simp only [AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, add_smul] rfl map_mul' := by intros apply NatTrans.ext ext dsimp - simp only [AddCommGroupCat.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, mul_smul] + simp only [AddCommGrp.ofHom_apply, DistribMulAction.toAddMonoidHom_apply, mul_smul] rfl #align ring_equiv_End_forget₂ ringEquivEndForget₂ diff --git a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean index 382d7f0b7e67c..215931b6e9643 100644 --- a/Mathlib/Algebra/Category/MonCat/Adjunctions.lean +++ b/Mathlib/Algebra/Category/MonCat/Adjunctions.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ import Mathlib.Algebra.Category.MonCat.Basic -import Mathlib.Algebra.Category.SemigroupCat.Basic +import Mathlib.Algebra.Category.Semigrp.Basic import Mathlib.Algebra.Group.WithOne.Basic import Mathlib.Algebra.FreeMonoid.Basic @@ -32,7 +32,7 @@ namespace MonCat /-- The functor of adjoining a neutral element `one` to a semigroup. -/ @[to_additive (attr := simps) "The functor of adjoining a neutral element `zero` to a semigroup"] -def adjoinOne : SemigroupCat.{u} ⥤ MonCat.{u} where +def adjoinOne : Semigrp.{u} ⥤ MonCat.{u} where obj S := MonCat.of (WithOne S) map := WithOne.map map_id _ := WithOne.map_id @@ -41,18 +41,18 @@ def adjoinOne : SemigroupCat.{u} ⥤ MonCat.{u} where #align adjoin_zero AddMonCat.adjoinZero @[to_additive] -instance hasForgetToSemigroup : HasForget₂ MonCat SemigroupCat where +instance hasForgetToSemigroup : HasForget₂ MonCat Semigrp where forget₂ := - { obj := fun M => SemigroupCat.of M + { obj := fun M => Semigrp.of M map := MonoidHom.toMulHom } set_option linter.uppercaseLean3 false in #align has_forget_to_Semigroup MonCat.hasForgetToSemigroup set_option linter.uppercaseLean3 false in #align has_forget_to_AddSemigroup AddMonCat.hasForgetToAddSemigroup -/-- The `adjoinOne`-forgetful adjunction from `SemigroupCat` to `MonCat`. -/ -@[to_additive "The `adjoinZero`-forgetful adjunction from `AddSemigroupCat` to `AddMonCat`"] -def adjoinOneAdj : adjoinOne ⊣ forget₂ MonCat.{u} SemigroupCat.{u} := +/-- The `adjoinOne`-forgetful adjunction from `Semigrp` to `MonCat`. -/ +@[to_additive "The `adjoinZero`-forgetful adjunction from `AddSemigrp` to `AddMonCat`"] +def adjoinOneAdj : adjoinOne ⊣ forget₂ MonCat.{u} Semigrp.{u} := Adjunction.mkOfHomEquiv { homEquiv := fun S M => WithOne.lift.symm homEquiv_naturality_left_symm := by diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index 1d59345c01eaf..db480bfc6eae7 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Johannes Hölzl, Yury Kudryashov -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.CategoryTheory.ConcreteCategory.ReflectsIso import Mathlib.Algebra.Ring.Equiv @@ -295,14 +295,14 @@ instance hasForgetToSemiRingCat : HasForget₂ RingCat SemiRingCat := set_option linter.uppercaseLean3 false in #align Ring.has_forget_to_SemiRing RingCat.hasForgetToSemiRingCat -instance hasForgetToAddCommGroupCat : HasForget₂ RingCat AddCommGroupCat where +instance hasForgetToAddCommGrp : HasForget₂ RingCat AddCommGrp where -- can't use BundledHom.mkHasForget₂, since AddCommGroup is an induced category forget₂ := - { obj := fun R => AddCommGroupCat.of R + { obj := fun R => AddCommGrp.of R -- Porting note: use `(_ := _)` similar to above. map := fun {R₁ R₂} f => RingHom.toAddMonoidHom (α := R₁) (β := R₂) f } set_option linter.uppercaseLean3 false in -#align Ring.has_forget_to_AddCommGroup RingCat.hasForgetToAddCommGroupCat +#align Ring.has_forget_to_AddCommGroup RingCat.hasForgetToAddCommGrp end RingCat @@ -722,7 +722,7 @@ set_option linter.uppercaseLean3 false in -- Porting note: This was the case in mathlib3, perhaps it is different now? attribute [local instance] reflectsIsomorphisms_forget₂ -example : (forget₂ RingCat AddCommGroupCat).ReflectsIsomorphisms := by infer_instance +example : (forget₂ RingCat AddCommGrp).ReflectsIsomorphisms := by infer_instance /-! `@[simp]` lemmas for `RingHom.comp` and categorical identities. diff --git a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean index 06647a9f71d27..d6261eacdf956 100644 --- a/Mathlib/Algebra/Category/Ring/FilteredColimits.lean +++ b/Mathlib/Algebra/Category/Ring/FilteredColimits.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Justus Springer -/ import Mathlib.Algebra.Category.Ring.Basic -import Mathlib.Algebra.Category.GroupCat.FilteredColimits +import Mathlib.Algebra.Category.Grp.FilteredColimits #align_import algebra.category.Ring.filtered_colimits from "leanprover-community/mathlib"@"c43486ecf2a5a17479a32ce09e4818924145e90e" @@ -268,8 +268,8 @@ set_option linter.uppercaseLean3 false in instance colimitRing : Ring.{max v u} <| R.{v, u} F := { (R F).str, - AddCommGroupCat.FilteredColimits.colimitAddCommGroup.{v, u} - (F ⋙ forget₂ RingCat AddCommGroupCat.{max v u}) with } + AddCommGrp.FilteredColimits.colimitAddCommGroup.{v, u} + (F ⋙ forget₂ RingCat AddCommGrp.{max v u}) with } set_option linter.uppercaseLean3 false in #align Ring.filtered_colimits.colimit_ring RingCat.FilteredColimits.colimitRing diff --git a/Mathlib/Algebra/Category/Ring/Limits.lean b/Mathlib/Algebra/Category/Ring/Limits.lean index d94dd52a715a2..d833e73410103 100644 --- a/Mathlib/Algebra/Category/Ring/Limits.lean +++ b/Mathlib/Algebra/Category/Ring/Limits.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import Mathlib.Algebra.Ring.Pi import Mathlib.Algebra.Category.Ring.Basic -import Mathlib.Algebra.Category.GroupCat.Limits +import Mathlib.Algebra.Category.Grp.Limits import Mathlib.Algebra.Ring.Subring.Basic #align_import algebra.category.Ring.limits from "leanprover-community/mathlib"@"c43486ecf2a5a17479a32ce09e4818924145e90e" @@ -361,11 +361,11 @@ set_option linter.uppercaseLean3 false in /-- The flat sections of a functor into `RingCat` form a subring of all sections. -/ def sectionsSubring : Subring (∀ j, F.obj j) := - letI f : J ⥤ AddGroupCat.{u} := - F ⋙ forget₂ RingCat.{u} AddCommGroupCat.{u} ⋙ - forget₂ AddCommGroupCat.{u} AddGroupCat.{u} + letI f : J ⥤ AddGrp.{u} := + F ⋙ forget₂ RingCat.{u} AddCommGrp.{u} ⋙ + forget₂ AddCommGrp.{u} AddGrp.{u} letI g : J ⥤ SemiRingCat.{u} := F ⋙ forget₂ RingCat.{u} SemiRingCat.{u} - { AddGroupCat.sectionsAddSubgroup (J := J) f, + { AddGrp.sectionsAddSubgroup (J := J) f, SemiRingCat.sectionsSubsemiring (J := J) g with carrier := (F ⋙ forget RingCat.{u}).sections } set_option linter.uppercaseLean3 false in @@ -458,19 +458,19 @@ set_option linter.uppercaseLean3 false in /-- An auxiliary declaration to speed up typechecking. -/ def forget₂AddCommGroupPreservesLimitsAux : - IsLimit ((forget₂ RingCat.{u} AddCommGroupCat).mapCone (limitCone.{v, u} F)) := by + IsLimit ((forget₂ RingCat.{u} AddCommGrp).mapCone (limitCone.{v, u} F)) := by -- Porting note: inline `f` would not compile - letI f := F ⋙ forget₂ RingCat.{u} AddCommGroupCat.{u} + letI f := F ⋙ forget₂ RingCat.{u} AddCommGrp.{u} letI : Small.{u} (Functor.sections (f ⋙ forget _)) := inferInstanceAs <| Small.{u} (Functor.sections (F ⋙ forget _)) - apply AddCommGroupCat.limitConeIsLimit.{v, u} f + apply AddCommGrp.limitConeIsLimit.{v, u} f set_option linter.uppercaseLean3 false in #align Ring.forget₂_AddCommGroup_preserves_limits_aux RingCat.forget₂AddCommGroupPreservesLimitsAux /-- The forgetful functor from rings to additive commutative groups preserves all limits. -/ instance forget₂AddCommGroupPreservesLimitsOfSize [UnivLE.{v, u}] : - PreservesLimitsOfSize.{v, v} (forget₂ RingCat.{u} AddCommGroupCat.{u}) where + PreservesLimitsOfSize.{v, v} (forget₂ RingCat.{u} AddCommGrp.{u}) where preservesLimitsOfShape {_ _} := { preservesLimit := fun {F} => preservesLimitOfPreservesLimitCone (limitConeIsLimit.{v, u} F) @@ -479,7 +479,7 @@ set_option linter.uppercaseLean3 false in #align Ring.forget₂_AddCommGroup_preserves_limits_of_size RingCat.forget₂AddCommGroupPreservesLimitsOfSize instance forget₂AddCommGroupPreservesLimits : - PreservesLimits (forget₂ RingCat AddCommGroupCat.{u}) := + PreservesLimits (forget₂ RingCat AddCommGrp.{u}) := RingCat.forget₂AddCommGroupPreservesLimitsOfSize.{u, u} set_option linter.uppercaseLean3 false in #align Ring.forget₂_AddCommGroup_preserves_limits RingCat.forget₂AddCommGroupPreservesLimits diff --git a/Mathlib/Algebra/Category/SemigroupCat/Basic.lean b/Mathlib/Algebra/Category/Semigrp/Basic.lean similarity index 72% rename from Mathlib/Algebra/Category/SemigroupCat/Basic.lean rename to Mathlib/Algebra/Category/Semigrp/Basic.lean index 0e72a54b94637..41bbd99d383b5 100644 --- a/Mathlib/Algebra/Category/SemigroupCat/Basic.lean +++ b/Mathlib/Algebra/Category/Semigrp/Basic.lean @@ -16,8 +16,8 @@ import Mathlib.CategoryTheory.Functor.ReflectsIso We introduce the bundled categories: * `MagmaCat` * `AddMagmaCat` -* `SemigroupCat` -* `AddSemigroupCat` +* `Semigrp` +* `AddSemigrp` along with the relevant forgetful functors between them. This closely follows `Mathlib.Algebra.Category.MonCat.Basic`. @@ -128,100 +128,100 @@ end MagmaCat /-- The category of semigroups and semigroup morphisms. -/ @[to_additive] -def SemigroupCat : Type (u + 1) := +def Semigrp : Type (u + 1) := Bundled Semigroup -#align Semigroup SemigroupCat -#align AddSemigroup AddSemigroupCat +#align Semigroup Semigrp +#align AddSemigroup AddSemigrp /-- The category of additive semigroups and semigroup morphisms. -/ -add_decl_doc AddSemigroupCat +add_decl_doc AddSemigrp -namespace SemigroupCat +namespace Semigrp @[to_additive] instance : BundledHom.ParentProjection @Semigroup.toMul := ⟨⟩ -deriving instance LargeCategory for SemigroupCat +deriving instance LargeCategory for Semigrp -- Porting note: deriving failed for `ConcreteCategory`, -- "default handlers have not been implemented yet" -- https://github.com/leanprover-community/mathlib4/issues/5020 -instance instConcreteCategory : ConcreteCategory SemigroupCat := +instance instConcreteCategory : ConcreteCategory Semigrp := BundledHom.concreteCategory (fun _ _ => _) -attribute [to_additive] instSemigroupCatLargeCategory SemigroupCat.instConcreteCategory +attribute [to_additive] instSemigrpLargeCategory Semigrp.instConcreteCategory @[to_additive] -instance : CoeSort SemigroupCat Type* where +instance : CoeSort Semigrp Type* where coe X := X.α -- Porting note: Hinting to Lean that `forget R` and `R` are the same -unif_hint forget_obj_eq_coe (R : SemigroupCat) where ⊢ - (forget SemigroupCat).obj R ≟ R -unif_hint _root_.AddSemigroupCat.forget_obj_eq_coe (R : AddSemigroupCat) where ⊢ - (forget AddSemigroupCat).obj R ≟ R +unif_hint forget_obj_eq_coe (R : Semigrp) where ⊢ + (forget Semigrp).obj R ≟ R +unif_hint _root_.AddSemigrp.forget_obj_eq_coe (R : AddSemigrp) where ⊢ + (forget AddSemigrp).obj R ≟ R @[to_additive] -instance (X : SemigroupCat) : Semigroup X := X.str +instance (X : Semigrp) : Semigroup X := X.str @[to_additive] -instance instFunLike (X Y : SemigroupCat) : FunLike (X ⟶ Y) X Y := +instance instFunLike (X Y : Semigrp) : FunLike (X ⟶ Y) X Y := inferInstanceAs <| FunLike (X →ₙ* Y) X Y @[to_additive] -instance instMulHomClass (X Y : SemigroupCat) : MulHomClass (X ⟶ Y) X Y := +instance instMulHomClass (X Y : Semigrp) : MulHomClass (X ⟶ Y) X Y := inferInstanceAs <| MulHomClass (X →ₙ* Y) X Y -/-- Construct a bundled `SemigroupCat` from the underlying type and typeclass. -/ +/-- Construct a bundled `Semigrp` from the underlying type and typeclass. -/ @[to_additive] -def of (M : Type u) [Semigroup M] : SemigroupCat := +def of (M : Type u) [Semigroup M] : Semigrp := Bundled.of M -#align Semigroup.of SemigroupCat.of -#align AddSemigroup.of AddSemigroupCat.of +#align Semigroup.of Semigrp.of +#align AddSemigroup.of AddSemigrp.of -/-- Construct a bundled `AddSemigroupCat` from the underlying type and typeclass. -/ -add_decl_doc AddSemigroupCat.of +/-- Construct a bundled `AddSemigrp` from the underlying type and typeclass. -/ +add_decl_doc AddSemigrp.of @[to_additive (attr := simp)] -theorem coe_of (R : Type u) [Semigroup R] : (SemigroupCat.of R : Type u) = R := +theorem coe_of (R : Type u) [Semigroup R] : (Semigrp.of R : Type u) = R := rfl -#align Semigroup.coe_of SemigroupCat.coe_of -#align AddSemigroup.coe_of AddSemigroupCat.coe_of +#align Semigroup.coe_of Semigrp.coe_of +#align AddSemigroup.coe_of AddSemigrp.coe_of @[to_additive (attr := simp)] lemma mulEquiv_coe_eq {X Y : Type _} [Semigroup X] [Semigroup Y] (e : X ≃* Y) : - (@DFunLike.coe (SemigroupCat.of X ⟶ SemigroupCat.of Y) _ (fun _ => (forget SemigroupCat).obj _) + (@DFunLike.coe (Semigrp.of X ⟶ Semigrp.of Y) _ (fun _ => (forget Semigrp).obj _) ConcreteCategory.instFunLike (e : X →ₙ* Y) : X → Y) = ↑e := rfl -/-- Typecheck a `MulHom` as a morphism in `SemigroupCat`. -/ +/-- Typecheck a `MulHom` as a morphism in `Semigrp`. -/ @[to_additive] def ofHom {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) : of X ⟶ of Y := f -#align Semigroup.of_hom SemigroupCat.ofHom -#align AddSemigroup.of_hom AddSemigroupCat.ofHom +#align Semigroup.of_hom Semigrp.ofHom +#align AddSemigroup.of_hom AddSemigrp.ofHom -/-- Typecheck an `AddHom` as a morphism in `AddSemigroupCat`. -/ -add_decl_doc AddSemigroupCat.ofHom +/-- Typecheck an `AddHom` as a morphism in `AddSemigrp`. -/ +add_decl_doc AddSemigrp.ofHom @[to_additive] -- Porting note: simp removed, simpNF says LHS simplifies to itself theorem ofHom_apply {X Y : Type u} [Semigroup X] [Semigroup Y] (f : X →ₙ* Y) (x : X) : ofHom f x = f x := rfl -#align Semigroup.of_hom_apply SemigroupCat.ofHom_apply -#align AddSemigroup.of_hom_apply AddSemigroupCat.ofHom_apply +#align Semigroup.of_hom_apply Semigrp.ofHom_apply +#align AddSemigroup.of_hom_apply AddSemigrp.ofHom_apply @[to_additive] -instance : Inhabited SemigroupCat := - ⟨SemigroupCat.of PEmpty⟩ +instance : Inhabited Semigrp := + ⟨Semigrp.of PEmpty⟩ @[to_additive] -instance hasForgetToMagmaCat : HasForget₂ SemigroupCat MagmaCat := +instance hasForgetToMagmaCat : HasForget₂ Semigrp MagmaCat := BundledHom.forget₂ _ _ -#align Semigroup.has_forget_to_Magma SemigroupCat.hasForgetToMagmaCat -#align AddSemigroup.has_forget_to_AddMagma AddSemigroupCat.hasForgetToAddMagmaCat +#align Semigroup.has_forget_to_Magma Semigrp.hasForgetToMagmaCat +#align AddSemigroup.has_forget_to_AddMagma AddSemigrp.hasForgetToAddMagmaCat -end SemigroupCat +end Semigrp variable {X Y : Type u} @@ -252,11 +252,11 @@ variable [Semigroup X] [Semigroup Y] @[to_additive (attr := simps) "Build an isomorphism in the category `AddSemigroup` from an `AddEquiv` between `AddSemigroup`s."] -def MulEquiv.toSemigroupCatIso (e : X ≃* Y) : SemigroupCat.of X ≅ SemigroupCat.of Y where +def MulEquiv.toSemigrpIso (e : X ≃* Y) : Semigrp.of X ≅ Semigrp.of Y where hom := e.toMulHom inv := e.symm.toMulHom -#align mul_equiv.to_Semigroup_iso MulEquiv.toSemigroupCatIso -#align add_equiv.to_AddSemigroup_iso AddEquiv.toAddSemigroupCatIso +#align mul_equiv.to_Semigroup_iso MulEquiv.toSemigrpIso +#align add_equiv.to_AddSemigroup_iso AddEquiv.toAddSemigrpIso end @@ -273,10 +273,10 @@ def magmaCatIsoToMulEquiv {X Y : MagmaCat} (i : X ≅ Y) : X ≃* Y := /-- Build a `MulEquiv` from an isomorphism in the category `Semigroup`. -/ @[to_additive "Build an `AddEquiv` from an isomorphism in the category `AddSemigroup`."] -def semigroupCatIsoToMulEquiv {X Y : SemigroupCat} (i : X ≅ Y) : X ≃* Y := +def semigrpIsoToMulEquiv {X Y : Semigrp} (i : X ≅ Y) : X ≃* Y := MulHom.toMulEquiv i.hom i.inv i.hom_inv_id i.inv_hom_id -#align category_theory.iso.Semigroup_iso_to_mul_equiv CategoryTheory.Iso.semigroupCatIsoToMulEquiv -#align category_theory.iso.Semigroup_iso_to_add_equiv CategoryTheory.Iso.addSemigroupCatIsoToAddEquiv +#align category_theory.iso.Semigroup_iso_to_mul_equiv CategoryTheory.Iso.semigrpIsoToMulEquiv +#align category_theory.iso.Semigroup_iso_to_add_equiv CategoryTheory.Iso.addSemigrpIsoToAddEquiv end CategoryTheory.Iso @@ -297,12 +297,12 @@ in `Semigroup` -/ @[to_additive "additive equivalences between `AddSemigroup`s are the same as (isomorphic to) isomorphisms in `AddSemigroup`"] -def mulEquivIsoSemigroupCatIso {X Y : Type u} [Semigroup X] [Semigroup Y] : - X ≃* Y ≅ SemigroupCat.of X ≅ SemigroupCat.of Y where - hom e := e.toSemigroupCatIso - inv i := i.semigroupCatIsoToMulEquiv -#align mul_equiv_iso_Semigroup_iso mulEquivIsoSemigroupCatIso -#align add_equiv_iso_AddSemigroup_iso addEquivIsoAddSemigroupCatIso +def mulEquivIsoSemigrpIso {X Y : Type u} [Semigroup X] [Semigroup Y] : + X ≃* Y ≅ Semigrp.of X ≅ Semigrp.of Y where + hom e := e.toSemigrpIso + inv i := i.semigrpIsoToMulEquiv +#align mul_equiv_iso_Semigroup_iso mulEquivIsoSemigrpIso +#align add_equiv_iso_AddSemigroup_iso addEquivIsoAddSemigrpIso @[to_additive] instance MagmaCat.forgetReflectsIsos : (forget MagmaCat.{u}).ReflectsIsomorphisms where @@ -314,19 +314,19 @@ instance MagmaCat.forgetReflectsIsos : (forget MagmaCat.{u}).ReflectsIsomorphism #align AddMagma.forget_reflects_isos AddMagmaCat.forgetReflectsIsos @[to_additive] -instance SemigroupCat.forgetReflectsIsos : (forget SemigroupCat.{u}).ReflectsIsomorphisms where +instance Semigrp.forgetReflectsIsos : (forget Semigrp.{u}).ReflectsIsomorphisms where reflects {X Y} f _ := by - let i := asIso ((forget SemigroupCat).map f) + let i := asIso ((forget Semigrp).map f) let e : X ≃* Y := { f, i.toEquiv with } - exact e.toSemigroupCatIso.isIso_hom -#align Semigroup.forget_reflects_isos SemigroupCat.forgetReflectsIsos -#align AddSemigroup.forget_reflects_isos AddSemigroupCat.forgetReflectsIsos + exact e.toSemigrpIso.isIso_hom +#align Semigroup.forget_reflects_isos Semigrp.forgetReflectsIsos +#align AddSemigroup.forget_reflects_isos AddSemigrp.forgetReflectsIsos -- Porting note: this was added in order to ensure that `forget₂ CommMonCat MonCat` -- automatically reflects isomorphisms -- we could have used `CategoryTheory.ConcreteCategory.ReflectsIso` alternatively @[to_additive] -instance SemigroupCat.forget₂_full : (forget₂ SemigroupCat MagmaCat).Full where +instance Semigrp.forget₂_full : (forget₂ Semigrp MagmaCat).Full where map_surjective f := ⟨f, rfl⟩ /-! @@ -335,4 +335,4 @@ we automatically obtain that the `forget₂` functors between our concrete categ reflect isomorphisms. -/ -example : (forget₂ SemigroupCat MagmaCat).ReflectsIsomorphisms := inferInstance +example : (forget₂ Semigrp MagmaCat).ReflectsIsomorphisms := inferInstance diff --git a/Mathlib/Algebra/HierarchyDesign.lean b/Mathlib/Algebra/HierarchyDesign.lean index 3a2af585eb184..1a9406427cc90 100644 --- a/Mathlib/Algebra/HierarchyDesign.lean +++ b/Mathlib/Algebra/HierarchyDesign.lean @@ -140,7 +140,7 @@ For many algebraic structures, particularly ones used in representation theory, etc., we also define "bundled" versions, which carry `category` instances. These bundled versions are usually named by appending `Cat`, -so for example we have `AddCommGroupCat` as a bundled `AddCommGroup`, and `TopCommRingCat` +so for example we have `AddCommGrp` as a bundled `AddCommGroup`, and `TopCommRingCat` (which bundles together `CommRing`, `TopologicalSpace`, and `TopologicalRing`). These bundled versions have many appealing features: @@ -148,9 +148,9 @@ These bundled versions have many appealing features: * a uniform notation (and definition) for isomorphisms `X ≅ Y` * a uniform API for subobjects, via the partial order `Subobject X` * interoperability with unbundled structures, via coercions to `Type` - (so if `G : AddCommGroupCat`, you can treat `G` as a type, + (so if `G : AddCommGrp`, you can treat `G` as a type, and it automatically has an `AddCommGroup` instance) - and lifting maps `AddCommGroupCat.of G`, when `G` is a type with an `AddCommGroup` instance. + and lifting maps `AddCommGrp.of G`, when `G` is a type with an `AddCommGroup` instance. If, for example you do the work of proving that a typeclass `Z` has a good notion of tensor product, you are strongly encouraged to provide the corresponding `MonoidalCategory` instance diff --git a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean index 721e276890554..df0f557290f4e 100644 --- a/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean +++ b/Mathlib/Algebra/Homology/HomotopyCategory/HomComplex.lean @@ -5,7 +5,7 @@ Authors: Joël Riou -/ import Mathlib.Algebra.Homology.Homotopy import Mathlib.Algebra.Ring.NegOnePow -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.Preadditive import Mathlib.Tactic.Linarith import Mathlib.CategoryTheory.Linear.LinearFunctor @@ -563,15 +563,15 @@ open HomComplex /-- The cochain complex of homomorphisms between two cochain complexes `F` and `G`. In degree `n : ℤ`, it consists of the abelian group `HomComplex.Cochain F G n`. -/ -- We also constructed the `d_apply` lemma using `@[simps]` --- until we made `AddCommGroupCat.coe_of` a simp lemma, +-- until we made `AddCommGrp.coe_of` a simp lemma, -- after which the simp normal form linter complains. -- It was not used a simp lemma in Mathlib. -- Possible solution: higher priority function coercions that remove the `of`? -- @[simp] @[simps! X] -def HomComplex : CochainComplex AddCommGroupCat ℤ where - X i := AddCommGroupCat.of (Cochain F G i) - d i j := AddCommGroupCat.ofHom (δ_hom ℤ F G i j) +def HomComplex : CochainComplex AddCommGrp ℤ where + X i := AddCommGrp.of (Cochain F G i) + d i j := AddCommGrp.ofHom (δ_hom ℤ F G i j) shape _ _ hij := by ext; apply δ_shape _ _ hij d_comp_d' _ _ _ _ _ := by ext; apply δ_δ diff --git a/Mathlib/Algebra/Homology/LocalCohomology.lean b/Mathlib/Algebra/Homology/LocalCohomology.lean index 0a8faf4c39ee5..d83baf548859b 100644 --- a/Mathlib/Algebra/Homology/LocalCohomology.lean +++ b/Mathlib/Algebra/Homology/LocalCohomology.lean @@ -96,7 +96,7 @@ variable {R : Type max u v} [CommRing R] {D : Type v} [SmallCategory D] lemma hasColimitDiagram (I : D ⥤ Ideal R) (i : ℕ) : HasColimit (diagram I i) := by - have : HasColimitsOfShape Dᵒᵖ (AddCommGroupCatMax.{u, v}) := inferInstance + have : HasColimitsOfShape Dᵒᵖ (AddCommGrpMax.{u, v}) := inferInstance infer_instance /- diff --git a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean index 7d335ef13a580..012292a6e2dfc 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Ab.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Ab.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ import Mathlib.Algebra.Homology.ShortComplex.ShortExact -import Mathlib.Algebra.Category.GroupCat.Abelian +import Mathlib.Algebra.Category.Grp.Abelian /-! # Homology and exactness of short complexes of abelian groups @@ -47,20 +47,20 @@ def abToCycles : S.X₁ →+ AddMonoidHom.ker S.g := given by a kernel and a quotient given by the `AddMonoidHom` API. -/ @[simps] def abLeftHomologyData : S.LeftHomologyData where - K := AddCommGroupCat.of (AddMonoidHom.ker S.g) - H := AddCommGroupCat.of ((AddMonoidHom.ker S.g) ⧸ AddMonoidHom.range S.abToCycles) + K := AddCommGrp.of (AddMonoidHom.ker S.g) + H := AddCommGrp.of ((AddMonoidHom.ker S.g) ⧸ AddMonoidHom.range S.abToCycles) i := (AddMonoidHom.ker S.g).subtype π := QuotientAddGroup.mk' _ wi := by ext ⟨_, hx⟩ exact hx - hi := AddCommGroupCat.kernelIsLimit _ + hi := AddCommGrp.kernelIsLimit _ wπ := by ext (x : S.X₁) erw [QuotientAddGroup.eq_zero_iff] rw [AddMonoidHom.mem_range] apply exists_apply_eq_apply - hπ := AddCommGroupCat.cokernelIsColimit (AddCommGroupCat.ofHom S.abToCycles) + hπ := AddCommGrp.cokernelIsColimit (AddCommGrp.ofHom S.abToCycles) @[simp] lemma abLeftHomologyData_f' : S.abLeftHomologyData.f' = S.abToCycles := rfl @@ -68,10 +68,10 @@ lemma abLeftHomologyData_f' : S.abLeftHomologyData.f' = S.abToCycles := rfl /-- Given a short complex `S` of abelian groups, this is the isomorphism between the abstract `S.cycles` of the homology API and the more concrete description as `AddMonoidHom.ker S.g`. -/ -noncomputable def abCyclesIso : S.cycles ≅ AddCommGroupCat.of (AddMonoidHom.ker S.g) := +noncomputable def abCyclesIso : S.cycles ≅ AddCommGrp.of (AddMonoidHom.ker S.g) := S.abLeftHomologyData.cyclesIso --- This was a simp lemma until we made `AddCommGroupCat.coe_of` a simp lemma, +-- This was a simp lemma until we made `AddCommGrp.coe_of` a simp lemma, -- after which the simp normal form linter complains. -- It was not used a simp lemma in Mathlib. -- Possible solution: higher priority function coercions that remove the `of`? @@ -87,13 +87,13 @@ the abstract `S.homology` of the homology API and the more explicit quotient of `AddMonoidHom.ker S.g` by the image of `S.abToCycles : S.X₁ →+ AddMonoidHom.ker S.g`. -/ noncomputable def abHomologyIso : S.homology ≅ - AddCommGroupCat.of ((AddMonoidHom.ker S.g) ⧸ AddMonoidHom.range S.abToCycles) := + AddCommGrp.of ((AddMonoidHom.ker S.g) ⧸ AddMonoidHom.range S.abToCycles) := S.abLeftHomologyData.homologyIso lemma exact_iff_surjective_abToCycles : S.Exact ↔ Function.Surjective S.abToCycles := by rw [S.abLeftHomologyData.exact_iff_epi_f', abLeftHomologyData_f', - AddCommGroupCat.epi_iff_surjective] + AddCommGrp.epi_iff_surjective] rfl lemma ab_exact_iff : @@ -122,11 +122,11 @@ lemma ab_exact_iff_range_eq_ker : S.Exact ↔ S.f.range = S.g.ker := by lemma ShortExact.ab_injective_f (hS : S.ShortExact) : Function.Injective S.f := - (AddCommGroupCat.mono_iff_injective _).1 hS.mono_f + (AddCommGrp.mono_iff_injective _).1 hS.mono_f lemma ShortExact.ab_surjective_g (hS : S.ShortExact) : Function.Surjective S.g := - (AddCommGroupCat.epi_iff_surjective _).1 hS.epi_g + (AddCommGrp.epi_iff_surjective _).1 hS.epi_g end ShortComplex diff --git a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean index 1eabdd1cc08ea..c3bb996a985c3 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/ConcreteCategory.lean @@ -42,7 +42,7 @@ variable [Preadditive C] [(forget₂ C Ab).Additive] [(forget₂ C Ab).Preserves lemma Preadditive.mono_iff_injective {X Y : C} (f : X ⟶ Y) : Mono f ↔ Function.Injective ((forget₂ C Ab).map f) := by - rw [← AddCommGroupCat.mono_iff_injective] + rw [← AddCommGrp.mono_iff_injective] constructor · intro infer_instance @@ -57,7 +57,7 @@ lemma Preadditive.mono_iff_injective' {X Y : C} (f : X ⟶ Y) : lemma Preadditive.epi_iff_surjective {X Y : C} (f : X ⟶ Y) : Epi f ↔ Function.Surjective ((forget₂ C Ab).map f) := by - rw [← AddCommGroupCat.epi_iff_surjective] + rw [← AddCommGrp.epi_iff_surjective] constructor · intro infer_instance diff --git a/Mathlib/Algebra/Module/CharacterModule.lean b/Mathlib/Algebra/Module/CharacterModule.lean index 915ca84dbf407..7dd0866babfc2 100644 --- a/Mathlib/Algebra/Module/CharacterModule.lean +++ b/Mathlib/Algebra/Module/CharacterModule.lean @@ -6,7 +6,7 @@ Authors: Jujian Zhang, Junyan Xu import Mathlib.Algebra.Module.LinearMap.Basic import Mathlib.Algebra.Category.ModuleCat.Basic -import Mathlib.Algebra.Category.GroupCat.Injective +import Mathlib.Algebra.Category.Grp.Injective import Mathlib.Topology.Instances.AddCircle import Mathlib.Topology.Instances.Rat import Mathlib.LinearAlgebra.Isomorphisms diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat.lean deleted file mode 100644 index 3e62136b0e772..0000000000000 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat.lean +++ /dev/null @@ -1,270 +0,0 @@ -/- -Copyright (c) 2021 Johan Commelin. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Johan Commelin, Riccardo Brasca --/ -import Mathlib.Analysis.Normed.Group.Constructions -import Mathlib.Analysis.Normed.Group.Hom -import Mathlib.CategoryTheory.Limits.Shapes.ZeroMorphisms -import Mathlib.CategoryTheory.ConcreteCategory.BundledHom -import Mathlib.CategoryTheory.Elementwise - -#align_import analysis.normed.group.SemiNormedGroup from "leanprover-community/mathlib"@"17ef379e997badd73e5eabb4d38f11919ab3c4b3" - -/-! -# The category of seminormed groups - -We define `SemiNormedGroupCat`, the category of seminormed groups and normed group homs between -them, as well as `SemiNormedGroupCat₁`, the subcategory of norm non-increasing morphisms. --/ - -set_option linter.uppercaseLean3 false - -noncomputable section - -universe u - -open CategoryTheory - -/-- The category of seminormed abelian groups and bounded group homomorphisms. -/ -def SemiNormedGroupCat : Type (u + 1) := - Bundled SeminormedAddCommGroup -#align SemiNormedGroup SemiNormedGroupCat - -namespace SemiNormedGroupCat - -instance bundledHom : BundledHom @NormedAddGroupHom where - toFun := @NormedAddGroupHom.toFun - id := @NormedAddGroupHom.id - comp := @NormedAddGroupHom.comp -#align SemiNormedGroup.bundled_hom SemiNormedGroupCat.bundledHom - -deriving instance LargeCategory for SemiNormedGroupCat - --- Porting note: deriving fails for ConcreteCategory, adding instance manually. --- See https://github.com/leanprover-community/mathlib4/issues/5020 --- deriving instance LargeCategory, ConcreteCategory for SemiRingCat -instance : ConcreteCategory SemiNormedGroupCat := by - dsimp [SemiNormedGroupCat] - infer_instance - -instance : CoeSort SemiNormedGroupCat Type* where - coe X := X.α - -/-- Construct a bundled `SemiNormedGroupCat` from the underlying type and typeclass. -/ -def of (M : Type u) [SeminormedAddCommGroup M] : SemiNormedGroupCat := - Bundled.of M -#align SemiNormedGroupCat.of SemiNormedGroupCat.of - -instance (M : SemiNormedGroupCat) : SeminormedAddCommGroup M := - M.str - --- Porting note (#10754): added instance -instance funLike {V W : SemiNormedGroupCat} : FunLike (V ⟶ W) V W where - coe := (forget SemiNormedGroupCat).map - coe_injective' := fun f g h => by cases f; cases g; congr - -instance toAddMonoidHomClass {V W : SemiNormedGroupCat} : AddMonoidHomClass (V ⟶ W) V W where - map_add f := f.map_add' - map_zero f := (AddMonoidHom.mk' f.toFun f.map_add').map_zero - --- Porting note (#10688): added to ease automation -@[ext] -lemma ext {M N : SemiNormedGroupCat} {f₁ f₂ : M ⟶ N} (h : ∀ (x : M), f₁ x = f₂ x) : f₁ = f₂ := - DFunLike.ext _ _ h - -@[simp] -theorem coe_of (V : Type u) [SeminormedAddCommGroup V] : (SemiNormedGroupCat.of V : Type u) = V := - rfl -#align SemiNormedGroup.coe_of SemiNormedGroupCat.coe_of - --- Porting note: marked with high priority to short circuit simplifier's path -@[simp (high)] -theorem coe_id (V : SemiNormedGroupCat) : (𝟙 V : V → V) = id := - rfl -#align SemiNormedGroup.coe_id SemiNormedGroupCat.coe_id - --- Porting note: marked with high priority to short circuit simplifier's path -@[simp (high)] -theorem coe_comp {M N K : SemiNormedGroupCat} (f : M ⟶ N) (g : N ⟶ K) : - (f ≫ g : M → K) = g ∘ f := - rfl -#align SemiNormedGroup.coe_comp SemiNormedGroupCat.coe_comp - -instance : Inhabited SemiNormedGroupCat := - ⟨of PUnit⟩ - -instance ofUnique (V : Type u) [SeminormedAddCommGroup V] [i : Unique V] : - Unique (SemiNormedGroupCat.of V) := - i -#align SemiNormedGroup.of_unique SemiNormedGroupCat.ofUnique - -instance {M N : SemiNormedGroupCat} : Zero (M ⟶ N) := - NormedAddGroupHom.zero - -@[simp] -theorem zero_apply {V W : SemiNormedGroupCat} (x : V) : (0 : V ⟶ W) x = 0 := - rfl -#align SemiNormedGroup.zero_apply SemiNormedGroupCat.zero_apply - -instance : Limits.HasZeroMorphisms.{u, u + 1} SemiNormedGroupCat where - -theorem isZero_of_subsingleton (V : SemiNormedGroupCat) [Subsingleton V] : Limits.IsZero V := by - refine ⟨fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩⟩ - · ext x; have : x = 0 := Subsingleton.elim _ _; simp only [this, map_zero] - · ext; apply Subsingleton.elim -#align SemiNormedGroup.is_zero_of_subsingleton SemiNormedGroupCat.isZero_of_subsingleton - -instance hasZeroObject : Limits.HasZeroObject SemiNormedGroupCat.{u} := - ⟨⟨of PUnit, isZero_of_subsingleton _⟩⟩ -#align SemiNormedGroup.has_zero_object SemiNormedGroupCat.hasZeroObject - -theorem iso_isometry_of_normNoninc {V W : SemiNormedGroupCat} (i : V ≅ W) (h1 : i.hom.NormNoninc) - (h2 : i.inv.NormNoninc) : Isometry i.hom := by - apply AddMonoidHomClass.isometry_of_norm - intro v - apply le_antisymm (h1 v) - calc - -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - ‖v‖ = ‖i.inv (i.hom v)‖ := by erw [Iso.hom_inv_id_apply] - _ ≤ ‖i.hom v‖ := h2 _ -#align SemiNormedGroup.iso_isometry_of_norm_noninc SemiNormedGroupCat.iso_isometry_of_normNoninc - -end SemiNormedGroupCat - -/-- `SemiNormedGroupCat₁` is a type synonym for `SemiNormedGroupCat`, -which we shall equip with the category structure consisting only of the norm non-increasing maps. --/ -def SemiNormedGroupCat₁ : Type (u + 1) := - Bundled SeminormedAddCommGroup -#align SemiNormedGroup₁ SemiNormedGroupCat₁ - -namespace SemiNormedGroupCat₁ - -instance : CoeSort SemiNormedGroupCat₁ Type* where - coe X := X.α - -instance : LargeCategory.{u} SemiNormedGroupCat₁ where - Hom X Y := { f : NormedAddGroupHom X Y // f.NormNoninc } - id X := ⟨NormedAddGroupHom.id X, NormedAddGroupHom.NormNoninc.id⟩ - comp {X Y Z} f g := ⟨g.1.comp f.1, g.2.comp f.2⟩ - --- Porting note (#10754): added instance -instance instFunLike (X Y : SemiNormedGroupCat₁) : FunLike (X ⟶ Y) X Y where - coe f := f.1.toFun - coe_injective' _ _ h := Subtype.val_inj.mp (NormedAddGroupHom.coe_injective h) - -@[ext] -theorem hom_ext {M N : SemiNormedGroupCat₁} (f g : M ⟶ N) (w : (f : M → N) = (g : M → N)) : - f = g := - Subtype.eq (NormedAddGroupHom.ext (congr_fun w)) -#align SemiNormedGroup₁.hom_ext SemiNormedGroupCat₁.hom_ext - -instance : ConcreteCategory.{u} SemiNormedGroupCat₁ where - forget := - { obj := fun X => X - map := fun f => f } - forget_faithful := { } - --- Porting note (#10754): added instance -instance toAddMonoidHomClass {V W : SemiNormedGroupCat₁} : AddMonoidHomClass (V ⟶ W) V W where - map_add f := f.1.map_add' - map_zero f := (AddMonoidHom.mk' f.1 f.1.map_add').map_zero - -/-- Construct a bundled `SemiNormedGroupCat₁` from the underlying type and typeclass. -/ -def of (M : Type u) [SeminormedAddCommGroup M] : SemiNormedGroupCat₁ := - Bundled.of M -#align SemiNormedGroup₁.of SemiNormedGroupCat₁.of - -instance (M : SemiNormedGroupCat₁) : SeminormedAddCommGroup M := - M.str - -/-- Promote a morphism in `SemiNormedGroupCat` to a morphism in `SemiNormedGroupCat₁`. -/ -def mkHom {M N : SemiNormedGroupCat} (f : M ⟶ N) (i : f.NormNoninc) : - SemiNormedGroupCat₁.of M ⟶ SemiNormedGroupCat₁.of N := - ⟨f, i⟩ -#align SemiNormedGroup₁.mk_hom SemiNormedGroupCat₁.mkHom - --- @[simp] -- Porting note: simpNF linter claims LHS simplifies with `SemiNormedGroupCat₁.coe_of` -theorem mkHom_apply {M N : SemiNormedGroupCat} (f : M ⟶ N) (i : f.NormNoninc) (x) : - mkHom f i x = f x := - rfl -#align SemiNormedGroup₁.mk_hom_apply SemiNormedGroupCat₁.mkHom_apply - -/-- Promote an isomorphism in `SemiNormedGroupCat` to an isomorphism in `SemiNormedGroupCat₁`. -/ -@[simps] -def mkIso {M N : SemiNormedGroupCat} (f : M ≅ N) (i : f.hom.NormNoninc) (i' : f.inv.NormNoninc) : - SemiNormedGroupCat₁.of M ≅ SemiNormedGroupCat₁.of N where - hom := mkHom f.hom i - inv := mkHom f.inv i' - hom_inv_id := by apply Subtype.eq; exact f.hom_inv_id - inv_hom_id := by apply Subtype.eq; exact f.inv_hom_id -#align SemiNormedGroup₁.mk_iso SemiNormedGroupCat₁.mkIso - -instance : HasForget₂ SemiNormedGroupCat₁ SemiNormedGroupCat where - forget₂ := - { obj := fun X => X - map := fun f => f.1 } - -@[simp] -theorem coe_of (V : Type u) [SeminormedAddCommGroup V] : (SemiNormedGroupCat₁.of V : Type u) = V := - rfl -#align SemiNormedGroup₁.coe_of SemiNormedGroupCat₁.coe_of - --- Porting note: marked with high priority to short circuit simplifier's path -@[simp (high)] -theorem coe_id (V : SemiNormedGroupCat₁) : ⇑(𝟙 V) = id := - rfl -#align SemiNormedGroup₁.coe_id SemiNormedGroupCat₁.coe_id - --- Porting note: marked with high priority to short circuit simplifier's path -@[simp (high)] -theorem coe_comp {M N K : SemiNormedGroupCat₁} (f : M ⟶ N) (g : N ⟶ K) : - (f ≫ g : M → K) = g ∘ f := - rfl -#align SemiNormedGroup₁.coe_comp SemiNormedGroupCat₁.coe_comp - --- Porting note: deleted `coe_comp'`, as we no longer have the relevant coercion. -#noalign SemiNormedGroup₁.coe_comp' - -instance : Inhabited SemiNormedGroupCat₁ := - ⟨of PUnit⟩ - -instance ofUnique (V : Type u) [SeminormedAddCommGroup V] [i : Unique V] : - Unique (SemiNormedGroupCat₁.of V) := - i -#align SemiNormedGroup₁.of_unique SemiNormedGroupCat₁.ofUnique - --- Porting note: extracted from `Limits.HasZeroMorphisms` instance below. -instance (X Y : SemiNormedGroupCat₁) : Zero (X ⟶ Y) where - zero := ⟨0, NormedAddGroupHom.NormNoninc.zero⟩ - -@[simp] -theorem zero_apply {V W : SemiNormedGroupCat₁} (x : V) : (0 : V ⟶ W) x = 0 := - rfl -#align SemiNormedGroup₁.zero_apply SemiNormedGroupCat₁.zero_apply - -instance : Limits.HasZeroMorphisms.{u, u + 1} SemiNormedGroupCat₁ where - -theorem isZero_of_subsingleton (V : SemiNormedGroupCat₁) [Subsingleton V] : Limits.IsZero V := by - refine ⟨fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩⟩ - · ext x; have : x = 0 := Subsingleton.elim _ _; simp only [this, map_zero] - · ext; apply Subsingleton.elim -#align SemiNormedGroup₁.is_zero_of_subsingleton SemiNormedGroupCat₁.isZero_of_subsingleton - -instance hasZeroObject : Limits.HasZeroObject SemiNormedGroupCat₁.{u} := - ⟨⟨of PUnit, isZero_of_subsingleton _⟩⟩ -#align SemiNormedGroup₁.has_zero_object SemiNormedGroupCat₁.hasZeroObject - -theorem iso_isometry {V W : SemiNormedGroupCat₁} (i : V ≅ W) : Isometry i.hom := by - change Isometry (⟨⟨i.hom, map_zero _⟩, fun _ _ => map_add _ _ _⟩ : V →+ W) - refine AddMonoidHomClass.isometry_of_norm _ ?_ - intro v - apply le_antisymm (i.hom.2 v) - calc - -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - ‖v‖ = ‖i.inv (i.hom v)‖ := by erw [Iso.hom_inv_id_apply] - _ ≤ ‖i.hom v‖ := i.inv.2 _ -#align SemiNormedGroup₁.iso_isometry SemiNormedGroupCat₁.iso_isometry - -end SemiNormedGroupCat₁ diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean new file mode 100644 index 0000000000000..bf096f8f6a26c --- /dev/null +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGrp.lean @@ -0,0 +1,270 @@ +/- +Copyright (c) 2021 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Riccardo Brasca +-/ +import Mathlib.Analysis.Normed.Group.Constructions +import Mathlib.Analysis.Normed.Group.Hom +import Mathlib.CategoryTheory.Limits.Shapes.ZeroMorphisms +import Mathlib.CategoryTheory.ConcreteCategory.BundledHom +import Mathlib.CategoryTheory.Elementwise + +#align_import analysis.normed.group.SemiNormedGroup from "leanprover-community/mathlib"@"17ef379e997badd73e5eabb4d38f11919ab3c4b3" + +/-! +# The category of seminormed groups + +We define `SemiNormedGrp`, the category of seminormed groups and normed group homs between +them, as well as `SemiNormedGrp₁`, the subcategory of norm non-increasing morphisms. +-/ + +set_option linter.uppercaseLean3 false + +noncomputable section + +universe u + +open CategoryTheory + +/-- The category of seminormed abelian groups and bounded group homomorphisms. -/ +def SemiNormedGrp : Type (u + 1) := + Bundled SeminormedAddCommGroup +#align SemiNormedGroup SemiNormedGrp + +namespace SemiNormedGrp + +instance bundledHom : BundledHom @NormedAddGroupHom where + toFun := @NormedAddGroupHom.toFun + id := @NormedAddGroupHom.id + comp := @NormedAddGroupHom.comp +#align SemiNormedGroup.bundled_hom SemiNormedGrp.bundledHom + +deriving instance LargeCategory for SemiNormedGrp + +-- Porting note: deriving fails for ConcreteCategory, adding instance manually. +-- See https://github.com/leanprover-community/mathlib4/issues/5020 +-- deriving instance LargeCategory, ConcreteCategory for SemiRingCat +instance : ConcreteCategory SemiNormedGrp := by + dsimp [SemiNormedGrp] + infer_instance + +instance : CoeSort SemiNormedGrp Type* where + coe X := X.α + +/-- Construct a bundled `SemiNormedGrp` from the underlying type and typeclass. -/ +def of (M : Type u) [SeminormedAddCommGroup M] : SemiNormedGrp := + Bundled.of M +#align SemiNormedGrp.of SemiNormedGrp.of + +instance (M : SemiNormedGrp) : SeminormedAddCommGroup M := + M.str + +-- Porting note (#10754): added instance +instance funLike {V W : SemiNormedGrp} : FunLike (V ⟶ W) V W where + coe := (forget SemiNormedGrp).map + coe_injective' := fun f g h => by cases f; cases g; congr + +instance toAddMonoidHomClass {V W : SemiNormedGrp} : AddMonoidHomClass (V ⟶ W) V W where + map_add f := f.map_add' + map_zero f := (AddMonoidHom.mk' f.toFun f.map_add').map_zero + +-- Porting note (#10688): added to ease automation +@[ext] +lemma ext {M N : SemiNormedGrp} {f₁ f₂ : M ⟶ N} (h : ∀ (x : M), f₁ x = f₂ x) : f₁ = f₂ := + DFunLike.ext _ _ h + +@[simp] +theorem coe_of (V : Type u) [SeminormedAddCommGroup V] : (SemiNormedGrp.of V : Type u) = V := + rfl +#align SemiNormedGroup.coe_of SemiNormedGrp.coe_of + +-- Porting note: marked with high priority to short circuit simplifier's path +@[simp (high)] +theorem coe_id (V : SemiNormedGrp) : (𝟙 V : V → V) = id := + rfl +#align SemiNormedGroup.coe_id SemiNormedGrp.coe_id + +-- Porting note: marked with high priority to short circuit simplifier's path +@[simp (high)] +theorem coe_comp {M N K : SemiNormedGrp} (f : M ⟶ N) (g : N ⟶ K) : + (f ≫ g : M → K) = g ∘ f := + rfl +#align SemiNormedGroup.coe_comp SemiNormedGrp.coe_comp + +instance : Inhabited SemiNormedGrp := + ⟨of PUnit⟩ + +instance ofUnique (V : Type u) [SeminormedAddCommGroup V] [i : Unique V] : + Unique (SemiNormedGrp.of V) := + i +#align SemiNormedGroup.of_unique SemiNormedGrp.ofUnique + +instance {M N : SemiNormedGrp} : Zero (M ⟶ N) := + NormedAddGroupHom.zero + +@[simp] +theorem zero_apply {V W : SemiNormedGrp} (x : V) : (0 : V ⟶ W) x = 0 := + rfl +#align SemiNormedGroup.zero_apply SemiNormedGrp.zero_apply + +instance : Limits.HasZeroMorphisms.{u, u + 1} SemiNormedGrp where + +theorem isZero_of_subsingleton (V : SemiNormedGrp) [Subsingleton V] : Limits.IsZero V := by + refine ⟨fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩⟩ + · ext x; have : x = 0 := Subsingleton.elim _ _; simp only [this, map_zero] + · ext; apply Subsingleton.elim +#align SemiNormedGroup.is_zero_of_subsingleton SemiNormedGrp.isZero_of_subsingleton + +instance hasZeroObject : Limits.HasZeroObject SemiNormedGrp.{u} := + ⟨⟨of PUnit, isZero_of_subsingleton _⟩⟩ +#align SemiNormedGroup.has_zero_object SemiNormedGrp.hasZeroObject + +theorem iso_isometry_of_normNoninc {V W : SemiNormedGrp} (i : V ≅ W) (h1 : i.hom.NormNoninc) + (h2 : i.inv.NormNoninc) : Isometry i.hom := by + apply AddMonoidHomClass.isometry_of_norm + intro v + apply le_antisymm (h1 v) + calc + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + ‖v‖ = ‖i.inv (i.hom v)‖ := by erw [Iso.hom_inv_id_apply] + _ ≤ ‖i.hom v‖ := h2 _ +#align SemiNormedGroup.iso_isometry_of_norm_noninc SemiNormedGrp.iso_isometry_of_normNoninc + +end SemiNormedGrp + +/-- `SemiNormedGrp₁` is a type synonym for `SemiNormedGrp`, +which we shall equip with the category structure consisting only of the norm non-increasing maps. +-/ +def SemiNormedGrp₁ : Type (u + 1) := + Bundled SeminormedAddCommGroup +#align SemiNormedGroup₁ SemiNormedGrp₁ + +namespace SemiNormedGrp₁ + +instance : CoeSort SemiNormedGrp₁ Type* where + coe X := X.α + +instance : LargeCategory.{u} SemiNormedGrp₁ where + Hom X Y := { f : NormedAddGroupHom X Y // f.NormNoninc } + id X := ⟨NormedAddGroupHom.id X, NormedAddGroupHom.NormNoninc.id⟩ + comp {X Y Z} f g := ⟨g.1.comp f.1, g.2.comp f.2⟩ + +-- Porting note (#10754): added instance +instance instFunLike (X Y : SemiNormedGrp₁) : FunLike (X ⟶ Y) X Y where + coe f := f.1.toFun + coe_injective' _ _ h := Subtype.val_inj.mp (NormedAddGroupHom.coe_injective h) + +@[ext] +theorem hom_ext {M N : SemiNormedGrp₁} (f g : M ⟶ N) (w : (f : M → N) = (g : M → N)) : + f = g := + Subtype.eq (NormedAddGroupHom.ext (congr_fun w)) +#align SemiNormedGroup₁.hom_ext SemiNormedGrp₁.hom_ext + +instance : ConcreteCategory.{u} SemiNormedGrp₁ where + forget := + { obj := fun X => X + map := fun f => f } + forget_faithful := { } + +-- Porting note (#10754): added instance +instance toAddMonoidHomClass {V W : SemiNormedGrp₁} : AddMonoidHomClass (V ⟶ W) V W where + map_add f := f.1.map_add' + map_zero f := (AddMonoidHom.mk' f.1 f.1.map_add').map_zero + +/-- Construct a bundled `SemiNormedGrp₁` from the underlying type and typeclass. -/ +def of (M : Type u) [SeminormedAddCommGroup M] : SemiNormedGrp₁ := + Bundled.of M +#align SemiNormedGroup₁.of SemiNormedGrp₁.of + +instance (M : SemiNormedGrp₁) : SeminormedAddCommGroup M := + M.str + +/-- Promote a morphism in `SemiNormedGrp` to a morphism in `SemiNormedGrp₁`. -/ +def mkHom {M N : SemiNormedGrp} (f : M ⟶ N) (i : f.NormNoninc) : + SemiNormedGrp₁.of M ⟶ SemiNormedGrp₁.of N := + ⟨f, i⟩ +#align SemiNormedGroup₁.mk_hom SemiNormedGrp₁.mkHom + +-- @[simp] -- Porting note: simpNF linter claims LHS simplifies with `SemiNormedGrp₁.coe_of` +theorem mkHom_apply {M N : SemiNormedGrp} (f : M ⟶ N) (i : f.NormNoninc) (x) : + mkHom f i x = f x := + rfl +#align SemiNormedGroup₁.mk_hom_apply SemiNormedGrp₁.mkHom_apply + +/-- Promote an isomorphism in `SemiNormedGrp` to an isomorphism in `SemiNormedGrp₁`. -/ +@[simps] +def mkIso {M N : SemiNormedGrp} (f : M ≅ N) (i : f.hom.NormNoninc) (i' : f.inv.NormNoninc) : + SemiNormedGrp₁.of M ≅ SemiNormedGrp₁.of N where + hom := mkHom f.hom i + inv := mkHom f.inv i' + hom_inv_id := by apply Subtype.eq; exact f.hom_inv_id + inv_hom_id := by apply Subtype.eq; exact f.inv_hom_id +#align SemiNormedGroup₁.mk_iso SemiNormedGrp₁.mkIso + +instance : HasForget₂ SemiNormedGrp₁ SemiNormedGrp where + forget₂ := + { obj := fun X => X + map := fun f => f.1 } + +@[simp] +theorem coe_of (V : Type u) [SeminormedAddCommGroup V] : (SemiNormedGrp₁.of V : Type u) = V := + rfl +#align SemiNormedGroup₁.coe_of SemiNormedGrp₁.coe_of + +-- Porting note: marked with high priority to short circuit simplifier's path +@[simp (high)] +theorem coe_id (V : SemiNormedGrp₁) : ⇑(𝟙 V) = id := + rfl +#align SemiNormedGroup₁.coe_id SemiNormedGrp₁.coe_id + +-- Porting note: marked with high priority to short circuit simplifier's path +@[simp (high)] +theorem coe_comp {M N K : SemiNormedGrp₁} (f : M ⟶ N) (g : N ⟶ K) : + (f ≫ g : M → K) = g ∘ f := + rfl +#align SemiNormedGroup₁.coe_comp SemiNormedGrp₁.coe_comp + +-- Porting note: deleted `coe_comp'`, as we no longer have the relevant coercion. +#noalign SemiNormedGroup₁.coe_comp' + +instance : Inhabited SemiNormedGrp₁ := + ⟨of PUnit⟩ + +instance ofUnique (V : Type u) [SeminormedAddCommGroup V] [i : Unique V] : + Unique (SemiNormedGrp₁.of V) := + i +#align SemiNormedGroup₁.of_unique SemiNormedGrp₁.ofUnique + +-- Porting note: extracted from `Limits.HasZeroMorphisms` instance below. +instance (X Y : SemiNormedGrp₁) : Zero (X ⟶ Y) where + zero := ⟨0, NormedAddGroupHom.NormNoninc.zero⟩ + +@[simp] +theorem zero_apply {V W : SemiNormedGrp₁} (x : V) : (0 : V ⟶ W) x = 0 := + rfl +#align SemiNormedGroup₁.zero_apply SemiNormedGrp₁.zero_apply + +instance : Limits.HasZeroMorphisms.{u, u + 1} SemiNormedGrp₁ where + +theorem isZero_of_subsingleton (V : SemiNormedGrp₁) [Subsingleton V] : Limits.IsZero V := by + refine ⟨fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩, fun X => ⟨⟨⟨0⟩, fun f => ?_⟩⟩⟩ + · ext x; have : x = 0 := Subsingleton.elim _ _; simp only [this, map_zero] + · ext; apply Subsingleton.elim +#align SemiNormedGroup₁.is_zero_of_subsingleton SemiNormedGrp₁.isZero_of_subsingleton + +instance hasZeroObject : Limits.HasZeroObject SemiNormedGrp₁.{u} := + ⟨⟨of PUnit, isZero_of_subsingleton _⟩⟩ +#align SemiNormedGroup₁.has_zero_object SemiNormedGrp₁.hasZeroObject + +theorem iso_isometry {V W : SemiNormedGrp₁} (i : V ≅ W) : Isometry i.hom := by + change Isometry (⟨⟨i.hom, map_zero _⟩, fun _ _ => map_add _ _ _⟩ : V →+ W) + refine AddMonoidHomClass.isometry_of_norm _ ?_ + intro v + apply le_antisymm (i.hom.2 v) + calc + -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 + ‖v‖ = ‖i.inv (i.hom v)‖ := by erw [Iso.hom_inv_id_apply] + _ ≤ ‖i.hom v‖ := i.inv.2 _ +#align SemiNormedGroup₁.iso_isometry SemiNormedGrp₁.iso_isometry + +end SemiNormedGrp₁ diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Completion.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean similarity index 64% rename from Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Completion.lean rename to Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean index d2425ea92cf2e..7d141242ade07 100644 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Completion.lean +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Completion.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca, Johan Commelin -/ -import Mathlib.Analysis.Normed.Group.SemiNormedGroupCat +import Mathlib.Analysis.Normed.Group.SemiNormedGrp import Mathlib.CategoryTheory.Preadditive.AdditiveFunctor import Mathlib.Analysis.Normed.Group.HomCompletion @@ -17,17 +17,17 @@ objects and morphisms). ## Main definitions -- `SemiNormedGroupCat.Completion : SemiNormedGroupCat ⥤ SemiNormedGroupCat` : the completion of a - seminormed group (defined as a functor on `SemiNormedGroupCat` to itself). -- `SemiNormedGroupCat.Completion.lift (f : V ⟶ W) : (Completion.obj V ⟶ W)` : a normed group hom +- `SemiNormedGrp.Completion : SemiNormedGrp ⥤ SemiNormedGrp` : the completion of a + seminormed group (defined as a functor on `SemiNormedGrp` to itself). +- `SemiNormedGrp.Completion.lift (f : V ⟶ W) : (Completion.obj V ⟶ W)` : a normed group hom from `V` to complete `W` extends ("lifts") to a seminormed group hom from the completion of `V` to `W`. ## Projects -1. Construct the category of complete seminormed groups, say `CompleteSemiNormedGroupCat` +1. Construct the category of complete seminormed groups, say `CompleteSemiNormedGrp` and promote the `Completion` functor below to a functor landing in this category. -2. Prove that the functor `Completion : SemiNormedGroupCat ⥤ CompleteSemiNormedGroupCat` +2. Prove that the functor `Completion : SemiNormedGrp ⥤ CompleteSemiNormedGrp` is left adjoint to the forgetful functor. -/ @@ -40,64 +40,64 @@ open UniformSpace MulOpposite CategoryTheory NormedAddGroupHom set_option linter.uppercaseLean3 false -namespace SemiNormedGroupCat +namespace SemiNormedGrp -/-- The completion of a seminormed group, as an endofunctor on `SemiNormedGroupCat`. -/ +/-- The completion of a seminormed group, as an endofunctor on `SemiNormedGrp`. -/ @[simps] -def completion : SemiNormedGroupCat.{u} ⥤ SemiNormedGroupCat.{u} where - obj V := SemiNormedGroupCat.of (Completion V) +def completion : SemiNormedGrp.{u} ⥤ SemiNormedGrp.{u} where + obj V := SemiNormedGrp.of (Completion V) map f := f.completion map_id _ := completion_id map_comp f g := (completion_comp f g).symm -#align SemiNormedGroup.Completion SemiNormedGroupCat.completion +#align SemiNormedGroup.Completion SemiNormedGrp.completion -instance completion_completeSpace {V : SemiNormedGroupCat} : CompleteSpace (completion.obj V) := +instance completion_completeSpace {V : SemiNormedGrp} : CompleteSpace (completion.obj V) := Completion.completeSpace _ -#align SemiNormedGroup.Completion_complete_space SemiNormedGroupCat.completion_completeSpace +#align SemiNormedGroup.Completion_complete_space SemiNormedGrp.completion_completeSpace /-- The canonical morphism from a seminormed group `V` to its completion. -/ @[simps] -def completion.incl {V : SemiNormedGroupCat} : V ⟶ completion.obj V where +def completion.incl {V : SemiNormedGrp} : V ⟶ completion.obj V where toFun v := (v : Completion V) map_add' := Completion.coe_add bound' := ⟨1, fun v => by simp⟩ -#align SemiNormedGroup.Completion.incl SemiNormedGroupCat.completion.incl +#align SemiNormedGroup.Completion.incl SemiNormedGrp.completion.incl -- These lemmas have always been bad (#7657), but leanprover/lean4#2644 made `simp` start noticing -attribute [nolint simpNF] SemiNormedGroupCat.completion.incl_apply +attribute [nolint simpNF] SemiNormedGrp.completion.incl_apply -theorem completion.norm_incl_eq {V : SemiNormedGroupCat} {v : V} : ‖completion.incl v‖ = ‖v‖ := +theorem completion.norm_incl_eq {V : SemiNormedGrp} {v : V} : ‖completion.incl v‖ = ‖v‖ := UniformSpace.Completion.norm_coe _ -#align SemiNormedGroup.Completion.norm_incl_eq SemiNormedGroupCat.completion.norm_incl_eq +#align SemiNormedGroup.Completion.norm_incl_eq SemiNormedGrp.completion.norm_incl_eq -theorem completion.map_normNoninc {V W : SemiNormedGroupCat} {f : V ⟶ W} (hf : f.NormNoninc) : +theorem completion.map_normNoninc {V W : SemiNormedGrp} {f : V ⟶ W} (hf : f.NormNoninc) : (completion.map f).NormNoninc := NormedAddGroupHom.NormNoninc.normNoninc_iff_norm_le_one.2 <| (NormedAddGroupHom.norm_completion f).le.trans <| NormedAddGroupHom.NormNoninc.normNoninc_iff_norm_le_one.1 hf -#align SemiNormedGroup.Completion.map_norm_noninc SemiNormedGroupCat.completion.map_normNoninc +#align SemiNormedGroup.Completion.map_norm_noninc SemiNormedGrp.completion.map_normNoninc -variable (V W : SemiNormedGroupCat) +variable (V W : SemiNormedGrp) /-- Given a normed group hom `V ⟶ W`, this defines the associated morphism from the completion of `V` to the completion of `W`. The difference from the definition obtained from the functoriality of completion is in that the map sending a morphism `f` to the associated morphism of completions is itself additive. -/ -def completion.mapHom (V W : SemiNormedGroupCat.{u}) : +def completion.mapHom (V W : SemiNormedGrp.{u}) : -- Porting note: cannot see instances through concrete cats - have (V W : SemiNormedGroupCat.{u}) : AddGroup (V ⟶ W) := + have (V W : SemiNormedGrp.{u}) : AddGroup (V ⟶ W) := inferInstanceAs <| AddGroup <| NormedAddGroupHom V W (V ⟶ W) →+ (completion.obj V ⟶ completion.obj W) := @AddMonoidHom.mk' _ _ (_) (_) completion.map fun f g => f.completion_add g -#align SemiNormedGroup.Completion.map_hom SemiNormedGroupCat.completion.mapHom +#align SemiNormedGroup.Completion.map_hom SemiNormedGrp.completion.mapHom -- @[simp] -- Porting note: removed simp since LHS simplifies and is not used -theorem completion.map_zero (V W : SemiNormedGroupCat) : completion.map (0 : V ⟶ W) = 0 := +theorem completion.map_zero (V W : SemiNormedGrp) : completion.map (0 : V ⟶ W) = 0 := -- Porting note: cannot see instances through concrete cats @AddMonoidHom.map_zero _ _ (_) (_) (completion.mapHom V W) -#align SemiNormedGroup.Completion.map_zero SemiNormedGroupCat.completion.map_zero +#align SemiNormedGroup.Completion.map_zero SemiNormedGrp.completion.map_zero -instance : Preadditive SemiNormedGroupCat.{u} where +instance : Preadditive SemiNormedGrp.{u} where homGroup P Q := inferInstanceAs <| AddCommGroup <| NormedAddGroupHom P Q add_comp _ Q _ f f' g := by ext x @@ -125,21 +125,21 @@ instance : Functor.Additive completion where /-- Given a normed group hom `f : V → W` with `W` complete, this provides a lift of `f` to the completion of `V`. The lemmas `lift_unique` and `lift_comp_incl` provide the api for the universal property of the completion. -/ -def completion.lift {V W : SemiNormedGroupCat} [CompleteSpace W] [T0Space W] (f : V ⟶ W) : +def completion.lift {V W : SemiNormedGrp} [CompleteSpace W] [T0Space W] (f : V ⟶ W) : completion.obj V ⟶ W where toFun := f.extension map_add' := f.extension.toAddMonoidHom.map_add' bound' := f.extension.bound' -#align SemiNormedGroup.Completion.lift SemiNormedGroupCat.completion.lift +#align SemiNormedGroup.Completion.lift SemiNormedGrp.completion.lift -theorem completion.lift_comp_incl {V W : SemiNormedGroupCat} [CompleteSpace W] [T0Space W] +theorem completion.lift_comp_incl {V W : SemiNormedGrp} [CompleteSpace W] [T0Space W] (f : V ⟶ W) : completion.incl ≫ completion.lift f = f := ext <| NormedAddGroupHom.extension_coe _ -#align SemiNormedGroup.Completion.lift_comp_incl SemiNormedGroupCat.completion.lift_comp_incl +#align SemiNormedGroup.Completion.lift_comp_incl SemiNormedGrp.completion.lift_comp_incl -theorem completion.lift_unique {V W : SemiNormedGroupCat} [CompleteSpace W] [T0Space W] +theorem completion.lift_unique {V W : SemiNormedGrp} [CompleteSpace W] [T0Space W] (f : V ⟶ W) (g : completion.obj V ⟶ W) : completion.incl ≫ g = f → g = completion.lift f := fun h => (NormedAddGroupHom.extension_unique _ fun v => ((ext_iff.1 h) v).symm).symm -#align SemiNormedGroup.Completion.lift_unique SemiNormedGroupCat.completion.lift_unique +#align SemiNormedGroup.Completion.lift_unique SemiNormedGrp.completion.lift_unique -end SemiNormedGroupCat +end SemiNormedGrp diff --git a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean similarity index 66% rename from Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean rename to Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean index aa0e721b37973..b54b499fb08db 100644 --- a/Mathlib/Analysis/Normed/Group/SemiNormedGroupCat/Kernels.lean +++ b/Mathlib/Analysis/Normed/Group/SemiNormedGrp/Kernels.lean @@ -3,24 +3,24 @@ Copyright (c) 2021 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca, Johan Commelin, Scott Morrison -/ -import Mathlib.Analysis.Normed.Group.SemiNormedGroupCat +import Mathlib.Analysis.Normed.Group.SemiNormedGrp import Mathlib.Analysis.Normed.Group.Quotient import Mathlib.CategoryTheory.Limits.Shapes.Kernels #align_import analysis.normed.group.SemiNormedGroup.kernels from "leanprover-community/mathlib"@"17ef379e997badd73e5eabb4d38f11919ab3c4b3" /-! -# Kernels and cokernels in SemiNormedGroupCat₁ and SemiNormedGroupCat +# Kernels and cokernels in SemiNormedGrp₁ and SemiNormedGrp -We show that `SemiNormedGroupCat₁` has cokernels +We show that `SemiNormedGrp₁` has cokernels (for which of course the `cokernel.π f` maps are norm non-increasing), -as well as the easier result that `SemiNormedGroupCat` has cokernels. We also show that -`SemiNormedGroupCat` has kernels. +as well as the easier result that `SemiNormedGrp` has cokernels. We also show that +`SemiNormedGrp` has kernels. So far, I don't see a way to state nicely what we really want: -`SemiNormedGroupCat` has cokernels, and `cokernel.π f` is norm non-increasing. +`SemiNormedGrp` has cokernels, and `cokernel.π f` is norm non-increasing. The problem is that the limits API doesn't promise you any particular model of the cokernel, -and in `SemiNormedGroupCat` one can always take a cokernel and rescale its norm +and in `SemiNormedGrp` one can always take a cokernel and rescale its norm (and hence making `cokernel.π f` arbitrarily large in norm), obtaining another categorical cokernel. -/ @@ -30,32 +30,32 @@ open CategoryTheory CategoryTheory.Limits universe u -namespace SemiNormedGroupCat₁ +namespace SemiNormedGrp₁ noncomputable section -/-- Auxiliary definition for `HasCokernels SemiNormedGroupCat₁`. -/ -def cokernelCocone {X Y : SemiNormedGroupCat₁.{u}} (f : X ⟶ Y) : Cofork f 0 := +/-- Auxiliary definition for `HasCokernels SemiNormedGrp₁`. -/ +def cokernelCocone {X Y : SemiNormedGrp₁.{u}} (f : X ⟶ Y) : Cofork f 0 := Cofork.ofπ - (@SemiNormedGroupCat₁.mkHom _ (SemiNormedGroupCat.of (Y ⧸ NormedAddGroupHom.range f.1)) + (@SemiNormedGrp₁.mkHom _ (SemiNormedGrp.of (Y ⧸ NormedAddGroupHom.range f.1)) f.1.range.normedMk (NormedAddGroupHom.isQuotientQuotient _).norm_le) (by ext x -- Porting note(https://github.com/leanprover-community/mathlib4/issues/5026): was -- simp only [comp_apply, Limits.zero_comp, NormedAddGroupHom.zero_apply, - -- SemiNormedGroupCat₁.mkHom_apply, SemiNormedGroupCat₁.zero_apply, + -- SemiNormedGrp₁.mkHom_apply, SemiNormedGrp₁.zero_apply, -- ← NormedAddGroupHom.mem_ker, f.1.range.ker_normedMk, f.1.mem_range] -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 - erw [Limits.zero_comp, comp_apply, SemiNormedGroupCat₁.mkHom_apply, - SemiNormedGroupCat₁.zero_apply, ← NormedAddGroupHom.mem_ker, f.1.range.ker_normedMk, + erw [Limits.zero_comp, comp_apply, SemiNormedGrp₁.mkHom_apply, + SemiNormedGrp₁.zero_apply, ← NormedAddGroupHom.mem_ker, f.1.range.ker_normedMk, f.1.mem_range] use x rfl) set_option linter.uppercaseLean3 false in -#align SemiNormedGroup₁.cokernel_cocone SemiNormedGroupCat₁.cokernelCocone +#align SemiNormedGroup₁.cokernel_cocone SemiNormedGrp₁.cokernelCocone -/-- Auxiliary definition for `HasCokernels SemiNormedGroupCat₁`. -/ -def cokernelLift {X Y : SemiNormedGroupCat₁.{u}} (f : X ⟶ Y) (s : CokernelCofork f) : +/-- Auxiliary definition for `HasCokernels SemiNormedGrp₁`. -/ +def cokernelLift {X Y : SemiNormedGrp₁.{u}} (f : X ⟶ Y) (s : CokernelCofork f) : (cokernelCocone f).pt ⟶ s.pt := by fconstructor -- The lift itself: @@ -68,9 +68,9 @@ def cokernelLift {X Y : SemiNormedGroupCat₁.{u}} (f : X ⟶ Y) (s : CokernelCo -- The lift has norm at most one: exact NormedAddGroupHom.lift_normNoninc _ _ _ s.π.2 set_option linter.uppercaseLean3 false in -#align SemiNormedGroup₁.cokernel_lift SemiNormedGroupCat₁.cokernelLift +#align SemiNormedGroup₁.cokernel_lift SemiNormedGrp₁.cokernelLift -instance : HasCokernels SemiNormedGroupCat₁.{u} where +instance : HasCokernels SemiNormedGrp₁.{u} where has_colimit f := HasColimit.mk { cocone := cokernelCocone f @@ -89,25 +89,25 @@ instance : HasCokernels SemiNormedGroupCat₁.{u} where (NormedAddGroupHom.lift_unique f.1.range _ _ _ (congr_arg Subtype.val w : _)) } -- Sanity check -example : HasCokernels SemiNormedGroupCat₁ := by infer_instance +example : HasCokernels SemiNormedGrp₁ := by infer_instance end -end SemiNormedGroupCat₁ +end SemiNormedGrp₁ -namespace SemiNormedGroupCat +namespace SemiNormedGrp section EqualizersAndKernels -- Porting note: these weren't needed in Lean 3 -instance {V W : SemiNormedGroupCat.{u}} : Sub (V ⟶ W) := +instance {V W : SemiNormedGrp.{u}} : Sub (V ⟶ W) := (inferInstance : Sub (NormedAddGroupHom V W)) -noncomputable instance {V W : SemiNormedGroupCat.{u}} : Norm (V ⟶ W) := +noncomputable instance {V W : SemiNormedGrp.{u}} : Norm (V ⟶ W) := (inferInstance : Norm (NormedAddGroupHom V W)) -noncomputable instance {V W : SemiNormedGroupCat.{u}} : NNNorm (V ⟶ W) := +noncomputable instance {V W : SemiNormedGrp.{u}} : NNNorm (V ⟶ W) := (inferInstance : NNNorm (NormedAddGroupHom V W)) /-- The equalizer cone for a parallel pair of morphisms of seminormed groups. -/ -def fork {V W : SemiNormedGroupCat.{u}} (f g : V ⟶ W) : Fork f g := +def fork {V W : SemiNormedGrp.{u}} (f g : V ⟶ W) : Fork f g := @Fork.ofι _ _ _ _ _ _ (of (f - g : NormedAddGroupHom V W).ker) (NormedAddGroupHom.incl (f - g).ker) <| by -- Porting note: not needed in mathlib3 @@ -118,9 +118,9 @@ def fork {V W : SemiNormedGroupCat.{u}} (f g : V ⟶ W) : Fork f g := NormedAddGroupHom.mem_ker, NormedAddGroupHom.coe_sub, Pi.sub_apply, sub_eq_zero] using this set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.fork SemiNormedGroupCat.fork +#align SemiNormedGroup.fork SemiNormedGrp.fork -instance hasLimit_parallelPair {V W : SemiNormedGroupCat.{u}} (f g : V ⟶ W) : +instance hasLimit_parallelPair {V W : SemiNormedGrp.{u}} (f g : V ⟶ W) : HasLimit (parallelPair f g) where exists_limit := Nonempty.intro @@ -135,22 +135,22 @@ instance hasLimit_parallelPair {V W : SemiNormedGroupCat.{u}} (f g : V ⟶ W) : -- Porting note: the `simp_rw` was was `rw [← h]` but motive is not type correct in mathlib4 ext x; dsimp; simp_rw [← h]; rfl} set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.has_limit_parallel_pair SemiNormedGroupCat.hasLimit_parallelPair +#align SemiNormedGroup.has_limit_parallel_pair SemiNormedGrp.hasLimit_parallelPair -instance : Limits.HasEqualizers.{u, u + 1} SemiNormedGroupCat := - @hasEqualizers_of_hasLimit_parallelPair SemiNormedGroupCat _ fun {_ _ f g} => - SemiNormedGroupCat.hasLimit_parallelPair f g +instance : Limits.HasEqualizers.{u, u + 1} SemiNormedGrp := + @hasEqualizers_of_hasLimit_parallelPair SemiNormedGrp _ fun {_ _ f g} => + SemiNormedGrp.hasLimit_parallelPair f g end EqualizersAndKernels section Cokernel --- PROJECT: can we reuse the work to construct cokernels in `SemiNormedGroupCat₁` here? +-- PROJECT: can we reuse the work to construct cokernels in `SemiNormedGrp₁` here? -- I don't see a way to do this that is less work than just repeating the relevant parts. -/-- Auxiliary definition for `HasCokernels SemiNormedGroupCat`. -/ +/-- Auxiliary definition for `HasCokernels SemiNormedGrp`. -/ noncomputable -def cokernelCocone {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : Cofork f 0 := - @Cofork.ofπ _ _ _ _ _ _ (SemiNormedGroupCat.of (Y ⧸ NormedAddGroupHom.range f)) f.range.normedMk +def cokernelCocone {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : Cofork f 0 := + @Cofork.ofπ _ _ _ _ _ _ (SemiNormedGrp.of (Y ⧸ NormedAddGroupHom.range f)) f.range.normedMk (by ext a simp only [comp_apply, Limits.zero_comp] @@ -158,7 +158,7 @@ def cokernelCocone {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : Cofork f 0 := -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 erw [comp_apply, NormedAddGroupHom.zero_apply] -- Porting note: Lean 3 didn't need this instance - letI : SeminormedAddCommGroup ((forget SemiNormedGroupCat).obj Y) := + letI : SeminormedAddCommGroup ((forget SemiNormedGrp).obj Y) := (inferInstance : SeminormedAddCommGroup Y) -- Porting note: again simp doesn't seem to be firing in the below line -- This used to be `rw`, but we need `erw` after leanprover/lean4#2644 @@ -166,11 +166,11 @@ def cokernelCocone {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : Cofork f 0 := -- This used to be `simp only [exists_apply_eq_apply]` before leanprover/lean4#2644 convert exists_apply_eq_apply f a) set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.cokernel_cocone SemiNormedGroupCat.cokernelCocone +#align SemiNormedGroup.cokernel_cocone SemiNormedGrp.cokernelCocone -/-- Auxiliary definition for `HasCokernels SemiNormedGroupCat`. -/ +/-- Auxiliary definition for `HasCokernels SemiNormedGrp`. -/ noncomputable -def cokernelLift {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) (s : CokernelCofork f) : +def cokernelLift {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) (s : CokernelCofork f) : (cokernelCocone f).pt ⟶ s.pt := NormedAddGroupHom.lift _ s.π (by @@ -180,11 +180,11 @@ def cokernelLift {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) (s : CokernelCofor -- This used to be the end of the proof before leanprover/lean4#2644 erw [zero_apply]) set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.cokernel_lift SemiNormedGroupCat.cokernelLift +#align SemiNormedGroup.cokernel_lift SemiNormedGrp.cokernelLift -/-- Auxiliary definition for `HasCokernels SemiNormedGroupCat`. -/ +/-- Auxiliary definition for `HasCokernels SemiNormedGrp`. -/ noncomputable -def isColimitCokernelCocone {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +def isColimitCokernelCocone {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : IsColimit (cokernelCocone f) := isColimitAux _ (cokernelLift f) (fun s => by @@ -197,80 +197,80 @@ def isColimitCokernelCocone {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : erw [zero_apply]) fun s m w => NormedAddGroupHom.lift_unique f.range _ _ _ w set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.is_colimit_cokernel_cocone SemiNormedGroupCat.isColimitCokernelCocone +#align SemiNormedGroup.is_colimit_cokernel_cocone SemiNormedGrp.isColimitCokernelCocone -instance : HasCokernels SemiNormedGroupCat.{u} where +instance : HasCokernels SemiNormedGrp.{u} where has_colimit f := HasColimit.mk { cocone := cokernelCocone f isColimit := isColimitCokernelCocone f } -- Sanity check -example : HasCokernels SemiNormedGroupCat := by infer_instance +example : HasCokernels SemiNormedGrp := by infer_instance section ExplicitCokernel /-- An explicit choice of cokernel, which has good properties with respect to the norm. -/ noncomputable -def explicitCokernel {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : SemiNormedGroupCat.{u} := +def explicitCokernel {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : SemiNormedGrp.{u} := (cokernelCocone f).pt set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel SemiNormedGroupCat.explicitCokernel +#align SemiNormedGroup.explicit_cokernel SemiNormedGrp.explicitCokernel /-- Descend to the explicit cokernel. -/ noncomputable -def explicitCokernelDesc {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : +def explicitCokernelDesc {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : explicitCokernel f ⟶ Z := (isColimitCokernelCocone f).desc (Cofork.ofπ g (by simp [w])) set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc SemiNormedGroupCat.explicitCokernelDesc +#align SemiNormedGroup.explicit_cokernel_desc SemiNormedGrp.explicitCokernelDesc /-- The projection from `Y` to the explicit cokernel of `X ⟶ Y`. -/ noncomputable -def explicitCokernelπ {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : Y ⟶ explicitCokernel f := +def explicitCokernelπ {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : Y ⟶ explicitCokernel f := (cokernelCocone f).ι.app WalkingParallelPair.one set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_π SemiNormedGroupCat.explicitCokernelπ +#align SemiNormedGroup.explicit_cokernel_π SemiNormedGrp.explicitCokernelπ -theorem explicitCokernelπ_surjective {X Y : SemiNormedGroupCat.{u}} {f : X ⟶ Y} : +theorem explicitCokernelπ_surjective {X Y : SemiNormedGrp.{u}} {f : X ⟶ Y} : Function.Surjective (explicitCokernelπ f) := surjective_quot_mk _ set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_π_surjective SemiNormedGroupCat.explicitCokernelπ_surjective +#align SemiNormedGroup.explicit_cokernel_π_surjective SemiNormedGrp.explicitCokernelπ_surjective @[simp, reassoc] -theorem comp_explicitCokernelπ {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +theorem comp_explicitCokernelπ {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : f ≫ explicitCokernelπ f = 0 := by convert (cokernelCocone f).w WalkingParallelPairHom.left simp set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.comp_explicit_cokernel_π SemiNormedGroupCat.comp_explicitCokernelπ +#align SemiNormedGroup.comp_explicit_cokernel_π SemiNormedGrp.comp_explicitCokernelπ -- Porting note: wasn't necessary in Lean 3. Is this a bug? attribute [simp] comp_explicitCokernelπ_assoc @[simp] -theorem explicitCokernelπ_apply_dom_eq_zero {X Y : SemiNormedGroupCat.{u}} {f : X ⟶ Y} (x : X) : +theorem explicitCokernelπ_apply_dom_eq_zero {X Y : SemiNormedGrp.{u}} {f : X ⟶ Y} (x : X) : (explicitCokernelπ f) (f x) = 0 := show (f ≫ explicitCokernelπ f) x = 0 by rw [comp_explicitCokernelπ]; rfl set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_π_apply_dom_eq_zero SemiNormedGroupCat.explicitCokernelπ_apply_dom_eq_zero +#align SemiNormedGroup.explicit_cokernel_π_apply_dom_eq_zero SemiNormedGrp.explicitCokernelπ_apply_dom_eq_zero @[simp, reassoc] -theorem explicitCokernelπ_desc {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelπ_desc {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : explicitCokernelπ f ≫ explicitCokernelDesc w = g := (isColimitCokernelCocone f).fac _ _ set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_π_desc SemiNormedGroupCat.explicitCokernelπ_desc +#align SemiNormedGroup.explicit_cokernel_π_desc SemiNormedGrp.explicitCokernelπ_desc @[simp] -theorem explicitCokernelπ_desc_apply {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelπ_desc_apply {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} {cond : f ≫ g = 0} (x : Y) : explicitCokernelDesc cond (explicitCokernelπ f x) = g x := show (explicitCokernelπ f ≫ explicitCokernelDesc cond) x = g x by rw [explicitCokernelπ_desc] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_π_desc_apply SemiNormedGroupCat.explicitCokernelπ_desc_apply +#align SemiNormedGroup.explicit_cokernel_π_desc_apply SemiNormedGrp.explicitCokernelπ_desc_apply -theorem explicitCokernelDesc_unique {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelDesc_unique {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) (e : explicitCokernel f ⟶ Z) (he : explicitCokernelπ f ≫ e = g) : e = explicitCokernelDesc w := by apply (isColimitCokernelCocone f).uniq (Cofork.ofπ g (by simp [w])) @@ -279,9 +279,9 @@ theorem explicitCokernelDesc_unique {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ simp · exact he set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_unique SemiNormedGroupCat.explicitCokernelDesc_unique +#align SemiNormedGroup.explicit_cokernel_desc_unique SemiNormedGrp.explicitCokernelDesc_unique -theorem explicitCokernelDesc_comp_eq_desc {X Y Z W : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelDesc_comp_eq_desc {X Y Z W : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} -- Porting note: renamed `cond` to `cond'` to avoid -- failed to rewrite using equation theorems for 'cond' {h : Z ⟶ W} {cond' : f ≫ g = 0} : @@ -291,17 +291,17 @@ theorem explicitCokernelDesc_comp_eq_desc {X Y Z W : SemiNormedGroupCat.{u}} {f refine explicitCokernelDesc_unique _ _ ?_ rw [← CategoryTheory.Category.assoc, explicitCokernelπ_desc] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_comp_eq_desc SemiNormedGroupCat.explicitCokernelDesc_comp_eq_desc +#align SemiNormedGroup.explicit_cokernel_desc_comp_eq_desc SemiNormedGrp.explicitCokernelDesc_comp_eq_desc @[simp] -theorem explicitCokernelDesc_zero {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} : +theorem explicitCokernelDesc_zero {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} : explicitCokernelDesc (show f ≫ (0 : Y ⟶ Z) = 0 from CategoryTheory.Limits.comp_zero) = 0 := Eq.symm <| explicitCokernelDesc_unique _ _ CategoryTheory.Limits.comp_zero set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_zero SemiNormedGroupCat.explicitCokernelDesc_zero +#align SemiNormedGroup.explicit_cokernel_desc_zero SemiNormedGrp.explicitCokernelDesc_zero @[ext] -theorem explicitCokernel_hom_ext {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} +theorem explicitCokernel_hom_ext {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} (e₁ e₂ : explicitCokernel f ⟶ Z) (h : explicitCokernelπ f ≫ e₁ = explicitCokernelπ f ≫ e₂) : e₁ = e₂ := by let g : Y ⟶ Z := explicitCokernelπ f ≫ e₂ @@ -311,9 +311,9 @@ theorem explicitCokernel_hom_ext {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} apply explicitCokernelDesc_unique exact h set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_hom_ext SemiNormedGroupCat.explicitCokernel_hom_ext +#align SemiNormedGroup.explicit_cokernel_hom_ext SemiNormedGrp.explicitCokernel_hom_ext -instance explicitCokernelπ.epi {X Y : SemiNormedGroupCat.{u}} {f : X ⟶ Y} : +instance explicitCokernelπ.epi {X Y : SemiNormedGrp.{u}} {f : X ⟶ Y} : Epi (explicitCokernelπ f) := by constructor intro Z g h H @@ -323,29 +323,29 @@ instance explicitCokernelπ.epi {X Y : SemiNormedGroupCat.{u}} {f : X ⟶ Y} : -- change (explicitCokernelπ f ≫ g) _ = _ rw [H] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_π.epi SemiNormedGroupCat.explicitCokernelπ.epi +#align SemiNormedGroup.explicit_cokernel_π.epi SemiNormedGrp.explicitCokernelπ.epi -theorem isQuotient_explicitCokernelπ {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +theorem isQuotient_explicitCokernelπ {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : NormedAddGroupHom.IsQuotient (explicitCokernelπ f) := NormedAddGroupHom.isQuotientQuotient _ set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.is_quotient_explicit_cokernel_π SemiNormedGroupCat.isQuotient_explicitCokernelπ +#align SemiNormedGroup.is_quotient_explicit_cokernel_π SemiNormedGrp.isQuotient_explicitCokernelπ -theorem normNoninc_explicitCokernelπ {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +theorem normNoninc_explicitCokernelπ {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : (explicitCokernelπ f).NormNoninc := (isQuotient_explicitCokernelπ f).norm_le set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.norm_noninc_explicit_cokernel_π SemiNormedGroupCat.normNoninc_explicitCokernelπ +#align SemiNormedGroup.norm_noninc_explicit_cokernel_π SemiNormedGrp.normNoninc_explicitCokernelπ open scoped NNReal -theorem explicitCokernelDesc_norm_le_of_norm_le {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} +theorem explicitCokernelDesc_norm_le_of_norm_le {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) (c : ℝ≥0) (h : ‖g‖ ≤ c) : ‖explicitCokernelDesc w‖ ≤ c := NormedAddGroupHom.lift_norm_le _ _ _ h set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_norm_le_of_norm_le SemiNormedGroupCat.explicitCokernelDesc_norm_le_of_norm_le +#align SemiNormedGroup.explicit_cokernel_desc_norm_le_of_norm_le SemiNormedGrp.explicitCokernelDesc_norm_le_of_norm_le -theorem explicitCokernelDesc_normNoninc {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelDesc_normNoninc {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} {cond : f ≫ g = 0} (hg : g.NormNoninc) : (explicitCokernelDesc cond).NormNoninc := by refine NormedAddGroupHom.NormNoninc.normNoninc_iff_norm_le_one.2 ?_ rw [← NNReal.coe_one] @@ -353,63 +353,63 @@ theorem explicitCokernelDesc_normNoninc {X Y Z : SemiNormedGroupCat.{u}} {f : X explicitCokernelDesc_norm_le_of_norm_le cond 1 (NormedAddGroupHom.NormNoninc.normNoninc_iff_norm_le_one.1 hg) set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_norm_noninc SemiNormedGroupCat.explicitCokernelDesc_normNoninc +#align SemiNormedGroup.explicit_cokernel_desc_norm_noninc SemiNormedGrp.explicitCokernelDesc_normNoninc -theorem explicitCokernelDesc_comp_eq_zero {X Y Z W : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelDesc_comp_eq_zero {X Y Z W : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} {h : Z ⟶ W} (cond : f ≫ g = 0) (cond2 : g ≫ h = 0) : explicitCokernelDesc cond ≫ h = 0 := by rw [← cancel_epi (explicitCokernelπ f), ← Category.assoc, explicitCokernelπ_desc] simp [cond2] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_comp_eq_zero SemiNormedGroupCat.explicitCokernelDesc_comp_eq_zero +#align SemiNormedGroup.explicit_cokernel_desc_comp_eq_zero SemiNormedGrp.explicitCokernelDesc_comp_eq_zero -theorem explicitCokernelDesc_norm_le {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelDesc_norm_le {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : ‖explicitCokernelDesc w‖ ≤ ‖g‖ := explicitCokernelDesc_norm_le_of_norm_le w ‖g‖₊ le_rfl set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_desc_norm_le SemiNormedGroupCat.explicitCokernelDesc_norm_le +#align SemiNormedGroup.explicit_cokernel_desc_norm_le SemiNormedGrp.explicitCokernelDesc_norm_le /-- The explicit cokernel is isomorphic to the usual cokernel. -/ noncomputable -def explicitCokernelIso {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +def explicitCokernelIso {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : explicitCokernel f ≅ cokernel f := (isColimitCokernelCocone f).coconePointUniqueUpToIso (colimit.isColimit _) set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_iso SemiNormedGroupCat.explicitCokernelIso +#align SemiNormedGroup.explicit_cokernel_iso SemiNormedGrp.explicitCokernelIso @[simp] -theorem explicitCokernelIso_hom_π {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +theorem explicitCokernelIso_hom_π {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : explicitCokernelπ f ≫ (explicitCokernelIso f).hom = cokernel.π _ := by simp [explicitCokernelπ, explicitCokernelIso, IsColimit.coconePointUniqueUpToIso] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_iso_hom_π SemiNormedGroupCat.explicitCokernelIso_hom_π +#align SemiNormedGroup.explicit_cokernel_iso_hom_π SemiNormedGrp.explicitCokernelIso_hom_π @[simp] -theorem explicitCokernelIso_inv_π {X Y : SemiNormedGroupCat.{u}} (f : X ⟶ Y) : +theorem explicitCokernelIso_inv_π {X Y : SemiNormedGrp.{u}} (f : X ⟶ Y) : cokernel.π f ≫ (explicitCokernelIso f).inv = explicitCokernelπ f := by simp [explicitCokernelπ, explicitCokernelIso] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_iso_inv_π SemiNormedGroupCat.explicitCokernelIso_inv_π +#align SemiNormedGroup.explicit_cokernel_iso_inv_π SemiNormedGrp.explicitCokernelIso_inv_π @[simp] -theorem explicitCokernelIso_hom_desc {X Y Z : SemiNormedGroupCat.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} +theorem explicitCokernelIso_hom_desc {X Y Z : SemiNormedGrp.{u}} {f : X ⟶ Y} {g : Y ⟶ Z} (w : f ≫ g = 0) : (explicitCokernelIso f).hom ≫ cokernel.desc f g w = explicitCokernelDesc w := by ext1 simp [explicitCokernelDesc, explicitCokernelπ, explicitCokernelIso, IsColimit.coconePointUniqueUpToIso] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel_iso_hom_desc SemiNormedGroupCat.explicitCokernelIso_hom_desc +#align SemiNormedGroup.explicit_cokernel_iso_hom_desc SemiNormedGrp.explicitCokernelIso_hom_desc /-- A special case of `CategoryTheory.Limits.cokernel.map` adapted to `explicitCokernel`. -/ -noncomputable def explicitCokernel.map {A B C D : SemiNormedGroupCat.{u}} +noncomputable def explicitCokernel.map {A B C D : SemiNormedGrp.{u}} {fab : A ⟶ B} {fbd : B ⟶ D} {fac : A ⟶ C} {fcd : C ⟶ D} (h : fab ≫ fbd = fac ≫ fcd) : explicitCokernel fab ⟶ explicitCokernel fcd := @explicitCokernelDesc _ _ _ fab (fbd ≫ explicitCokernelπ _) <| by simp [reassoc_of% h] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_cokernel.map SemiNormedGroupCat.explicitCokernel.map +#align SemiNormedGroup.explicit_cokernel.map SemiNormedGrp.explicitCokernel.map /-- A special case of `CategoryTheory.Limits.cokernel.map_desc` adapted to `explicitCokernel`. -/ -theorem ExplicitCoker.map_desc {A B C D B' D' : SemiNormedGroupCat.{u}} +theorem ExplicitCoker.map_desc {A B C D B' D' : SemiNormedGrp.{u}} {fab : A ⟶ B} {fbd : B ⟶ D} {fac : A ⟶ C} {fcd : C ⟶ D} {h : fab ≫ fbd = fac ≫ fcd} {fbb' : B ⟶ B'} {fdd' : D ⟶ D'} {condb : fab ≫ fbb' = 0} {condd : fcd ≫ fdd' = 0} {g : B' ⟶ D'} (h' : fbb' ≫ g = fbd ≫ fdd') : @@ -418,10 +418,10 @@ theorem ExplicitCoker.map_desc {A B C D B' D' : SemiNormedGroupCat.{u}} simp [← cancel_epi (explicitCokernelπ fab), ← Category.assoc, Category.assoc, explicitCokernelπ_desc, h'] set_option linter.uppercaseLean3 false in -#align SemiNormedGroup.explicit_coker.map_desc SemiNormedGroupCat.ExplicitCoker.map_desc +#align SemiNormedGroup.explicit_coker.map_desc SemiNormedGrp.ExplicitCoker.map_desc end ExplicitCokernel end Cokernel -end SemiNormedGroupCat +end SemiNormedGrp diff --git a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean index 5249e60a9c9f6..7124c23ab3b9f 100644 --- a/Mathlib/CategoryTheory/Galois/Prorepresentability.lean +++ b/Mathlib/CategoryTheory/Galois/Prorepresentability.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Christian Merten. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Christian Merten -/ -import Mathlib.Algebra.Category.GroupCat.Limits +import Mathlib.Algebra.Category.Grp.Limits import Mathlib.CategoryTheory.CofilteredSystem import Mathlib.CategoryTheory.Galois.Decomposition import Mathlib.CategoryTheory.Limits.FunctorCategory @@ -178,8 +178,8 @@ open PointedGaloisObject /-- The diagram sending each pointed Galois object to its automorphism group as an object of `C`. -/ @[simps] -noncomputable def autGaloisSystem : PointedGaloisObject F ⥤ GroupCat.{u₂} where - obj := fun A ↦ GroupCat.of <| Aut (A : C) +noncomputable def autGaloisSystem : PointedGaloisObject F ⥤ Grp.{u₂} where + obj := fun A ↦ Grp.of <| Aut (A : C) map := fun {A B} f ↦ (autMapHom f : Aut (A : C) →* Aut (B : C)) map_id := fun A ↦ by ext (σ : Aut A.obj) @@ -198,7 +198,7 @@ noncomputable instance : Group (AutGalois F) := /-- The canonical projection from `AutGalois F` to the `C`-automorphism group of each pointed Galois object. -/ noncomputable def AutGalois.π (A : PointedGaloisObject F) : AutGalois F →* Aut (A : C) := - GroupCat.sectionsπMonoidHom (autGaloisSystem F) A + Grp.sectionsπMonoidHom (autGaloisSystem F) A /- Not a `simp` lemma, because we usually don't want to expose the internals here. -/ lemma AutGalois.π_apply (A : PointedGaloisObject F) (x : AutGalois F) : @@ -253,7 +253,7 @@ We first establish the isomorphism between `End F` and `AutGalois F`, from which - `endEquivAutGalois : End F ≅ AutGalois F`: this is the composition of `endEquivSectionsFibers` with: - `(incl F ⋙ F).sections ≅ (autGaloisSystem F ⋙ forget GroupCat).sections` + `(incl F ⋙ F).sections ≅ (autGaloisSystem F ⋙ forget Grp).sections` which is induced from the level-wise equivalence `Aut A ≃ F.obj A` for a Galois object `A`. @@ -293,7 +293,7 @@ lemma endEquivSectionsFibers_π (f : End F) (A : PointedGaloisObject F) : /-- Functorial isomorphism `Aut A ≅ F.obj A` for Galois objects `A`. -/ noncomputable def autIsoFibers : - autGaloisSystem F ⋙ forget GroupCat ≅ incl F ⋙ F' := + autGaloisSystem F ⋙ forget Grp ≅ incl F ⋙ F' := NatIso.ofComponents (fun A ↦ ((evaluationEquivOfIsGalois F A A.pt).toIso)) (fun {A B} f ↦ by ext (φ : Aut A.obj) diff --git a/Mathlib/CategoryTheory/Linear/Yoneda.lean b/Mathlib/CategoryTheory/Linear/Yoneda.lean index b72aec118fa8b..0270836948027 100644 --- a/Mathlib/CategoryTheory/Linear/Yoneda.lean +++ b/Mathlib/CategoryTheory/Linear/Yoneda.lean @@ -73,7 +73,7 @@ theorem whiskering_linearYoneda : @[simp] theorem whiskering_linearYoneda₂ : - linearYoneda R C ⋙ (whiskeringRight _ _ _).obj (forget₂ (ModuleCat.{v} R) AddCommGroupCat.{v}) = + linearYoneda R C ⋙ (whiskeringRight _ _ _).obj (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) = preadditiveYoneda := rfl #align category_theory.whiskering_linear_yoneda₂ CategoryTheory.whiskering_linearYoneda₂ @@ -87,7 +87,7 @@ theorem whiskering_linearCoyoneda : @[simp] theorem whiskering_linearCoyoneda₂ : linearCoyoneda R C ⋙ - (whiskeringRight _ _ _).obj (forget₂ (ModuleCat.{v} R) AddCommGroupCat.{v}) = + (whiskeringRight _ _ _).obj (forget₂ (ModuleCat.{v} R) AddCommGrp.{v}) = preadditiveCoyoneda := rfl #align category_theory.whiskering_linear_coyoneda₂ CategoryTheory.whiskering_linearCoyoneda₂ diff --git a/Mathlib/CategoryTheory/Monoidal/Mon_.lean b/Mathlib/CategoryTheory/Monoidal/Mon_.lean index 03bba6f31ab8a..9829ab8934f72 100644 --- a/Mathlib/CategoryTheory/Monoidal/Mon_.lean +++ b/Mathlib/CategoryTheory/Monoidal/Mon_.lean @@ -606,7 +606,7 @@ Projects: available in mathlib3#3463) * More generally, check that `Mon_ (Mon_ C) ≌ CommMon_ C` when `C` is braided. * Check that `Mon_ TopCat ≌ [bundled topological monoids]`. -* Check that `Mon_ AddCommGroupCat ≌ RingCat`. +* Check that `Mon_ AddCommGrp ≌ RingCat`. (We've already got `Mon_ (ModuleCat R) ≌ AlgebraCat R`, in `Mathlib.CategoryTheory.Monoidal.Internal.Module`.) * Can you transport this monoidal structure to `RingCat` or `AlgebraCat R`? diff --git a/Mathlib/CategoryTheory/Preadditive/Generator.lean b/Mathlib/CategoryTheory/Preadditive/Generator.lean index 6e1ed46b1f049..21b06de38950a 100644 --- a/Mathlib/CategoryTheory/Preadditive/Generator.lean +++ b/Mathlib/CategoryTheory/Preadditive/Generator.lean @@ -55,14 +55,14 @@ theorem isSeparator_iff_faithful_preadditiveCoyoneda (G : C) : IsSeparator G ↔ (preadditiveCoyoneda.obj (op G)).Faithful := by rw [isSeparator_iff_faithful_coyoneda_obj, ← whiskering_preadditiveCoyoneda, Functor.comp_obj, whiskeringRight_obj_obj] - exact ⟨fun h => Functor.Faithful.of_comp _ (forget AddCommGroupCat), + exact ⟨fun h => Functor.Faithful.of_comp _ (forget AddCommGrp), fun h => Functor.Faithful.comp _ _⟩ #align category_theory.is_separator_iff_faithful_preadditive_coyoneda CategoryTheory.isSeparator_iff_faithful_preadditiveCoyoneda theorem isSeparator_iff_faithful_preadditiveCoyonedaObj (G : C) : IsSeparator G ↔ (preadditiveCoyonedaObj (op G)).Faithful := by rw [isSeparator_iff_faithful_preadditiveCoyoneda, preadditiveCoyoneda_obj] - exact ⟨fun h => Functor.Faithful.of_comp _ (forget₂ _ AddCommGroupCat.{v}), + exact ⟨fun h => Functor.Faithful.of_comp _ (forget₂ _ AddCommGrp.{v}), fun h => Functor.Faithful.comp _ _⟩ #align category_theory.is_separator_iff_faithful_preadditive_coyoneda_obj CategoryTheory.isSeparator_iff_faithful_preadditiveCoyonedaObj @@ -70,14 +70,14 @@ theorem isCoseparator_iff_faithful_preadditiveYoneda (G : C) : IsCoseparator G ↔ (preadditiveYoneda.obj G).Faithful := by rw [isCoseparator_iff_faithful_yoneda_obj, ← whiskering_preadditiveYoneda, Functor.comp_obj, whiskeringRight_obj_obj] - exact ⟨fun h => Functor.Faithful.of_comp _ (forget AddCommGroupCat), + exact ⟨fun h => Functor.Faithful.of_comp _ (forget AddCommGrp), fun h => Functor.Faithful.comp _ _⟩ #align category_theory.is_coseparator_iff_faithful_preadditive_yoneda CategoryTheory.isCoseparator_iff_faithful_preadditiveYoneda theorem isCoseparator_iff_faithful_preadditiveYonedaObj (G : C) : IsCoseparator G ↔ (preadditiveYonedaObj G).Faithful := by rw [isCoseparator_iff_faithful_preadditiveYoneda, preadditiveYoneda_obj] - exact ⟨fun h => Functor.Faithful.of_comp _ (forget₂ _ AddCommGroupCat.{v}), + exact ⟨fun h => Functor.Faithful.of_comp _ (forget₂ _ AddCommGrp.{v}), fun h => Functor.Faithful.comp _ _⟩ #align category_theory.is_coseparator_iff_faithful_preadditive_yoneda_obj CategoryTheory.isCoseparator_iff_faithful_preadditiveYonedaObj diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean index 937d9e1307a29..59d4008d32b1f 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Basic.lean @@ -5,7 +5,7 @@ Authors: Markus Himmel -/ import Mathlib.CategoryTheory.Preadditive.Opposite import Mathlib.Algebra.Category.ModuleCat.Basic -import Mathlib.Algebra.Category.GroupCat.Preadditive +import Mathlib.Algebra.Category.Grp.Preadditive #align_import category_theory.preadditive.yoneda.basic from "leanprover-community/mathlib"@"09f981f72d43749f1fa072deade828d9c1e185bb" @@ -52,14 +52,14 @@ object `X` to the group of morphisms `X ⟶ Y`. At each point, we get an additio structure, see `preadditiveYonedaObj`. -/ @[simps] -def preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGroupCat.{v} where +def preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGrp.{v} where obj Y := preadditiveYonedaObj Y ⋙ forget₂ _ _ map f := { app := fun X => { toFun := fun g => g ≫ f map_zero' := Limits.zero_comp map_add' := fun g g' => add_comp _ _ _ _ _ _ } - naturality := fun X X' g => AddCommGroupCat.ext fun x => Category.assoc _ _ _ } + naturality := fun X X' g => AddCommGrp.ext fun x => Category.assoc _ _ _ } map_id _ := by ext; dsimp; simp map_comp f g := by ext; dsimp; simp #align category_theory.preadditive_yoneda CategoryTheory.preadditiveYoneda @@ -81,7 +81,7 @@ object `Y` to the group of morphisms `X ⟶ Y`. At each point, we get an additio structure, see `preadditiveCoyonedaObj`. -/ @[simps] -def preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGroupCat.{v} where +def preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGrp.{v} where obj X := preadditiveCoyonedaObj X ⋙ forget₂ _ _ map f := { app := fun Y => @@ -89,7 +89,7 @@ def preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGroupCat.{v} where map_zero' := Limits.comp_zero map_add' := fun g g' => comp_add _ _ _ _ _ _ } naturality := fun Y Y' g => - AddCommGroupCat.ext fun x => Eq.symm <| Category.assoc _ _ _ } + AddCommGrp.ext fun x => Eq.symm <| Category.assoc _ _ _ } map_id _ := by ext; dsimp; simp map_comp f g := by ext; dsimp; simp #align category_theory.preadditive_coyoneda CategoryTheory.preadditiveCoyoneda @@ -116,7 +116,7 @@ Yoneda embedding. @[simp] theorem whiskering_preadditiveYoneda : preadditiveYoneda ⋙ - (whiskeringRight Cᵒᵖ AddCommGroupCat (Type v)).obj (forget AddCommGroupCat) = + (whiskeringRight Cᵒᵖ AddCommGrp (Type v)).obj (forget AddCommGrp) = yoneda := rfl #align category_theory.whiskering_preadditive_yoneda CategoryTheory.whiskering_preadditiveYoneda @@ -127,33 +127,33 @@ Yoneda embedding. @[simp] theorem whiskering_preadditiveCoyoneda : preadditiveCoyoneda ⋙ - (whiskeringRight C AddCommGroupCat (Type v)).obj (forget AddCommGroupCat) = + (whiskeringRight C AddCommGrp (Type v)).obj (forget AddCommGrp) = coyoneda := rfl #align category_theory.whiskering_preadditive_coyoneda CategoryTheory.whiskering_preadditiveCoyoneda -instance full_preadditiveYoneda : (preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGroupCat).Full := +instance full_preadditiveYoneda : (preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGrp).Full := let _ : Functor.Full (preadditiveYoneda ⋙ - (whiskeringRight Cᵒᵖ AddCommGroupCat (Type v)).obj (forget AddCommGroupCat)) := + (whiskeringRight Cᵒᵖ AddCommGrp (Type v)).obj (forget AddCommGrp)) := Yoneda.yoneda_full Functor.Full.of_comp_faithful preadditiveYoneda - ((whiskeringRight Cᵒᵖ AddCommGroupCat (Type v)).obj (forget AddCommGroupCat)) + ((whiskeringRight Cᵒᵖ AddCommGrp (Type v)).obj (forget AddCommGrp)) #align category_theory.preadditive_yoneda_full CategoryTheory.full_preadditiveYoneda -instance full_preadditiveCoyoneda : (preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGroupCat).Full := +instance full_preadditiveCoyoneda : (preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGrp).Full := let _ : Functor.Full (preadditiveCoyoneda ⋙ - (whiskeringRight C AddCommGroupCat (Type v)).obj (forget AddCommGroupCat)) := + (whiskeringRight C AddCommGrp (Type v)).obj (forget AddCommGrp)) := Coyoneda.coyoneda_full Functor.Full.of_comp_faithful preadditiveCoyoneda - ((whiskeringRight C AddCommGroupCat (Type v)).obj (forget AddCommGroupCat)) + ((whiskeringRight C AddCommGrp (Type v)).obj (forget AddCommGrp)) #align category_theory.preadditive_coyoneda_full CategoryTheory.full_preadditiveCoyoneda -instance faithful_preadditiveYoneda : (preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGroupCat).Faithful := +instance faithful_preadditiveYoneda : (preadditiveYoneda : C ⥤ Cᵒᵖ ⥤ AddCommGrp).Faithful := Functor.Faithful.of_comp_eq whiskering_preadditiveYoneda #align category_theory.preadditive_yoneda_faithful CategoryTheory.faithful_preadditiveYoneda instance faithful_preadditiveCoyoneda : - (preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGroupCat).Faithful := + (preadditiveCoyoneda : Cᵒᵖ ⥤ C ⥤ AddCommGrp).Faithful := Functor.Faithful.of_comp_eq whiskering_preadditiveCoyoneda #align category_theory.preadditive_coyoneda_faithful CategoryTheory.faithful_preadditiveCoyoneda diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Injective.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Injective.lean index f94443b25637a..987980162f5dd 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Injective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Injective.lean @@ -5,7 +5,7 @@ Authors: Markus Himmel, Scott Morrison -/ import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic import Mathlib.CategoryTheory.Preadditive.Injective -import Mathlib.Algebra.Category.GroupCat.EpiMono +import Mathlib.Algebra.Category.Grp.EpiMono import Mathlib.Algebra.Category.ModuleCat.EpiMono #align_import category_theory.preadditive.yoneda.injective from "leanprover-community/mathlib"@"f8d8465c3c392a93b9ed226956e26dee00975946" @@ -33,7 +33,7 @@ theorem injective_iff_preservesEpimorphisms_preadditiveYoneda_obj (J : C) : Injective J ↔ (preadditiveYoneda.obj J).PreservesEpimorphisms := by rw [injective_iff_preservesEpimorphisms_yoneda_obj] refine - ⟨fun h : (preadditiveYoneda.obj J ⋙ (forget AddCommGroupCat)).PreservesEpimorphisms => ?_, ?_⟩ + ⟨fun h : (preadditiveYoneda.obj J ⋙ (forget AddCommGrp)).PreservesEpimorphisms => ?_, ?_⟩ · exact Functor.preservesEpimorphisms_of_preserves_of_reflects (preadditiveYoneda.obj J) (forget _) · intro diff --git a/Mathlib/CategoryTheory/Preadditive/Yoneda/Projective.lean b/Mathlib/CategoryTheory/Preadditive/Yoneda/Projective.lean index 049878e82ee07..836398dc7163c 100644 --- a/Mathlib/CategoryTheory/Preadditive/Yoneda/Projective.lean +++ b/Mathlib/CategoryTheory/Preadditive/Yoneda/Projective.lean @@ -5,7 +5,7 @@ Authors: Markus Himmel, Scott Morrison -/ import Mathlib.CategoryTheory.Preadditive.Yoneda.Basic import Mathlib.CategoryTheory.Preadditive.Projective -import Mathlib.Algebra.Category.GroupCat.EpiMono +import Mathlib.Algebra.Category.Grp.EpiMono #align_import category_theory.preadditive.yoneda.projective from "leanprover-community/mathlib"@"f8d8465c3c392a93b9ed226956e26dee00975946" @@ -32,7 +32,7 @@ theorem projective_iff_preservesEpimorphisms_preadditiveCoyoneda_obj (P : C) : Projective P ↔ (preadditiveCoyoneda.obj (op P)).PreservesEpimorphisms := by rw [projective_iff_preservesEpimorphisms_coyoneda_obj] refine ⟨fun h : (preadditiveCoyoneda.obj (op P) ⋙ - forget AddCommGroupCat).PreservesEpimorphisms => ?_, ?_⟩ + forget AddCommGrp).PreservesEpimorphisms => ?_, ?_⟩ · exact Functor.preservesEpimorphisms_of_preserves_of_reflects (preadditiveCoyoneda.obj (op P)) (forget _) · intro @@ -43,7 +43,7 @@ theorem projective_iff_preservesEpimorphisms_preadditiveCoyoneda_obj' (P : C) : Projective P ↔ (preadditiveCoyoneda.obj (op P)).PreservesEpimorphisms := by rw [projective_iff_preservesEpimorphisms_coyoneda_obj] refine ⟨fun h : (preadditiveCoyoneda.obj (op P) ⋙ - forget AddCommGroupCat).PreservesEpimorphisms => ?_, ?_⟩ + forget AddCommGrp).PreservesEpimorphisms => ?_, ?_⟩ · exact Functor.preservesEpimorphisms_of_preserves_of_reflects (preadditiveCoyoneda.obj (op P)) (forget _) · intro diff --git a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean index 3028cd0ab1e68..52563bacbae85 100644 --- a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean +++ b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean @@ -3,14 +3,14 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic /-! The cohomology of a sheaf of groups in degree 1 In this file, we shall define the cohomology in degree 1 of a sheaf of groups (TODO). -Currently, given a presheaf of groups `G : Cᵒᵖ ⥤ GroupCat` and a family +Currently, given a presheaf of groups `G : Cᵒᵖ ⥤ Grp` and a family of objects `U : I → C`, we define 1-cochains/1-cocycles/H^1 with values in `G` over `U`. (This definition neither requires the assumption that `G` is a sheaf, nor that `U` covers the terminal object.) @@ -48,7 +48,7 @@ variable {C : Type u} [Category.{v} C] namespace PresheafOfGroups -variable (G : Cᵒᵖ ⥤ GroupCat.{w}) {X : C} {I : Type w'} (U : I → C) +variable (G : Cᵒᵖ ⥤ Grp.{w}) {X : C} {I : Type w'} (U : I → C) /-- A zero cochain consists of a family of sections. -/ def ZeroCochain := ∀ (i : I), G.obj (Opposite.op (U i)) @@ -68,7 +68,7 @@ lemma mul_apply (γ₁ γ₂ : ZeroCochain G U) (i : I) : (γ₁ * γ₂) i = γ end Cochain₀ -/-- A 1-cochain of a presheaf of groups `G : Cᵒᵖ ⥤ GroupCat` on a family `U : I → C` of objects +/-- A 1-cochain of a presheaf of groups `G : Cᵒᵖ ⥤ Grp` on a family `U : I → C` of objects consists of the data of an element in `G.obj (Opposite.op T)` whenever we have elements `i` and `j` in `I` and maps `a : T ⟶ U i` and `b : T ⟶ U j`, and it must satisfy a compatibility with respect to precomposition. (When the binary product of `U i` and `U j` exists, this @@ -189,7 +189,7 @@ end OneCocycle variable (G U) in /-- The cohomology in degree 1 of a presheaf of groups -`G : Cᵒᵖ ⥤ GroupCat` on a family of objects `U : I → C`. -/ +`G : Cᵒᵖ ⥤ Grp` on a family of objects `U : I → C`. -/ def H1 := Quot (OneCocycle.IsCohomologous (G := G) (U := U)) /-- The cohomology class of a 1-cocycle. -/ diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index 53ca9ed1cc130..15e4dd5d4d5f8 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -42,11 +42,11 @@ from `M` to `𝕜` is a sheaf of commutative rings, the *structure sheaf* of `M` ## TODO -There are variants of `smoothSheafCommGroup.compLeft` for `GroupCat`, `RingCat`, `CommRingCat`; +There are variants of `smoothSheafCommGroup.compLeft` for `Grp`, `RingCat`, `CommRingCat`; this is just boilerplate and can be added as needed. Similarly, there are variants of `smoothSheafCommRing.forgetStalk` and `smoothSheafCommRing.eval` -for `GroupCat`, `CommGroupCat` and `RingCat` which can be added as needed. +for `Grp`, `CommGrp` and `RingCat` which can be added as needed. Currently there is a universe restriction: one can consider the sheaf of smooth functions from `M` to `N` only if `M` and `N` are in the same universe. For example, since `ℂ` is in `Type`, we can @@ -157,9 +157,9 @@ noncomputable instance (U : (Opens (TopCat.of M))ᵒᵖ) : -/ @[to_additive "The presheaf of smooth functions from `M` to `G`, for `G` an additive Lie group, as a presheaf of additive groups."] -noncomputable def smoothPresheafGroup : TopCat.Presheaf GroupCat.{u} (TopCat.of M) := - { obj := fun U ↦ GroupCat.of ((smoothSheaf IM I M G).presheaf.obj U) - map := fun h ↦ GroupCat.ofHom <| +noncomputable def smoothPresheafGroup : TopCat.Presheaf Grp.{u} (TopCat.of M) := + { obj := fun U ↦ Grp.of ((smoothSheaf IM I M G).presheaf.obj U) + map := fun h ↦ Grp.ofHom <| SmoothMap.restrictMonoidHom IM I G <| CategoryTheory.leOfHom h.unop map_id := fun _ ↦ rfl map_comp := fun _ _ ↦ rfl } @@ -168,10 +168,10 @@ noncomputable def smoothPresheafGroup : TopCat.Presheaf GroupCat.{u} (TopCat.of groups. -/ @[to_additive "The sheaf of smooth functions from `M` to `G`, for `G` an additive Lie group, as a sheaf of additive groups."] -noncomputable def smoothSheafGroup : TopCat.Sheaf GroupCat.{u} (TopCat.of M) := +noncomputable def smoothSheafGroup : TopCat.Sheaf Grp.{u} (TopCat.of M) := { val := smoothPresheafGroup IM I M G cond := by - rw [CategoryTheory.Presheaf.isSheaf_iff_isSheaf_forget _ _ (CategoryTheory.forget GroupCat)] + rw [CategoryTheory.Presheaf.isSheaf_iff_isSheaf_forget _ _ (CategoryTheory.forget Grp)] exact CategoryTheory.Sheaf.cond (smoothSheaf IM I M G) } end LieGroup @@ -188,9 +188,9 @@ open Manifold in presheaf of abelian groups. -/ @[to_additive "The presheaf of smooth functions from `M` to `A`, for `A` an additive abelian Lie group, as a presheaf of additive abelian groups."] -noncomputable def smoothPresheafCommGroup : TopCat.Presheaf CommGroupCat.{u} (TopCat.of M) := - { obj := fun U ↦ CommGroupCat.of ((smoothSheaf IM I M A).presheaf.obj U) - map := fun h ↦ CommGroupCat.ofHom <| +noncomputable def smoothPresheafCommGroup : TopCat.Presheaf CommGrp.{u} (TopCat.of M) := + { obj := fun U ↦ CommGrp.of ((smoothSheaf IM I M A).presheaf.obj U) + map := fun h ↦ CommGrp.ofHom <| SmoothMap.restrictMonoidHom IM I A <| CategoryTheory.leOfHom h.unop map_id := fun _ ↦ rfl map_comp := fun _ _ ↦ rfl } @@ -199,11 +199,11 @@ noncomputable def smoothPresheafCommGroup : TopCat.Presheaf CommGroupCat.{u} (To sheaf of abelian groups. -/ @[to_additive "The sheaf of smooth functions from `M` to `A`, for `A` an abelian additive Lie group, as a sheaf of abelian additive groups."] -noncomputable def smoothSheafCommGroup : TopCat.Sheaf CommGroupCat.{u} (TopCat.of M) := +noncomputable def smoothSheafCommGroup : TopCat.Sheaf CommGrp.{u} (TopCat.of M) := { val := smoothPresheafCommGroup IM I M A cond := by rw [CategoryTheory.Presheaf.isSheaf_iff_isSheaf_forget _ _ - (CategoryTheory.forget CommGroupCat)] + (CategoryTheory.forget CommGrp)] exact CategoryTheory.Sheaf.cond (smoothSheaf IM I M A) } /-- For a manifold `M` and a smooth homomorphism `φ` between abelian Lie groups `A`, `A'`, the @@ -215,7 +215,7 @@ to `smoothSheafAddCommGroup IM I' M A'`."] def smoothSheafCommGroup.compLeft (φ : A →* A') (hφ : Smooth I I' φ) : smoothSheafCommGroup IM I M A ⟶ smoothSheafCommGroup IM I' M A' := CategoryTheory.Sheaf.Hom.mk <| - { app := fun _ ↦ CommGroupCat.ofHom <| SmoothMap.compLeftMonoidHom _ _ φ hφ + { app := fun _ ↦ CommGrp.ofHom <| SmoothMap.compLeftMonoidHom _ _ φ hφ naturality := fun _ _ _ ↦ rfl } end CommLieGroup diff --git a/Mathlib/RepresentationTheory/Action/Basic.lean b/Mathlib/RepresentationTheory/Action/Basic.lean index d5ba5f65642c7..6fdfca4d1ea2b 100644 --- a/Mathlib/RepresentationTheory/Action/Basic.lean +++ b/Mathlib/RepresentationTheory/Action/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import Mathlib.Algebra.Category.GroupCat.Basic +import Mathlib.Algebra.Category.Grp.Basic import Mathlib.CategoryTheory.SingleObj import Mathlib.CategoryTheory.Limits.FunctorCategory import Mathlib.CategoryTheory.Limits.Preserves.Basic @@ -53,7 +53,7 @@ set_option linter.uppercaseLean3 false in /-- When a group acts, we can lift the action to the group of automorphisms. -/ @[simps] -def ρAut {G : GroupCat.{u}} (A : Action V (MonCat.of G)) : G ⟶ GroupCat.of (Aut A.V) where +def ρAut {G : Grp.{u}} (A : Action V (MonCat.of G)) : G ⟶ Grp.of (Aut A.V) where toFun g := { hom := A.ρ g inv := A.ρ (g⁻¹ : G) @@ -77,13 +77,13 @@ set_option linter.uppercaseLean3 false in #align Action.inhabited' Action.inhabited' /-- The trivial representation of a group. -/ -def trivial : Action AddCommGroupCat G where - V := AddCommGroupCat.of PUnit +def trivial : Action AddCommGrp G where + V := AddCommGrp.of PUnit ρ := 1 set_option linter.uppercaseLean3 false in #align Action.trivial Action.trivial -instance : Inhabited (Action AddCommGroupCat G) := +instance : Inhabited (Action AddCommGrp G) := ⟨trivial G⟩ end diff --git a/Mathlib/RepresentationTheory/Action/Monoidal.lean b/Mathlib/RepresentationTheory/Action/Monoidal.lean index 059efdcf6235a..3ceb960aa09b0 100644 --- a/Mathlib/RepresentationTheory/Action/Monoidal.lean +++ b/Mathlib/RepresentationTheory/Action/Monoidal.lean @@ -263,7 +263,7 @@ theorem functorCategoryMonoidalEquivalence.inverse_map {A B : SingleObj G ⥤ V} set_option linter.uppercaseLean3 false in #align Action.functor_category_monoidal_equivalence.inverse_map Action.functorCategoryMonoidalEquivalence.inverse_map -variable (H : GroupCat.{u}) +variable (H : Grp.{u}) instance [RightRigidCategory V] : RightRigidCategory (SingleObj (H : MonCat.{u}) ⥤ V) := by change RightRigidCategory (SingleObj H ⥤ V); infer_instance diff --git a/Mathlib/RepresentationTheory/Character.lean b/Mathlib/RepresentationTheory/Character.lean index 09b66e8de7bb3..57b65f29b9589 100644 --- a/Mathlib/RepresentationTheory/Character.lean +++ b/Mathlib/RepresentationTheory/Character.lean @@ -113,7 +113,7 @@ end Group section Orthogonality -variable {G : GroupCat.{u}} [IsAlgClosed k] +variable {G : Grp.{u}} [IsAlgClosed k] open scoped Classical diff --git a/Mathlib/RepresentationTheory/FdRep.lean b/Mathlib/RepresentationTheory/FdRep.lean index 02031dd14dafe..3bba834b4873e 100644 --- a/Mathlib/RepresentationTheory/FdRep.lean +++ b/Mathlib/RepresentationTheory/FdRep.lean @@ -156,7 +156,7 @@ variable {k G : Type u} [Field k] [Group G] -- Verify that the right rigid structure is available when the monoid is a group. noncomputable instance : RightRigidCategory (FdRep k G) := by - change RightRigidCategory (Action (FGModuleCat k) (GroupCat.of G)); infer_instance + change RightRigidCategory (Action (FGModuleCat k) (Grp.of G)); infer_instance end FdRep diff --git a/Mathlib/RepresentationTheory/Invariants.lean b/Mathlib/RepresentationTheory/Invariants.lean index c4660451c57c8..90f471b55590e 100644 --- a/Mathlib/RepresentationTheory/Invariants.lean +++ b/Mathlib/RepresentationTheory/Invariants.lean @@ -128,7 +128,7 @@ open CategoryTheory Action section Rep -variable {k : Type u} [CommRing k] {G : GroupCat.{u}} +variable {k : Type u} [CommRing k] {G : Grp.{u}} theorem mem_invariants_iff_comm {X Y : Rep k G} (f : X.V →ₗ[k] Y.V) (g : G) : (linHom X.ρ Y.ρ) g f = f ↔ f.comp (X.ρ g) = (Y.ρ g).comp f := by @@ -156,7 +156,7 @@ end Rep section FdRep -variable {k : Type u} [Field k] {G : GroupCat.{u}} +variable {k : Type u} [Field k] {G : Grp.{u}} /-- The invariants of the representation `linHom X.ρ Y.ρ` correspond to the representation homomorphisms from `X` to `Y`. -/ diff --git a/Mathlib/Tactic/ToAdditive.lean b/Mathlib/Tactic/ToAdditive.lean index b05aa9d8303e3..076d4b45f690c 100644 --- a/Mathlib/Tactic/ToAdditive.lean +++ b/Mathlib/Tactic/ToAdditive.lean @@ -942,6 +942,8 @@ def nameDict : String → List String | "units" => ["add", "Units"] | "cyclic" => ["add", "Cyclic"] | "rootable" => ["divisible"] + | "semigrp" => ["add", "Semigrp"] + | "grp" => ["add", "Grp"] | "commute" => ["add", "Commute"] | "semiconj" => ["add", "Semiconj"] | "zpowers" => ["zmultiples"] diff --git a/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean b/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean index 7141d0c75f333..e404a7778c16c 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/UniqueGluing.lean @@ -16,7 +16,7 @@ We provide an alternative formulation of the sheaf condition in terms of unique We work with sheaves valued in a concrete category `C` admitting all limits, whose forgetful functor `C ⥤ Type` preserves limits and reflects isomorphisms. The usual categories of algebraic -structures, such as `MonCat`, `AddCommGroupCat`, `RingCat`, `CommRingCat` etc. are all examples of +structures, such as `MonCat`, `AddCommGrp`, `RingCat`, `CommRingCat` etc. are all examples of this kind of category. A presheaf `F : presheaf C X` satisfies the sheaf condition if and only if, for every diff --git a/Mathlib/Topology/VectorBundle/Basic.lean b/Mathlib/Topology/VectorBundle/Basic.lean index abecd9d916c16..eb79504902d7a 100644 --- a/Mathlib/Topology/VectorBundle/Basic.lean +++ b/Mathlib/Topology/VectorBundle/Basic.lean @@ -625,7 +625,7 @@ instance topologicalSpaceFiber (x : B) : TopologicalSpace (Z.Fiber x) := Z.toFiberBundleCore.topologicalSpaceFiber x #align vector_bundle_core.topological_space_fiber VectorBundleCore.topologicalSpaceFiber --- Porting note: fixed: used to assume both `[NormedAddCommGroup F]` and `[AddCommGroupCat F]` +-- Porting note: fixed: used to assume both `[NormedAddCommGroup F]` and `[AddCommGrp F]` instance addCommGroupFiber (x : B) : AddCommGroup (Z.Fiber x) := inferInstanceAs (AddCommGroup F) #align vector_bundle_core.add_comm_group_fiber VectorBundleCore.addCommGroupFiber diff --git a/scripts/nolints.json b/scripts/nolints.json index 0162ea9bdfb4c..2db1d2315d023 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -69,7 +69,7 @@ ["docBlame", "AddMagmaCat.forget_obj_eq_coe"], ["docBlame", "AddMonoidAlgebra.mapDomain"], ["docBlame", "AddMonoidAlgebra.single"], - ["docBlame", "AddSemigroupCat.forget_obj_eq_coe"], + ["docBlame", "AddSemigrp.forget_obj_eq_coe"], ["docBlame", "AffineBasis.toFun"], ["docBlame", "AffineEquiv.linear"], ["docBlame", "AffineMap.linear"], @@ -360,7 +360,7 @@ ["docBlame", "RootableBy.root"], ["docBlame", "SchwartzMap.toFun"], ["docBlame", "SemiRingCat.forget_obj_eq_coe"], - ["docBlame", "SemigroupCat.forget_obj_eq_coe"], + ["docBlame", "Semigrp.forget_obj_eq_coe"], ["docBlame", "Set.«term{_|_}»"], ["docBlame", "Shrink.rec"], ["docBlame", "SlashAction.map"],