Documentation

Mathlib.Order.Nucleus

Nucleus #

Locales are the dual concept to frames. Locale theory is a branch of point-free topology, where intuitively locales are like topological spaces which may or may not have enough points. Sublocales of a locale generalize the concept of subspaces in topology to the point-free setting.

A nucleus is an endomorphism of a frame which corresponds to a sublocale.

References #

https://ncatlab.org/nlab/show/sublocale https://ncatlab.org/nlab/show/nucleus

structure Nucleus (X : Type u_2) [SemilatticeInf X] extends InfHom X X :
Type u_2

A nucleus is an inflationary idempotent inf-preserving endomorphism of a semilattice. In a frame, nuclei correspond to sublocales.

Instances For
    class NucleusClass (F : Type u_2) (X : Type u_3) [SemilatticeInf X] [FunLike F X X] extends InfHomClass F X X :

    NucleusClass F X states that F is a type of nuclei.

    • map_inf (f : F) (a b : X) : f (a b) = f a f b
    • idempotent (x : X) (f : F) : f (f x) f x

      A nucleus is idempotent.

    • le_apply (x : X) (f : F) : x f x

      A nucleus is inflationary.

    Instances
      instance Nucleus.instFunLike {X : Type u_1} [CompleteLattice X] :
      Equations
      theorem Nucleus.toFun_eq_coe {X : Type u_1} [CompleteLattice X] (n : Nucleus X) :
      n.toFun = n
      @[simp]
      theorem Nucleus.coe_toInfHom {X : Type u_1} [CompleteLattice X] (n : Nucleus X) :
      n.toInfHom = n
      @[simp]
      theorem Nucleus.coe_mk {X : Type u_1} [CompleteLattice X] (f : InfHom X X) (h1 : ∀ (x : X), f.toFun (f.toFun x) f.toFun x) (h2 : ∀ (x : X), x f.toFun x) :
      { toInfHom := f, idempotent' := h1, le_apply' := h2 } = f

      Every Nucleus is a ClosureOperator.

      Equations
      Instances For
        theorem Nucleus.idempotent {X : Type u_1} [CompleteLattice X] {n : Nucleus X} {x : X} :
        n (n x) = n x
        theorem Nucleus.le_apply {X : Type u_1} [CompleteLattice X] {n : Nucleus X} {x : X} :
        x n x
        theorem Nucleus.monotone {X : Type u_1} [CompleteLattice X] {n : Nucleus X} :
        theorem Nucleus.map_inf {X : Type u_1} [CompleteLattice X] {n : Nucleus X} {x y : X} :
        n (x y) = n x n y
        theorem Nucleus.ext {X : Type u_1} [CompleteLattice X] {m n : Nucleus X} (h : ∀ (a : X), m a = n a) :
        m = n

        A Nucleus preserves ⊤.

        @[simp]
        theorem Nucleus.coe_le_coe {X : Type u_1} [CompleteLattice X] {n m : Nucleus X} :
        m n m n
        @[simp]
        theorem Nucleus.coe_lt_coe {X : Type u_1} [CompleteLattice X] {n m : Nucleus X} :
        m < n m < n
        instance Nucleus.instBot {X : Type u_1} [CompleteLattice X] :

        The smallest Nucleus is the identity.

        Equations
        • Nucleus.instBot = { bot := { toFun := fun (x : X) => x, map_inf' := , idempotent' := , le_apply' := } }
        instance Nucleus.instTop {X : Type u_1} [CompleteLattice X] :

        The biggest Nucleus sends everything to .

        Equations
        • Nucleus.instTop = { top := { toFun := , map_inf' := , idempotent' := , le_apply' := } }
        @[simp]
        theorem Nucleus.coe_bot {X : Type u_1} [CompleteLattice X] :
        = id
        @[simp]
        theorem Nucleus.coe_top {X : Type u_1} [CompleteLattice X] :
        =
        @[simp]
        theorem Nucleus.bot_apply {X : Type u_1} [CompleteLattice X] (x : X) :
        x = x
        @[simp]
        theorem Nucleus.top_apply {X : Type u_1} [CompleteLattice X] (x : X) :
        Equations
        • Nucleus.instInfSet = { sInf := fun (s : Set (Nucleus X)) => { toFun := fun (x : X) => fs, f x, map_inf' := , idempotent' := , le_apply' := } }
        @[simp]
        theorem Nucleus.sInf_apply {X : Type u_1} [CompleteLattice X] (s : Set (Nucleus X)) (x : X) :
        (sInf s) x = js, j x
        @[simp]
        theorem Nucleus.iInf_apply {X : Type u_1} [CompleteLattice X] {ι : Type u_2} (f : ιNucleus X) (x : X) :
        (iInf f) x = ⨅ (j : ι), (f j) x
        @[simp]
        theorem Nucleus.inf_apply {X : Type u_1} [CompleteLattice X] (m n : Nucleus X) (x : X) :
        (m n) x = m x n x
        theorem Nucleus.map_himp_le {X : Type u_1} [Order.Frame X] {n : Nucleus X} {x y : X} :
        n (x y) x n y
        theorem Nucleus.map_himp_apply {X : Type u_1} [Order.Frame X] (n : Nucleus X) (x y : X) :
        n (x n y) = x n y
        instance Nucleus.instHImp {X : Type u_1} [Order.Frame X] :
        Equations
        • Nucleus.instHImp = { himp := fun (m n : Nucleus X) => { toFun := fun (x : X) => ⨅ (y : X), ⨅ (_ : y x), m y n y, map_inf' := , idempotent' := , le_apply' := } }
        @[simp]
        theorem Nucleus.himp_apply {X : Type u_1} [Order.Frame X] (m n : Nucleus X) (x : X) :
        (m n) x = ⨅ (y : X), ⨅ (_ : y x), m y n y