Skip to content

Commit

Permalink
chore: classify added lemma porting notes (#10791)
Browse files Browse the repository at this point in the history
Classifies by adding number (#10756) to porting notes claiming `added lemma`.
  • Loading branch information
pitmonticone committed Feb 21, 2024
1 parent ba55c68 commit 10e929a
Show file tree
Hide file tree
Showing 5 changed files with 18 additions and 18 deletions.
12 changes: 6 additions & 6 deletions Mathlib/Algebra/Category/GroupCat/Basic.lean
Expand Up @@ -68,18 +68,18 @@ instance {X Y : GroupCat} : CoeFun (X ⟢ Y) fun _ => X β†’ Y where
instance instFunLike (X Y : GroupCat) : FunLike (X ⟢ Y) X Y :=
show FunLike (X β†’* Y) X Y from inferInstance

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_id {X : GroupCat} : (πŸ™ X : X β†’ X) = id := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_comp {X Y Z : GroupCat} {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

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : X ⟢ Y) : (forget GroupCat).map f = (f : X β†’ Y) := rfl

@[to_additive (attr := ext)]
Expand Down Expand Up @@ -217,18 +217,18 @@ instance {X Y : CommGroupCat} : CoeFun (X ⟢ Y) fun _ => X β†’ Y where
instance instFunLike (X Y : CommGroupCat) : FunLike (X ⟢ Y) X Y :=
show FunLike (X β†’* Y) X Y from inferInstance

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_id {X : CommGroupCat} : (πŸ™ X : X β†’ X) = id := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_comp {X Y Z : CommGroupCat} {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

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma forget_map {X Y : CommGroupCat} (f : X ⟢ Y) :
(forget CommGroupCat).map f = (f : X β†’ Y) :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/GroupWithZeroCat.lean
Expand Up @@ -72,7 +72,7 @@ instance groupWithZeroConcreteCategory : ConcreteCategory GroupWithZeroCat where
map := fun f => f.toFun }
forget_faithful := ⟨fun h => DFunLike.coe_injective h⟩

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : X ⟢ Y) : (forget GroupWithZeroCat).map f = f := rfl
instance hasForgetToBipointed : HasForgetβ‚‚ GroupWithZeroCat Bipointed where
forgetβ‚‚ :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Algebra/Category/ModuleCat/Basic.lean
Expand Up @@ -227,7 +227,7 @@ theorem comp_def (f : M ⟢ N) (g : N ⟢ U) : f ≫ g = g.comp f :=
rfl
#align Module.comp_def ModuleCat.comp_def

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : M ⟢ N) : (forget (ModuleCat R)).map f = (f : M β†’ N) := rfl

end ModuleCat
Expand Down
12 changes: 6 additions & 6 deletions Mathlib/Algebra/Category/MonCat/Basic.lean
Expand Up @@ -93,15 +93,15 @@ instance instFunLike (X Y : MonCat) : FunLike (X ⟢ Y) X Y :=
instance instMonoidHomClass (X Y : MonCat) : MonoidHomClass (X ⟢ Y) X Y :=
inferInstanceAs <| MonoidHomClass (X β†’* Y) X Y

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_id {X : MonCat} : (πŸ™ X : X β†’ X) = id := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_comp {X Y Z : MonCat} {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g : X β†’ Z) = g ∘ f := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)] lemma forget_map (f : X ⟢ Y) : (forget MonCat).map f = f := rfl

@[to_additive (attr := ext)]
Expand Down Expand Up @@ -215,15 +215,15 @@ instance {X Y : CommMonCat} : CoeFun (X ⟢ Y) fun _ => X β†’ Y where
instance instFunLike (X Y : CommMonCat) : FunLike (X ⟢ Y) X Y :=
show FunLike (X β†’* Y) X Y by infer_instance

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_id {X : CommMonCat} : (πŸ™ X : X β†’ X) = id := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma coe_comp {X Y Z : CommMonCat} {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g : X β†’ Z) = g ∘ f := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[to_additive (attr := simp)]
lemma forget_map {X Y : CommMonCat} (f : X ⟢ Y) :
(forget CommMonCat).map f = (f : X β†’ Y) :=
Expand Down
8 changes: 4 additions & 4 deletions Mathlib/Algebra/Category/Ring/Basic.lean
Expand Up @@ -91,7 +91,7 @@ lemma coe_id {X : SemiRingCat} : (πŸ™ X : X β†’ X) = id := rfl
-- Porting note (#10756): added lemma
lemma coe_comp {X Y Z : SemiRingCat} {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g : X β†’ Z) = g ∘ f := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : X ⟢ Y) : (forget SemiRingCat).map f = (f : X β†’ Y) := rfl

lemma ext {X Y : SemiRingCat} {f g : X ⟢ Y} (w : βˆ€ x : X, f x = g x) : f = g :=
Expand Down Expand Up @@ -215,7 +215,7 @@ lemma coe_id {X : RingCat} : (πŸ™ X : X β†’ X) = id := rfl
-- Porting note (#10756): added lemma
lemma coe_comp {X Y Z : RingCat} {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g : X β†’ Z) = g ∘ f := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : X ⟢ Y) : (forget RingCat).map f = (f : X β†’ Y) := rfl

lemma ext {X Y : RingCat} {f g : X ⟢ Y} (w : βˆ€ x : X, f x = g x) : f = g :=
Expand Down Expand Up @@ -321,7 +321,7 @@ lemma coe_id {X : CommSemiRingCat} : (πŸ™ X : X β†’ X) = id := rfl
-- Porting note (#10756): added lemma
lemma coe_comp {X Y Z : CommSemiRingCat} {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g : X β†’ Z) = g ∘ f := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : X ⟢ Y) : (forget CommSemiRingCat).map f = (f : X β†’ Y) := rfl

lemma ext {X Y : CommSemiRingCat} {f g : X ⟢ Y} (w : βˆ€ x : X, f x = g x) : f = g :=
Expand Down Expand Up @@ -441,7 +441,7 @@ lemma coe_id {X : CommRingCat} : (πŸ™ X : X β†’ X) = id := rfl
-- Porting note (#10756): added lemma
lemma coe_comp {X Y Z : CommRingCat} {f : X ⟢ Y} {g : Y ⟢ Z} : (f ≫ g : X β†’ Z) = g ∘ f := rfl

-- porting note: added
-- porting note (#10756): added lemma
@[simp] lemma forget_map (f : X ⟢ Y) : (forget CommRingCat).map f = (f : X β†’ Y) := rfl

lemma ext {X Y : CommRingCat} {f g : X ⟢ Y} (w : βˆ€ x : X, f x = g x) : f = g :=
Expand Down

0 comments on commit 10e929a

Please sign in to comment.