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 : α}
[AddLeftMono α]
(h : 0 ≤ a)
:
|a| = a
theorem
abs_of_pos
{α : Type u_1}
[Lattice α]
[AddGroup α]
{a : α}
[AddLeftMono α]
(h : 0 < a)
:
|a| = a
theorem
abs_le_abs_of_nonneg
{α : Type u_1}
[Lattice α]
[AddGroup α]
{a : α}
{b : α}
[AddLeftMono α]
(ha : 0 ≤ a)
(hab : a ≤ b)
:
|a| ≤ |b|
@[simp]
theorem
abs_nonneg
{α : Type u_1}
[Lattice α]
[AddGroup α]
[AddLeftMono α]
[AddRightMono α]
(a : α)
:
0 ≤ |a|
@[simp]
@[simp]
theorem
abs_abs
{α : Type u_1}
[Lattice α]
[AddGroup α]
[AddLeftMono α]
[AddRightMono α]
(a : α)
:
|(|a|)| = |a|
The absolute value satisfies the triangle inequality.
theorem
abs_abs_sub_abs_le
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
:
theorem
sup_sub_inf_eq_abs_sub
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
:
theorem
two_nsmul_sup_eq_add_add_abs_sub
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
:
theorem
two_nsmul_inf_eq_add_sub_abs_sub
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
:
theorem
abs_sub_sup_add_abs_sub_inf
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
(c : α)
:
theorem
abs_sup_sub_sup_le_abs
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
(c : α)
:
theorem
abs_inf_sub_inf_le_abs
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(a : α)
(b : α)
(c : α)
:
theorem
Birkhoff_inequalities
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
[AddLeftMono α]
(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]
@[simp]
theorem
abs_pos_of_pos
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
(h : 0 < a)
:
0 < |a|
theorem
one_lt_mabs_pos_of_one_lt
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
(h : 1 < a)
:
theorem
abs_pos_of_neg
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
(h : a < 0)
:
0 < |a|
theorem
one_lt_mabs_of_lt_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
(h : a < 1)
:
theorem
abs_ne_zero
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
[AddRightMono α]
:
theorem
mabs_ne_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
[MulRightMono α]
:
@[simp]
theorem
abs_eq_zero
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
[AddRightMono α]
:
@[simp]
theorem
mabs_eq_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
[MulRightMono α]
:
@[simp]
theorem
abs_nonpos_iff
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
[AddRightMono α]
:
@[simp]
theorem
mabs_le_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
[MulRightMono α]
:
theorem
abs_le_abs_of_nonpos
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
{b : α}
[AddRightMono α]
(ha : a ≤ 0)
(hab : b ≤ a)
:
|a| ≤ |b|
theorem
mabs_le_mabs_of_le_one
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
{b : α}
[MulRightMono α]
(ha : a ≤ 1)
(hab : b ≤ a)
:
theorem
abs_lt
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
{b : α}
[AddRightMono α]
:
theorem
mabs_lt
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
{b : α}
[MulRightMono α]
:
theorem
neg_lt_of_abs_lt
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
{a : α}
{b : α}
[AddRightMono α]
(h : |a| < b)
:
theorem
inv_lt_of_mabs_lt
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
{a : α}
{b : α}
[MulRightMono α]
(h : mabs a < b)
:
theorem
max_sub_min_eq_abs'
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
[AddRightMono α]
(a : α)
(b : α)
:
theorem
max_div_min_eq_mabs'
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
[MulRightMono α]
(a : α)
(b : α)
:
theorem
max_sub_min_eq_abs
{α : Type u_1}
[AddGroup α]
[LinearOrder α]
[AddLeftMono α]
[AddRightMono α]
(a : α)
(b : α)
:
theorem
max_div_min_eq_mabs
{α : Type u_1}
[Group α]
[LinearOrder α]
[MulLeftMono α]
[MulRightMono α]
(a : α)
(b : α)
:
A set s
in a lattice ordered group is solid if for all x ∈ s
and all y ∈ α
such that
|y| ≤ |x|
, then y ∈ s
.
Equations
- LatticeOrderedAddCommGroup.IsSolid s = ∀ ⦃x : α⦄, x ∈ s → ∀ ⦃y : α⦄, |y| ≤ |x| → y ∈ s
Instances For
def
LatticeOrderedAddCommGroup.solidClosure
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
(s : Set α)
:
Set α
The solid closure of a subset s
is the smallest superset of s
that is solid.
Equations
- LatticeOrderedAddCommGroup.solidClosure s = {y : α | ∃ (x : α), x ∈ s ∧ |y| ≤ |x|}
Instances For
theorem
LatticeOrderedAddCommGroup.isSolid_solidClosure
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
(s : Set α)
:
theorem
LatticeOrderedAddCommGroup.solidClosure_min
{α : Type u_1}
[Lattice α]
[AddCommGroup α]
{s : Set α}
{t : Set α}
(hst : s ⊆ t)
(ht : LatticeOrderedAddCommGroup.IsSolid t)
:
@[deprecated neg_le_abs]
Alias of neg_le_abs
.
@[deprecated neg_abs_le]
Alias of neg_abs_le
.