Documentation

Mathlib.Order.BourbakiWitt

Bourbaki-Witt Theorem #

This file proves the Bourbaki-Witt Theorem.

Main definitions #

Main statements #

References #

The proof used can be found in [Lan02]

structure NonemptyChain (α : Type u_2) [LE α] :
Type u_2

The type of nonempty chains of an order

Instances For
    theorem NonemptyChain.ext_iff {α : Type u_2} {inst✝ : LE α} {x y : NonemptyChain α} :
    theorem NonemptyChain.ext {α : Type u_2} {inst✝ : LE α} {x y : NonemptyChain α} (carrier : x.carrier = y.carrier) :
    x = y
    instance instSetLikeNonemptyChain {α : Type u_2} [LE α] :
    Equations
    class ChainCompletePartialOrder (α : Type u_2) extends PartialOrder α :
    Type u_2

    A chain complete partial order (CCPO) is a nonempty partial order such that every nonempty chain has a supremum (which we call cSup)

    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.
      structure ChainCompletePartialOrder.IsAdmissible {α : Type u_1} [ChainCompletePartialOrder α] (x : α) (f : αα) (s : Set α) :

      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

      • image_self_subset_self : f '' s s

        The image of an admissible set under f is a subset of itself

      • cSup_mem (c : NonemptyChain α) : c scSup c s

        If a chain is a subset of an admissible set, its cSup is a member of the admissible set

      Instances For
        theorem ChainCompletePartialOrder.ici_isAdmissible {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} (le_map : ∀ (x : α), x f x) :
        @[reducible, inline]
        abbrev ChainCompletePartialOrder.bot {α : Type u_1} [ChainCompletePartialOrder α] (x : α) (f : αα) :
        Set α

        The bottom admissible set with base point x and inflationary function f

        Equations
        Instances For
          theorem ChainCompletePartialOrder.bot_isAdmissible {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} (le_map : ∀ (x : α), x f x) :
          IsAdmissible x f (bot x f)
          theorem ChainCompletePartialOrder.subset_bot_iff {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} {s : Set α} (h : IsAdmissible x f s) :
          s bot x f s = bot x f
          theorem ChainCompletePartialOrder.map_mem_bot {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} {y : α} (le_map : ∀ (x : α), x f x) (h : y bot x f) :
          f y bot x f
          structure ChainCompletePartialOrder.IsExtremePt {α : Type u_1} [ChainCompletePartialOrder α] (x : α) (f : αα) (y : α) :

          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]

          • mem_bot : y bot x f
          • map_le_of_mem_of_lt {z : α} (h : z bot x f) (h' : z < y) : f z y
          Instances For
            theorem ChainCompletePartialOrder.IsExtremePt.bot_eq_of_le_or_map_le {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} {y : α} (le_map : ∀ (x : α), x f x) (hy : IsExtremePt x f y) :
            {z : α | z bot x f (z y f y z)} = bot x f

            If y is an extreme point and f is inflationary, then there are no element between y and f y.

            theorem ChainCompletePartialOrder.IsExtremePt.setOf_isExtremePt_isAdmissible {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} (le_map : ∀ (x : α), x f x) :
            IsAdmissible x f {y : α | IsExtremePt x f y}
            theorem ChainCompletePartialOrder.IsExtremePt.setOf_isExtremePt_eq_bot {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} (le_map : ∀ (x : α), x f x) :
            {y : α | IsExtremePt x f y} = bot x f
            theorem ChainCompletePartialOrder.IsExtremePt.mem_bot_iff_isExtremePt {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} {y : α} (le_map : ∀ (x : α), x f x) :
            y bot x f IsExtremePt x f y
            theorem ChainCompletePartialOrder.IsExtremePt.bot_isChain {α : Type u_1} [ChainCompletePartialOrder α] {x : α} {f : αα} (le_map : ∀ (x : α), x f x) :
            IsChain (fun (x1 x2 : α) => x1 x2) (bot x f)