Documentation

Mathlib.Algebra.Prime.Defs

Prime and irreducible elements. #

In this file we define the predicate Prime p saying that an element of a commutative monoid with zero is prime. Namely, Prime p means that p isn't zero, it isn't a unit, and p ∣ a * b → p ∣ a ∨ p ∣ b for all a, b;

In decomposition monoids (e.g., , ), this predicate is equivalent to Irreducible (see irreducible_iff_prime), however this is not true in general.

Main definitions #

Main results #

def Prime {M : Type u_1} [CommMonoidWithZero M] (p : M) :

An element p of a commutative monoid with zero (e.g., a ring) is called prime, if it's not zero, not a unit, and p ∣ a * b → p ∣ a ∨ p ∣ b for all a, b.

Equations
theorem Prime.ne_zero {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) :
p 0
theorem Prime.not_unit {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) :
theorem Prime.not_dvd_one {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) :
¬p 1
theorem Prime.ne_one {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) :
p 1
theorem Prime.dvd_or_dvd {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) {a b : M} (h : p a * b) :
p a p b
theorem Prime.dvd_mul {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) {a b : M} :
p a * b p a p b
theorem Prime.isPrimal {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) :
theorem Prime.not_dvd_mul {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) {a b : M} (ha : ¬p a) (hb : ¬p b) :
¬p a * b
theorem Prime.dvd_of_dvd_pow {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) {a : M} {n : } (h : p a ^ n) :
p a
theorem Prime.dvd_pow_iff_dvd {M : Type u_1} [CommMonoidWithZero M] {p : M} (hp : Prime p) {a : M} {n : } (hn : n 0) :
p a ^ n p a
@[simp]
@[simp]
theorem not_prime_one {M : Type u_1} [CommMonoidWithZero M] :
structure Irreducible {M : Type u_1} [Monoid M] (p : M) :

Irreducible p states that p is non-unit and only factors into units.

We explicitly avoid stating that p is non-zero, this would require a semiring. Assuming only a monoid allows us to reuse irreducible for associated elements.

  • not_unit : ¬IsUnit p

    p is not a unit

  • isUnit_or_isUnit' (a b : M) : p = a * bIsUnit a IsUnit b

    if p factors then one factor is a unit

theorem Irreducible.not_dvd_one {M : Type u_1} [CommMonoid M] {p : M} (hp : Irreducible p) :
¬p 1
theorem Irreducible.isUnit_or_isUnit {M : Type u_1} [Monoid M] {p : M} (hp : Irreducible p) {a b : M} (h : p = a * b) :
theorem irreducible_iff {M : Type u_1} [Monoid M] {p : M} :
Irreducible p ¬IsUnit p ∀ (a b : M), p = a * bIsUnit a IsUnit b
@[simp]
theorem not_irreducible_one {M : Type u_1} [Monoid M] :
theorem Irreducible.ne_one {M : Type u_1} [Monoid M] {p : M} :
Irreducible pp 1
theorem Irreducible.ne_zero {M : Type u_1} [MonoidWithZero M] {p : M} :
Irreducible pp 0
theorem of_irreducible_mul {M : Type u_2} [Monoid M] {x y : M} :
Irreducible (x * y)IsUnit x IsUnit y
theorem irreducible_or_factor {M : Type u_2} [Monoid M] (x : M) (h : ¬IsUnit x) :
theorem Irreducible.dvd_symm {M : Type u_1} [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) :
p qq p

If p and q are irreducible, then p ∣ q implies q ∣ p.

theorem Irreducible.dvd_comm {M : Type u_1} [Monoid M] {p q : M} (hp : Irreducible p) (hq : Irreducible q) :
p q q p
theorem Irreducible.prime_of_isPrimal {M : Type u_1} [CommMonoidWithZero M] {a : M} (irr : Irreducible a) (primal : IsPrimal a) :
theorem Prime.irreducible {M : Type u_1} [CancelCommMonoidWithZero M] {p : M} (hp : Prime p) :