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 Commute.sq_sub_sq {R : Type x} [Ring R] {a : R} {b : R} (h : Commute a b) :
a ^ 2 - b ^ 2 = (a + b) * (a - b)
theorem Commute.sq_eq_sq_iff_eq_or_eq_neg {R : Type x} [Ring R] {a : R} {b : R} [NoZeroDivisors R] (h : Commute a b) :
a ^ 2 = b ^ 2 a = b a = -b
theorem neg_one_pow_eq_or (R : Type x) [Monoid R] [HasDistribNeg R] (n : ) :
(-1) ^ n = 1 (-1) ^ n = -1
theorem neg_pow {R : Type x} [Monoid R] [HasDistribNeg R] (a : R) (n : ) :
(-a) ^ n = (-1) ^ n * a ^ n
theorem neg_pow' {R : Type x} [Monoid R] [HasDistribNeg R] (a : R) (n : ) :
(-a) ^ n = a ^ n * (-1) ^ n
theorem neg_pow_bit0 {R : Type x} [Monoid R] [HasDistribNeg R] (a : R) (n : ) :
(-a) ^ bit0 n = a ^ bit0 n
@[simp]
theorem neg_pow_bit1 {R : Type x} [Monoid R] [HasDistribNeg R] (a : R) (n : ) :
(-a) ^ bit1 n = -a ^ bit1 n
theorem neg_sq {R : Type x} [Monoid R] [HasDistribNeg R] (a : R) :
(-a) ^ 2 = a ^ 2
theorem neg_one_sq {R : Type x} [Monoid R] [HasDistribNeg R] :
(-1) ^ 2 = 1
theorem neg_pow_two {R : Type x} [Monoid R] [HasDistribNeg R] (a : R) :
(-a) ^ 2 = a ^ 2

Alias of neg_sq.

theorem neg_one_pow_two {R : Type x} [Monoid R] [HasDistribNeg R] :
(-1) ^ 2 = 1

Alias of neg_one_sq.

@[simp]
theorem neg_one_pow_mul_eq_zero_iff {R : Type x} [Ring R] {a : R} {n : } :
(-1) ^ n * a = 0 a = 0
@[simp]
theorem mul_neg_one_pow_eq_zero_iff {R : Type x} [Ring R] {a : R} {n : } :
a * (-1) ^ n = 0 a = 0
@[simp]
theorem sq_eq_one_iff {R : Type x} [Ring R] {a : R} [NoZeroDivisors R] :
a ^ 2 = 1 a = 1 a = -1
theorem sq_ne_one_iff {R : Type x} [Ring R] {a : R} [NoZeroDivisors R] :
a ^ 2 1 a 1 a -1
theorem neg_one_pow_eq_pow_mod_two {R : Type x} [Ring R] (n : ) :
(-1) ^ n = (-1) ^ (n % 2)
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 sq_sub_sq {R : Type x} [CommRing R] (a : R) (b : R) :
a ^ 2 - b ^ 2 = (a + b) * (a - b)
theorem pow_two_sub_pow_two {R : Type x} [CommRing R] (a : R) (b : R) :
a ^ 2 - b ^ 2 = (a + b) * (a - b)

Alias of sq_sub_sq.

theorem sub_sq {R : Type x} [CommRing R] (a : R) (b : R) :
(a - b) ^ 2 = a ^ 2 - 2 * a * b + b ^ 2
theorem sub_pow_two {R : Type x} [CommRing R] (a : R) (b : R) :
(a - b) ^ 2 = a ^ 2 - 2 * a * b + b ^ 2

Alias of sub_sq.

theorem sub_sq' {R : Type x} [CommRing R] (a : R) (b : R) :
(a - b) ^ 2 = a ^ 2 + b ^ 2 - 2 * a * b
theorem sq_eq_sq_iff_eq_or_eq_neg {R : Type x} [CommRing R] [NoZeroDivisors R] {a : R} {b : R} :
a ^ 2 = b ^ 2 a = b a = -b
theorem eq_or_eq_neg_of_sq_eq_sq {R : Type x} [CommRing R] [NoZeroDivisors R] (a : R) (b : R) :
a ^ 2 = b ^ 2a = b a = -b
theorem Units.sq_eq_sq_iff_eq_or_eq_neg {R : Type x} [CommRing R] [NoZeroDivisors R] {a : Rˣ} {b : Rˣ} :
a ^ 2 = b ^ 2 a = b a = -b
theorem Units.eq_or_eq_neg_of_sq_eq_sq {R : Type x} [CommRing R] [NoZeroDivisors R] (a : Rˣ) (b : Rˣ) (h : a ^ 2 = b ^ 2) :
a = b a = -b
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.

@[instance 100]
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