Documentation

Mathlib.Order.CompleteLattice.Defs

Definition of complete lattices #

This file contains the definition of complete lattices with suprema/infima of arbitrary sets.

Main definitions #

Naming conventions #

In lemma names,

Notation #

instance OrderDual.supSet (α : Type u_8) [InfSet α] :
Equations
instance OrderDual.infSet (α : Type u_8) [SupSet α] :
Equations
class CompleteSemilatticeSup (α : Type u_8) extends PartialOrder α, SupSet α :
Type u_8

Note that we rarely use CompleteSemilatticeSup (in fact, any such object is always a CompleteLattice, so it's usually best to start there).

Nevertheless it is sometimes a useful intermediate step in constructions.

Instances
theorem le_sSup {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a : α} :
a sa sSup s
theorem sSup_le {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a : α} :
(∀ bs, b a)sSup s a
theorem isLUB_sSup {α : Type u_1} [CompleteSemilatticeSup α] (s : Set α) :
IsLUB s (sSup s)
theorem isLUB_iff_sSup_eq {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a : α} :
IsLUB s a sSup s = a
theorem IsLUB.sSup_eq {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a : α} :
IsLUB s asSup s = a

Alias of the forward direction of isLUB_iff_sSup_eq.

theorem le_sSup_of_le {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a b : α} (hb : b s) (h : a b) :
a sSup s
theorem sSup_le_sSup {α : Type u_1} [CompleteSemilatticeSup α] {s t : Set α} (h : s t) :
@[simp]
theorem sSup_le_iff {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a : α} :
sSup s a bs, b a
theorem le_sSup_iff {α : Type u_1} [CompleteSemilatticeSup α] {s : Set α} {a : α} :
a sSup s bupperBounds s, a b
theorem le_iSup_iff {α : Type u_1} {ι : Sort u_4} [CompleteSemilatticeSup α] {a : α} {s : ια} :
a iSup s ∀ (b : α), (∀ (i : ι), s i b)a b
class CompleteSemilatticeInf (α : Type u_8) extends PartialOrder α, InfSet α :
Type u_8

Note that we rarely use CompleteSemilatticeInf (in fact, any such object is always a CompleteLattice, so it's usually best to start there).

Nevertheless it is sometimes a useful intermediate step in constructions.

Instances
theorem sInf_le {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a : α} :
a ssInf s a
theorem le_sInf {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a : α} :
(∀ bs, a b)a sInf s
theorem isGLB_sInf {α : Type u_1} [CompleteSemilatticeInf α] (s : Set α) :
IsGLB s (sInf s)
theorem isGLB_iff_sInf_eq {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a : α} :
IsGLB s a sInf s = a
theorem IsGLB.sInf_eq {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a : α} :
IsGLB s asInf s = a

Alias of the forward direction of isGLB_iff_sInf_eq.

theorem sInf_le_of_le {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a b : α} (hb : b s) (h : b a) :
sInf s a
theorem sInf_le_sInf {α : Type u_1} [CompleteSemilatticeInf α] {s t : Set α} (h : s t) :
@[simp]
theorem le_sInf_iff {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a : α} :
a sInf s bs, a b
theorem sInf_le_iff {α : Type u_1} [CompleteSemilatticeInf α] {s : Set α} {a : α} :
sInf s a blowerBounds s, b a
theorem iInf_le_iff {α : Type u_1} {ι : Sort u_4} [CompleteSemilatticeInf α] {a : α} {s : ια} :
iInf s a ∀ (b : α), (∀ (i : ι), b s i)b a
class CompleteLattice (α : Type u_8) extends Lattice α, CompleteSemilatticeSup α, CompleteSemilatticeInf α, Top α, Bot α :
Type u_8

A complete lattice is a bounded lattice which has suprema and infima for every subset.

Instances
@[instance 100]
Equations
def completeLatticeOfInf (α : Type u_8) [H1 : PartialOrder α] [H2 : InfSet α] (isGLB_sInf : ∀ (s : Set α), IsGLB s (sInf s)) :

Create a CompleteLattice from a PartialOrder and InfSet that returns the greatest lower bound of a set. Usually this constructor provides poor definitional equalities. If other fields are known explicitly, they should be provided; for example, if inf is known explicitly, construct the CompleteLattice instance as

instance : CompleteLattice my_T where
  inf := better_inf
  le_inf := ...
  inf_le_right := ...
  inf_le_left := ...
  -- don't care to fix sup, sSup, bot, top
  __ := completeLatticeOfInf my_T _
Equations
  • One or more equations did not get rendered due to their size.

Any CompleteSemilatticeInf is in fact a CompleteLattice.

Note that this construction has bad definitional properties: see the doc-string on completeLatticeOfInf.

Equations
def completeLatticeOfSup (α : Type u_8) [H1 : PartialOrder α] [H2 : SupSet α] (isLUB_sSup : ∀ (s : Set α), IsLUB s (sSup s)) :

Create a CompleteLattice from a PartialOrder and SupSet that returns the least upper bound of a set. Usually this constructor provides poor definitional equalities. If other fields are known explicitly, they should be provided; for example, if inf is known explicitly, construct the CompleteLattice instance as

instance : CompleteLattice my_T where
  inf := better_inf
  le_inf := ...
  inf_le_right := ...
  inf_le_left := ...
  -- don't care to fix sup, sInf, bot, top
  __ := completeLatticeOfSup my_T _
Equations
  • One or more equations did not get rendered due to their size.

Any CompleteSemilatticeSup is in fact a CompleteLattice.

Note that this construction has bad definitional properties: see the doc-string on completeLatticeOfSup.

Equations
class CompleteLinearOrder (α : Type u_8) extends CompleteLattice α, BiheytingAlgebra α, Ord α :
Type u_8

A complete linear order is a linear order whose lattice structure is complete.

Instances
Equations
  • One or more equations did not get rendered due to their size.
Equations
  • One or more equations did not get rendered due to their size.
Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem toDual_sSup {α : Type u_1} [SupSet α] (s : Set α) :
@[simp]
theorem toDual_sInf {α : Type u_1} [InfSet α] (s : Set α) :
@[simp]
theorem ofDual_sSup {α : Type u_1} [InfSet α] (s : Set αᵒᵈ) :
@[simp]
theorem ofDual_sInf {α : Type u_1} [SupSet α] (s : Set αᵒᵈ) :
@[simp]
theorem toDual_iSup {α : Type u_1} {ι : Sort u_4} [SupSet α] (f : ια) :
OrderDual.toDual (⨆ (i : ι), f i) = ⨅ (i : ι), OrderDual.toDual (f i)
@[simp]
theorem toDual_iInf {α : Type u_1} {ι : Sort u_4} [InfSet α] (f : ια) :
OrderDual.toDual (⨅ (i : ι), f i) = ⨆ (i : ι), OrderDual.toDual (f i)
@[simp]
theorem ofDual_iSup {α : Type u_1} {ι : Sort u_4} [InfSet α] (f : ιαᵒᵈ) :
OrderDual.ofDual (⨆ (i : ι), f i) = ⨅ (i : ι), OrderDual.ofDual (f i)
@[simp]
theorem ofDual_iInf {α : Type u_1} {ι : Sort u_4} [SupSet α] (f : ιαᵒᵈ) :
OrderDual.ofDual (⨅ (i : ι), f i) = ⨆ (i : ι), OrderDual.ofDual (f i)
theorem lt_sSup_iff {α : Type u_1} [CompleteLinearOrder α] {s : Set α} {b : α} :
b < sSup s as, b < a
theorem sInf_lt_iff {α : Type u_1} [CompleteLinearOrder α] {s : Set α} {b : α} :
sInf s < b as, a < b
theorem sSup_eq_top {α : Type u_1} [CompleteLinearOrder α] {s : Set α} :
sSup s = b < , as, b < a
theorem sInf_eq_bot {α : Type u_1} [CompleteLinearOrder α] {s : Set α} :
sInf s = b > , as, a < b
theorem lt_iSup_iff {α : Type u_1} {ι : Sort u_4} [CompleteLinearOrder α] {a : α} {f : ια} :
a < iSup f ∃ (i : ι), a < f i
theorem iInf_lt_iff {α : Type u_1} {ι : Sort u_4} [CompleteLinearOrder α] {a : α} {f : ια} :
iInf f < a ∃ (i : ι), f i < a
theorem lt_biSup_iff {α : Type u_1} {β : Type u_2} [CompleteLinearOrder α] {a : α} {s : Set β} {f : βα} :
a < is, f i is, a < f i
theorem biInf_lt_iff {α : Type u_1} {β : Type u_2} [CompleteLinearOrder α] {a : α} {s : Set β} {f : βα} :
is, f i < a is, f i < a