1
1
/-
2
2
Copyright (c) 2020 Kenny Lau. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
- Authors: Kenny Lau, Devon Tuma
4
+ Authors: Kenny Lau, Devon Tuma, Oliver Nash
5
5
-/
6
6
import Mathlib.GroupTheory.Submonoid.Operations
7
7
import Mathlib.GroupTheory.Submonoid.Membership
@@ -13,7 +13,8 @@ import Mathlib.GroupTheory.Subgroup.MulOpposite
13
13
# Non-zero divisors and smul-divisors
14
14
15
15
In this file we define the submonoid `nonZeroDivisors` and `nonZeroSMulDivisors` of a
16
- `MonoidWithZero`.
16
+ `MonoidWithZero`. We also define `nonZeroDivisorsLeft` and `nonZeroDivisorsRight` for
17
+ non-commutative monoids.
17
18
18
19
## Notations
19
20
@@ -27,6 +28,50 @@ your own code.
27
28
28
29
-/
29
30
31
+ variable (M₀ : Type *) [MonoidWithZero M₀]
32
+
33
+ /-- The collection of elements of a `MonoidWithZero` that are not left zero divisors form a
34
+ `Submonoid`. -/
35
+ def nonZeroDivisorsLeft : Submonoid M₀ where
36
+ carrier := {x | ∀ y, y * x = 0 → y = 0 }
37
+ one_mem' := by simp
38
+ mul_mem' {x} {y} hx hy := fun z hz ↦ hx _ <| hy _ (mul_assoc z x y ▸ hz)
39
+
40
+ @[simp] lemma mem_nonZeroDivisorsLeft_iff {x : M₀} :
41
+ x ∈ nonZeroDivisorsLeft M₀ ↔ ∀ y, y * x = 0 → y = 0 :=
42
+ Iff.rfl
43
+
44
+ /-- The collection of elements of a `MonoidWithZero` that are not right zero divisors form a
45
+ `Submonoid`. -/
46
+ def nonZeroDivisorsRight : Submonoid M₀ where
47
+ carrier := {x | ∀ y, x * y = 0 → y = 0 }
48
+ one_mem' := by simp
49
+ mul_mem' := fun {x} {y} hx hy z hz ↦ hy _ (hx _ ((mul_assoc x y z).symm ▸ hz))
50
+
51
+ @[simp] lemma mem_nonZeroDivisorsRight_iff {x : M₀} :
52
+ x ∈ nonZeroDivisorsRight M₀ ↔ ∀ y, x * y = 0 → y = 0 :=
53
+ Iff.rfl
54
+
55
+ lemma nonZeroDivisorsLeft_eq_right (M₀ : Type *) [CommMonoidWithZero M₀] :
56
+ nonZeroDivisorsLeft M₀ = nonZeroDivisorsRight M₀ := by
57
+ ext x; simp [mul_comm x]
58
+
59
+ @[simp] lemma coe_nonZeroDivisorsLeft_eq [NoZeroDivisors M₀] [Nontrivial M₀] :
60
+ nonZeroDivisorsLeft M₀ = {x : M₀ | x ≠ 0 } := by
61
+ ext x
62
+ simp only [SetLike.mem_coe, mem_nonZeroDivisorsLeft_iff, mul_eq_zero, forall_eq_or_imp, true_and,
63
+ Set.mem_setOf_eq]
64
+ refine' ⟨fun h ↦ _, fun hx y hx' ↦ by contradiction⟩
65
+ contrapose! h
66
+ exact ⟨1 , h, one_ne_zero⟩
67
+
68
+ @[simp] lemma coe_nonZeroDivisorsRight_eq [NoZeroDivisors M₀] [Nontrivial M₀] :
69
+ nonZeroDivisorsRight M₀ = {x : M₀ | x ≠ 0 } := by
70
+ ext x
71
+ simp only [SetLike.mem_coe, mem_nonZeroDivisorsRight_iff, mul_eq_zero, Set.mem_setOf_eq]
72
+ refine' ⟨fun h ↦ _, fun hx y hx' ↦ by aesop⟩
73
+ contrapose! h
74
+ exact ⟨1 , Or.inl h, one_ne_zero⟩
30
75
31
76
/-- The submonoid of non-zero-divisors of a `MonoidWithZero` `R`. -/
32
77
def nonZeroDivisors (R : Type *) [MonoidWithZero R] : Submonoid R where
0 commit comments