Documentation

Mathlib.Algebra.GCDMonoid.Basic

Monoids with normalization functions, gcd, and lcm #

This file defines extra structures on CancelCommMonoidWithZeros, including IsDomains.

Main Definitions #

For the NormalizedGCDMonoid instances on and , see Mathlib.Algebra.GCDMonoid.Nat.

Implementation Notes #

TODO #

Tags #

divisibility, gcd, lcm, normalize

Normalization monoid: multiplying with normUnit gives a normal form for associated elements.

Instances

    Chooses an element of each associate class, by multiplying by normUnit

    Equations
    • normalize = { toFun := fun (x : α) => x * (normUnit x), map_zero' := , map_one' := , map_mul' := }
    @[simp]
    theorem normUnit_mul_normUnit {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] (a : α) :
    normUnit (a * (normUnit a)) = 1
    theorem normalize_eq_normalize {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] {a b : α} (hab : a b) (hba : b a) :
    theorem dvd_antisymm_of_normalize_eq {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] {a b : α} (ha : normalize a = a) (hb : normalize b = b) (hab : a b) (hba : b a) :
    a = b
    theorem Associated.eq_of_normalized {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] {a b : α} (h : Associated a b) (ha : normalize a = a) (hb : normalize b = b) :
    a = b
    @[simp]
    theorem dvd_normalize_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] {a b : α} :
    @[simp]
    theorem normalize_dvd_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] {a b : α} :

    Maps an element of Associates back to the normalized element of its associate class

    Equations
    class GCDMonoid (α : Type u_2) [CancelCommMonoidWithZero α] :
    Type u_2

    GCD monoid: a CancelCommMonoidWithZero with gcd (greatest common divisor) and lcm (least common multiple) operations, determined up to a unit. The type class focuses on gcd and we derive the corresponding lcm facts from gcd.

    • gcd : ααα

      The greatest common divisor between two elements.

    • lcm : ααα

      The least common multiple between two elements.

    • gcd_dvd_left (a b : α) : gcd a b a

      The GCD is a divisor of the first element.

    • gcd_dvd_right (a b : α) : gcd a b b

      The GCD is a divisor of the second element.

    • dvd_gcd {a b c : α} : a ca ba gcd c b

      Any common divisor of both elements is a divisor of the GCD.

    • gcd_mul_lcm (a b : α) : Associated (gcd a b * lcm a b) (a * b)

      The product of two elements is Associated with the product of their GCD and LCM.

    • lcm_zero_left (a : α) : lcm 0 a = 0

      0 is left-absorbing.

    • lcm_zero_right (a : α) : lcm a 0 = 0

      0 is right-absorbing.

    Instances

      Normalized GCD monoid: a CancelCommMonoidWithZero with normalization and gcd (greatest common divisor) and lcm (least common multiple) operations. In this setting gcd and lcm form a bounded lattice on the associated elements where gcd is the infimum, lcm is the supremum, 1 is bottom, and 0 is top. The type class focuses on gcd and we derive the corresponding lcm facts from gcd.

      Instances
        @[simp]
        theorem normalize_gcd {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) :
        normalize (gcd a b) = gcd a b
        theorem gcd_mul_lcm {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        Associated (gcd a b * lcm a b) (a * b)
        theorem dvd_gcd_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b c : α) :
        a gcd b c a b a c
        theorem gcd_comm {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) :
        gcd a b = gcd b a
        theorem gcd_comm' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        Associated (gcd a b) (gcd b a)
        theorem gcd_assoc {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (m n k : α) :
        gcd (gcd m n) k = gcd m (gcd n k)
        theorem gcd_assoc' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        Associated (gcd (gcd m n) k) (gcd m (gcd n k))
        theorem gcd_eq_normalize {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {a b c : α} (habc : gcd a b c) (hcab : c gcd a b) :
        @[simp]
        theorem gcd_zero_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        theorem gcd_zero_left' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a : α) :
        Associated (gcd 0 a) a
        @[simp]
        theorem gcd_zero_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        theorem gcd_zero_right' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a : α) :
        Associated (gcd a 0) a
        @[simp]
        theorem gcd_eq_zero_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        gcd a b = 0 a = 0 b = 0
        theorem gcd_ne_zero_of_left {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b : α} (ha : a 0) :
        gcd a b 0
        theorem gcd_ne_zero_of_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b : α} (hb : b 0) :
        gcd a b 0
        @[simp]
        theorem gcd_one_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        gcd 1 a = 1
        @[simp]
        theorem isUnit_gcd_one_left {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a : α) :
        IsUnit (gcd 1 a)
        theorem gcd_one_left' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a : α) :
        Associated (gcd 1 a) 1
        @[simp]
        theorem gcd_one_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        gcd a 1 = 1
        @[simp]
        theorem isUnit_gcd_one_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a : α) :
        IsUnit (gcd a 1)
        theorem gcd_one_right' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a : α) :
        Associated (gcd a 1) 1
        theorem gcd_dvd_gcd {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b c d : α} (hab : a b) (hcd : c d) :
        gcd a c gcd b d
        theorem Associated.gcd {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) :
        Associated (gcd a₁ b₁) (gcd a₂ b₂)
        @[simp]
        theorem gcd_same {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        @[simp]
        theorem gcd_mul_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b c : α) :
        gcd (a * b) (a * c) = normalize a * gcd b c
        theorem gcd_mul_left' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b c : α) :
        Associated (gcd (a * b) (a * c)) (a * gcd b c)
        @[simp]
        theorem gcd_mul_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b c : α) :
        gcd (b * a) (c * a) = gcd b c * normalize a
        @[simp]
        theorem gcd_mul_right' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b c : α) :
        Associated (gcd (b * a) (c * a)) (gcd b c * a)
        theorem gcd_eq_left_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) (h : normalize a = a) :
        gcd a b = a a b
        theorem gcd_eq_right_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) (h : normalize b = b) :
        gcd a b = b b a
        theorem gcd_dvd_gcd_mul_left {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        gcd m n gcd (k * m) n
        theorem gcd_dvd_gcd_mul_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        gcd m n gcd (m * k) n
        theorem gcd_dvd_gcd_mul_left_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        gcd m n gcd m (k * n)
        theorem gcd_dvd_gcd_mul_right_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        gcd m n gcd m (n * k)
        theorem Associated.gcd_eq_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {m n : α} (h : Associated m n) (k : α) :
        gcd m k = gcd n k
        theorem Associated.gcd_eq_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {m n : α} (h : Associated m n) (k : α) :
        gcd k m = gcd k n
        theorem dvd_gcd_mul_of_dvd_mul {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {m n k : α} (H : k m * n) :
        k gcd k m * n
        theorem dvd_gcd_mul_iff_dvd_mul {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {m n k : α} :
        k gcd k m * n k m * n
        theorem dvd_mul_gcd_of_dvd_mul {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {m n k : α} (H : k m * n) :
        k m * gcd k n
        theorem dvd_mul_gcd_iff_dvd_mul {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {m n k : α} :
        k m * gcd k n k m * n

        Represent a divisor of m * n as a product of a divisor of m and a divisor of n.

        Note: In general, this representation is highly non-unique.

        See Nat.dvdProdDvdOfDvdProd for a constructive version on .

        theorem gcd_mul_dvd_mul_gcd {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (k m n : α) :
        gcd k (m * n) gcd k m * gcd k n
        theorem gcd_pow_right_dvd_pow_gcd {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b : α} {k : } :
        gcd a (b ^ k) gcd a b ^ k
        theorem gcd_pow_left_dvd_pow_gcd {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b : α} {k : } :
        gcd (a ^ k) b gcd a b ^ k
        theorem pow_dvd_of_mul_eq_pow {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b c d₁ d₂ : α} (ha : a 0) (hab : IsUnit (gcd a b)) {k : } (h : a * b = c ^ k) (hc : c = d₁ * d₂) (hd₁ : d₁ a) :
        d₁ ^ k 0 d₁ ^ k a
        theorem exists_associated_pow_of_mul_eq_pow {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b c : α} (hab : IsUnit (gcd a b)) {k : } (h : a * b = c ^ k) :
        (d : α), Associated (d ^ k) a
        theorem exists_eq_pow_of_mul_eq_pow {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] [Subsingleton αˣ] {a b c : α} (hab : IsUnit (gcd a b)) {k : } (h : a * b = c ^ k) :
        (d : α), a = d ^ k
        theorem gcd_greatest {α : Type u_2} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {a b d : α} (hda : d a) (hdb : d b) (hd : ∀ (e : α), e ae be d) :
        theorem gcd_greatest_associated {α : Type u_2} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b d : α} (hda : d a) (hdb : d b) (hd : ∀ (e : α), e ae be d) :
        Associated d (gcd a b)
        theorem isUnit_gcd_of_eq_mul_gcd {α : Type u_2} [CancelCommMonoidWithZero α] [GCDMonoid α] {x y x' y' : α} (ex : x = gcd x y * x') (ey : y = gcd x y * y') (h : gcd x y 0) :
        IsUnit (gcd x' y')
        theorem extract_gcd {α : Type u_2} [CancelCommMonoidWithZero α] [GCDMonoid α] (x y : α) :
        (x' : α), (y' : α), x = gcd x y * x' y = gcd x y * y' IsUnit (gcd x' y')
        theorem associated_gcd_left_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {x y : α} :
        Associated x (gcd x y) x y
        theorem associated_gcd_right_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {x y : α} :
        Associated y (gcd x y) y x
        theorem Irreducible.isUnit_gcd_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {x y : α} (hx : Irreducible x) :
        IsUnit (gcd x y) ¬x y
        theorem Irreducible.gcd_eq_one_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {x y : α} (hx : Irreducible x) :
        gcd x y = 1 ¬x y
        theorem gcd_neg' {α : Type u_1} [CancelCommMonoidWithZero α] [HasDistribNeg α] [GCDMonoid α] {a b : α} :
        Associated (gcd a (-b)) (gcd a b)
        theorem gcd_neg {α : Type u_1} [CancelCommMonoidWithZero α] [HasDistribNeg α] [NormalizedGCDMonoid α] {a b : α} :
        gcd a (-b) = gcd a b
        theorem neg_gcd' {α : Type u_1} [CancelCommMonoidWithZero α] [HasDistribNeg α] [GCDMonoid α] {a b : α} :
        Associated (gcd (-a) b) (gcd a b)
        theorem neg_gcd {α : Type u_1} [CancelCommMonoidWithZero α] [HasDistribNeg α] [NormalizedGCDMonoid α] {a b : α} :
        gcd (-a) b = gcd a b
        theorem lcm_dvd_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b c : α} :
        lcm a b c a c b c
        theorem dvd_lcm_left {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        a lcm a b
        theorem dvd_lcm_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        b lcm a b
        theorem lcm_dvd {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b c : α} (hab : a b) (hcb : c b) :
        lcm a c b
        @[simp]
        theorem lcm_eq_zero_iff {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        lcm a b = 0 a = 0 b = 0
        @[simp]
        theorem normalize_lcm {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) :
        normalize (lcm a b) = lcm a b
        theorem lcm_comm {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) :
        lcm a b = lcm b a
        theorem lcm_comm' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (a b : α) :
        Associated (lcm a b) (lcm b a)
        theorem lcm_assoc {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (m n k : α) :
        lcm (lcm m n) k = lcm m (lcm n k)
        theorem lcm_assoc' {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        Associated (lcm (lcm m n) k) (lcm m (lcm n k))
        theorem lcm_eq_normalize {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {a b c : α} (habc : lcm a b c) (hcab : c lcm a b) :
        theorem lcm_dvd_lcm {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a b c d : α} (hab : a b) (hcd : c d) :
        lcm a c lcm b d
        theorem Associated.lcm {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] {a₁ a₂ b₁ b₂ : α} (ha : Associated a₁ a₂) (hb : Associated b₁ b₂) :
        Associated (lcm a₁ b₁) (lcm a₂ b₂)
        @[simp]
        theorem lcm_units_coe_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (u : αˣ) (a : α) :
        lcm (↑u) a = normalize a
        @[simp]
        theorem lcm_units_coe_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) (u : αˣ) :
        lcm a u = normalize a
        @[simp]
        theorem lcm_one_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        @[simp]
        theorem lcm_one_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        @[simp]
        theorem lcm_same {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a : α) :
        @[simp]
        theorem lcm_eq_one_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) :
        lcm a b = 1 a 1 b 1
        @[simp]
        theorem lcm_mul_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b c : α) :
        lcm (a * b) (a * c) = normalize a * lcm b c
        @[simp]
        theorem lcm_mul_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b c : α) :
        lcm (b * a) (c * a) = lcm b c * normalize a
        theorem lcm_eq_left_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) (h : normalize a = a) :
        lcm a b = a b a
        theorem lcm_eq_right_iff {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] (a b : α) (h : normalize b = b) :
        lcm a b = b a b
        theorem lcm_dvd_lcm_mul_left {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        lcm m n lcm (k * m) n
        theorem lcm_dvd_lcm_mul_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        lcm m n lcm (m * k) n
        theorem lcm_dvd_lcm_mul_left_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        lcm m n lcm m (k * n)
        theorem lcm_dvd_lcm_mul_right_right {α : Type u_1} [CancelCommMonoidWithZero α] [GCDMonoid α] (m n k : α) :
        lcm m n lcm m (n * k)
        theorem lcm_eq_of_associated_left {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {m n : α} (h : Associated m n) (k : α) :
        lcm m k = lcm n k
        theorem lcm_eq_of_associated_right {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizedGCDMonoid α] {m n : α} (h : Associated m n) (k : α) :
        lcm k m = lcm k n
        @[instance 100]
        Equations
        @[simp]
        theorem normUnit_eq_one {α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] (x : α) :
        @[simp]
        theorem normalize_eq {α : Type u_1} [CancelCommMonoidWithZero α] [Subsingleton αˣ] (x : α) :

        If a monoid's only unit is 1, then it is isomorphic to its associates.

        Equations
        theorem gcd_eq_of_dvd_sub_right {α : Type u_1} [CommRing α] [IsDomain α] [NormalizedGCDMonoid α] {a b c : α} (h : a b - c) :
        gcd a b = gcd a c
        theorem gcd_eq_of_dvd_sub_left {α : Type u_1} [CommRing α] [IsDomain α] [NormalizedGCDMonoid α] {a b c : α} (h : a b - c) :
        gcd b a = gcd c a

        Define NormalizationMonoid on a structure from a MonoidHom inverse to Associates.mk.

        Equations
        noncomputable def gcdMonoidOfGCD {α : Type u_1} [CancelCommMonoidWithZero α] [DecidableEq α] (gcd : ααα) (gcd_dvd_left : ∀ (a b : α), gcd a b a) (gcd_dvd_right : ∀ (a b : α), gcd a b b) (dvd_gcd : ∀ {a b c : α}, a ca ba gcd c b) :

        Define GCDMonoid on a structure just from the gcd and its properties.

        Equations
        • One or more equations did not get rendered due to their size.
        noncomputable def normalizedGCDMonoidOfGCD {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] [DecidableEq α] (gcd : ααα) (gcd_dvd_left : ∀ (a b : α), gcd a b a) (gcd_dvd_right : ∀ (a b : α), gcd a b b) (dvd_gcd : ∀ {a b c : α}, a ca ba gcd c b) (normalize_gcd : ∀ (a b : α), normalize (gcd a b) = gcd a b) :

        Define NormalizedGCDMonoid on a structure just from the gcd and its properties.

        Equations
        • One or more equations did not get rendered due to their size.
        noncomputable def gcdMonoidOfLCM {α : Type u_1} [CancelCommMonoidWithZero α] [DecidableEq α] (lcm : ααα) (dvd_lcm_left : ∀ (a b : α), a lcm a b) (dvd_lcm_right : ∀ (a b : α), b lcm a b) (lcm_dvd : ∀ {a b c : α}, c ab alcm c b a) :

        Define GCDMonoid on a structure just from the lcm and its properties.

        Equations
        • One or more equations did not get rendered due to their size.
        noncomputable def normalizedGCDMonoidOfLCM {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] [DecidableEq α] (lcm : ααα) (dvd_lcm_left : ∀ (a b : α), a lcm a b) (dvd_lcm_right : ∀ (a b : α), b lcm a b) (lcm_dvd : ∀ {a b c : α}, c ab alcm c b a) (normalize_lcm : ∀ (a b : α), normalize (lcm a b) = lcm a b) :

        Define NormalizedGCDMonoid on a structure just from the lcm and its properties.

        Equations
        • One or more equations did not get rendered due to their size.
        noncomputable def gcdMonoidOfExistsGCD {α : Type u_1} [CancelCommMonoidWithZero α] [DecidableEq α] (h : ∀ (a b : α), (c : α), ∀ (d : α), d a d b d c) :

        Define a GCDMonoid structure on a monoid just from the existence of a gcd.

        Equations
        noncomputable def normalizedGCDMonoidOfExistsGCD {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] [DecidableEq α] (h : ∀ (a b : α), (c : α), ∀ (d : α), d a d b d c) :

        Define a NormalizedGCDMonoid structure on a monoid just from the existence of a gcd.

        Equations
        noncomputable def gcdMonoidOfExistsLCM {α : Type u_1} [CancelCommMonoidWithZero α] [DecidableEq α] (h : ∀ (a b : α), (c : α), ∀ (d : α), a d b d c d) :

        Define a GCDMonoid structure on a monoid just from the existence of an lcm.

        Equations
        noncomputable def normalizedGCDMonoidOfExistsLCM {α : Type u_1} [CancelCommMonoidWithZero α] [NormalizationMonoid α] [DecidableEq α] (h : ∀ (a b : α), (c : α), ∀ (d : α), a d b d c d) :

        Define a NormalizedGCDMonoid structure on a monoid just from the existence of an lcm.

        Equations
        @[instance 100]
        Equations
        • One or more equations did not get rendered due to their size.
        @[simp]
        theorem CommGroupWithZero.coe_normUnit (G₀ : Type u_2) [CommGroupWithZero G₀] [DecidableEq G₀] {a : G₀} (h0 : a 0) :
        (normUnit a) = a⁻¹
        theorem CommGroupWithZero.normalize_eq_one (G₀ : Type u_2) [CommGroupWithZero G₀] [DecidableEq G₀] {a : G₀} (h0 : a 0) :
        Equations
        • One or more equations did not get rendered due to their size.