Skip to content

Commit ee3fc23

Browse files
committed
chore(Algebra/Group/Aut): Do not import MonoidWithZero (#15430)
Move what needs it to a new file `Algebra.GroupWithZero.Action.Basic` which will eventually contain more content coming from `Algebra.SMulWithZero`.
1 parent 34f13c7 commit ee3fc23

File tree

11 files changed

+187
-103
lines changed

11 files changed

+187
-103
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -290,6 +290,7 @@ import Mathlib.Algebra.Group.WithOne.Basic
290290
import Mathlib.Algebra.Group.WithOne.Defs
291291
import Mathlib.Algebra.Group.ZeroOne
292292
import Mathlib.Algebra.GroupPower.IterateHom
293+
import Mathlib.Algebra.GroupWithZero.Action.Basic
293294
import Mathlib.Algebra.GroupWithZero.Action.Defs
294295
import Mathlib.Algebra.GroupWithZero.Action.Opposite
295296
import Mathlib.Algebra.GroupWithZero.Action.Pi

Mathlib/Algebra/Group/Aut.lean

Lines changed: 2 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,6 @@ Copyright (c) 2018 Johannes Hölzl. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov
55
-/
6-
import Mathlib.Algebra.GroupWithZero.Action.Defs
76
import Mathlib.GroupTheory.Perm.Basic
87

98
/-!
@@ -27,7 +26,6 @@ MulAut, AddAut
2726
-/
2827

2928
assert_not_exists MonoidWithZero
30-
assert_not_exists Ring
3129

3230
variable {A : Type*} {M : Type*} {G : Type*}
3331

@@ -101,12 +99,10 @@ def toPerm : MulAut M →* Equiv.Perm M where
10199
/-- The tautological action by `MulAut M` on `M`.
102100
103101
This generalizes `Function.End.applyMulAction`. -/
104-
instance applyMulDistribMulAction {M} [Monoid M] : MulDistribMulAction (MulAut M) M where
102+
instance applyMulAction {M} [Monoid M] : MulAction (MulAut M) M where
105103
smul := (· <| ·)
106104
one_smul _ := rfl
107105
mul_smul _ _ _ := rfl
108-
smul_one := map_one
109-
smul_mul := map_mul
110106

111107
@[simp]
112108
protected theorem smul_def {M} [Monoid M] (f : MulAut M) (a : M) : f • a = f a :=
@@ -208,10 +204,8 @@ def toPerm : AddAut A →* Equiv.Perm A where
208204
/-- The tautological action by `AddAut A` on `A`.
209205
210206
This generalizes `Function.End.applyMulAction`. -/
211-
instance applyDistribMulAction {A} [AddMonoid A] : DistribMulAction (AddAut A) A where
207+
instance applyMulAction {A} [AddMonoid A] : MulAction (AddAut A) A where
212208
smul := (· <| ·)
213-
smul_zero := map_zero
214-
smul_add := map_add
215209
one_smul _ := rfl
216210
mul_smul _ _ _ := rfl
217211

Lines changed: 175 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,175 @@
1+
/-
2+
Copyright (c) 2018 Chris Hughes. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Hughes, Yury Kudryashov
5+
-/
6+
import Mathlib.Algebra.Group.Action.Basic
7+
import Mathlib.Algebra.Group.Action.Prod
8+
import Mathlib.Algebra.Group.Aut
9+
import Mathlib.Algebra.GroupWithZero.Action.Defs
10+
import Mathlib.Algebra.GroupWithZero.Prod
11+
import Mathlib.Algebra.SMulWithZero
12+
13+
/-!
14+
# Definitions of group actions
15+
16+
This file defines a hierarchy of group action type-classes on top of the previously defined
17+
notation classes `SMul` and its additive version `VAdd`:
18+
19+
* `MulAction M α` and its additive version `AddAction G P` are typeclasses used for
20+
actions of multiplicative and additive monoids and groups; they extend notation classes
21+
`SMul` and `VAdd` that are defined in `Algebra.Group.Defs`;
22+
* `DistribMulAction M A` is a typeclass for an action of a multiplicative monoid on
23+
an additive monoid such that `a • (b + c) = a • b + a • c` and `a • 0 = 0`.
24+
25+
The hierarchy is extended further by `Module`, defined elsewhere.
26+
27+
Also provided are typeclasses for faithful and transitive actions, and typeclasses regarding the
28+
interaction of different group actions,
29+
30+
* `SMulCommClass M N α` and its additive version `VAddCommClass M N α`;
31+
* `IsScalarTower M N α` and its additive version `VAddAssocClass M N α`;
32+
* `IsCentralScalar M α` and its additive version `IsCentralVAdd M N α`.
33+
34+
## Notation
35+
36+
- `a • b` is used as notation for `SMul.smul a b`.
37+
- `a +ᵥ b` is used as notation for `VAdd.vadd a b`.
38+
39+
## Implementation details
40+
41+
This file should avoid depending on other parts of `GroupTheory`, to avoid import cycles.
42+
More sophisticated lemmas belong in `GroupTheory.GroupAction`.
43+
44+
## Tags
45+
46+
group action
47+
-/
48+
49+
-- TODO:
50+
-- assert_not_exists Ring
51+
52+
open Function
53+
54+
variable {G G₀ A M N M₀ N₀ R α : Type*}
55+
56+
section GroupWithZero
57+
variable [GroupWithZero G₀] [MulAction G₀ α] {a : G₀}
58+
59+
protected lemma MulAction.bijective₀ (ha : a ≠ 0) : Bijective (a • · : α → α) :=
60+
MulAction.bijective <| Units.mk0 a ha
61+
62+
protected lemma MulAction.injective₀ (ha : a ≠ 0) : Injective (a • · : α → α) :=
63+
(MulAction.bijective₀ ha).injective
64+
65+
protected lemma MulAction.surjective₀ (ha : a ≠ 0) : Surjective (a • · : α → α) :=
66+
(MulAction.bijective₀ ha).surjective
67+
68+
end GroupWithZero
69+
70+
section DistribMulAction
71+
variable [Group G] [Monoid M] [AddMonoid A] [DistribMulAction M A]
72+
variable (A)
73+
74+
/-- Each element of the group defines an additive monoid isomorphism.
75+
76+
This is a stronger version of `MulAction.toPerm`. -/
77+
@[simps (config := { simpRhs := true })]
78+
def DistribMulAction.toAddEquiv [DistribMulAction G A] (x : G) : A ≃+ A where
79+
__ := toAddMonoidHom A x
80+
__ := MulAction.toPermHom G A x
81+
82+
variable (G)
83+
84+
/-- Each element of the group defines an additive monoid isomorphism.
85+
86+
This is a stronger version of `MulAction.toPermHom`. -/
87+
@[simps]
88+
def DistribMulAction.toAddAut [DistribMulAction G A] : G →* AddAut A where
89+
toFun := toAddEquiv _
90+
map_one' := AddEquiv.ext (one_smul _)
91+
map_mul' _ _ := AddEquiv.ext (mul_smul _ _)
92+
93+
end DistribMulAction
94+
95+
/-- Scalar multiplication as a monoid homomorphism with zero. -/
96+
@[simps]
97+
def smulMonoidWithZeroHom [MonoidWithZero M₀] [MulZeroOneClass N₀] [MulActionWithZero M₀ N₀]
98+
[IsScalarTower M₀ N₀ N₀] [SMulCommClass M₀ N₀ N₀] : M₀ × N₀ →*₀ N₀ :=
99+
{ smulMonoidHom with map_zero' := smul_zero _ }
100+
101+
section MulDistribMulAction
102+
variable [Group G] [Monoid M] [MulDistribMulAction G M]
103+
variable (M)
104+
105+
/-- Each element of the group defines a multiplicative monoid isomorphism.
106+
107+
This is a stronger version of `MulAction.toPerm`. -/
108+
@[simps (config := { simpRhs := true })]
109+
def MulDistribMulAction.toMulEquiv (x : G) : M ≃* M :=
110+
{ MulDistribMulAction.toMonoidHom M x, MulAction.toPermHom G M x with }
111+
112+
variable (G)
113+
114+
/-- Each element of the group defines a multiplicative monoid isomorphism.
115+
116+
This is a stronger version of `MulAction.toPermHom`. -/
117+
@[simps]
118+
def MulDistribMulAction.toMulAut : G →* MulAut M where
119+
toFun := MulDistribMulAction.toMulEquiv M
120+
map_one' := MulEquiv.ext (one_smul _)
121+
map_mul' _ _ := MulEquiv.ext (mul_smul _ _)
122+
123+
end MulDistribMulAction
124+
125+
126+
namespace MulAut
127+
128+
/-- The tautological action by `MulAut M` on `M`.
129+
130+
This generalizes `Function.End.applyMulAction`. -/
131+
instance applyMulDistribMulAction [Monoid M] : MulDistribMulAction (MulAut M) M where
132+
smul := (· <| ·)
133+
one_smul _ := rfl
134+
mul_smul _ _ _ := rfl
135+
smul_one := map_one
136+
smul_mul := map_mul
137+
138+
end MulAut
139+
140+
namespace AddAut
141+
142+
/-- The tautological action by `AddAut A` on `A`.
143+
144+
This generalizes `Function.End.applyMulAction`. -/
145+
instance applyDistribMulAction [AddMonoid A] : DistribMulAction (AddAut A) A where
146+
smul := (· <| ·)
147+
smul_zero := map_zero
148+
smul_add := map_add
149+
one_smul _ := rfl
150+
mul_smul _ _ _ := rfl
151+
152+
end AddAut
153+
154+
section Arrow
155+
variable [Group G] [MulAction G A] [Monoid M]
156+
157+
attribute [local instance] arrowAction
158+
159+
/-- When `M` is a monoid, `ArrowAction` is additionally a `MulDistribMulAction`. -/
160+
def arrowMulDistribMulAction : MulDistribMulAction G (A → M) where
161+
smul_one _ := rfl
162+
smul_mul _ _ _ := rfl
163+
164+
attribute [local instance] arrowMulDistribMulAction
165+
166+
/-- Given groups `G H` with `G` acting on `A`, `G` acts by
167+
multiplicative automorphisms on `A → H`. -/
168+
@[simps!] def mulAutArrow : G →* MulAut (A → M) := MulDistribMulAction.toMulAut _ _
169+
170+
end Arrow
171+
172+
lemma IsUnit.smul_sub_iff_sub_inv_smul [Group G] [Monoid R] [AddGroup R] [DistribMulAction G R]
173+
[IsScalarTower G R R] [SMulCommClass G R R] (r : G) (a : R) :
174+
IsUnit (r • (1 : R) - a) ↔ IsUnit (1 - r⁻¹ • a) := by
175+
rw [← isUnit_smul_iff r (1 - r⁻¹ • a), smul_sub, smul_inv_smul]

Mathlib/Algebra/Module/Equiv/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,12 +4,12 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro, Anne Baanen,
55
Frédéric Dupuis, Heather Macbeth
66
-/
7-
import Mathlib.Algebra.Module.Equiv.Defs
87
import Mathlib.Algebra.Field.Defs
8+
import Mathlib.Algebra.GroupWithZero.Action.Basic
9+
import Mathlib.Algebra.Module.Equiv.Defs
910
import Mathlib.Algebra.Module.Hom
1011
import Mathlib.Algebra.Module.LinearMap.End
1112
import Mathlib.Algebra.Module.Pi
12-
import Mathlib.GroupTheory.GroupAction.Group
1313

1414
/-!
1515
# Further results on (semi)linear equivalences.

Mathlib/Algebra/Ring/Action/Group.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,8 +3,8 @@ Copyright (c) 2020 Kenny Lau. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Kenny Lau
55
-/
6+
import Mathlib.Algebra.GroupWithZero.Action.Basic
67
import Mathlib.Algebra.Ring.Action.Basic
7-
import Mathlib.GroupTheory.GroupAction.Group
88
import Mathlib.Algebra.Ring.Equiv
99

1010
/-!

Mathlib/Algebra/Ring/AddAut.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@ Copyright (c) 2022 Yury Kudryashov. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury Kudryashov
55
-/
6-
import Mathlib.GroupTheory.GroupAction.Group
6+
import Mathlib.Algebra.GroupWithZero.Action.Basic
77
import Mathlib.Algebra.Module.Defs
88

99
/-!

Mathlib/Algebra/SMulWithZero.lean

Lines changed: 1 addition & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -4,10 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Damiano Testa
55
-/
66
import Mathlib.Algebra.Group.Action.Opposite
7-
import Mathlib.Algebra.Group.Action.Prod
87
import Mathlib.Algebra.GroupWithZero.Action.Defs
8+
import Mathlib.Algebra.GroupWithZero.Hom
99
import Mathlib.Algebra.GroupWithZero.Opposite
10-
import Mathlib.Algebra.GroupWithZero.Prod
1110
import Mathlib.Algebra.Ring.Defs
1211

1312
/-!
@@ -201,12 +200,6 @@ theorem smul_inv₀ [SMulCommClass α β β] [IsScalarTower α β β] (c : α) (
201200

202201
end GroupWithZero
203202

204-
/-- Scalar multiplication as a monoid homomorphism with zero. -/
205-
@[simps]
206-
def smulMonoidWithZeroHom {α β : Type*} [MonoidWithZero α] [MulZeroOneClass β]
207-
[MulActionWithZero α β] [IsScalarTower α β β] [SMulCommClass α β β] : α × β →*₀ β :=
208-
{ smulMonoidHom with map_zero' := smul_zero _ }
209-
210203
-- This instance seems a bit incongruous in this file, but `#find_home!` told me to put it here.
211204
instance NonUnitalNonAssocSemiring.toDistribSMul [NonUnitalNonAssocSemiring R] :
212205
DistribSMul R R where

Mathlib/Data/DFinsupp/Basic.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Kenny Lau
55
-/
66
import Mathlib.Algebra.BigOperators.GroupWithZero.Finset
7+
import Mathlib.Algebra.Group.Action.Prod
78
import Mathlib.Algebra.Group.Submonoid.Membership
89
import Mathlib.Algebra.GroupWithZero.Action.Pi
910
import Mathlib.Algebra.Module.LinearMap.Defs

Mathlib/Data/Set/Pointwise/SMul.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin, Floris van Doorn
55
-/
66
import Mathlib.Algebra.Group.Pi.Basic
7+
import Mathlib.Algebra.GroupWithZero.Action.Basic
78
import Mathlib.Algebra.Module.Defs
89
import Mathlib.Data.Set.Pairwise.Basic
910
import Mathlib.Data.Set.Pointwise.Basic

0 commit comments

Comments
 (0)