Skip to content

Commit 07e59b7

Browse files
committed
chore: Move RingTheory/NonZeroDivisors under Algebra instead. (#8685)
Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com>
1 parent 5b80db7 commit 07e59b7

File tree

10 files changed

+55
-77
lines changed

10 files changed

+55
-77
lines changed

Mathlib.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -217,6 +217,7 @@ import Mathlib.Algebra.GroupWithZero.Defs
217217
import Mathlib.Algebra.GroupWithZero.Divisibility
218218
import Mathlib.Algebra.GroupWithZero.InjSurj
219219
import Mathlib.Algebra.GroupWithZero.NeZero
220+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
220221
import Mathlib.Algebra.GroupWithZero.Power
221222
import Mathlib.Algebra.GroupWithZero.Semiconj
222223
import Mathlib.Algebra.GroupWithZero.Units.Basic
@@ -2210,7 +2211,6 @@ import Mathlib.GroupTheory.Submonoid.Membership
22102211
import Mathlib.GroupTheory.Submonoid.MulOpposite
22112212
import Mathlib.GroupTheory.Submonoid.Operations
22122213
import Mathlib.GroupTheory.Submonoid.Pointwise
2213-
import Mathlib.GroupTheory.Submonoid.ZeroDivisors
22142214
import Mathlib.GroupTheory.Subsemigroup.Basic
22152215
import Mathlib.GroupTheory.Subsemigroup.Center
22162216
import Mathlib.GroupTheory.Subsemigroup.Centralizer
@@ -3068,7 +3068,6 @@ import Mathlib.RingTheory.Nilpotent
30683068
import Mathlib.RingTheory.Noetherian
30693069
import Mathlib.RingTheory.NonUnitalSubring.Basic
30703070
import Mathlib.RingTheory.NonUnitalSubsemiring.Basic
3071-
import Mathlib.RingTheory.NonZeroDivisors
30723071
import Mathlib.RingTheory.Norm
30733072
import Mathlib.RingTheory.Nullstellensatz
30743073
import Mathlib.RingTheory.OreLocalization.Basic

Mathlib/RingTheory/NonZeroDivisors.lean renamed to Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean

Lines changed: 47 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
/-
22
Copyright (c) 2020 Kenny Lau. All rights reserved.
33
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
55
-/
66
import Mathlib.GroupTheory.Submonoid.Operations
77
import Mathlib.GroupTheory.Submonoid.Membership
@@ -13,7 +13,8 @@ import Mathlib.GroupTheory.Subgroup.MulOpposite
1313
# Non-zero divisors and smul-divisors
1414
1515
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.
1718
1819
## Notations
1920
@@ -27,6 +28,50 @@ your own code.
2728
2829
-/
2930

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⟩
3075

3176
/-- The submonoid of non-zero-divisors of a `MonoidWithZero` `R`. -/
3277
def nonZeroDivisors (R : Type*) [MonoidWithZero R] : Submonoid R where

Mathlib/GroupTheory/Submonoid/ZeroDivisors.lean

Lines changed: 0 additions & 66 deletions
This file was deleted.

Mathlib/LinearAlgebra/SesquilinearForm.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,7 +7,7 @@ import Mathlib.Algebra.Module.LinearMap
77
import Mathlib.LinearAlgebra.Basis.Bilinear
88
import Mathlib.LinearAlgebra.BilinearMap
99
import Mathlib.Algebra.EuclideanDomain.Instances
10-
import Mathlib.RingTheory.NonZeroDivisors
10+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
1111

1212
#align_import linear_algebra.sesquilinear_form from "leanprover-community/mathlib"@"87c54600fe3cdc7d32ff5b50873ac724d86aef8d"
1313

Mathlib/RingTheory/Ideal/Operations.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@ import Mathlib.Data.Nat.Choose.Sum
99
import Mathlib.LinearAlgebra.Basis.Bilinear
1010
import Mathlib.RingTheory.Coprime.Lemmas
1111
import Mathlib.RingTheory.Ideal.Basic
12-
import Mathlib.RingTheory.NonZeroDivisors
12+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
1313

1414
#align_import ring_theory.ideal.operations from "leanprover-community/mathlib"@"e7f0ddbf65bd7181a85edb74b64bdc35ba4bdc74"
1515

Mathlib/RingTheory/Localization/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -7,7 +7,7 @@ import Mathlib.Algebra.Algebra.Tower
77
import Mathlib.Algebra.Ring.Equiv
88
import Mathlib.GroupTheory.MonoidLocalization
99
import Mathlib.RingTheory.Ideal.Basic
10-
import Mathlib.RingTheory.NonZeroDivisors
10+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
1111
import Mathlib.Tactic.Ring
1212

1313
#align_import ring_theory.localization.basic from "leanprover-community/mathlib"@"b69c9a770ecf37eb21f7b8cf4fa00de3b62694ec"

Mathlib/RingTheory/Localization/Integral.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@ import Mathlib.RingTheory.Ideal.LocalRing
1010
import Mathlib.RingTheory.IntegralClosure
1111
import Mathlib.RingTheory.Localization.FractionRing
1212
import Mathlib.RingTheory.Localization.Integer
13-
import Mathlib.RingTheory.NonZeroDivisors
13+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
1414

1515
#align_import ring_theory.localization.integral from "leanprover-community/mathlib"@"831c494092374cfe9f50591ed0ac81a25efc5b86"
1616

Mathlib/RingTheory/Nilpotent.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ Authors: Oliver Nash
55
-/
66
import Mathlib.Data.Nat.Choose.Sum
77
import Mathlib.Algebra.Algebra.Bilinear
8-
import Mathlib.GroupTheory.Submonoid.ZeroDivisors
8+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
99
import Mathlib.RingTheory.Ideal.Operations
1010
import Mathlib.Algebra.GeomSum
1111
import Mathlib.LinearAlgebra.Matrix.ToLin

Mathlib/RingTheory/OreLocalization/Basic.lean

Lines changed: 1 addition & 1 deletion
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: Jakob von Raumer, Kevin Klinge
55
-/
66
import Mathlib.GroupTheory.MonoidLocalization
7-
import Mathlib.RingTheory.NonZeroDivisors
7+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
88
import Mathlib.RingTheory.OreLocalization.OreSet
99
import Mathlib.Tactic.NoncommRing
1010

Mathlib/RingTheory/Polynomial/ScaleRoots.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@ Copyright (c) 2020 Anne Baanen. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Anne Baanen, Devon Tuma
55
-/
6-
import Mathlib.RingTheory.NonZeroDivisors
6+
import Mathlib.Algebra.GroupWithZero.NonZeroDivisors
77
import Mathlib.Data.Polynomial.AlgebraMap
88

99
#align_import ring_theory.polynomial.scale_roots from "leanprover-community/mathlib"@"40ac1b258344e0c2b4568dc37bfad937ec35a727"

0 commit comments

Comments
 (0)