Skip to content

Commit c8969e4

Browse files
committed
chore(Algebra/Group): Do not import GroupWithZero (#11202)
I am claiming that anything within the `Algebra.Group` folder should be additivisable, to the exception of `MonoidHom.End` maybe. This is not the case of `NeZero`, `MonoidWithZero` and `MonoidWithZeroHom` which were all imported to prove a few lemmas. Those lemmas are easily moved to another file.
1 parent 35909ea commit c8969e4

File tree

6 files changed

+53
-46
lines changed

6 files changed

+53
-46
lines changed

Mathlib/Algebra/Group/Equiv/Basic.lean

Lines changed: 0 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -149,22 +149,6 @@ instance (priority := 100) instMonoidHomClass
149149

150150
variable [EquivLike F α β]
151151

152-
-- See note [lower instance priority]
153-
instance (priority := 100) toZeroHomClass
154-
[MulZeroClass α] [MulZeroClass β] [MulEquivClass F α β] :
155-
ZeroHomClass F α β where
156-
map_zero := fun e =>
157-
calc
158-
e 0 = e 0 * e (EquivLike.inv e 0) := by rw [← map_mul, zero_mul]
159-
_ = 0 := by simp
160-
161-
-- See note [lower instance priority]
162-
instance (priority := 100) toMonoidWithZeroHomClass
163-
[MulZeroOneClass α] [MulZeroOneClass β] [MulEquivClass F α β] :
164-
MonoidWithZeroHomClass F α β :=
165-
{ MulEquivClass.instMonoidHomClass F, MulEquivClass.toZeroHomClass F with }
166-
#align mul_equiv_class.to_monoid_with_zero_hom_class MulEquivClass.toMonoidWithZeroHomClass
167-
168152
variable {F}
169153

170154
@[to_additive (attr := simp)]

Mathlib/Algebra/Group/Hom/Basic.lean

Lines changed: 4 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -5,8 +5,7 @@ Authors: Patrick Massot, Kevin Buzzard, Scott Morrison, Johan Commelin, Chris Hu
55
Johannes Hölzl, Yury Kudryashov
66
-/
77
import Mathlib.Algebra.Group.Basic
8-
import Mathlib.Algebra.GroupWithZero.Hom
9-
import Mathlib.Algebra.NeZero
8+
import Mathlib.Algebra.Group.Hom.Defs
109

1110
#align_import algebra.hom.group from "leanprover-community/mathlib"@"a148d797a1094ab554ad4183a4ad6f130358ef64"
1211

@@ -15,6 +14,9 @@ import Mathlib.Algebra.NeZero
1514
1615
-/
1716

17+
-- `NeZero` cannot be additivised, hence its theory should be developed outside of the
18+
-- `Algebra.Group` folder.
19+
assert_not_exists NeZero
1820

1921
variable {α β M N P : Type*}
2022

@@ -24,23 +26,6 @@ variable {G : Type*} {H : Type*}
2426
-- groups
2527
variable {F : Type*}
2628

27-
namespace NeZero
28-
29-
theorem of_map {R M} [Zero R] [Zero M] [FunLike F R M] [ZeroHomClass F R M]
30-
(f : F) {r : R} [neZero : NeZero (f r)] : NeZero r :=
31-
fun h => ne (f r) <| by rw [h]; exact ZeroHomClass.map_zero f⟩
32-
#align ne_zero.of_map NeZero.of_map
33-
34-
theorem of_injective {R M} [Zero R] {r : R} [NeZero r] [Zero M] [FunLike F R M]
35-
[ZeroHomClass F R M] {f : F}
36-
(hf : Function.Injective f) : NeZero (f r) :=
37-
by
38-
rw [← ZeroHomClass.map_zero f]
39-
exact hf.ne NeZero.out⟩
40-
#align ne_zero.of_injective NeZero.of_injective
41-
42-
end NeZero
43-
4429
section DivisionCommMonoid
4530

4631
variable [DivisionCommMonoid α]
@@ -261,11 +246,3 @@ lemma comp_div (f : G →* H) (g h : M →* G) : f.comp (g / h) = f.comp g / f.c
261246
ext; simp only [Function.comp_apply, div_apply, map_div, coe_comp]
262247

263248
end InvDiv
264-
265-
/-- Given two monoid with zero morphisms `f`, `g` to a commutative monoid, `f * g` is the monoid
266-
with zero morphism sending `x` to `f x * g x`. -/
267-
instance [MulZeroOneClass M] [CommMonoidWithZero N] : Mul (M →*₀ N) :=
268-
fun f g => { (f * g : M →* N) with
269-
toFun := fun a => f a * g a,
270-
map_zero' := by dsimp only []; rw [map_zero, zero_mul] }⟩
271-
-- Porting note: why do we need `dsimp` here?

Mathlib/Algebra/Group/Prod.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,7 @@ Authors: Simon Hudon, Patrick Massot, Yury Kudryashov
55
-/
66
import Mathlib.Algebra.Group.Opposite
77
import Mathlib.Algebra.Group.Units.Hom
8+
import Mathlib.Algebra.GroupWithZero.Hom
89
import Mathlib.Algebra.GroupWithZero.Units.Basic
910

1011
#align_import algebra.group.prod from "leanprover-community/mathlib"@"cd391184c85986113f8c00844cfe6dda1d34be3d"

Mathlib/Algebra/GroupWithZero/Hom.lean

Lines changed: 46 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,8 +3,9 @@ Copyright (c) 2020 Patrick Massot. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Eric Wieser
55
-/
6-
import Mathlib.Algebra.Group.Hom.Defs
6+
import Mathlib.Algebra.Group.Equiv.Basic
77
import Mathlib.Algebra.GroupWithZero.Defs
8+
import Mathlib.Algebra.NeZero
89

910
#align_import algebra.hom.group from "leanprover-community/mathlib"@"a148d797a1094ab554ad4183a4ad6f130358ef64"
1011

@@ -35,6 +36,19 @@ monoid homomorphism
3536

3637
open Function
3738

39+
namespace NeZero
40+
variable {F α β : Type*} [Zero α] [Zero β] [FunLike F α β] [ZeroHomClass F α β] {a : α}
41+
42+
lemma of_map (f : F) [neZero : NeZero (f a)] : NeZero a :=
43+
fun h ↦ ne (f a) <| by rw [h]; exact ZeroHomClass.map_zero f⟩
44+
#align ne_zero.of_map NeZero.of_map
45+
46+
lemma of_injective {f : F} (hf : Injective f) [NeZero a] : NeZero (f a) :=
47+
by rw [← ZeroHomClass.map_zero f]; exact hf.ne NeZero.out⟩
48+
#align ne_zero.of_injective NeZero.of_injective
49+
50+
end NeZero
51+
3852
variable {F α β γ δ : Type*} [MulZeroOneClass α] [MulZeroOneClass β] [MulZeroOneClass γ]
3953
[MulZeroOneClass δ]
4054

@@ -234,3 +248,34 @@ lemma toZeroHom_injective : Injective (toZeroHom : (α →*₀ β) → ZeroHom
234248

235249
-- Unlike the other homs, `MonoidWithZeroHom` does not have a `1` or `0`
236250
instance : Inhabited (α →*₀ α) := ⟨id α⟩
251+
252+
/-- Given two monoid with zero morphisms `f`, `g` to a commutative monoid with zero, `f * g` is the
253+
monoid with zero morphism sending `x` to `f x * g x`. -/
254+
instance {β} [CommMonoidWithZero β] : Mul (α →*₀ β) where
255+
mul f g :=
256+
{ (f * g : α →* β) with
257+
map_zero' := by dsimp; rw [map_zero, zero_mul] }
258+
259+
end MonoidWithZeroHom
260+
261+
/-! ### Equivalences -/
262+
263+
namespace MulEquivClass
264+
variable {F α β : Type*} [EquivLike F α β]
265+
266+
-- See note [lower instance priority]
267+
instance (priority := 100) toZeroHomClass [MulZeroClass α] [MulZeroClass β] [MulEquivClass F α β] :
268+
ZeroHomClass F α β where
269+
map_zero f :=
270+
calc
271+
f 0 = f 0 * f (EquivLike.inv f 0) := by rw [← map_mul, zero_mul]
272+
_ = 0 := by simp
273+
274+
-- See note [lower instance priority]
275+
instance (priority := 100) toMonoidWithZeroHomClass
276+
[MulZeroOneClass α] [MulZeroOneClass β] [MulEquivClass F α β] :
277+
MonoidWithZeroHomClass F α β :=
278+
{ MulEquivClass.instMonoidHomClass F, MulEquivClass.toZeroHomClass with }
279+
#align mul_equiv_class.to_monoid_with_zero_hom_class MulEquivClass.toMonoidWithZeroHomClass
280+
281+
end MulEquivClass

Mathlib/Algebra/GroupWithZero/Units/Lemmas.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
-/
66
import Mathlib.Algebra.Group.Commute.Units
7-
import Mathlib.Algebra.Group.Hom.Basic
87
import Mathlib.Algebra.Group.Units.Hom
98
import Mathlib.Algebra.GroupWithZero.Commute
9+
import Mathlib.Algebra.GroupWithZero.Hom
1010
import Mathlib.Algebra.GroupWithZero.Units.Basic
1111
import Mathlib.GroupTheory.GroupAction.Units
1212

Mathlib/Algebra/Order/Hom/Monoid.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yaël Dillies
55
-/
6-
import Mathlib.Algebra.Group.Hom.Basic
6+
import Mathlib.Algebra.GroupWithZero.Hom
77
import Mathlib.Algebra.Order.Group.Instances
88
import Mathlib.Algebra.Order.Monoid.WithZero.Defs
99
import Mathlib.Order.Hom.Basic

0 commit comments

Comments
 (0)