Documentation

Mathlib.Topology.Sheaves.LocalPredicate

Functions satisfying a local predicate form a sheaf. #

At this stage, in Mathlib/Topology/Sheaves/SheafOfFunctions.lean we've proved that not-necessarily-continuous functions from a topological space into some type (or type family) form a sheaf.

Why do the continuous functions form a sheaf? The point is just that continuity is a local condition, so one can use the lifting condition for functions to provide a candidate lift, then verify that the lift is actually continuous by using the factorisation condition for the lift (which guarantees that on each open set it agrees with the functions being lifted, which were assumed to be continuous).

This file abstracts this argument to work for any collection of dependent functions on a topological space satisfying a "local predicate".

As an application, we check that continuity is a local predicate in this sense, and provide

A sheaf constructed in this way has a natural map stalkToFiber from the stalks to the types in the ambient type family.

We give conditions sufficient to show that this map is injective and/or surjective.

structure TopCat.PrelocalPredicate {X : TopCat} (T : XType v) :

Given a topological space X : TopCat and a type family T : X → Type, a P : PrelocalPredicate T consists of:

  • a family of predicates P.pred, one for each U : Opens X, of the form (Π x : U, T x) → Prop
  • a proof that if f : Π x : V, T x satisfies the predicate on V : Opens X, then the restriction of f to any open subset U also satisfies the predicate.
  • pred : {U : TopologicalSpace.Opens X} → ((x : U) → T x)Prop

    The underlying predicate of a prelocal predicate

  • res : ∀ {U V : TopologicalSpace.Opens X} (i : U V) (f : (x : V) → T x), self.pred fself.pred fun (x : U) => f ((fun (x : U) => { val := x, property := }) x)

    The underlying predicate should be invariant under restriction

Instances For
    @[simp]
    theorem TopCat.continuousPrelocal_pred (X : TopCat) (T : TopCat) :
    ∀ {x : TopologicalSpace.Opens X} (f : xT), (TopCat.continuousPrelocal X T).pred f = Continuous f
    def TopCat.continuousPrelocal (X : TopCat) (T : TopCat) :
    TopCat.PrelocalPredicate fun (x : X) => T

    Continuity is a "prelocal" predicate on functions to a fixed topological space T.

    Equations
    Instances For

      Satisfying the inhabited linter.

      Equations
      structure TopCat.LocalPredicate {X : TopCat} (T : XType v) extends TopCat.PrelocalPredicate :

      Given a topological space X : TopCat and a type family T : X → Type, a P : LocalPredicate T consists of:

      • a family of predicates P.pred, one for each U : Opens X, of the form (Π x : U, T x) → Prop
      • a proof that if f : Π x : V, T x satisfies the predicate on V : Opens X, then the restriction of f to any open subset U also satisfies the predicate, and
      • a proof that given some f : Π x : U, T x, if for every x : U we can find an open set x ∈ V ≤ U so that the restriction of f to V satisfies the predicate, then f itself satisfies the predicate.
      • pred : {U : TopologicalSpace.Opens X} → ((x : U) → T x)Prop
      • res : ∀ {U V : TopologicalSpace.Opens X} (i : U V) (f : (x : V) → T x), self.pred fself.pred fun (x : U) => f ((fun (x : U) => { val := x, property := }) x)
      • locality : ∀ {U : TopologicalSpace.Opens X} (f : (x : U) → T x), (∀ (x : U), ∃ (V : TopologicalSpace.Opens X) (_ : x V) (i : V U), self.pred fun (x : V) => f ((fun (x : V) => { val := x, property := }) x))self.pred f

        A local predicate must be local --- provided that it is locally satisfied, it is also globally satisfied

      Instances For
        def TopCat.continuousLocal (X : TopCat) (T : TopCat) :
        TopCat.LocalPredicate fun (x : X) => T

        Continuity is a "local" predicate on functions to a fixed topological space T.

        Equations
        Instances For
          instance TopCat.inhabitedLocalPredicate (X : TopCat) (T : TopCat) :
          Inhabited (TopCat.LocalPredicate fun (x : X) => T)

          Satisfying the inhabited linter.

          Equations

          Given a P : PrelocalPredicate, we can always construct a LocalPredicate by asking that the condition from P holds locally near every point.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            theorem TopCat.PrelocalPredicate.sheafifyOf {X : TopCat} {T : XType v} {P : TopCat.PrelocalPredicate T} {U : TopologicalSpace.Opens X} {f : (x : U) → T x} (h : P.pred f) :
            @[simp]
            theorem TopCat.subpresheafToTypes_map_coe {X : TopCat} {T : XType v} (P : TopCat.PrelocalPredicate T) {U : (TopologicalSpace.Opens X)ᵒᵖ} {V : (TopologicalSpace.Opens X)ᵒᵖ} (i : U V) (f : { f : (x : U.unop) → T x // P.pred f }) (x : V.unop) :
            ((TopCat.subpresheafToTypes P).map i f) x = f { val := x, property := }
            @[simp]
            theorem TopCat.subpresheafToTypes_obj {X : TopCat} {T : XType v} (P : TopCat.PrelocalPredicate T) (U : (TopologicalSpace.Opens X)ᵒᵖ) :
            (TopCat.subpresheafToTypes P).obj U = { f : (x : U.unop) → T x // P.pred f }

            The subpresheaf of dependent functions on X satisfying the "pre-local" predicate P.

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

              The natural transformation including the subpresheaf of functions satisfying a local predicate into the presheaf of all functions.

              Equations
              Instances For

                The functions satisfying a local predicate satisfy the sheaf condition.

                @[simp]
                theorem TopCat.subsheafToTypes_val {X : TopCat} {T : XType v} (P : TopCat.LocalPredicate T) :
                (TopCat.subsheafToTypes P).val = TopCat.subpresheafToTypes P.toPrelocalPredicate

                The subsheaf of the sheaf of all dependently typed functions satisfying the local predicate P.

                Equations
                Instances For

                  There is a canonical map from the stalk to the original fiber, given by evaluating sections.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    theorem TopCat.stalkToFiber_surjective {X : TopCat} {T : XType v} (P : TopCat.LocalPredicate T) (x : X) (w : ∀ (t : T x), ∃ (U : TopologicalSpace.OpenNhds x) (f : (y : U.obj) → T y) (_ : P.pred f), f { val := x, property := } = t) :

                    The stalkToFiber map is surjective at x if every point in the fiber T x has an allowed section passing through it.

                    theorem TopCat.stalkToFiber_injective {X : TopCat} {T : XType v} (P : TopCat.LocalPredicate T) (x : X) (w : ∀ (U V : TopologicalSpace.OpenNhds x) (fU : (y : U.obj) → T y), P.pred fU∀ (fV : (y : V.obj) → T y), P.pred fVfU { val := x, property := } = fV { val := x, property := }∃ (W : TopologicalSpace.OpenNhds x) (iU : W U) (iV : W V), ∀ (w : W.obj), fU ((fun (x_4 : W.obj) => { val := x_4, property := }) w) = fV ((fun (x_4 : W.obj) => { val := x_4, property := }) w)) :

                    The stalkToFiber map is injective at x if any two allowed sections which agree at x agree on some neighborhood of x.

                    Some repackaging: the presheaf of functions satisfying continuousPrelocal is just the same thing as the presheaf of continuous functions.

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

                      The sheaf of continuous functions on X with values in a space T.

                      Equations
                      Instances For