Lemmas about the interaction of power operations with order in terms of CovariantClass #
theorem
Left.one_le_pow_of_le
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
{a : M}
(ha : 1 ≤ a)
(n : ℕ)
 :
theorem
Left.nsmul_nonneg
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
{a : M}
(ha : 0 ≤ a)
(n : ℕ)
 :
theorem
Left.pow_le_one_of_le
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
{a : M}
(ha : a ≤ 1)
(n : ℕ)
 :
theorem
Left.nsmul_nonpos
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
{a : M}
(ha : a ≤ 0)
(n : ℕ)
 :
theorem
one_le_pow_of_one_le'
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
{a : M}
(ha : 1 ≤ a)
(n : ℕ)
 :
Alias of Left.one_le_pow_of_le.
theorem
nsmul_nonneg
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
{a : M}
(ha : 0 ≤ a)
(n : ℕ)
 :
theorem
pow_le_one'
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
{a : M}
(ha : a ≤ 1)
(n : ℕ)
 :
Alias of Left.pow_le_one_of_le.
theorem
nsmul_nonpos
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
{a : M}
(ha : a ≤ 0)
(n : ℕ)
 :
theorem
pow_lt_one'
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
{a : M}
{n : ℕ}
(h : a < 1)
(hn : n ≠ 0)
 :
Alias of Left.pow_lt_one_of_lt.
theorem
pow_right_monotone
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
{a : M}
(ha : 1 ≤ a)
 :
theorem
nsmul_left_monotone
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
{a : M}
(ha : 0 ≤ a)
 :
theorem
pow_right_strictMono'
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftStrictMono M]
{a : M}
(ha : 1 < a)
 :
StrictMono fun (x : ℕ) => a ^ x
theorem
nsmul_left_strictMono
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
{a : M}
(ha : 0 < a)
 :
StrictMono fun (x : ℕ) => x • a
theorem
Right.one_le_pow_of_le
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulRightMono M]
{x : M}
(hx : 1 ≤ x)
{n : ℕ}
 :
theorem
Right.nsmul_nonneg
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddRightMono M]
{x : M}
(hx : 0 ≤ x)
{n : ℕ}
 :
theorem
Right.pow_le_one_of_le
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulRightMono M]
{x : M}
(hx : x ≤ 1)
{n : ℕ}
 :
theorem
Right.nsmul_nonpos
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddRightMono M]
{x : M}
(hx : x ≤ 0)
{n : ℕ}
 :
theorem
StrictMono.pow_const
{β : Type u_1}
{M : Type u_3}
[Monoid M]
[Preorder M]
[Preorder β]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{f : β → M}
(hf : StrictMono f)
{n : ℕ}
 :
n ≠ 0 → StrictMono fun (x : β) => f x ^ n
theorem
StrictMono.const_nsmul
{β : Type u_1}
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[Preorder β]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{f : β → M}
(hf : StrictMono f)
{n : ℕ}
 :
n ≠ 0 → StrictMono fun (x : β) => n • f x
theorem
pow_left_strictMono
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{n : ℕ}
(hn : n ≠ 0)
 :
StrictMono fun (x : M) => x ^ n
See also pow_left_strictMonoOn₀.
theorem
nsmul_right_strictMono
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{n : ℕ}
(hn : n ≠ 0)
 :
StrictMono fun (x : M) => n • x
theorem
pow_lt_pow_left'
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{n : ℕ}
(hn : n ≠ 0)
{a b : M}
(hab : a < b)
 :
theorem
nsmul_lt_nsmul_right
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{n : ℕ}
(hn : n ≠ 0)
{a b : M}
(hab : a < b)
 :
theorem
pow_le_pow_left'
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
[MulRightMono M]
{a b : M}
(hab : a ≤ b)
(i : ℕ)
 :
theorem
nsmul_le_nsmul_right
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
[AddRightMono M]
{a b : M}
(hab : a ≤ b)
(i : ℕ)
 :
theorem
Monotone.pow_const
{β : Type u_1}
{M : Type u_3}
[Monoid M]
[Preorder M]
[Preorder β]
[MulLeftMono M]
[MulRightMono M]
{f : β → M}
(hf : Monotone f)
(n : ℕ)
 :
theorem
Monotone.const_nsmul
{β : Type u_1}
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[Preorder β]
[AddLeftMono M]
[AddRightMono M]
{f : β → M}
(hf : Monotone f)
(n : ℕ)
 :
theorem
pow_left_mono
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
[MulRightMono M]
(n : ℕ)
 :
theorem
nsmul_right_mono
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
[AddRightMono M]
(n : ℕ)
 :
theorem
pow_le_pow
{M : Type u_3}
[Monoid M]
[Preorder M]
[MulLeftMono M]
[MulRightMono M]
{a b : M}
(hab : a ≤ b)
(ht : 1 ≤ b)
{m n : ℕ}
(hmn : m ≤ n)
 :
theorem
nsmul_le_nsmul
{M : Type u_3}
[AddMonoid M]
[Preorder M]
[AddLeftMono M]
[AddRightMono M]
{a b : M}
(hab : a ≤ b)
(ht : 0 ≤ b)
{m n : ℕ}
(hmn : m ≤ n)
 :
theorem
le_pow_sup
{M : Type u_3}
[Monoid M]
[SemilatticeSup M]
[MulLeftMono M]
[MulRightMono M]
{a b : M}
{n : ℕ}
 :
theorem
pow_inf_le
{M : Type u_3}
[Monoid M]
[SemilatticeInf M]
[MulLeftMono M]
[MulRightMono M]
{a b : M}
{n : ℕ}
 :
theorem
one_le_pow_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
nsmul_nonneg_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
pow_le_one_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
nsmul_nonpos_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
one_lt_pow_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
nsmul_pos_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
pow_lt_one_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
nsmul_neg_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
pow_eq_one_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
nsmul_eq_zero_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
{x : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
pow_le_pow_iff_right'
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
{a : M}
{m n : ℕ}
(ha : 1 < a)
 :
theorem
nsmul_le_nsmul_iff_left
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
{a : M}
{m n : ℕ}
(ha : 0 < a)
 :
theorem
pow_lt_pow_iff_right'
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
{a : M}
{m n : ℕ}
(ha : 1 < a)
 :
theorem
nsmul_lt_nsmul_iff_left
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
{a : M}
{m n : ℕ}
(ha : 0 < a)
 :
theorem
lt_of_pow_lt_pow_left'
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
[MulRightMono M]
{a b : M}
(n : ℕ)
 :
theorem
lt_of_nsmul_lt_nsmul_right
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
[AddRightMono M]
{a b : M}
(n : ℕ)
 :
theorem
min_lt_of_mul_lt_sq
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
[MulRightMono M]
{a b c : M}
(h : a * b < c ^ 2)
 :
theorem
min_lt_of_add_lt_two_nsmul
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
[AddRightMono M]
{a b c : M}
(h : a + b < 2 • c)
 :
theorem
lt_max_of_sq_lt_mul
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftMono M]
[MulRightMono M]
{a b c : M}
(h : a ^ 2 < b * c)
 :
theorem
lt_max_of_two_nsmul_lt_add
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftMono M]
[AddRightMono M]
{a b c : M}
(h : 2 • a < b + c)
 :
theorem
pow_le_pow_iff_left
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{a b : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
nsmul_le_nsmul_iff_right
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{a b : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
le_of_pow_le_pow_left'
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{a b : M}
{n : ℕ}
(hn : n ≠ 0)
 :
Alias of the forward direction of pow_le_pow_iff_left.
theorem
le_of_nsmul_le_nsmul_right
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{a b : M}
{n : ℕ}
(hn : n ≠ 0)
 :
theorem
min_le_of_mul_le_sq
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{a b c : M}
(h : a * b ≤ c ^ 2)
 :
theorem
min_le_of_add_le_two_nsmul
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{a b c : M}
(h : a + b ≤ 2 • c)
 :
theorem
le_max_of_sq_le_mul
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
{a b c : M}
(h : a ^ 2 ≤ b * c)
 :
theorem
le_max_of_two_nsmul_le_add
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
{a b c : M}
(h : 2 • a ≤ b + c)
 :
theorem
Left.pow_lt_one_iff'
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
{n : ℕ}
{x : M}
(hn : 0 < n)
 :
theorem
Left.nsmul_neg_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
{n : ℕ}
{x : M}
(hn : 0 < n)
 :
theorem
Left.pow_lt_one_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
{n : ℕ}
{x : M}
(hn : 0 < n)
 :
theorem
Right.pow_lt_one_iff
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulRightStrictMono M]
{n : ℕ}
{x : M}
(hn : 0 < n)
 :
theorem
Right.nsmul_neg_iff
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddRightStrictMono M]
{n : ℕ}
{x : M}
(hn : 0 < n)
 :
instance
instIsMulTorsionFreeOfMulLeftStrictMonoOfMulRightStrictMono
{M : Type u_3}
[Monoid M]
[LinearOrder M]
[MulLeftStrictMono M]
[MulRightStrictMono M]
 :
instance
instIsAddTorsionFreeOfAddLeftStrictMonoOfAddRightStrictMono
{M : Type u_3}
[AddMonoid M]
[LinearOrder M]
[AddLeftStrictMono M]
[AddRightStrictMono M]
 :
theorem
one_le_zpow
{G : Type u_2}
[DivInvMonoid G]
[Preorder G]
[MulLeftMono G]
{x : G}
(H : 1 ≤ x)
{n : ℤ}
(hn : 0 ≤ n)
 :
theorem
zsmul_nonneg
{G : Type u_2}
[SubNegMonoid G]
[Preorder G]
[AddLeftMono G]
{x : G}
(H : 0 ≤ x)
{n : ℤ}
(hn : 0 ≤ n)
 :
theorem
one_lt_zpow
{G : Type u_2}
[DivInvMonoid G]
[Preorder G]
[MulLeftMono G]
{x : G}
(hx : 1 < x)
{n : ℤ}
(hn : 0 < n)
 :