Documentation

Mathlib.Algebra.Group.Units.Basic

Units (i.e., invertible elements) of a monoid #

An element of a Monoid is a unit if it has a two-sided inverse. This file contains the basic lemmas on units in a monoid, especially focussing on singleton types and unique types.

TODO #

The results here should be used to golf the basic Group lemmas.

theorem unique_zero {α : Type u_1} [Unique α] [Zero α] :
default = 0
theorem unique_one {α : Type u_1} [Unique α] [One α] :
default = 1
@[simp]
theorem AddUnits.add_neg_cancel_right {α : Type u} [AddMonoid α] (a : α) (b : AddUnits α) :
a + b + (-b) = a
@[simp]
theorem Units.mul_inv_cancel_right {α : Type u} [Monoid α] (a : α) (b : αˣ) :
a * b * b⁻¹ = a
@[simp]
theorem AddUnits.neg_add_cancel_right {α : Type u} [AddMonoid α] (a : α) (b : AddUnits α) :
a + (-b) + b = a
@[simp]
theorem Units.inv_mul_cancel_right {α : Type u} [Monoid α] (a : α) (b : αˣ) :
a * b⁻¹ * b = a
@[simp]
theorem AddUnits.add_right_inj {α : Type u} [AddMonoid α] (a : AddUnits α) {b : α} {c : α} :
a + b = a + c b = c
@[simp]
theorem Units.mul_right_inj {α : Type u} [Monoid α] (a : αˣ) {b : α} {c : α} :
a * b = a * c b = c
@[simp]
theorem AddUnits.add_left_inj {α : Type u} [AddMonoid α] (a : AddUnits α) {b : α} {c : α} :
b + a = c + a b = c
@[simp]
theorem Units.mul_left_inj {α : Type u} [Monoid α] (a : αˣ) {b : α} {c : α} :
b * a = c * a b = c
theorem AddUnits.eq_add_neg_iff_add_eq {α : Type u} [AddMonoid α] (c : AddUnits α) {a : α} {b : α} :
a = b + (-c) a + c = b
theorem Units.eq_mul_inv_iff_mul_eq {α : Type u} [Monoid α] (c : αˣ) {a : α} {b : α} :
a = b * c⁻¹ a * c = b
theorem AddUnits.eq_neg_add_iff_add_eq {α : Type u} [AddMonoid α] (b : AddUnits α) {a : α} {c : α} :
a = (-b) + c b + a = c
theorem Units.eq_inv_mul_iff_mul_eq {α : Type u} [Monoid α] (b : αˣ) {a : α} {c : α} :
a = b⁻¹ * c b * a = c
theorem AddUnits.add_neg_eq_iff_eq_add {α : Type u} [AddMonoid α] (b : AddUnits α) {a : α} {c : α} :
a + (-b) = c a = c + b
theorem Units.mul_inv_eq_iff_eq_mul {α : Type u} [Monoid α] (b : αˣ) {a : α} {c : α} :
a * b⁻¹ = c a = c * b
theorem AddUnits.neg_eq_of_add_eq_zero_left {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} (h : a + u = 0) :
(-u) = a
theorem Units.inv_eq_of_mul_eq_one_left {α : Type u} [Monoid α] {u : αˣ} {a : α} (h : a * u = 1) :
u⁻¹ = a
theorem AddUnits.neg_eq_of_add_eq_zero_right {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} (h : u + a = 0) :
(-u) = a
theorem Units.inv_eq_of_mul_eq_one_right {α : Type u} [Monoid α] {u : αˣ} {a : α} (h : u * a = 1) :
u⁻¹ = a
theorem AddUnits.eq_neg_of_add_eq_zero_left {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} (h : u + a = 0) :
a = (-u)
theorem Units.eq_inv_of_mul_eq_one_left {α : Type u} [Monoid α] {u : αˣ} {a : α} (h : u * a = 1) :
a = u⁻¹
theorem AddUnits.eq_neg_of_add_eq_zero_right {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} (h : a + u = 0) :
a = (-u)
theorem Units.eq_inv_of_mul_eq_one_right {α : Type u} [Monoid α] {u : αˣ} {a : α} (h : a * u = 1) :
a = u⁻¹
@[simp]
theorem AddUnits.add_neg_eq_zero {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} :
a + (-u) = 0 a = u
@[simp]
theorem Units.mul_inv_eq_one {α : Type u} [Monoid α] {u : αˣ} {a : α} :
a * u⁻¹ = 1 a = u
@[simp]
theorem AddUnits.neg_add_eq_zero {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} :
(-u) + a = 0 u = a
@[simp]
theorem Units.inv_mul_eq_one {α : Type u} [Monoid α] {u : αˣ} {a : α} :
u⁻¹ * a = 1 u = a
theorem AddUnits.add_eq_zero_iff_eq_neg {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} :
a + u = 0 a = (-u)
theorem Units.mul_eq_one_iff_eq_inv {α : Type u} [Monoid α] {u : αˣ} {a : α} :
a * u = 1 a = u⁻¹
theorem AddUnits.add_eq_zero_iff_neg_eq {α : Type u} [AddMonoid α] {u : AddUnits α} {a : α} :
u + a = 0 (-u) = a
theorem Units.mul_eq_one_iff_inv_eq {α : Type u} [Monoid α] {u : αˣ} {a : α} :
u * a = 1 u⁻¹ = a
theorem AddUnits.neg_unique {α : Type u} [AddMonoid α] {u₁ : AddUnits α} {u₂ : AddUnits α} (h : u₁ = u₂) :
(-u₁) = (-u₂)
theorem Units.inv_unique {α : Type u} [Monoid α] {u₁ : αˣ} {u₂ : αˣ} (h : u₁ = u₂) :
u₁⁻¹ = u₂⁻¹
@[simp]
theorem divp_left_inj {α : Type u} [Monoid α] (u : αˣ) {a : α} {b : α} :
a /ₚ u = b /ₚ u a = b
theorem divp_eq_iff_mul_eq {α : Type u} [Monoid α] {x : α} {u : αˣ} {y : α} :
x /ₚ u = y y * u = x
theorem eq_divp_iff_mul_eq {α : Type u} [Monoid α] {x : α} {u : αˣ} {y : α} :
x = y /ₚ u x * u = y
theorem divp_eq_one_iff_eq {α : Type u} [Monoid α] {a : α} {u : αˣ} :
a /ₚ u = 1 a = u
theorem inv_eq_one_divp' {α : Type u} [Monoid α] (u : αˣ) :
(1 / u) = 1 /ₚ u

Used for field_simp to deal with inverses of units. This form of the lemma is essential since field_simp likes to use inv_eq_one_div to rewrite ↑u⁻¹ = ↑(1 / u).

theorem AddLeftCancelMonoid.eq_zero_of_add_right {α : Type u} [AddLeftCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
a = 0
theorem LeftCancelMonoid.eq_one_of_mul_right {α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
a = 1
theorem AddLeftCancelMonoid.eq_zero_of_add_left {α : Type u} [AddLeftCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
b = 0
theorem LeftCancelMonoid.eq_one_of_mul_left {α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
b = 1
@[simp]
theorem AddLeftCancelMonoid.add_eq_zero {α : Type u} [AddLeftCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b = 0 a = 0 b = 0
@[simp]
theorem LeftCancelMonoid.mul_eq_one {α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b = 1 a = 1 b = 1
theorem AddLeftCancelMonoid.add_ne_zero {α : Type u} [AddLeftCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b 0 a 0 b 0
theorem LeftCancelMonoid.mul_ne_one {α : Type u} [LeftCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b 1 a 1 b 1
theorem AddRightCancelMonoid.eq_zero_of_add_right {α : Type u} [AddRightCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
a = 0
theorem RightCancelMonoid.eq_one_of_mul_right {α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
a = 1
theorem AddRightCancelMonoid.eq_zero_of_add_left {α : Type u} [AddRightCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
b = 0
theorem RightCancelMonoid.eq_one_of_mul_left {α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
b = 1
@[simp]
theorem AddRightCancelMonoid.add_eq_zero {α : Type u} [AddRightCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b = 0 a = 0 b = 0
@[simp]
theorem RightCancelMonoid.mul_eq_one {α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b = 1 a = 1 b = 1
theorem AddRightCancelMonoid.add_ne_zero {α : Type u} [AddRightCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b 0 a 0 b 0
theorem RightCancelMonoid.mul_ne_one {α : Type u} [RightCancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b 1 a 1 b 1
theorem eq_zero_of_add_right' {α : Type u} [AddCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
a = 0
theorem eq_one_of_mul_right' {α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
a = 1
theorem eq_zero_of_add_left' {α : Type u} [AddCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
b = 0
theorem eq_one_of_mul_left' {α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
b = 1
theorem add_eq_zero' {α : Type u} [AddCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b = 0 a = 0 b = 0
theorem mul_eq_one' {α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b = 1 a = 1 b = 1
theorem add_ne_zero' {α : Type u} [AddCancelMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b 0 a 0 b 0
theorem mul_ne_one' {α : Type u} [CancelMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b 1 a 1 b 1
theorem divp_mul_eq_mul_divp {α : Type u} [CommMonoid α] (x : α) (y : α) (u : αˣ) :
x /ₚ u * y = x * y /ₚ u
theorem divp_eq_divp_iff {α : Type u} [CommMonoid α] {x : α} {y : α} {ux : αˣ} {uy : αˣ} :
x /ₚ ux = y /ₚ uy x * uy = y * ux
theorem divp_mul_divp {α : Type u} [CommMonoid α] (x : α) (y : α) (ux : αˣ) (uy : αˣ) :
x /ₚ ux * (y /ₚ uy) = x * y /ₚ (ux * uy)
theorem eq_zero_of_add_right {α : Type u} [AddCommMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
a = 0
theorem eq_one_of_mul_right {α : Type u} [CommMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
a = 1
theorem eq_zero_of_add_left {α : Type u} [AddCommMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} (h : a + b = 0) :
b = 0
theorem eq_one_of_mul_left {α : Type u} [CommMonoid α] [Subsingleton αˣ] {a : α} {b : α} (h : a * b = 1) :
b = 1
@[simp]
theorem add_eq_zero {α : Type u} [AddCommMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b = 0 a = 0 b = 0
@[simp]
theorem mul_eq_one {α : Type u} [CommMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b = 1 a = 1 b = 1
theorem add_ne_zero {α : Type u} [AddCommMonoid α] [Subsingleton (AddUnits α)] {a : α} {b : α} :
a + b 0 a 0 b 0
theorem mul_ne_one {α : Type u} [CommMonoid α] [Subsingleton αˣ] {a : α} {b : α} :
a * b 1 a 1 b 1

IsUnit predicate #

theorem isUnit_of_subsingleton {M : Type u_1} [Monoid M] [Subsingleton M] (a : M) :
instance instCanLiftAddUnitsValIsAddUnit {M : Type u_1} [AddMonoid M] :
CanLift M (AddUnits M) AddUnits.val IsAddUnit
Equations
  • =
instance instCanLiftUnitsValIsUnit {M : Type u_1} [Monoid M] :
CanLift M Mˣ Units.val IsUnit
Equations
  • =

A subsingleton AddMonoid has a unique additive unit.

Equations
  • instUniqueAddUnitsOfSubsingleton = { toInhabited := AddUnits.instInhabited, uniq := }

A subsingleton Monoid has a unique unit.

Equations
  • instUniqueUnitsOfSubsingleton = { toInhabited := Units.instInhabited, uniq := }
theorem units_eq_one {M : Type u_1} [Monoid M] [Subsingleton Mˣ] (u : Mˣ) :
u = 1
theorem IsAddUnit.add_left_inj {M : Type u_1} [AddMonoid M] {a : M} {b : M} {c : M} (h : IsAddUnit a) :
b + a = c + a b = c
theorem IsUnit.mul_left_inj {M : Type u_1} [Monoid M] {a : M} {b : M} {c : M} (h : IsUnit a) :
b * a = c * a b = c
theorem IsAddUnit.add_right_inj {M : Type u_1} [AddMonoid M] {a : M} {b : M} {c : M} (h : IsAddUnit a) :
a + b = a + c b = c
theorem IsUnit.mul_right_inj {M : Type u_1} [Monoid M] {a : M} {b : M} {c : M} (h : IsUnit a) :
a * b = a * c b = c
theorem IsAddUnit.add_left_cancel {M : Type u_1} [AddMonoid M] {a : M} {b : M} {c : M} (h : IsAddUnit a) :
a + b = a + cb = c
theorem IsUnit.mul_left_cancel {M : Type u_1} [Monoid M] {a : M} {b : M} {c : M} (h : IsUnit a) :
a * b = a * cb = c
theorem IsAddUnit.add_right_cancel {M : Type u_1} [AddMonoid M] {a : M} {b : M} {c : M} (h : IsAddUnit b) :
a + b = c + ba = c
theorem IsUnit.mul_right_cancel {M : Type u_1} [Monoid M] {a : M} {b : M} {c : M} (h : IsUnit b) :
a * b = c * ba = c
theorem IsAddUnit.add_right_injective {M : Type u_1} [AddMonoid M] {a : M} (h : IsAddUnit a) :
Function.Injective fun (x : M) => a + x
theorem IsUnit.mul_right_injective {M : Type u_1} [Monoid M] {a : M} (h : IsUnit a) :
Function.Injective fun (x : M) => a * x
theorem IsAddUnit.add_left_injective {M : Type u_1} [AddMonoid M] {b : M} (h : IsAddUnit b) :
Function.Injective fun (x : M) => x + b
theorem IsUnit.mul_left_injective {M : Type u_1} [Monoid M] {b : M} (h : IsUnit b) :
Function.Injective fun (x : M) => x * b
theorem IsUnit.isUnit_iff_mulLeft_bijective {M : Type u_1} [Monoid M] {a : M} :
IsUnit a Function.Bijective fun (x : M) => a * x
theorem IsUnit.isUnit_iff_mulRight_bijective {M : Type u_1} [Monoid M] {a : M} :
IsUnit a Function.Bijective fun (x : M) => x * a
@[simp]
theorem IsAddUnit.add_neg_cancel_right {α : Type u} [SubtractionMonoid α] {b : α} (h : IsAddUnit b) (a : α) :
a + b + -b = a
@[simp]
theorem IsUnit.mul_inv_cancel_right {α : Type u} [DivisionMonoid α] {b : α} (h : IsUnit b) (a : α) :
a * b * b⁻¹ = a
@[simp]
theorem IsAddUnit.neg_add_cancel_right {α : Type u} [SubtractionMonoid α] {b : α} (h : IsAddUnit b) (a : α) :
a + -b + b = a
@[simp]
theorem IsUnit.inv_mul_cancel_right {α : Type u} [DivisionMonoid α] {b : α} (h : IsUnit b) (a : α) :
a * b⁻¹ * b = a
theorem IsAddUnit.eq_add_neg_iff_add_eq {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit c) :
a = b + -c a + c = b
theorem IsUnit.eq_mul_inv_iff_mul_eq {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit c) :
a = b * c⁻¹ a * c = b
theorem IsAddUnit.eq_neg_add_iff_add_eq {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit b) :
a = -b + c b + a = c
theorem IsUnit.eq_inv_mul_iff_mul_eq {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit b) :
a = b⁻¹ * c b * a = c
theorem IsAddUnit.neg_add_eq_iff_eq_add {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit a) :
-a + b = c b = a + c
theorem IsUnit.inv_mul_eq_iff_eq_mul {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit a) :
a⁻¹ * b = c b = a * c
theorem IsAddUnit.add_neg_eq_iff_eq_add {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit b) :
a + -b = c a = c + b
theorem IsUnit.mul_inv_eq_iff_eq_mul {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit b) :
a * b⁻¹ = c a = c * b
theorem IsAddUnit.add_neg_eq_zero {α : Type u} [SubtractionMonoid α] {a : α} {b : α} (h : IsAddUnit b) :
a + -b = 0 a = b
theorem IsUnit.mul_inv_eq_one {α : Type u} [DivisionMonoid α] {a : α} {b : α} (h : IsUnit b) :
a * b⁻¹ = 1 a = b
theorem IsAddUnit.neg_add_eq_zero {α : Type u} [SubtractionMonoid α] {a : α} {b : α} (h : IsAddUnit a) :
-a + b = 0 a = b
theorem IsUnit.inv_mul_eq_one {α : Type u} [DivisionMonoid α] {a : α} {b : α} (h : IsUnit a) :
a⁻¹ * b = 1 a = b
theorem IsAddUnit.add_eq_zero_iff_eq_neg {α : Type u} [SubtractionMonoid α] {a : α} {b : α} (h : IsAddUnit b) :
a + b = 0 a = -b
theorem IsUnit.mul_eq_one_iff_eq_inv {α : Type u} [DivisionMonoid α] {a : α} {b : α} (h : IsUnit b) :
a * b = 1 a = b⁻¹
theorem IsAddUnit.add_eq_zero_iff_neg_eq {α : Type u} [SubtractionMonoid α] {a : α} {b : α} (h : IsAddUnit a) :
a + b = 0 -a = b
theorem IsUnit.mul_eq_one_iff_inv_eq {α : Type u} [DivisionMonoid α] {a : α} {b : α} (h : IsUnit a) :
a * b = 1 a⁻¹ = b
@[simp]
theorem IsAddUnit.sub_add_cancel {α : Type u} [SubtractionMonoid α] {b : α} (h : IsAddUnit b) (a : α) :
a - b + b = a
@[simp]
theorem IsUnit.div_mul_cancel {α : Type u} [DivisionMonoid α] {b : α} (h : IsUnit b) (a : α) :
a / b * b = a
@[simp]
theorem IsAddUnit.add_sub_cancel_right {α : Type u} [SubtractionMonoid α] {b : α} (h : IsAddUnit b) (a : α) :
a + b - b = a
@[simp]
theorem IsUnit.mul_div_cancel_right {α : Type u} [DivisionMonoid α] {b : α} (h : IsUnit b) (a : α) :
a * b / b = a
theorem IsAddUnit.add_zero_sub_cancel {α : Type u} [SubtractionMonoid α] {a : α} (h : IsAddUnit a) :
a + (0 - a) = 0
theorem IsUnit.mul_one_div_cancel {α : Type u} [DivisionMonoid α] {a : α} (h : IsUnit a) :
a * (1 / a) = 1
theorem IsAddUnit.zero_sub_add_cancel {α : Type u} [SubtractionMonoid α] {a : α} (h : IsAddUnit a) :
0 - a + a = 0
theorem IsUnit.one_div_mul_cancel {α : Type u} [DivisionMonoid α] {a : α} (h : IsUnit a) :
1 / a * a = 1
theorem IsAddUnit.sub_left_inj {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit c) :
a - c = b - c a = b
theorem IsUnit.div_left_inj {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit c) :
a / c = b / c a = b
theorem IsAddUnit.sub_eq_iff {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit b) :
a - b = c a = c + b
theorem IsUnit.div_eq_iff {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit b) :
a / b = c a = c * b
theorem IsAddUnit.eq_sub_iff {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit c) :
a = b - c a + c = b
theorem IsUnit.eq_div_iff {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit c) :
a = b / c a * c = b
theorem IsAddUnit.sub_eq_of_eq_add {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit b) :
a = c + ba - b = c
theorem IsUnit.div_eq_of_eq_mul {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit b) :
a = c * ba / b = c
theorem IsAddUnit.eq_sub_of_add_eq {α : Type u} [SubtractionMonoid α] {a : α} {b : α} {c : α} (h : IsAddUnit c) :
a + c = ba = b - c
theorem IsUnit.eq_div_of_mul_eq {α : Type u} [DivisionMonoid α] {a : α} {b : α} {c : α} (h : IsUnit c) :
a * c = ba = b / c
theorem IsAddUnit.sub_eq_zero_iff_eq {α : Type u} [SubtractionMonoid α] {a : α} {b : α} (h : IsAddUnit b) :
a - b = 0 a = b
theorem IsUnit.div_eq_one_iff_eq {α : Type u} [DivisionMonoid α] {a : α} {b : α} (h : IsUnit b) :
a / b = 1 a = b
@[deprecated sub_add_cancel_right]
theorem IsAddUnit.sub_add_left {α : Type u} [SubtractionMonoid α] {a : α} {b : α} (h : IsAddUnit b) :
b - (a + b) = 0 - a
@[deprecated div_mul_cancel_right]
theorem IsUnit.div_mul_left {α : Type u} [DivisionMonoid α] {a : α} {b : α} (h : IsUnit b) :
b / (a * b) = 1 / a
theorem IsAddUnit.add_add_sub {α : Type u} [SubtractionMonoid α] {b : α} (a : α) (h : IsAddUnit b) :
a + b + (0 - b) = a
theorem IsUnit.mul_mul_div {α : Type u} [DivisionMonoid α] {b : α} (a : α) (h : IsUnit b) :
a * b * (1 / b) = a
@[deprecated sub_add_cancel_left]
theorem IsAddUnit.sub_add_right {α : Type u} [SubtractionCommMonoid α] {a : α} (h : IsAddUnit a) (b : α) :
a - (a + b) = 0 - b
@[deprecated div_mul_cancel_left]
theorem IsUnit.div_mul_right {α : Type u} [DivisionCommMonoid α] {a : α} (h : IsUnit a) (b : α) :
a / (a * b) = 1 / b
theorem IsAddUnit.add_sub_cancel_left {α : Type u} [SubtractionCommMonoid α] {a : α} (h : IsAddUnit a) (b : α) :
a + b - a = b
theorem IsUnit.mul_div_cancel_left {α : Type u} [DivisionCommMonoid α] {a : α} (h : IsUnit a) (b : α) :
a * b / a = b
theorem IsAddUnit.add_sub_cancel {α : Type u} [SubtractionCommMonoid α] {a : α} (h : IsAddUnit a) (b : α) :
a + (b - a) = b
theorem IsUnit.mul_div_cancel {α : Type u} [DivisionCommMonoid α] {a : α} (h : IsUnit a) (b : α) :
a * (b / a) = b
theorem IsAddUnit.add_eq_add_of_sub_eq_sub {α : Type u} [SubtractionCommMonoid α] {b : α} {d : α} (hb : IsAddUnit b) (hd : IsAddUnit d) (a : α) (c : α) (h : a - b = c - d) :
a + d = c + b
theorem IsUnit.mul_eq_mul_of_div_eq_div {α : Type u} [DivisionCommMonoid α] {b : α} {d : α} (hb : IsUnit b) (hd : IsUnit d) (a : α) (c : α) (h : a / b = c / d) :
a * d = c * b
theorem IsAddUnit.sub_eq_sub_iff {α : Type u} [SubtractionCommMonoid α] {a : α} {b : α} {c : α} {d : α} (hb : IsAddUnit b) (hd : IsAddUnit d) :
a - b = c - d a + d = c + b
theorem IsUnit.div_eq_div_iff {α : Type u} [DivisionCommMonoid α] {a : α} {b : α} {c : α} {d : α} (hb : IsUnit b) (hd : IsUnit d) :
a / b = c / d a * d = c * b
theorem IsAddUnit.sub_sub_cancel {α : Type u} [SubtractionCommMonoid α] {a : α} {b : α} (h : IsAddUnit a) :
a - (a - b) = b
theorem IsUnit.div_div_cancel {α : Type u} [DivisionCommMonoid α] {a : α} {b : α} (h : IsUnit a) :
a / (a / b) = b
theorem IsAddUnit.sub_sub_cancel_left {α : Type u} [SubtractionCommMonoid α] {a : α} {b : α} (h : IsAddUnit a) :
a - b - a = -b
theorem IsUnit.div_div_cancel_left {α : Type u} [DivisionCommMonoid α] {a : α} {b : α} (h : IsUnit a) :
a / b / a = b⁻¹
@[deprecated IsUnit.mul_div_cancel]
theorem IsUnit.mul_div_cancel' {α : Type u} [DivisionCommMonoid α] {a : α} (h : IsUnit a) (b : α) :
a * (b / a) = b

Alias of IsUnit.mul_div_cancel.

@[deprecated IsAddUnit.add_sub_cancel]
theorem IsAddUnit.add_sub_cancel' {α : Type u} [SubtractionCommMonoid α] {a : α} (h : IsAddUnit a) (b : α) :
a + (b - a) = b

Alias of IsAddUnit.add_sub_cancel.