Documentation

Mathlib.Algebra.Ring.Commute

Semirings and rings #

This file gives lemmas about semirings, rings and domains. This is analogous to Mathlib.Algebra.Group.Basic, the difference being that the former is about + and * separately, while the present file is about their interaction.

For the definitions of semirings and rings see Mathlib.Algebra.Ring.Defs.

@[simp]
theorem Commute.add_right {R : Type x} [Distrib R] {a : R} {b : R} {c : R} :
Commute a bCommute a cCommute a (b + c)
@[simp]
theorem Commute.add_left {R : Type x} [Distrib R] {a : R} {b : R} {c : R} :
Commute a cCommute b cCommute (a + b) c
@[deprecated]
theorem Commute.bit0_right {R : Type x} [Distrib R] {x : R} {y : R} (h : Commute x y) :
@[deprecated]
theorem Commute.bit0_left {R : Type x} [Distrib R] {x : R} {y : R} (h : Commute x y) :
@[deprecated]
theorem Commute.bit1_right {R : Type x} [NonAssocSemiring R] {x : R} {y : R} (h : Commute x y) :
@[deprecated]
theorem Commute.bit1_left {R : Type x} [NonAssocSemiring R] {x : R} {y : R} (h : Commute x y) :
theorem Commute.mul_self_sub_mul_self_eq {R : Type x} [NonUnitalNonAssocRing R] {a : R} {b : R} (h : Commute a b) :
a * a - b * b = (a + b) * (a - b)

Representation of a difference of two squares of commuting elements as a product.

theorem Commute.mul_self_sub_mul_self_eq' {R : Type x} [NonUnitalNonAssocRing R] {a : R} {b : R} (h : Commute a b) :
a * a - b * b = (a - b) * (a + b)
theorem Commute.mul_self_eq_mul_self_iff {R : Type x} [NonUnitalNonAssocRing R] [NoZeroDivisors R] {a : R} {b : R} (h : Commute a b) :
a * a = b * b a = b a = -b
theorem Commute.neg_right {R : Type x} [Mul R] [HasDistribNeg R] {a : R} {b : R} :
Commute a bCommute a (-b)
@[simp]
theorem Commute.neg_right_iff {R : Type x} [Mul R] [HasDistribNeg R] {a : R} {b : R} :
Commute a (-b) Commute a b
theorem Commute.neg_left {R : Type x} [Mul R] [HasDistribNeg R] {a : R} {b : R} :
Commute a bCommute (-a) b
@[simp]
theorem Commute.neg_left_iff {R : Type x} [Mul R] [HasDistribNeg R] {a : R} {b : R} :
Commute (-a) b Commute a b
theorem Commute.neg_one_right {R : Type x} [MulOneClass R] [HasDistribNeg R] (a : R) :
Commute a (-1)
theorem Commute.neg_one_left {R : Type x} [MulOneClass R] [HasDistribNeg R] (a : R) :
Commute (-1) a
@[simp]
theorem Commute.sub_right {R : Type x} [NonUnitalNonAssocRing R] {a : R} {b : R} {c : R} :
Commute a bCommute a cCommute a (b - c)
@[simp]
theorem Commute.sub_left {R : Type x} [NonUnitalNonAssocRing R] {a : R} {b : R} {c : R} :
Commute a cCommute b cCommute (a - b) c
theorem mul_self_sub_mul_self {R : Type x} [CommRing R] (a : R) (b : R) :
a * a - b * b = (a + b) * (a - b)

Representation of a difference of two squares in a commutative ring as a product.

theorem mul_self_sub_one {R : Type x} [NonAssocRing R] (a : R) :
a * a - 1 = (a + 1) * (a - 1)
theorem mul_self_eq_mul_self_iff {R : Type x} [CommRing R] [NoZeroDivisors R] {a : R} {b : R} :
a * a = b * b a = b a = -b
theorem mul_self_eq_one_iff {R : Type x} [NonAssocRing R] [NoZeroDivisors R] {a : R} :
a * a = 1 a = 1 a = -1
theorem Units.inv_eq_self_iff {R : Type x} [Ring R] [NoZeroDivisors R] (u : Rˣ) :
u⁻¹ = u u = 1 u = -1

In the unit group of an integral domain, a unit is its own inverse iff the unit is one or one's additive inverse.

Equations
  • Ring.instBracket = { bracket := fun (x y : R) => x * y - y * x }
theorem Ring.lie_def {R : Type x} [NonUnitalNonAssocRing R] (x : R) (y : R) :
x, y = x * y - y * x
theorem commute_iff_lie_eq {R : Type x} [NonUnitalNonAssocRing R] {x : R} {y : R} :
Commute x y x, y = 0
theorem Commute.lie_eq {R : Type x} [NonUnitalNonAssocRing R] {x : R} {y : R} (h : Commute x y) :
x, y = 0