Documentation

Mathlib.AlgebraicGeometry.Gluing

Gluing Schemes #

Given a family of gluing data of schemes, we may glue them together. Also see the section about "locally directed" gluing, which is a special case where the conditions are easier to check.

Main definitions #

Main results #

Implementation details #

All the hard work is done in AlgebraicGeometry/PresheafedSpace/Gluing.lean where we glue presheafed spaces, sheafed spaces, and locally ringed spaces.

A family of gluing data consists of

  1. An index type J
  2. A scheme U i for each i : J.
  3. A scheme V i j for each i j : J. (Note that this is J × J → Scheme rather than J → J → Scheme to connect to the limits library easier.)
  4. An open immersion f i j : V i j ⟶ U i for each i j : ι.
  5. A transition map t i j : V i j ⟶ V j i for each i j : ι. such that
  6. f i i is an isomorphism.
  7. t i i is the identity.
  8. V i j ×[U i] V i k ⟶ V i j ⟶ V j i factors through V j k ×[U j] V j i ⟶ V j i via some t' : V i j ×[U i] V i k ⟶ V j k ×[U j] V j i.
  9. t' i j k ≫ t' j k i ≫ t' k i j = 𝟙 _.

We can then glue the schemes U i together by identifying V i j with V j i, such that the U i's are open subschemes of the glued space.

Instances For
    @[reducible, inline]

    The glue data of locally ringed spaces associated to a family of glue data of schemes.

    Equations
    Instances For

      (Implementation). The glued scheme of a glue data. This should not be used outside this file. Use AlgebraicGeometry.Scheme.GlueData.glued instead.

      Equations
      Instances For
        @[reducible, inline]

        The glued scheme of a glued space.

        Equations
        Instances For
          @[reducible, inline]

          The immersion from D.U i into the glued space.

          Equations
          Instances For
            @[reducible, inline]

            The gluing as sheafed spaces is isomorphic to the gluing as presheafed spaces.

            Equations
            Instances For
              @[simp]

              Promoted to higher priority to short circuit simplifier.

              The pullback cone spanned by V i j ⟶ U i and V i j ⟶ U j. This is a pullback diagram (vPullbackConeIsLimit).

              Equations
              Instances For

                The following diagram is a pullback, i.e. Vᵢⱼ is the intersection of Uᵢ and Uⱼ in X.

                Vᵢⱼ ⟶ Uᵢ
                 |      |
                 ↓      ↓
                 Uⱼ ⟶ X
                
                Equations
                Instances For

                  The underlying topological space of the glued scheme is isomorphic to the gluing of the underlying spaces

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    def AlgebraicGeometry.Scheme.GlueData.Rel (D : GlueData) (a b : (i : D.J) × (D.U i)) :

                    An equivalence relation on Σ i, D.U i that holds iff 𝖣.ι i x = 𝖣.ι j y. See AlgebraicGeometry.Scheme.GlueData.ι_eq_iff.

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

                      The open cover of the glued space given by the glue data.

                      Equations
                      Instances For

                        (Implementation) the transition maps in the glue data associated with an open cover.

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

                          The glue data associated with an open cover. The canonical isomorphism 𝒰.gluedCover.glued ⟶ X is provided by 𝒰.fromGlued.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            @[simp]
                            theorem AlgebraicGeometry.Scheme.Cover.gluedCover_U {X : Scheme} (𝒰 : X.OpenCover) (j : 𝒰.J) :
                            (gluedCover 𝒰).U j = 𝒰.obj j
                            @[simp]
                            @[simp]
                            theorem AlgebraicGeometry.Scheme.Cover.gluedCover_V {X : Scheme} (𝒰 : X.OpenCover) (x✝ : 𝒰.J × 𝒰.J) :
                            (gluedCover 𝒰).V x✝ = match x✝ with | (x, y) => CategoryTheory.Limits.pullback (𝒰.map x) (𝒰.map y)
                            @[simp]
                            theorem AlgebraicGeometry.Scheme.Cover.gluedCover_t {X : Scheme} (𝒰 : X.OpenCover) (x✝ x✝¹ : 𝒰.J) :
                            (gluedCover 𝒰).t x✝ x✝¹ = (CategoryTheory.Limits.pullbackSymmetry (𝒰.map x✝) (𝒰.map x✝¹)).hom
                            @[simp]
                            theorem AlgebraicGeometry.Scheme.Cover.gluedCover_t' {X : Scheme} (𝒰 : X.OpenCover) (x y z : 𝒰.J) :
                            (gluedCover 𝒰).t' x y z = gluedCoverT' 𝒰 x y z
                            @[simp]
                            theorem AlgebraicGeometry.Scheme.Cover.gluedCover_f {X : Scheme} (𝒰 : X.OpenCover) (x✝ x✝¹ : 𝒰.J) :
                            (gluedCover 𝒰).f x✝ x✝¹ = CategoryTheory.Limits.pullback.fst (𝒰.map x✝) (𝒰.map x✝¹)

                            The canonical morphism from the gluing of an open cover of X into X. This is an isomorphism, as witnessed by an IsIso instance.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              def AlgebraicGeometry.Scheme.Cover.glueMorphisms {X : Scheme} (𝒰 : X.OpenCover) {Y : Scheme} (f : (x : 𝒰.J) → 𝒰.obj x Y) (hf : ∀ (x y : 𝒰.J), CategoryTheory.CategoryStruct.comp (CategoryTheory.Limits.pullback.fst (𝒰.map x) (𝒰.map y)) (f x) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Limits.pullback.snd (𝒰.map x) (𝒰.map y)) (f y)) :
                              X Y

                              Given an open cover of X, and a morphism 𝒰.obj x ⟶ Y for each open subscheme in the cover, such that these morphisms are compatible in the intersection (pullback), we may glue the morphisms together into a morphism X ⟶ Y.

                              Note: If X is exactly (defeq to) the gluing of U i, then using Multicoequalizer.desc suffices.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                theorem AlgebraicGeometry.Scheme.Cover.hom_ext {X : Scheme} (𝒰 : X.OpenCover) {Y : Scheme} (f₁ f₂ : X Y) (h : ∀ (x : 𝒰.J), CategoryTheory.CategoryStruct.comp (𝒰.map x) f₁ = CategoryTheory.CategoryStruct.comp (𝒰.map x) f₂) :
                                f₁ = f₂
                                @[simp]
                                theorem AlgebraicGeometry.Scheme.Cover.ι_glueMorphisms {X : Scheme} (𝒰 : X.OpenCover) {Y : Scheme} (f : (x : 𝒰.J) → 𝒰.obj x Y) (hf : ∀ (x y : 𝒰.J), CategoryTheory.CategoryStruct.comp (CategoryTheory.Limits.pullback.fst (𝒰.map x) (𝒰.map y)) (f x) = CategoryTheory.CategoryStruct.comp (CategoryTheory.Limits.pullback.snd (𝒰.map x) (𝒰.map y)) (f y)) (x : 𝒰.J) :

                                Locally directed gluing #

                                We say that a diagram of open immersions is "locally directed" if for any V, W ⊆ U in the diagram, V ∩ W is a union of elements in the diagram. Equivalently, for every x ∈ U in the diagram, the set of elements containing x is directed (and hence the name).

                                For such a diagram, we can glue them directly since the gluing conditions are always satisfied. The intended usage is to provide the following instances:

                                noncomputable def AlgebraicGeometry.Scheme.IsLocallyDirected.V {J : Type w} [CategoryTheory.Category.{v, w} J] (F : CategoryTheory.Functor J Scheme) [∀ {i j : J} (f : i j), IsOpenImmersion (F.map f)] (i j : J) :
                                (F.obj i).Opens

                                (Implemetation detail) The intersection V in the glue data associated to a locally directed diagram.

                                Equations
                                Instances For

                                  (Implemetation detail) The inclusion map V i j ⟶ F j in the glue data associated to a locally directed diagram.

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

                                    (Implemetation detail) The transition map V i j ⟶ V j i in the glue data associated to a locally directed diagram.

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

                                      (Implemetation detail) The glue data associated to a locally directed diagram.

                                      One usually does not want to use this directly, and instead use the generic colimit API.

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

                                        (Implemetation detail) The cocone associated to a locally directed diagram.

                                        One usually does not want to use this directly, and instead use the generic colimit API.

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

                                          (Implemetation detail) The cocone associated to a locally directed diagram is a colimit.

                                          One usually does not want to use this directly, and instead use the generic colimit API.

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

                                            (Implemetation detail) The cocone associated to a locally directed diagram is a colimit as locally ringed spaces.

                                            One usually does not want to use this directly, and instead use the generic colimit API.

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

                                              The open cover of the colimit of a locally directed diagram by the components.

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