mathlib documentation

algebra.group_power.lemmas

Lemmas about power operations on monoids and groups #

This file contains lemmas about monoid.pow, group.pow, nsmul, zsmul which require additional imports besides those available in algebra.group_power.basic.

(Additive) monoid #

@[simp]
theorem nsmul_one {A : Type y} [add_monoid_with_one A] (n : ) :
n 1 = n
@[protected, instance]
def invertible_pow {M : Type u} [monoid M] (m : M) [invertible m] (n : ) :
Equations
theorem inv_of_pow {M : Type u} [monoid M] (m : M) [invertible m] (n : ) [invertible (m ^ n)] :
(m ^ n) = m ^ n
theorem is_unit.pow {M : Type u} [monoid M] {m : M} (n : ) :
is_unit mis_unit (m ^ n)
@[simp]
theorem is_unit_pow_succ_iff {M : Type u} [monoid M] {m : M} {n : } :
is_unit (m ^ (n + 1)) is_unit m
theorem is_unit_pos_pow_iff {M : Type u} [monoid M] {m : M} {n : } (h : 0 < n) :
def invertible_of_pow_succ_eq_one {M : Type u} [monoid M] (x : M) (n : ) (hx : x ^ n.succ = 1) :

If x ^ n.succ = 1 then x has an inverse, x^n.

Equations
def invertible_of_pow_eq_one {M : Type u} [monoid M] (x : M) (n : ) (hx : x ^ n = 1) (hn : 0 < n) :

If x ^ n = 1 then x has an inverse, x^(n - 1).

Equations
theorem is_unit_of_pow_eq_one {M : Type u} [monoid M] (x : M) (n : ) (hx : x ^ n = 1) (hn : 0 < n) :
theorem smul_pow {M : Type u} {N : Type v} [monoid M] [monoid N] [mul_action M N] [is_scalar_tower M N N] [smul_comm_class M N N] (k : M) (x : N) (p : ) :
(k x) ^ p = k ^ p x ^ p
@[simp]
theorem smul_pow' {M : Type u} {N : Type v} [monoid M] [monoid N] [mul_distrib_mul_action M N] (x : M) (m : N) (n : ) :
x m ^ n = (x m) ^ n
theorem zsmul_one {A : Type y} [add_group_with_one A] (n : ) :
n 1 = n
theorem mul_zsmul' {α : Type u_1} [subtraction_monoid α] (a : α) (m n : ) :
(m * n) a = n m a
theorem zpow_mul {α : Type u_1} [division_monoid α] (a : α) (m n : ) :
a ^ (m * n) = (a ^ m) ^ n
theorem zpow_mul' {α : Type u_1} [division_monoid α] (a : α) (m n : ) :
a ^ (m * n) = (a ^ n) ^ m
theorem mul_zsmul {α : Type u_1} [subtraction_monoid α] (a : α) (m n : ) :
(m * n) a = m n a
theorem bit0_zsmul {α : Type u_1} [subtraction_monoid α] (a : α) (n : ) :
bit0 n a = n a + n a
theorem zpow_bit0 {α : Type u_1} [division_monoid α] (a : α) (n : ) :
a ^ bit0 n = a ^ n * a ^ n
theorem zpow_bit0' {α : Type u_1} [division_monoid α] (a : α) (n : ) :
a ^ bit0 n = (a * a) ^ n
theorem bit0_zsmul' {α : Type u_1} [subtraction_monoid α] (a : α) (n : ) :
bit0 n a = n (a + a)
@[simp]
theorem zpow_bit0_neg {α : Type u_1} [division_monoid α] [has_distrib_neg α] (x : α) (n : ) :
(-x) ^ bit0 n = x ^ bit0 n
theorem zpow_add_one {G : Type w} [group G] (a : G) (n : ) :
a ^ (n + 1) = a ^ n * a
theorem add_one_zsmul {G : Type w} [add_group G] (a : G) (n : ) :
(n + 1) a = n a + a
theorem zpow_sub_one {G : Type w} [group G] (a : G) (n : ) :
a ^ (n - 1) = a ^ n * a⁻¹
theorem zsmul_sub_one {G : Type w} [add_group G] (a : G) (n : ) :
(n - 1) a = n a + -a
theorem zpow_add {G : Type w} [group G] (a : G) (m n : ) :
a ^ (m + n) = a ^ m * a ^ n
theorem add_zsmul {G : Type w} [add_group G] (a : G) (m n : ) :
(m + n) a = m a + n a
theorem mul_self_zpow {G : Type w} [group G] (b : G) (m : ) :
b * b ^ m = b ^ (m + 1)
theorem add_zsmul_self {G : Type w} [add_group G] (b : G) (m : ) :
b + m b = (m + 1) b
theorem add_self_zsmul {G : Type w} [add_group G] (b : G) (m : ) :
m b + b = (m + 1) b
theorem mul_zpow_self {G : Type w} [group G] (b : G) (m : ) :
b ^ m * b = b ^ (m + 1)
theorem sub_zsmul {G : Type w} [add_group G] (a : G) (m n : ) :
(m - n) a = m a + -(n a)
theorem zpow_sub {G : Type w} [group G] (a : G) (m n : ) :
a ^ (m - n) = a ^ m * (a ^ n)⁻¹
theorem one_add_zsmul {G : Type w} [add_group G] (a : G) (i : ) :
(1 + i) a = a + i a
theorem zpow_one_add {G : Type w} [group G] (a : G) (i : ) :
a ^ (1 + i) = a * a ^ i
theorem zsmul_add_comm {G : Type w} [add_group G] (a : G) (i j : ) :
i a + j a = j a + i a
theorem zpow_mul_comm {G : Type w} [group G] (a : G) (i j : ) :
a ^ i * a ^ j = a ^ j * a ^ i
theorem zpow_bit1 {G : Type w} [group G] (a : G) (n : ) :
a ^ bit1 n = a ^ n * a ^ n * a
theorem bit1_zsmul {G : Type w} [add_group G] (a : G) (n : ) :
bit1 n a = n a + n a + a

zpow/zsmul and an order #

Those lemmas are placed here (rather than in algebra.group_power.order with their friends) because they require facts from data.int.basic.

theorem one_lt_zpow' {α : Type u_1} [ordered_comm_group α] {a : α} (ha : 1 < a) {k : } (hk : 0 < k) :
1 < a ^ k
theorem zsmul_pos {α : Type u_1} [ordered_add_comm_group α] {a : α} (ha : 0 < a) {k : } (hk : 0 < k) :
0 < k a
theorem zsmul_strict_mono_left {α : Type u_1} [ordered_add_comm_group α] {a : α} (ha : 0 < a) :
strict_mono (λ (n : ), n a)
theorem zpow_strict_mono_right {α : Type u_1} [ordered_comm_group α] {a : α} (ha : 1 < a) :
strict_mono (λ (n : ), a ^ n)
theorem zsmul_mono_left {α : Type u_1} [ordered_add_comm_group α] {a : α} (ha : 0 a) :
monotone (λ (n : ), n a)
theorem zpow_mono_right {α : Type u_1} [ordered_comm_group α] {a : α} (ha : 1 a) :
monotone (λ (n : ), a ^ n)
theorem zsmul_le_zsmul {α : Type u_1} [ordered_add_comm_group α] {m n : } {a : α} (ha : 0 a) (h : m n) :
m a n a
theorem zpow_le_zpow {α : Type u_1} [ordered_comm_group α] {m n : } {a : α} (ha : 1 a) (h : m n) :
a ^ m a ^ n
theorem zsmul_lt_zsmul {α : Type u_1} [ordered_add_comm_group α] {m n : } {a : α} (ha : 0 < a) (h : m < n) :
m a < n a
theorem zpow_lt_zpow {α : Type u_1} [ordered_comm_group α] {m n : } {a : α} (ha : 1 < a) (h : m < n) :
a ^ m < a ^ n
theorem zsmul_le_zsmul_iff {α : Type u_1} [ordered_add_comm_group α] {m n : } {a : α} (ha : 0 < a) :
m a n a m n
theorem zpow_le_zpow_iff {α : Type u_1} [ordered_comm_group α] {m n : } {a : α} (ha : 1 < a) :
a ^ m a ^ n m n
theorem zpow_lt_zpow_iff {α : Type u_1} [ordered_comm_group α] {m n : } {a : α} (ha : 1 < a) :
a ^ m < a ^ n m < n
theorem zsmul_lt_zsmul_iff {α : Type u_1} [ordered_add_comm_group α] {m n : } {a : α} (ha : 0 < a) :
m a < n a m < n
theorem zsmul_strict_mono_right (α : Type u_1) [ordered_add_comm_group α] {n : } (hn : 0 < n) :
strict_mono (λ (_x : α), n _x)
theorem zpow_strict_mono_left (α : Type u_1) [ordered_comm_group α] {n : } (hn : 0 < n) :
strict_mono (λ (_x : α), _x ^ n)
theorem zpow_mono_left (α : Type u_1) [ordered_comm_group α] {n : } (hn : 0 n) :
monotone (λ (_x : α), _x ^ n)
theorem zsmul_mono_right (α : Type u_1) [ordered_add_comm_group α] {n : } (hn : 0 n) :
monotone (λ (_x : α), n _x)
theorem zpow_le_zpow' {α : Type u_1} [ordered_comm_group α] {n : } {a b : α} (hn : 0 n) (h : a b) :
a ^ n b ^ n
theorem zsmul_le_zsmul' {α : Type u_1} [ordered_add_comm_group α] {n : } {a b : α} (hn : 0 n) (h : a b) :
n a n b
theorem zsmul_lt_zsmul' {α : Type u_1} [ordered_add_comm_group α] {n : } {a b : α} (hn : 0 < n) (h : a < b) :
n a < n b
theorem zpow_lt_zpow' {α : Type u_1} [ordered_comm_group α] {n : } {a b : α} (hn : 0 < n) (h : a < b) :
a ^ n < b ^ n
theorem zsmul_le_zsmul_iff' {α : Type u_1} [linear_ordered_add_comm_group α] {n : } (hn : 0 < n) {a b : α} :
n a n b a b
theorem zpow_le_zpow_iff' {α : Type u_1} [linear_ordered_comm_group α] {n : } (hn : 0 < n) {a b : α} :
a ^ n b ^ n a b
theorem zpow_lt_zpow_iff' {α : Type u_1} [linear_ordered_comm_group α] {n : } (hn : 0 < n) {a b : α} :
a ^ n < b ^ n a < b
theorem zsmul_lt_zsmul_iff' {α : Type u_1} [linear_ordered_add_comm_group α] {n : } (hn : 0 < n) {a b : α} :
n a < n b a < b
@[nolint]
theorem zpow_left_injective {α : Type u_1} [linear_ordered_comm_group α] {n : } (hn : n 0) :
function.injective (λ (_x : α), _x ^ n)
@[nolint]
theorem zsmul_right_injective {α : Type u_1} [linear_ordered_add_comm_group α] {n : } (hn : n 0) :
function.injective (λ (_x : α), n _x)

See also smul_right_injective. TODO: provide a no_zero_smul_divisors instance. We can't do that here because importing that definition would create import cycles.

theorem zsmul_right_inj {α : Type u_1} [linear_ordered_add_comm_group α] {n : } {a b : α} (hn : n 0) :
n a = n b a = b
theorem zpow_left_inj {α : Type u_1} [linear_ordered_comm_group α] {n : } {a b : α} (hn : n 0) :
a ^ n = b ^ n a = b
theorem zpow_eq_zpow_iff' {α : Type u_1} [linear_ordered_comm_group α] {n : } {a b : α} (hn : n 0) :
a ^ n = b ^ n a = b

Alias of zsmul_right_inj, for ease of discovery alongside zsmul_le_zsmul_iff' and zsmul_lt_zsmul_iff'.

theorem zsmul_eq_zsmul_iff' {α : Type u_1} [linear_ordered_add_comm_group α] {n : } {a b : α} (hn : n 0) :
n a = n b a = b

Alias of zsmul_right_inj, for ease of discovery alongside zsmul_le_zsmul_iff' and zsmul_lt_zsmul_iff'.

theorem abs_nsmul {α : Type u_1} [linear_ordered_add_comm_group α] (n : ) (a : α) :
|n a| = n |a|
theorem abs_zsmul {α : Type u_1} [linear_ordered_add_comm_group α] (n : ) (a : α) :
|n a| = |n| |a|
theorem abs_add_eq_add_abs_le {α : Type u_1} [linear_ordered_add_comm_group α] {a b : α} (hle : a b) :
|a + b| = |a| + |b| 0 a 0 b a 0 b 0
theorem abs_add_eq_add_abs_iff {α : Type u_1} [linear_ordered_add_comm_group α] (a b : α) :
|a + b| = |a| + |b| 0 a 0 b a 0 b 0
@[simp]
theorem with_bot.coe_nsmul {A : Type y} [add_monoid A] (a : A) (n : ) :
(n a) = n a
theorem nsmul_eq_mul' {R : Type u₁} [non_assoc_semiring R] (a : R) (n : ) :
n a = a * n
@[simp]
theorem nsmul_eq_mul {R : Type u₁} [non_assoc_semiring R] (n : ) (a : R) :
n a = n * a
@[protected, instance]

Note that add_comm_monoid.nat_smul_comm_class requires stronger assumptions on R.

@[protected, instance]

Note that add_comm_monoid.nat_is_scalar_tower requires stronger assumptions on R.

@[simp, norm_cast]
theorem nat.cast_pow {R : Type u₁} [semiring R] (n m : ) :
(n ^ m) = n ^ m
@[simp, norm_cast]
theorem int.coe_nat_pow (n m : ) :
(n ^ m) = n ^ m
theorem int.nat_abs_pow (n : ) (k : ) :
(n ^ k).nat_abs = n.nat_abs ^ k
theorem bit0_mul {R : Type u₁} [non_unital_non_assoc_ring R] {n r : R} :
bit0 n * r = 2 (n * r)
theorem mul_bit0 {R : Type u₁} [non_unital_non_assoc_ring R] {n r : R} :
r * bit0 n = 2 (r * n)
theorem bit1_mul {R : Type u₁} [non_assoc_ring R] {n r : R} :
bit1 n * r = 2 (n * r) + r
theorem mul_bit1 {R : Type u₁} [non_assoc_ring R] {n r : R} :
r * bit1 n = 2 (r * n) + r
@[simp]
theorem zsmul_eq_mul {R : Type u₁} [ring R] (a : R) (n : ) :
n a = n * a
theorem zsmul_eq_mul' {R : Type u₁} [ring R] (a : R) (n : ) :
n a = a * n
@[protected, instance]

Note that add_comm_group.int_smul_comm_class requires stronger assumptions on R.

@[protected, instance]

Note that add_comm_group.int_is_scalar_tower requires stronger assumptions on R.

theorem zsmul_int_int (a b : ) :
a b = a * b
theorem zsmul_int_one (n : ) :
n 1 = n
@[simp, norm_cast]
theorem int.cast_pow {R : Type u₁} [ring R] (n : ) (m : ) :
(n ^ m) = n ^ m
theorem neg_one_pow_eq_pow_mod_two {R : Type u₁} [ring R] {n : } :
(-1) ^ n = (-1) ^ (n % 2)
theorem one_add_mul_le_pow' {R : Type u₁} [ordered_semiring R] {a : R} (Hsq : 0 a * a) (Hsq' : 0 (1 + a) * (1 + a)) (H : 0 2 + a) (n : ) :
1 + n * a (1 + a) ^ n

Bernoulli's inequality. This version works for semirings but requires additional hypotheses 0 ≤ a * a and 0 ≤ (1 + a) * (1 + a).

theorem pow_le_pow_of_le_one {R : Type u₁} [ordered_semiring R] {a : R} (h : 0 a) (ha : a 1) {i j : } (hij : i j) :
a ^ j a ^ i
theorem pow_le_of_le_one {R : Type u₁} [ordered_semiring R] {a : R} (h₀ : 0 a) (h₁ : a 1) {n : } (hn : n 0) :
a ^ n a
theorem sq_le {R : Type u₁} [ordered_semiring R] {a : R} (h₀ : 0 a) (h₁ : a 1) :
a ^ 2 a
theorem sign_cases_of_C_mul_pow_nonneg {R : Type u₁} [linear_ordered_semiring R] {C r : R} (h : ∀ (n : ), 0 C * r ^ n) :
C = 0 0 < C 0 r
@[simp]
theorem abs_pow {R : Type u₁} [linear_ordered_ring R] (a : R) (n : ) :
|a ^ n| = |a| ^ n
@[simp]
theorem pow_bit1_neg_iff {R : Type u₁} [linear_ordered_ring R] {a : R} {n : } :
a ^ bit1 n < 0 a < 0
@[simp]
theorem pow_bit1_nonneg_iff {R : Type u₁} [linear_ordered_ring R] {a : R} {n : } :
0 a ^ bit1 n 0 a
@[simp]
theorem pow_bit1_nonpos_iff {R : Type u₁} [linear_ordered_ring R] {a : R} {n : } :
a ^ bit1 n 0 a 0
@[simp]
theorem pow_bit1_pos_iff {R : Type u₁} [linear_ordered_ring R] {a : R} {n : } :
0 < a ^ bit1 n 0 < a
theorem strict_mono_pow_bit1 {R : Type u₁} [linear_ordered_ring R] (n : ) :
strict_mono (λ (a : R), a ^ bit1 n)
theorem one_add_mul_le_pow {R : Type u₁} [linear_ordered_ring R] {a : R} (H : -2 a) (n : ) :
1 + n * a (1 + a) ^ n

Bernoulli's inequality for n : ℕ, -2 ≤ a.

theorem one_add_mul_sub_le_pow {R : Type u₁} [linear_ordered_ring R] {a : R} (H : -1 a) (n : ) :
1 + n * (a - 1) a ^ n

Bernoulli's inequality reformulated to estimate a^n.

theorem int.units_pow_two (u : ˣ) :
u ^ 2 = 1

Alias of int.units_sq.

theorem int.units_pow_eq_pow_mod_two (u : ˣ) (n : ) :
u ^ n = u ^ (n % 2)
@[simp]
theorem int.nat_abs_sq (x : ) :
(x.nat_abs) ^ 2 = x ^ 2
theorem int.nat_abs_pow_two (x : ) :
(x.nat_abs) ^ 2 = x ^ 2

Alias of int.nat_abs_sq.

theorem int.abs_le_self_sq (a : ) :
(a.nat_abs) a ^ 2
theorem int.abs_le_self_pow_two (a : ) :
(a.nat_abs) a ^ 2

Alias of int.abs_le_self_sq.

theorem int.le_self_sq (b : ) :
b b ^ 2
theorem int.le_self_pow_two (b : ) :
b b ^ 2

Alias of int.le_self_sq.

def powers_hom (M : Type u) [monoid M] :

Monoid homomorphisms from multiplicative are defined by the image of multiplicative.of_add 1.

Equations
def zpowers_hom (G : Type w) [group G] :

Monoid homomorphisms from multiplicative are defined by the image of multiplicative.of_add 1.

Equations
def multiples_hom (A : Type y) [add_monoid A] :
A ( →+ A)

Additive homomorphisms from are defined by the image of 1.

Equations
def zmultiples_hom (A : Type y) [add_group A] :
A ( →+ A)

Additive homomorphisms from are defined by the image of 1.

Equations
@[simp]
theorem powers_hom_apply {M : Type u} [monoid M] (x : M) (n : multiplicative ) :
@[simp]
@[simp]
theorem zpowers_hom_apply {G : Type w} [group G] (x : G) (n : multiplicative ) :
@[simp]
@[simp]
theorem multiples_hom_apply {A : Type y} [add_monoid A] (x : A) (n : ) :
((multiples_hom A) x) n = n x
@[simp]
theorem multiples_hom_symm_apply {A : Type y} [add_monoid A] (f : →+ A) :
@[simp]
theorem zmultiples_hom_apply {A : Type y} [add_group A] (x : A) (n : ) :
((zmultiples_hom A) x) n = n x
@[simp]
theorem zmultiples_hom_symm_apply {A : Type y} [add_group A] (f : →+ A) :
@[ext]
theorem monoid_hom.ext_mnat {M : Type u} [monoid M] ⦃f g : multiplicative →* M⦄ (h : f (multiplicative.of_add 1) = g (multiplicative.of_add 1)) :
f = g

monoid_hom.ext_mint is defined in data.int.cast

theorem add_monoid_hom.apply_nat {M : Type u} [add_monoid M] (f : →+ M) (n : ) :
f n = n f 1

add_monoid_hom.ext_nat is defined in data.nat.cast

theorem add_monoid_hom.apply_int {M : Type u} [add_group M] (f : →+ M) (n : ) :
f n = n f 1

add_monoid_hom.ext_int is defined in data.int.cast

def powers_mul_hom (M : Type u) [comm_monoid M] :

If M is commutative, powers_hom is a multiplicative equivalence.

Equations
def zpowers_mul_hom (G : Type w) [comm_group G] :

If M is commutative, zpowers_hom is a multiplicative equivalence.

Equations
def multiples_add_hom (A : Type y) [add_comm_monoid A] :
A ≃+ ( →+ A)

If M is commutative, multiples_hom is an additive equivalence.

Equations
def zmultiples_add_hom (A : Type y) [add_comm_group A] :
A ≃+ ( →+ A)

If M is commutative, zmultiples_hom is an additive equivalence.

Equations
@[simp]
theorem powers_mul_hom_apply {M : Type u} [comm_monoid M] (x : M) (n : multiplicative ) :
@[simp]
theorem zpowers_mul_hom_apply {G : Type w} [comm_group G] (x : G) (n : multiplicative ) :
@[simp]
theorem multiples_add_hom_apply {A : Type y} [add_comm_monoid A] (x : A) (n : ) :
@[simp]
theorem multiples_add_hom_symm_apply {A : Type y} [add_comm_monoid A] (f : →+ A) :
@[simp]
theorem zmultiples_add_hom_apply {A : Type y} [add_comm_group A] (x : A) (n : ) :
@[simp]
theorem zmultiples_add_hom_symm_apply {A : Type y} [add_comm_group A] (f : →+ A) :

Commutativity (again) #

Facts about semiconj_by and commute that require zpow or zsmul, or the fact that integer multiplication equals semiring multiplication.

@[simp]
theorem semiconj_by.cast_nat_mul_right {R : Type u₁} [semiring R] {a x y : R} (h : semiconj_by a x y) (n : ) :
semiconj_by a (n * x) (n * y)
@[simp]
theorem semiconj_by.cast_nat_mul_left {R : Type u₁} [semiring R] {a x y : R} (h : semiconj_by a x y) (n : ) :
semiconj_by (n * a) x y
@[simp]
theorem semiconj_by.cast_nat_mul_cast_nat_mul {R : Type u₁} [semiring R] {a x y : R} (h : semiconj_by a x y) (m n : ) :
semiconj_by (m * a) (n * x) (n * y)
@[simp]
theorem add_semiconj_by.add_units_zsmul_right {M : Type u} [add_monoid M] {a : M} {x y : add_units M} (h : add_semiconj_by a x y) (m : ) :
@[simp]
theorem semiconj_by.units_zpow_right {M : Type u} [monoid M] {a : M} {x y : Mˣ} (h : semiconj_by a x y) (m : ) :
semiconj_by a (x ^ m) (y ^ m)
@[simp]
theorem semiconj_by.cast_int_mul_right {R : Type u₁} [ring R] {a x y : R} (h : semiconj_by a x y) (m : ) :
semiconj_by a (m * x) (m * y)
@[simp]
theorem semiconj_by.cast_int_mul_left {R : Type u₁} [ring R] {a x y : R} (h : semiconj_by a x y) (m : ) :
semiconj_by (m * a) x y
@[simp]
theorem semiconj_by.cast_int_mul_cast_int_mul {R : Type u₁} [ring R] {a x y : R} (h : semiconj_by a x y) (m n : ) :
semiconj_by (m * a) (n * x) (n * y)
@[simp]
theorem commute.cast_nat_mul_right {R : Type u₁} [semiring R] {a b : R} (h : commute a b) (n : ) :
commute a (n * b)
@[simp]
theorem commute.cast_nat_mul_left {R : Type u₁} [semiring R] {a b : R} (h : commute a b) (n : ) :
commute (n * a) b
@[simp]
theorem commute.cast_nat_mul_cast_nat_mul {R : Type u₁} [semiring R] {a b : R} (h : commute a b) (m n : ) :
commute (m * a) (n * b)
@[simp]
theorem commute.self_cast_nat_mul {R : Type u₁} [semiring R] {a : R} (n : ) :
commute a (n * a)
@[simp]
theorem commute.cast_nat_mul_self {R : Type u₁} [semiring R] {a : R} (n : ) :
commute (n * a) a
@[simp]
theorem commute.self_cast_nat_mul_cast_nat_mul {R : Type u₁} [semiring R] {a : R} (m n : ) :
commute (m * a) (n * a)
@[simp]
theorem commute.units_zpow_right {M : Type u} [monoid M] {a : M} {u : Mˣ} (h : commute a u) (m : ) :
commute a (u ^ m)
@[simp]
theorem add_commute.add_units_zsmul_right {M : Type u} [add_monoid M] {a : M} {u : add_units M} (h : add_commute a u) (m : ) :
@[simp]
theorem commute.units_zpow_left {M : Type u} [monoid M] {u : Mˣ} {a : M} (h : commute u a) (m : ) :
commute (u ^ m) a
@[simp]
theorem add_commute.add_units_zsmul_left {M : Type u} [add_monoid M] {u : add_units M} {a : M} (h : add_commute u a) (m : ) :
@[simp]
theorem commute.cast_int_mul_right {R : Type u₁} [ring R] {a b : R} (h : commute a b) (m : ) :
commute a (m * b)
@[simp]
theorem commute.cast_int_mul_left {R : Type u₁} [ring R] {a b : R} (h : commute a b) (m : ) :
commute (m * a) b
theorem commute.cast_int_mul_cast_int_mul {R : Type u₁} [ring R] {a b : R} (h : commute a b) (m n : ) :
commute (m * a) (n * b)
@[simp]
theorem commute.cast_int_left {R : Type u₁} [ring R] (a : R) (m : ) :
@[simp]
theorem commute.cast_int_right {R : Type u₁} [ring R] (a : R) (m : ) :
@[simp]
theorem commute.self_cast_int_mul {R : Type u₁} [ring R] (a : R) (n : ) :
commute a (n * a)
@[simp]
theorem commute.cast_int_mul_self {R : Type u₁} [ring R] (a : R) (n : ) :
commute (n * a) a
theorem commute.self_cast_int_mul_cast_int_mul {R : Type u₁} [ring R] (a : R) (m n : ) :
commute (m * a) (n * a)
theorem units.conj_pow {M : Type u} [monoid M] (u : Mˣ) (x : M) (n : ) :
(u * x * u⁻¹) ^ n = u * x ^ n * u⁻¹
theorem units.conj_pow' {M : Type u} [monoid M] (u : Mˣ) (x : M) (n : ) :
(u⁻¹ * x * u) ^ n = u⁻¹ * x ^ n * u
@[simp]
theorem mul_opposite.op_pow {M : Type u} [monoid M] (x : M) (n : ) :

Moving to the opposite monoid commutes with taking powers.

@[simp]
theorem mul_opposite.unop_pow {M : Type u} [monoid M] (x : Mᵐᵒᵖ) (n : ) :
@[simp]
theorem mul_opposite.op_zpow {M : Type u} [div_inv_monoid M] (x : M) (z : ) :

Moving to the opposite group or group_with_zero commutes with taking powers.

@[simp]
theorem mul_opposite.unop_zpow {M : Type u} [div_inv_monoid M] (x : Mᵐᵒᵖ) (z : ) :