Absolute values in ordered groups #
The absolute value of an element in a group which is also a lattice is its supremum with its
negation. This generalizes the usual absolute value on real numbers (|x| = max x (-x)
).
Notations #
|a|
: The absolute value of an elementa
of an additive lattice ordered group|a|ₘ
: The absolute value of an elementa
of a multiplicative lattice ordered group
theorem
abs_nsmul
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
(n : ℕ)
(a : G)
:
theorem
inv_le_of_mabs_le
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b : G}
(h : mabs a ≤ b)
:
theorem
neg_le_of_abs_le
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
(h : |a| ≤ b)
:
theorem
le_of_mabs_le
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b : G}
(h : mabs a ≤ b)
:
theorem
le_of_abs_le
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
(h : |a| ≤ b)
:
The triangle inequality in LinearOrderedCommGroup
s.
The triangle inequality in LinearOrderedAddCommGroup
s.
theorem
mabs_div_le_iff
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
:
theorem
abs_sub_le_iff
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
:
theorem
mabs_div_lt_iff
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
:
theorem
abs_sub_lt_iff
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
:
theorem
div_le_of_mabs_div_le_left
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
(h : mabs (a / b) ≤ c)
:
theorem
sub_le_of_abs_sub_le_left
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
(h : |a - b| ≤ c)
:
theorem
div_le_of_mabs_div_le_right
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
(h : mabs (a / b) ≤ c)
:
theorem
sub_le_of_abs_sub_le_right
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
(h : |a - b| ≤ c)
:
theorem
div_lt_of_mabs_div_lt_left
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
(h : mabs (a / b) < c)
:
theorem
sub_lt_of_abs_sub_lt_left
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
(h : |a - b| < c)
:
theorem
div_lt_of_mabs_div_lt_right
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
(h : mabs (a / b) < c)
:
theorem
sub_lt_of_abs_sub_lt_right
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
(h : |a - b| < c)
:
theorem
mabs_div_mabs_le_mabs_div
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
(a b : G)
:
theorem
abs_sub_abs_le_abs_sub
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
(a b : G)
:
theorem
mabs_mabs_div_mabs_le_mabs_div
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
(a b : G)
:
theorem
abs_abs_sub_abs_le_abs_sub
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
(a b : G)
:
theorem
mabs_div_le_of_one_le_of_le
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b n : G}
(one_le_a : 1 ≤ a)
(a_le_n : a ≤ n)
(one_le_b : 1 ≤ b)
(b_le_n : b ≤ n)
:
|a / b|ₘ ≤ n
if 1 ≤ a ≤ n
and 1 ≤ b ≤ n
.
theorem
abs_sub_le_of_nonneg_of_le
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b n : G}
(one_le_a : 0 ≤ a)
(a_le_n : a ≤ n)
(one_le_b : 0 ≤ b)
(b_le_n : b ≤ n)
:
|a - b| ≤ n
if 0 ≤ a ≤ n
and 0 ≤ b ≤ n
.
theorem
mabs_div_lt_of_one_le_of_lt
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b n : G}
(one_le_a : 1 ≤ a)
(a_lt_n : a < n)
(one_le_b : 1 ≤ b)
(b_lt_n : b < n)
:
|a - b| < n
if 0 ≤ a < n
and 0 ≤ b < n
.
theorem
abs_sub_lt_of_nonneg_of_lt
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b n : G}
(one_le_a : 0 ≤ a)
(a_lt_n : a < n)
(one_le_b : 0 ≤ b)
(b_lt_n : b < n)
:
|a / b|ₘ < n
if 1 ≤ a < n
and 1 ≤ b < n
.
theorem
abs_eq
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
(hb : 0 ≤ b)
:
theorem
mabs_le_max_mabs_mabs
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b c : G}
(hab : a ≤ b)
(hbc : b ≤ c)
:
theorem
abs_le_max_abs_abs
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b c : G}
(hab : a ≤ b)
(hbc : b ≤ c)
:
theorem
eq_of_mabs_div_eq_one
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b : G}
(h : mabs (a / b) = 1)
:
theorem
eq_of_abs_sub_eq_zero
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
(h : |a - b| = 0)
:
theorem
abs_sub_le
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
(a b c : G)
:
theorem
mabs_mul_three
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
(a b c : G)
:
theorem
abs_add_three
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
(a b c : G)
:
theorem
mabs_div_le_of_le_of_le
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b lb ub : G}
(hal : lb ≤ a)
(hau : a ≤ ub)
(hbl : lb ≤ b)
(hbu : b ≤ ub)
:
theorem
abs_sub_le_of_le_of_le
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b lb ub : G}
(hal : lb ≤ a)
(hau : a ≤ ub)
(hbl : lb ≤ b)
(hbu : b ≤ ub)
:
@[deprecated abs_sub_le_of_le_of_le (since := "2025-03-02")]
theorem
dist_bdd_within_interval
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b lb ub : G}
(hal : lb ≤ a)
(hau : a ≤ ub)
(hbl : lb ≤ b)
(hbu : b ≤ ub)
:
Alias of abs_sub_le_of_le_of_le
.
theorem
eq_of_mabs_div_le_one
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b : G}
(h : mabs (a / b) ≤ 1)
:
theorem
eq_of_abs_sub_nonpos
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
(h : |a - b| ≤ 0)
:
theorem
eq_of_mabs_div_lt_all
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{x y : G}
(h : ∀ (ε : G), ε > 1 → mabs (x / y) < ε)
:
theorem
eq_of_abs_sub_lt_all
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{x y : G}
(h : ∀ (ε : G), ε > 0 → |x - y| < ε)
:
theorem
eq_of_mabs_div_le_all
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
[DenselyOrdered G]
{x y : G}
(h : ∀ (ε : G), ε > 1 → mabs (x / y) ≤ ε)
:
theorem
eq_of_abs_sub_le_all
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
[DenselyOrdered G]
{x y : G}
(h : ∀ (ε : G), ε > 0 → |x - y| ≤ ε)
:
theorem
mabs_div_le_one
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
{a b : G}
:
theorem
abs_sub_nonpos
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
:
theorem
abs_sub_pos
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a b : G}
:
@[simp]
@[simp]
theorem
abs_eq_self
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a : G}
:
@[simp]
@[simp]
theorem
abs_eq_neg_self
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{a : G}
:
For an element a
of an additive linear ordered group,
either |a| = a
and 0 ≤ a
, or |a| = -a
and a < 0
.
Use cases on this lemma to automate linarith in inequalities
@[simp]
theorem
max_one_mul_max_inv_one_eq_mabs_self
{G : Type u_1}
[CommGroup G]
[LinearOrder G]
[IsOrderedMonoid G]
(a : G)
:
@[simp]
theorem
max_zero_add_max_neg_zero_eq_abs_self
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
(a : G)
:
theorem
apply_abs_le_mul_of_one_le'
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{H : Type u_2}
[MulOneClass H]
[LE H]
[MulLeftMono H]
[MulRightMono H]
{f : G → H}
{a : G}
(h₁ : 1 ≤ f a)
(h₂ : 1 ≤ f (-a))
:
theorem
apply_abs_le_add_of_nonneg'
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{H : Type u_2}
[AddZeroClass H]
[LE H]
[AddLeftMono H]
[AddRightMono H]
{f : G → H}
{a : G}
(h₁ : 0 ≤ f a)
(h₂ : 0 ≤ f (-a))
:
theorem
apply_abs_le_mul_of_one_le
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{H : Type u_2}
[MulOneClass H]
[LE H]
[MulLeftMono H]
[MulRightMono H]
{f : G → H}
(h : ∀ (x : G), 1 ≤ f x)
(a : G)
:
theorem
apply_abs_le_add_of_nonneg
{G : Type u_1}
[AddCommGroup G]
[LinearOrder G]
[IsOrderedAddMonoid G]
{H : Type u_2}
[AddZeroClass H]
[LE H]
[AddLeftMono H]
[AddRightMono H]
{f : G → H}
(h : ∀ (x : G), 0 ≤ f x)
(a : G)
: