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.
Instances For
    @[simp]

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

    Instances For

      Satisfying the inhabited linter.

      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.
      Instances For

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

        Instances For

          Satisfying the inhabited linter.

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

          Instances For
            theorem TopCat.PrelocalPredicate.sheafifyOf {X : TopCat} {T : XType v} {P : TopCat.PrelocalPredicate T} {U : TopologicalSpace.Opens X} {f : (x : { x // x U }) → T x} (h : TopCat.PrelocalPredicate.pred P 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 // TopCat.PrelocalPredicate.pred P f }) (x : { x // x V.unop }) :
            ↑((TopCat.subpresheafToTypes P).map i f) x = f { val := x, property := (_ : x U.unop) }

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

            Instances For

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

              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
                def TopCat.subsheafToTypes {X : TopCat} {T : XType v} (P : TopCat.LocalPredicate T) :

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

                Instances For

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

                  Instances For
                    theorem TopCat.stalkToFiber_surjective {X : TopCat} {T : XType v} (P : TopCat.LocalPredicate T) (x : X) (w : ∀ (t : T x), U f x, f { val := x, property := (_ : x U.obj) } = 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 : { x // x U.obj }) → T y), TopCat.PrelocalPredicate.pred P.toPrelocalPredicate fU∀ (fV : (y : { x // x V.obj }) → T y), TopCat.PrelocalPredicate.pred P.toPrelocalPredicate fVfU { val := x, property := (_ : x U.obj) } = fV { val := x, property := (_ : x V.obj) }W iU iV, ∀ (w : { x // x W.obj }), fU ((fun x => { val := x, property := (_ : x U.obj) }) w) = fV ((fun x => { val := x, property := (_ : x V.obj) }) 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.

                    Instances For

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

                      Instances For