Documentation

Mathlib.Algebra.Group.Units.Opposite

Units in multiplicative and additive opposites #

The units of the opposites are equivalent to the opposites of the units.

Equations
  • One or more equations did not get rendered due to their size.
Instances For

    The additive units of the additive opposites are equivalent to the additive opposites of the additive units.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem Units.coe_unop_opEquiv {M : Type u_2} [Monoid M] (u : Mᵐᵒᵖˣ) :
      (MulOpposite.unop (Units.opEquiv u)) = MulOpposite.unop u
      @[simp]
      theorem AddUnits.coe_unop_opEquiv {M : Type u_2} [AddMonoid M] (u : AddUnits Mᵃᵒᵖ) :
      (AddOpposite.unop (AddUnits.opEquiv u)) = AddOpposite.unop u
      @[simp]
      theorem Units.coe_opEquiv_symm {M : Type u_2} [Monoid M] (u : Mˣᵐᵒᵖ) :
      (Units.opEquiv.symm u) = MulOpposite.op (MulOpposite.unop u)
      @[simp]
      theorem AddUnits.coe_opEquiv_symm {M : Type u_2} [AddMonoid M] (u : (AddUnits M)ᵃᵒᵖ) :
      (AddUnits.opEquiv.symm u) = AddOpposite.op (AddOpposite.unop u)
      theorem IsUnit.op {M : Type u_2} [Monoid M] {m : M} (h : IsUnit m) :
      theorem IsAddUnit.op {M : Type u_2} [AddMonoid M] {m : M} (h : IsAddUnit m) :
      theorem IsUnit.unop {M : Type u_2} [Monoid M] {m : Mᵐᵒᵖ} (h : IsUnit m) :
      @[simp]
      theorem isUnit_op {M : Type u_2} [Monoid M] {m : M} :
      @[simp]
      theorem isAddUnit_op {M : Type u_2} [AddMonoid M] {m : M} :
      @[simp]
      theorem isUnit_unop {M : Type u_2} [Monoid M] {m : Mᵐᵒᵖ} :