Documentation

Mathlib.Topology.Sheaves.Alexandrov

Let X be a preorder , and consider the associated Alexandrov topology on X. Given a functor F : X ⥤ C to a complete category, we can extend F to a presheaf on (the topological space) X by taking the right Kan extension along the canonical functor X ⥤ (Opens X)ᵒᵖ sending x : X to the principal open {y | x ≤ y} in the Alexandrov topology.

This file proves that this presheaf is a sheaf.

Given x : X, this is the principal open subset of X generated by x.

Equations
Instances For

    The functor sending x : X to the principal open associated with x.

    Equations
    Instances For
      @[simp]
      theorem Alexandrov.principals_map (X : Type v) [TopologicalSpace X] [Preorder X] [Topology.IsUpperSet X] {x y : X} (f : x y) :
      (principals X).map f = .hom.op
      theorem Alexandrov.exists_le_of_le_sup {X : Type v} [TopologicalSpace X] [Preorder X] [Topology.IsUpperSet X] {ι : Type v} {x : X} (Us : ιTopologicalSpace.Opens X) (h : principalOpen x iSup Us) :
      ∃ (i : ι), principalOpen x Us i
      @[reducible, inline]

      Given a structured arrow f with domain U : Opens X over principals X, this functor sends f to its "generator", an element of X.

      Equations
      Instances For

        Given a structured arrow f with domain iSup Us over principals X, where Us is a family of Opens X, this functor sends f to the principal open associated with it, considered as an object in the full subcategory of all V : Opens X such that V ≤ Us i for some i.

        This definition is primarily meant to be used in lowerCone, and isLimit below.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          @[simp]
          theorem Alexandrov.projSup_map {X : Type v} [TopologicalSpace X] [Preorder X] [Topology.IsUpperSet X] {ι : Type v} (Us : ιTopologicalSpace.Opens X) {X✝ Y✝ : CategoryTheory.StructuredArrow (Opposite.op (iSup Us)) (principals X)} (e : X✝ Y✝) :
          (projSup Us).map e = .hom.op

          This is an auxiliary definition which is only meant to be used in isLimit below.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For

            This is the main construction in this file showing that the right Kan extension of F : X ⥤ C along principals : X ⥤ (Opens X)ᵒᵖ is a sheaf, by showing that a certain cone is a limit cone.

            See isSheaf_principalsKanExtension for the main application.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For

              The main theorem of this file. If X is a topological space and preorder whose topology is the UpperSet topology associated with the preorder, F : X ⥤ C is a functor into a complete category from the preorder category, and P : (Opens X)ᵒᵖ ⥤ C denotes the right Kan extension of F along the functor X ⥤ (Open X)ᵒᵖ which sends x : X to {y | x ≤ y}, then P is a sheaf.