Documentation

Mathlib.Algebra.Order.Group.Lattice

Lattice ordered groups #

Lattice ordered groups were introduced by [Birkhoff][birkhoff1942]. They form the algebraic underpinnings of vector lattices, Banach lattices, AL-space, AM-space etc.

A lattice ordered group is a type α satisfying:

This file establishes basic properties of lattice ordered groups. It is shown that when the group is commutative, the lattice is distributive. This also holds in the non-commutative case ([Birkhoff][birkhoff1942],[Fuchs][fuchs1963]) but we do not yet have the machinery to establish this in mathlib.

References #

Tags #

lattice, order, group

theorem mul_sup {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] (a b c : α) :
c * (ab) = c * ac * b
theorem add_sup {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] (a b c : α) :
c + ab = (c + a) ⊔ (c + b)
theorem sup_mul {α : Type u_1} [Lattice α] [Group α] [MulRightMono α] (a b c : α) :
(ab) * c = a * cb * c
theorem sup_add {α : Type u_1} [Lattice α] [AddGroup α] [AddRightMono α] (a b c : α) :
ab + c = (a + c) ⊔ (b + c)
theorem mul_inf {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] (a b c : α) :
c * (ab) = c * ac * b
theorem add_inf {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] (a b c : α) :
c + ab = (c + a) ⊓ (c + b)
theorem inf_mul {α : Type u_1} [Lattice α] [Group α] [MulRightMono α] (a b c : α) :
(ab) * c = a * cb * c
theorem inf_add {α : Type u_1} [Lattice α] [AddGroup α] [AddRightMono α] (a b c : α) :
ab + c = (a + c) ⊓ (b + c)
theorem sup_div {α : Type u_1} [Lattice α] [Group α] [MulRightMono α] (a b c : α) :
(ab) / c = a / cb / c
theorem sup_sub {α : Type u_1} [Lattice α] [AddGroup α] [AddRightMono α] (a b c : α) :
ab - c = (a - c) ⊔ (b - c)
theorem inf_div {α : Type u_1} [Lattice α] [Group α] [MulRightMono α] (a b c : α) :
(ab) / c = a / cb / c
theorem inf_sub {α : Type u_1} [Lattice α] [AddGroup α] [AddRightMono α] (a b c : α) :
ab - c = (a - c) ⊓ (b - c)
theorem inv_sup {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] [MulRightMono α] (a b : α) :
(ab)⁻¹ = a⁻¹b⁻¹
theorem neg_sup {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] [AddRightMono α] (a b : α) :
-(ab) = -a-b
theorem inv_inf {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] [MulRightMono α] (a b : α) :
(ab)⁻¹ = a⁻¹b⁻¹
theorem neg_inf {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] [AddRightMono α] (a b : α) :
-(ab) = -a-b
theorem div_sup {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] [MulRightMono α] (a b c : α) :
c / (ab) = c / ac / b
theorem sub_sup {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] [AddRightMono α] (a b c : α) :
c - ab = (c - a) ⊓ (c - b)
theorem div_inf {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] [MulRightMono α] (a b c : α) :
c / (ab) = c / ac / b
theorem sub_inf {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] [AddRightMono α] (a b c : α) :
c - ab = (c - a) ⊔ (c - b)
theorem pow_two_semiclosed {α : Type u_1} [Lattice α] [Group α] [MulLeftMono α] [MulRightMono α] {a : α} (ha : 1 a ^ 2) :
1 a
theorem nsmul_two_semiclosed {α : Type u_1} [Lattice α] [AddGroup α] [AddLeftMono α] [AddRightMono α] {a : α} (ha : 0 2 a) :
0 a
theorem inf_mul_sup {α : Type u_1} [Lattice α] [CommGroup α] [MulLeftMono α] (a b : α) :
(ab) * (ab) = a * b
theorem inf_add_sup {α : Type u_1} [Lattice α] [AddCommGroup α] [AddLeftMono α] (a b : α) :
ab + ab = a + b

Every lattice ordered commutative group is a distributive lattice.

Equations

Every lattice ordered commutative additive group is a distributive lattice

Equations