Bourbaki-Witt Theorem #
This file proves the Bourbaki-Witt Theorem.
Main definitions #
- class
ChainCompletePartialOrder: A nonempty partial order is a chain complete partial order such that every nonempty chain has a supremum
Main statements #
nonempty_fixedPoints_of_inflationary: The Bourbaki-Witt Theorem : If $X$ is a chain complete partial order and $f : X → X$ is inflationary (i.e. ∀ x, x ≤ f x), then $f$ has a fixed point
References #
The proof used can be found in [Lan02]
Equations
- instSetLikeNonemptyChain = { coe := NonemptyChain.carrier, coe_injective' := ⋯ }
A chain complete partial order (CCPO) is a nonempty partial order such that every
nonempty chain has a supremum (which we call cSup)
- cSup : NonemptyChain α → α
The supremum of a nonempty chain
cSupis an upper bound of the nonempty chaincSupis a lower bound of the set of upper bounds of the nonempty chain
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.
An admissible set for given x : α and f : α → α has x, the base point, as a least element
and is closed under applying f and cSup.
- base_isLeast : IsLeast s x
The base point is the least element of an admissible set
The image of an admissible set under
fis a subset of itself- cSup_mem (c : NonemptyChain α) : ↑c ⊆ s → cSup c ∈ s
If a chain is a subset of an admissible set, its
cSupis a member of the admissible set
Instances For
The bottom admissible set with base point x and inflationary function f
Equations
- ChainCompletePartialOrder.bot x f = ⋂₀ {s : Set α | ChainCompletePartialOrder.IsAdmissible x f s}
Instances For
y is an extreme point for x : α and f : α → α if it is in the bottom admissible set and
y is larger than f z for any z < y in the bottom admissible set.
This definition comes from [Lan02]
Instances For
If y is an extreme point and f is inflationary, then there are no element between y and
f y.