⊤ and ⊥, bounded lattices and variants #
This file defines top and bottom elements (greatest and least elements) of a type, the bounded
variants of different kinds of lattices, sets up the typeclass hierarchy between them and provides
instances for Prop and fun.
Main declarations #
<Top/Bot> α: Typeclasses to declare the⊤/⊥notation.Order<Top/Bot> α: Order with a top/bottom element.BoundedOrder α: Order with a top and bottom element.
Top, bottom element #
An order is (noncomputably) either an OrderTop or a NoTopOrder. Use as
cases topOrderOrNoTopOrder α.
Equations
Instances For
An order is (noncomputably) either an OrderBot or a NoBotOrder. Use as
cases botOrderOrNoBotOrder α.
Equations
Instances For
Alias of ne_top_of_lt.
Alias of lt_top_of_lt.
@[simp]
@[simp]
@[simp]
@[simp]
Alias of the forward direction of isMax_iff_eq_top.
Alias of the forward direction of isTop_iff_eq_top.
theorem
ne_top_of_le_ne_top
{α : Type u}
[PartialOrder α]
[OrderTop α]
{a b : α}
(hb : b ≠ ⊤)
(hab : a ≤ b)
:
theorem
ne_bot_of_le_ne_bot
{α : Type u}
[PartialOrder α]
[OrderBot α]
{a b : α}
(hb : b ≠ ⊥)
(hab : b ≤ a)
:
theorem
OrderTop.ext_top
{α : Type u_1}
{hA : PartialOrder α}
(A : OrderTop α)
{hB : PartialOrder α}
(B : OrderTop α)
(H : ∀ (x y : α), x ≤ y ↔ x ≤ y)
:
theorem
OrderBot.ext_bot
{α : Type u_1}
{hA : PartialOrder α}
(A : OrderBot α)
{hB : PartialOrder α}
(B : OrderBot α)
(H : ∀ (x y : α), x ≤ y ↔ x ≤ y)
:
Equations
- OrderDual.instOrderTop α = { toTop := inferInstanceAs (Top αᵒᵈ), le_top := ⋯ }
Equations
- OrderDual.instOrderBot α = { toBot := inferInstanceAs (Bot αᵒᵈ), bot_le := ⋯ }
@[deprecated not_bot_lt_iff (since := "2025-12-03")]
theorem
eq_bot_of_minimal
{α : Type u}
[PartialOrder α]
[OrderBot α]
{a : α}
(h : ∀ (b : α), ¬b < a)
:
Bounded order #
Equations
- OrderDual.instBoundedOrder α = { toOrderTop := inferInstanceAs (OrderTop αᵒᵈ), toOrderBot := inferInstanceAs (OrderBot αᵒᵈ) }
Function lattices #
instance
Pi.instOrderBot
{ι : Type u_1}
{α' : ι → Type u_2}
[(i : ι) → LE (α' i)]
[(i : ι) → OrderBot (α' i)]
:
OrderBot ((i : ι) → α' i)
Equations
- Pi.instOrderBot = { toBot := Pi.instBotForall, bot_le := ⋯ }
instance
Pi.instOrderTop
{ι : Type u_1}
{α' : ι → Type u_2}
[(i : ι) → LE (α' i)]
[(i : ι) → OrderTop (α' i)]
:
OrderTop ((i : ι) → α' i)
Equations
- Pi.instOrderTop = { toTop := Pi.instTopForall, le_top := ⋯ }
instance
Pi.instBoundedOrder
{ι : Type u_1}
{α' : ι → Type u_2}
[(i : ι) → LE (α' i)]
[(i : ι) → BoundedOrder (α' i)]
:
BoundedOrder ((i : ι) → α' i)
Equations
- Pi.instBoundedOrder = { toOrderTop := inferInstanceAs (OrderTop ((i : ι) → α' i)), toOrderBot := inferInstanceAs (OrderBot ((i : ι) → α' i)) }
@[reducible, inline]
abbrev
BoundedOrder.lift
{α : Type u}
{β : Type v}
[LE α]
[Top α]
[Bot α]
[LE β]
[BoundedOrder β]
(f : α → β)
(map_le : ∀ (a b : α), f a ≤ f b → a ≤ b)
(map_top : f ⊤ = ⊤)
(map_bot : f ⊥ = ⊥)
:
Pullback a BoundedOrder.
Equations
- BoundedOrder.lift f map_le map_top map_bot = { toOrderTop := OrderTop.lift f map_le map_top, toOrderBot := OrderBot.lift f map_le map_bot }
Instances For
Subtype, order dual, product lattices #
@[reducible, inline]
abbrev
Subtype.boundedOrder
{α : Type u}
{p : α → Prop}
[LE α]
[BoundedOrder α]
(hbot : p ⊥)
(htop : p ⊤)
:
BoundedOrder (Subtype p)
A subtype remains a bounded order if the property holds at ⊥ and ⊤.
Equations
- Subtype.boundedOrder hbot htop = { toOrderTop := Subtype.orderTop htop, toOrderBot := Subtype.orderBot hbot }
Instances For
instance
Prod.instBoundedOrder
(α : Type u)
(β : Type v)
[LE α]
[LE β]
[BoundedOrder α]
[BoundedOrder β]
:
BoundedOrder (α × β)
Equations
- Prod.instBoundedOrder α β = { toOrderTop := inferInstanceAs (OrderTop (α × β)), toOrderBot := inferInstanceAs (OrderBot (α × β)) }
Equations
Equations
Equations
- ULift.instBoundedOrder = { toOrderTop := ULift.instOrderTop, toOrderBot := ULift.instOrderBot }
@[simp]
@[simp]
@[simp]
Equations
- Bool.instBoundedOrder = { top := true, le_top := Bool.le_true, bot := false, bot_le := Bool.false_le }