Regular elements #
We introduce left-regular, right-regular and regular elements, along with their to_additive
analogues add-left-regular, add-right-regular and add-regular elements.
By definition, a regular element in a commutative ring is a non-zero divisor.
Lemma isRegular_of_ne_zero implies that every non-zero element of an integral domain is regular.
Since it assumes that the ring is a CancelMonoidWithZero it applies also, for instance, to ℕ.
The lemmas in Section MulZeroClass show that the 0 element is (left/right-)regular if and
only if the MulZeroClass is trivial.  This is useful when figuring out stopping conditions for
regular sequences: if 0 is ever an element of a regular sequence, then we can extend the sequence
by adding one further 0.
The final goal is to develop part of the API to prove, eventually, results about non-zero-divisors.
An add-left-regular element is an element c such that addition
on the left by c is injective.
Equations
- IsAddLeftRegular c = Function.Injective fun (x : R) => c + x
Instances For
A left-regular element is an element c such that multiplication on the left by c
is injective.
Equations
- IsLeftRegular c = Function.Injective fun (x : R) => c * x
Instances For
An add-right-regular element is an element c such that addition
on the right by c is injective.
Equations
- IsAddRightRegular c = Function.Injective fun (x : R) => x + c
Instances For
A right-regular element is an element c such that multiplication on the right by c
is injective.
Equations
- IsRightRegular c = Function.Injective fun (x : R) => x * c
Instances For
An add-regular element is an element c such that addition by c both on the left and
on the right is injective.
- left : IsAddLeftRegular cAn add-regular element cis left-regular
- right : IsAddRightRegular cAn add-regular element cis right-regular
Instances For
An add-regular element c is left-regular
An add-regular element c is right-regular
A regular element is an element c such that multiplication by c both on the left and
on the right is injective.
- left : IsLeftRegular cA regular element cis left-regular
- right : IsRightRegular cA regular element cis right-regular
Instances For
A regular element c is left-regular
A regular element c is right-regular
Equations
- ⋯ = ⋯
Instances For
Equations
- ⋯ = ⋯
Instances For
In an additive semigroup, the sum of add-left-regular elements is add-left.regular.
In a semigroup, the product of left-regular elements is left-regular.
In an additive semigroup, the sum of add-right-regular elements is add-right-regular.
In a semigroup, the product of right-regular elements is right-regular.
In an additive semigroup, the sum of add-regular elements is add-regular.
If an element b becomes add-left-regular after adding to it on the left
an add-left-regular element, then b is add-left-regular.
If an element b becomes left-regular after multiplying it on the left by a left-regular
element, then b is left-regular.
An element is add-left-regular if and only if adding to it on the left an add-left-regular element is add-left-regular.
An element is left-regular if and only if multiplying it on the left by a left-regular element is left-regular.
If an element b becomes add-right-regular after adding to it on the right
an add-right-regular element, then b is add-right-regular.
If an element b becomes right-regular after multiplying it on the right by a right-regular
element, then b is right-regular.
An element is add-right-regular if and only if adding it on the right to an add-right-regular element is add-right-regular.
An element is right-regular if and only if multiplying it on the right with a right-regular element is right-regular.
Two elements a and b are add-regular if and only if both sums a + b and
b + a are add-regular.
The "most used" implication of add_and_add_iff, with split
hypotheses, instead of ∧.
The element 0 is left-regular if and only if R is trivial.
The element 0 is right-regular if and only if R is trivial.
The element 0 is regular if and only if R is trivial.
The element 0 is left-regular if and only if R is trivial.
In a non-trivial MulZeroClass, the 0 element is not left-regular.
The element 0 is right-regular if and only if R is trivial.
In a non-trivial MulZeroClass, the 0 element is not right-regular.
The element 0 is regular if and only if R is trivial.
A left-regular element of a Nontrivial MulZeroClass is non-zero.
A right-regular element of a Nontrivial MulZeroClass is non-zero.
A regular element of a Nontrivial MulZeroClass is non-zero.
In a non-trivial ring, the element 0 is not left-regular -- with typeclasses.
In a non-trivial ring, the element 0 is not right-regular -- with typeclasses.
In a non-trivial ring, the element 0 is not regular -- with typeclasses.
If adding 0 on either side is the identity, 0 is regular.
If multiplying by 1 on either side is the identity, 1 is regular.
A sum is add-regular if and only if the summands are.
A product is regular if and only if the factors are.
An element admitting a left additive opposite is add-left-regular.
An element admitting a left inverse is left-regular.
An element admitting a right additive opposite is add-right-regular.
An element admitting a right inverse is right-regular.
If R is an additive monoid, an element in add_units R is add-regular.
An additive unit in an additive monoid is add-regular.
If all additions cancel on the left then every element is add-left-regular.
If all multiplications cancel on the left then every element is left-regular.
If all additions cancel on the right then every element is add-right-regular.
If all multiplications cancel on the right then every element is right-regular.
If all additions cancel then every element is add-regular.
If all multiplications cancel then every element is regular.
Non-zero elements of an integral domain are regular.
In a non-trivial integral domain, an element is regular iff it is non-zero.