Skip to content

Commit a8db3aa

Browse files
committed
feat: protect Subgroup.subtype (#3254)
Mathlib 3: leanprover-community/mathlib3#18712
1 parent f13648c commit a8db3aa

File tree

1 file changed

+6
-6
lines changed

1 file changed

+6
-6
lines changed

Mathlib/GroupTheory/Subgroup/Basic.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Kexing Ying
55
66
! This file was ported from Lean 3 source module group_theory.subgroup.basic
7-
! leanprover-community/mathlib commit c10e724be91096453ee3db13862b9fb9a992fef2
7+
! leanprover-community/mathlib commit 6b60020790e39e77bfd633ba3d5562ff82e52c79
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -286,13 +286,13 @@ instance (priority := 75) toLinearOrderedCommGroup {G : Type _} [LinearOrderedCo
286286
/-- The natural group hom from a subgroup of group `G` to `G`. -/
287287
@[to_additive (attr := coe)
288288
"The natural group hom from an additive subgroup of `AddGroup` `G` to `G`."]
289-
def subtype : H →* G :=
289+
protected def subtype : H →* G :=
290290
⟨⟨((↑) : H → G), rfl⟩, fun _ _ => rfl⟩
291291
#align subgroup_class.subtype SubgroupClass.subtype
292292
#align add_subgroup_class.subtype AddSubgroupClass.subtype
293293

294294
@[to_additive (attr := simp)]
295-
theorem coeSubtype : (subtype H : H → G) = ((↑) : H → G) := by
295+
theorem coeSubtype : (SubgroupClass.subtype H : H → G) = ((↑) : H → G) := by
296296
rfl
297297
#align subgroup_class.coe_subtype SubgroupClass.coeSubtype
298298
#align add_subgroup_class.coe_subtype AddSubgroupClass.coeSubtype
@@ -354,7 +354,7 @@ theorem coe_inclusion {H K : S} {h : H ≤ K} (a : H) : (inclusion h a : G) = a
354354

355355
@[to_additive (attr := simp)]
356356
theorem subtype_comp_inclusion {H K : S} (hH : H ≤ K) :
357-
(subtype K).comp (inclusion hH) = subtype H := by
357+
(SubgroupClass.subtype K).comp (inclusion hH) = SubgroupClass.subtype H := by
358358
ext
359359
simp only [MonoidHom.comp_apply, coeSubtype, coe_inclusion]
360360
#align subgroup_class.subtype_comp_inclusion SubgroupClass.subtype_comp_inclusion
@@ -790,7 +790,7 @@ instance toLinearOrderedCommGroup {G : Type _} [LinearOrderedCommGroup G] (H : S
790790

791791
/-- The natural group hom from a subgroup of group `G` to `G`. -/
792792
@[to_additive "The natural group hom from an `AddSubgroup` of `AddGroup` `G` to `G`."]
793-
def subtype : H →* G :=
793+
protected def subtype : H →* G :=
794794
⟨⟨((↑) : H → G), rfl⟩, fun _ _ => rfl⟩
795795
#align subgroup.subtype Subgroup.subtype
796796
#align add_subgroup.subtype AddSubgroup.subtype
@@ -802,7 +802,7 @@ theorem coeSubtype : ⇑ H.subtype = ((↑) : H → G) :=
802802
#align add_subgroup.coe_subtype AddSubgroup.coeSubtype
803803

804804
@[to_additive]
805-
theorem subtype_injective : Function.Injective (subtype H) :=
805+
theorem subtype_injective : Function.Injective (Subgroup.subtype H) :=
806806
Subtype.coe_injective
807807
#align subgroup.subtype_injective Subgroup.subtype_injective
808808
#align add_subgroup.subtype_injective AddSubgroup.subtype_injective

0 commit comments

Comments
 (0)