Associated elements. #
In this file we define an equivalence relation Associated
saying that two elements of a monoid differ by a multiplication by a unit.
Then we show that the quotient type Associates
is a monoid
and prove basic properties of this quotient.
theorem
Associated.trans
{M : Type u_1}
[Monoid M]
{x y z : M}
:
Associated x y → Associated y z → Associated x z
The setoid of the relation x ~ᵤ y
iff there is a unit u
such that x * u = y
Equations
- Associated.setoid M = { r := Associated, iseqv := ⋯ }
Instances For
theorem
Associated.map
{M : Type u_2}
{N : Type u_3}
[Monoid M]
[Monoid N]
{F : Type u_4}
[FunLike F M N]
[MonoidHomClass F M N]
(f : F)
{x y : M}
(ha : Associated x y)
:
Associated (f x) (f y)
@[simp]
@[simp]
theorem
associated_one_of_mul_eq_one
{M : Type u_1}
[CommMonoid M]
{a : M}
(b : M)
(hab : a * b = 1)
:
Associated a 1
theorem
associated_one_of_associated_mul_one
{M : Type u_1}
[CommMonoid M]
{a b : M}
:
Associated (a * b) 1 → Associated a 1
theorem
associated_mul_unit_left
{N : Type u_2}
[Monoid N]
(a u : N)
(hu : IsUnit u)
:
Associated (a * u) a
theorem
associated_unit_mul_left
{N : Type u_2}
[CommMonoid N]
(a u : N)
(hu : IsUnit u)
:
Associated (u * a) a
theorem
associated_mul_unit_right
{N : Type u_2}
[Monoid N]
(a u : N)
(hu : IsUnit u)
:
Associated a (a * u)
theorem
associated_unit_mul_right
{N : Type u_2}
[CommMonoid N]
(a u : N)
(hu : IsUnit u)
:
Associated a (u * a)
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Associated.mul_left
{M : Type u_1}
[Monoid M]
(a : M)
{b c : M}
(h : Associated b c)
:
Associated (a * b) (a * c)
theorem
Associated.mul_right
{M : Type u_1}
[CommMonoid M]
{a b : M}
(h : Associated a b)
(c : M)
:
Associated (a * c) (b * c)
theorem
Associated.mul_mul
{M : Type u_1}
[CommMonoid M]
{a₁ a₂ b₁ b₂ : M}
(h₁ : Associated a₁ b₁)
(h₂ : Associated a₂ b₂)
:
Associated (a₁ * a₂) (b₁ * b₂)
theorem
Associated.pow_pow
{M : Type u_1}
[CommMonoid M]
{a b : M}
{n : ℕ}
(h : Associated a b)
:
Associated (a ^ n) (b ^ n)
theorem
associated_of_dvd_dvd
{M : Type u_1}
[CancelMonoidWithZero M]
{a b : M}
(hab : a ∣ b)
(hba : b ∣ a)
:
Associated a b
instance
instDecidableRelAssociatedOfDvd
{M : Type u_1}
[CancelMonoidWithZero M]
[DecidableRel fun (x1 x2 : M) => x1 ∣ x2]
:
DecidableRel fun (x1 x2 : M) => Associated x1 x2
Equations
- instDecidableRelAssociatedOfDvd x✝¹ x✝ = decidable_of_iff (x✝¹ ∣ x✝ ∧ x✝ ∣ x✝¹) ⋯
theorem
Associated.prime
{M : Type u_1}
[CommMonoidWithZero M]
{p q : M}
(h : Associated p q)
(hp : Prime p)
:
Prime q
theorem
Irreducible.associated_of_dvd
{M : Type u_1}
[Monoid M]
{p q : M}
(p_irr : Irreducible p)
(q_irr : Irreducible q)
(dvd : p ∣ q)
:
Associated p q
theorem
Irreducible.dvd_irreducible_iff_associated
{M : Type u_1}
[Monoid M]
{p q : M}
(pp : Irreducible p)
(qp : Irreducible q)
:
theorem
Prime.associated_of_dvd
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p q : M}
(p_prime : Prime p)
(q_prime : Prime q)
(dvd : p ∣ q)
:
Associated p q
theorem
Prime.dvd_prime_iff_associated
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p q : M}
(pp : Prime p)
(qp : Prime q)
:
theorem
Irreducible.isUnit_iff_not_associated_of_dvd
{M : Type u_1}
[Monoid M]
{x y : M}
(hx : Irreducible x)
(hy : y ∣ x)
:
theorem
Associated.irreducible
{M : Type u_1}
[Monoid M]
{p q : M}
(h : Associated p q)
(hp : Irreducible p)
:
theorem
Associated.of_mul_left
{M : Type u_1}
[CancelCommMonoidWithZero M]
{a b c d : M}
(h : Associated (a * b) (c * d))
(h₁ : Associated a c)
(ha : a ≠ 0)
:
Associated b d
theorem
Associated.of_mul_right
{M : Type u_1}
[CancelCommMonoidWithZero M]
{a b c d : M}
:
Associated (a * b) (c * d) → Associated b d → b ≠ 0 → Associated a c
theorem
Associated.of_pow_associated_of_prime
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p₁ p₂ : M}
{k₁ k₂ : ℕ}
(hp₁ : Prime p₁)
(hp₂ : Prime p₂)
(hk₁ : 0 < k₁)
(h : Associated (p₁ ^ k₁) (p₂ ^ k₂))
:
Associated p₁ p₂
theorem
Associated.of_pow_associated_of_prime'
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p₁ p₂ : M}
{k₁ k₂ : ℕ}
(hp₁ : Prime p₁)
(hp₂ : Prime p₂)
(hk₂ : 0 < k₂)
(h : Associated (p₁ ^ k₁) (p₂ ^ k₂))
:
Associated p₁ p₂
theorem
Irreducible.isRelPrime_iff_not_dvd
{M : Type u_1}
[Monoid M]
{p n : M}
(hp : Irreducible p)
:
See also Irreducible.coprime_iff_not_dvd
.
theorem
prime_dvd_prime_iff_eq
{M : Type u_2}
[CancelCommMonoidWithZero M]
[Subsingleton Mˣ]
{p q : M}
(pp : Prime p)
(qp : Prime q)
:
theorem
eq_of_prime_pow_eq
{R : Type u_2}
[CancelCommMonoidWithZero R]
[Subsingleton Rˣ]
{p₁ p₂ : R}
{k₁ k₂ : ℕ}
(hp₁ : Prime p₁)
(hp₂ : Prime p₂)
(hk₁ : 0 < k₁)
(h : p₁ ^ k₁ = p₂ ^ k₂)
:
theorem
eq_of_prime_pow_eq'
{R : Type u_2}
[CancelCommMonoidWithZero R]
[Subsingleton Rˣ]
{p₁ p₂ : R}
{k₁ k₂ : ℕ}
(hp₁ : Prime p₁)
(hp₂ : Prime p₂)
(hk₁ : 0 < k₂)
(h : p₁ ^ k₁ = p₂ ^ k₂)
:
@[reducible, inline]
The quotient of a monoid by the Associated
relation. Two elements x
and y
are associated iff there is a unit u
such that x * u = y
. There is a natural
monoid structure on Associates M
.
Equations
- Associates M = Quotient (Associated.setoid M)
Instances For
@[reducible, inline]
The canonical quotient map from a monoid M
into the Associates
of M
Equations
- Associates.mk a = ⟦a⟧
Instances For
Equations
- Associates.instInhabited = { default := ⟦1⟧ }
@[simp]
theorem
Associates.mk_quot_out
{M : Type u_1}
[Monoid M]
(a : M)
:
Associated (Quot.out (Associates.mk a)) a
Equations
- Associates.instOne = { one := ⟦1⟧ }
@[simp]
Equations
- Associates.instBot = { bot := 1 }
instance
Associates.instUniqueOfSubsingleton
{M : Type u_1}
[Monoid M]
[Subsingleton M]
:
Unique (Associates M)
Equations
- Associates.instUniqueOfSubsingleton = { default := 1, uniq := ⋯ }
Equations
- Associates.instMul = { mul := Quotient.map₂ (fun (x1 x2 : M) => x1 * x2) ⋯ }
Equations
Equations
Associates.mk
as a MonoidHom
.
Equations
- Associates.mkMonoidHom = { toFun := Associates.mk, map_one' := ⋯, map_mul' := ⋯ }
Instances For
@[simp]
theorem
Associates.associated_map_mk
{M : Type u_1}
[CommMonoid M]
{f : Associates M →* M}
(hinv : Function.RightInverse (⇑f) Associates.mk)
(a : M)
:
Associated a (f (Associates.mk a))
Equations
- Associates.uniqueUnits = { toInhabited := Units.instInhabited, uniq := ⋯ }
@[simp]
theorem
Associates.mul_mono
{M : Type u_1}
[CommMonoid M]
{a b c d : Associates M}
(h₁ : a ≤ b)
(h₂ : c ≤ d)
:
Equations
@[simp]
theorem
Associates.dvd_of_mk_le_mk
{M : Type u_1}
[CommMonoid M]
{a b : M}
:
Associates.mk a ≤ Associates.mk b → a ∣ b
theorem
Associates.mk_le_mk_of_dvd
{M : Type u_1}
[CommMonoid M]
{a b : M}
:
a ∣ b → Associates.mk a ≤ Associates.mk b
@[simp]
@[simp]
@[simp]
Equations
- Associates.instZero = { zero := ⟦0⟧ }
Equations
- Associates.instTopOfZero = { top := 0 }
@[simp]
instance
Associates.instNontrivial
{M : Type u_1}
[MonoidWithZero M]
[Nontrivial M]
:
Nontrivial (Associates M)
Equations
Equations
@[simp]
Equations
instance
Associates.instDecidableRelDvd
{M : Type u_1}
[CommMonoidWithZero M]
[DecidableRel fun (x1 x2 : M) => x1 ∣ x2]
:
DecidableRel fun (x1 x2 : Associates M) => x1 ∣ x2
Equations
- a.instDecidableRelDvd b = Quotient.recOnSubsingleton₂ a b fun (x x_1 : M) => decidable_of_iff' (x ∣ x_1) ⋯
theorem
Associates.Prime.le_or_le
{M : Type u_1}
[CommMonoidWithZero M]
{p : Associates M}
(hp : Prime p)
{a b : Associates M}
(h : p ≤ a * b)
:
@[simp]
@[simp]
@[simp]
theorem
Associates.dvdNotUnit_of_lt
{M : Type u_1}
[CommMonoidWithZero M]
{a b : Associates M}
(hlt : a < b)
:
DvdNotUnit a b
Equations
theorem
associates_irreducible_iff_prime
{M : Type u_1}
[CancelCommMonoidWithZero M]
[DecompositionMonoid M]
{p : Associates M}
:
theorem
Associates.le_of_mul_le_mul_left
{M : Type u_1}
[CancelCommMonoidWithZero M]
(a b c : Associates M)
(ha : a ≠ 0)
:
theorem
Associates.one_or_eq_of_le_of_prime
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p m : Associates M}
(hp : Prime p)
(hle : m ≤ p)
:
theorem
Associates.dvdNotUnit_iff_lt
{M : Type u_1}
[CancelCommMonoidWithZero M]
{a b : Associates M}
:
theorem
dvdNotUnit_of_dvdNotUnit_associated
{M : Type u_1}
[CommMonoidWithZero M]
[Nontrivial M]
{p q r : M}
(h : DvdNotUnit p q)
(h' : Associated q r)
:
DvdNotUnit p r
theorem
isUnit_of_associated_mul
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p b : M}
(h : Associated (p * b) p)
(hp : p ≠ 0)
:
IsUnit b
theorem
DvdNotUnit.not_associated
{M : Type u_1}
[CancelCommMonoidWithZero M]
{p q : M}
(h : DvdNotUnit p q)
:
¬Associated p q