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
mabs a
is the absolute value of a
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
abs a
is the absolute value of a
Equations
- One or more equations did not get rendered due to their size.
Instances For
Unexpander for the notation |a|ₘ
for mabs a
.
Tries to add discretionary parentheses in unparseable cases.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Unexpander for the notation |a|
for abs a
.
Tries to add discretionary parentheses in unparseable cases.
Equations
- One or more equations did not get rendered due to their size.
Instances For
theorem
abs_of_nonneg
{α : Type u_1}
[Lattice α]
[AddGroup α]
{a : α}
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(h : 0 ≤ a)
:
|a| = a
theorem
abs_of_pos
{α : Type u_1}
[Lattice α]
[AddGroup α]
{a : α}
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(h : 0 < a)
:
|a| = a
@[simp]
theorem
abs_nonneg
{α : Type u_1}
[Lattice α]
[AddGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
0 ≤ |a|
@[simp]
theorem
one_le_mabs
{α : Type u_1}
[Lattice α]
[Group α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
@[simp]
theorem
abs_abs
{α : Type u_1}
[Lattice α]
[AddGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
|(|a|)| = |a|
theorem
abs_add_le
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
The absolute value satisfies the triangle inequality.
theorem
abs_abs_sub_abs_le
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
sup_sub_inf_eq_abs_sub
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
two_nsmul_sup_eq_add_add_abs_sub
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
two_nsmul_inf_eq_add_sub_abs_sub
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
abs_sup_sub_sup_le_abs
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
(c : α)
:
theorem
abs_inf_sub_inf_le_abs
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
(c : α)
:
theorem
abs_by_cases
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
{a : α}
(P : α → Prop)
(h1 : P a)
(h2 : P (-a))
:
P |a|
theorem
mabs_by_cases
{α : Type u_1}
[Group α]
[LinearOrder α]
{a : α}
(P : α → Prop)
(h1 : P a)
(h2 : P a⁻¹)
:
P (mabs a)
theorem
eq_or_eq_neg_of_abs_eq
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
{a : α}
{b : α}
(h : |a| = b)
:
@[simp]
theorem
one_lt_mabs
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
:
theorem
abs_pos_of_pos
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
(h : 0 < a)
:
0 < |a|
theorem
one_lt_mabs_pos_of_one_lt
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
(h : 1 < a)
:
theorem
abs_pos_of_neg
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
(h : a < 0)
:
0 < |a|
theorem
one_lt_mabs_of_lt_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
(h : a < 1)
:
theorem
neg_abs_le
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
theorem
inv_mabs_le
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
theorem
add_abs_nonneg
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
theorem
one_le_mul_mabs
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
theorem
neg_abs_le_neg
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
theorem
inv_mabs_le_inv
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
theorem
abs_ne_zero
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
:
theorem
mabs_ne_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
:
@[simp]
theorem
abs_eq_zero
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
:
@[simp]
theorem
mabs_eq_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
:
@[simp]
theorem
abs_nonpos_iff
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
:
@[simp]
theorem
mabs_le_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
:
theorem
abs_le_abs_of_nonpos
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
{b : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(ha : a ≤ 0)
(hab : b ≤ a)
:
|a| ≤ |b|
theorem
mabs_le_mabs_of_le_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
{b : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(ha : a ≤ 1)
(hab : b ≤ a)
:
theorem
abs_lt
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
{b : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
:
theorem
mabs_lt
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
{b : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
:
theorem
neg_lt_of_abs_lt
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
{b : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(h : |a| < b)
:
theorem
inv_lt_of_mabs_lt
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
{a : α}
{b : α}
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(h : mabs a < b)
:
theorem
max_sub_min_eq_abs'
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
max_div_min_eq_mabs'
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
max_sub_min_eq_abs
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
max_div_min_eq_mabs
{α : Type u_1}
[Group α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
[CovariantClass α α (Function.swap fun (x x_1 : α) => x * x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
(b : α)
:
theorem
neg_le_of_abs_le
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
(h : |a| ≤ b)
:
theorem
le_of_abs_le
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
(h : |a| ≤ b)
:
a ≤ b
theorem
apply_abs_le_add_of_nonneg'
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{β : Type u_2}
[AddZeroClass β]
[Preorder β]
[CovariantClass β β (fun (x x_1 : β) => x + x_1) fun (x x_1 : β) => x ≤ x_1]
[CovariantClass β β (Function.swap fun (x x_1 : β) => x + x_1) fun (x x_1 : β) => x ≤ x_1]
{f : α → β}
{a : α}
(h₁ : 0 ≤ f a)
(h₂ : 0 ≤ f (-a))
:
theorem
apply_abs_le_mul_of_one_le'
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{β : Type u_2}
[MulOneClass β]
[Preorder β]
[CovariantClass β β (fun (x x_1 : β) => x * x_1) fun (x x_1 : β) => x ≤ x_1]
[CovariantClass β β (Function.swap fun (x x_1 : β) => x * x_1) fun (x x_1 : β) => x ≤ x_1]
{f : α → β}
{a : α}
(h₁ : 1 ≤ f a)
(h₂ : 1 ≤ f (-a))
:
theorem
apply_abs_le_add_of_nonneg
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{β : Type u_2}
[AddZeroClass β]
[Preorder β]
[CovariantClass β β (fun (x x_1 : β) => x + x_1) fun (x x_1 : β) => x ≤ x_1]
[CovariantClass β β (Function.swap fun (x x_1 : β) => x + x_1) fun (x x_1 : β) => x ≤ x_1]
{f : α → β}
(h : ∀ (x : α), 0 ≤ f x)
(a : α)
:
theorem
apply_abs_le_mul_of_one_le
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{β : Type u_2}
[MulOneClass β]
[Preorder β]
[CovariantClass β β (fun (x x_1 : β) => x * x_1) fun (x x_1 : β) => x ≤ x_1]
[CovariantClass β β (Function.swap fun (x x_1 : β) => x * x_1) fun (x x_1 : β) => x ≤ x_1]
{f : α → β}
(h : ∀ (x : α), 1 ≤ f x)
(a : α)
:
The triangle inequality in LinearOrderedAddCommGroup
s.
theorem
sub_le_of_abs_sub_le_left
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
{c : α}
(h : |a - b| ≤ c)
:
theorem
sub_le_of_abs_sub_le_right
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
{c : α}
(h : |a - b| ≤ c)
:
theorem
sub_lt_of_abs_sub_lt_left
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
{c : α}
(h : |a - b| < c)
:
theorem
sub_lt_of_abs_sub_lt_right
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
{c : α}
(h : |a - b| < c)
:
theorem
abs_le_max_abs_abs
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
{c : α}
(hab : a ≤ b)
(hbc : b ≤ c)
:
theorem
eq_of_abs_sub_eq_zero
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
(h : |a - b| = 0)
:
a = b
theorem
eq_of_abs_sub_nonpos
{α : Type u_1}
[LinearOrderedAddCommGroup α]
{a : α}
{b : α}
(h : |a - b| ≤ 0)
:
a = b
@[simp]
@[simp]
@[simp]
theorem
max_zero_add_max_neg_zero_eq_abs_self
{α : Type u_1}
[LinearOrderedAddCommGroup α]
(a : α)
:
@[deprecated neg_le_abs]
Alias of neg_le_abs
.
@[deprecated neg_abs_le]
theorem
neg_abs_le_self
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[CovariantClass α α (fun (x x_1 : α) => x + x_1) fun (x x_1 : α) => x ≤ x_1]
(a : α)
:
Alias of neg_abs_le
.