-
Notifications
You must be signed in to change notification settings - Fork 337
/
Center.lean
80 lines (57 loc) · 2.12 KB
/
Center.lean
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
/-
Copyright (c) 2021 Eric Wieser. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Eric Wieser, Jireh Loreaux
-/
import Mathlib.Algebra.Group.Center
import Mathlib.Algebra.Group.Subsemigroup.Operations
/-!
# Centers of semigroups, as subsemigroups.
## Main definitions
* `Subsemigroup.center`: the center of a semigroup
* `AddSubsemigroup.center`: the center of an additive semigroup
We provide `Submonoid.center`, `AddSubmonoid.center`, `Subgroup.center`, `AddSubgroup.center`,
`Subsemiring.center`, and `Subring.center` in other files.
## References
* [Cabrera García and Rodríguez Palacios, Non-associative normed algebras. Volume 1]
[cabreragarciarodriguezpalacios2014]
-/
-- Guard against import creep
assert_not_exists Finset
/-! ### `Set.center` as a `Subsemigroup`. -/
variable (M)
namespace Subsemigroup
section Mul
variable [Mul M]
/-- The center of a semigroup `M` is the set of elements that commute with everything in `M` -/
@[to_additive
"The center of a semigroup `M` is the set of elements that commute with everything in `M`"]
def center : Subsemigroup M where
carrier := Set.center M
mul_mem' := Set.mul_mem_center
-- Porting note: `coe_center` is now redundant
variable {M}
/-- The center of a magma is commutative and associative. -/
@[to_additive "The center of an additive magma is commutative and associative."]
instance center.commSemigroup : CommSemigroup (center M) where
mul_assoc _ b _ := Subtype.ext <| b.2.mid_assoc _ _
mul_comm a _ := Subtype.ext <| a.2.comm _
end Mul
section Semigroup
variable {M} [Semigroup M]
@[to_additive]
theorem mem_center_iff {z : M} : z ∈ center M ↔ ∀ g, g * z = z * g := by
rw [← Semigroup.mem_center_iff]
exact Iff.rfl
@[to_additive]
instance decidableMemCenter (a) [Decidable <| ∀ b : M, b * a = a * b] :
Decidable (a ∈ center M) :=
decidable_of_iff' _ Semigroup.mem_center_iff
end Semigroup
section CommSemigroup
variable [CommSemigroup M]
@[to_additive (attr := simp)]
theorem center_eq_top : center M = ⊤ :=
SetLike.coe_injective (Set.center_eq_univ M)
end CommSemigroup
end Subsemigroup