Documentation

Mathlib.AlgebraicGeometry.Limits

(Co)Limits of Schemes #

We construct various limits and colimits in the category of schemes.

TODO #

Spec ℤ is the terminal object in the category of schemes.

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

    Spec ℤ is the terminal object in the category of schemes.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      noncomputable def AlgebraicGeometry.Scheme.emptyTo (X : Scheme) :

      The map from the empty scheme.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem AlgebraicGeometry.Scheme.emptyTo_base (X : Scheme) :
        X.emptyTo.base = TopCat.ofHom { toFun := fun (x : PEmpty.{u + 1}) => x.elim, continuous_toFun := }
        Equations

        The empty scheme is the initial object in the category of schemes.

        Equations
        Instances For
          @[instance 100]
          @[instance 100]

          Spec 0 is the initial object in the category of schemes.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[instance 100]
            @[instance 100]

            This is true in general for finite discrete schemes. See below.

            The images of each component in the coproduct is disjoint.

            theorem AlgebraicGeometry.isEmpty_of_commSq_sigmaι_of_ne {σ : Type v} {g : σScheme} [Small.{u, v} σ] {i j : σ} {Z : Scheme} {a : Z g i} {b : Z g j} (h : CategoryTheory.CommSq a b (CategoryTheory.Limits.Sigma.ι g i) (CategoryTheory.Limits.Sigma.ι g j)) (hij : i j) :
            IsEmpty Z
            noncomputable def AlgebraicGeometry.sigmaOpenCover {σ : Type v} (g : σScheme) [Small.{u, v} σ] :

            The cover of ∐ X by the Xᵢ.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              @[simp]
              theorem AlgebraicGeometry.sigmaOpenCover_X {σ : Type v} (g : σScheme) [Small.{u, v} σ] (a✝ : σ) :
              (sigmaOpenCover g).X a✝ = g a✝
              noncomputable def AlgebraicGeometry.sigmaMk {ι : Type u} (f : ιScheme) :
              (i : ι) × (f i) ≃ₜ ↥( f)

              The underlying topological space of the coproduct is homeomorphic to the disjoint union.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem AlgebraicGeometry.sigmaMk_mk {ι : Type u} (f : ιScheme) (i : ι) (x : (f i)) :
                theorem AlgebraicGeometry.isOpenImmersion_sigmaDesc {σ : Type v} (g : σScheme) [Small.{u, v} σ] {X : Scheme} (α : (i : σ) → g i X) [∀ (i : σ), IsOpenImmersion (α i)] ( : Pairwise (Function.onFun Disjoint fun (x : σ) => Set.range (CategoryTheory.ConcreteCategory.hom (α x).base))) :
                theorem AlgebraicGeometry.nonempty_isColimit_cofanMk_of {σ : Type v} [Small.{u, v} σ] {X : σScheme} {S : Scheme} (f : (i : σ) → X i S) [∀ (i : σ), IsOpenImmersion (f i)] (hcov : ⨆ (i : σ), Scheme.Hom.opensRange (f i) = ) (hdisj : Pairwise (Function.onFun Disjoint fun (x : σ) => Scheme.Hom.opensRange (f x))) :

                S is the disjoint union of Xᵢ if the Xᵢ are covering, pairwise disjoint open subschemes of S.

                (Implementation Detail) The coproduct of the two schemes is given by indexed coproducts over WalkingPair.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  noncomputable def AlgebraicGeometry.coprodMk (X Y : Scheme) :
                  X Y ≃ₜ ↥(X ⨿ Y)

                  The underlying topological space of the coproduct is homeomorphic to the disjoint union

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    noncomputable def AlgebraicGeometry.coprodOpenCover (X Y : Scheme) :

                    The open cover of the coproduct of two schemes.

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

                      If X and Y are open disjoint and covering open subschemes of S, S is the disjoint union of X and Y.

                      The sections on coproducts of schemes are the (categorical) product of the sections on the components

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        noncomputable def AlgebraicGeometry.coprodSpec (R S : Type u) [CommRing R] [CommRing S] :
                        Spec { carrier := R, commRing := inst✝ } ⨿ Spec { carrier := S, commRing := inst✝¹ } Spec { carrier := R × S, commRing := Prod.instCommRing }

                        The map Spec R ⨿ Spec S ⟶ Spec (R × S). This is an isomorphism as witnessed by an IsIso instance provided below.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          theorem AlgebraicGeometry.coprodSpec_coprodMk (R S : Type u) [CommRing R] [CommRing S] (x : (Spec { carrier := R, commRing := inst✝ }) (Spec { carrier := S, commRing := inst✝¹ })) :
                          (CategoryTheory.ConcreteCategory.hom (coprodSpec R S).base) ((coprodMk (Spec { carrier := R, commRing := inst✝ }) (Spec { carrier := S, commRing := inst✝¹ })) x) = (PrimeSpectrum.primeSpectrumProd R S).symm x
                          theorem AlgebraicGeometry.coprodSpec_apply (R S : Type u) [CommRing R] [CommRing S] (x : ↥(Spec { carrier := R, commRing := inst✝ } ⨿ Spec { carrier := S, commRing := inst✝¹ })) :
                          (CategoryTheory.ConcreteCategory.hom (coprodSpec R S).base) x = (PrimeSpectrum.primeSpectrumProd R S).symm ((coprodMk (Spec { carrier := R, commRing := inst✝ }) (Spec { carrier := S, commRing := inst✝¹ })).symm x)
                          theorem AlgebraicGeometry.isIso_stalkMap_coprodSpec (R S : Type u) [CommRing R] [CommRing S] (x : ↥(Spec { carrier := R, commRing := inst✝ } ⨿ Spec { carrier := S, commRing := inst✝¹ })) :
                          noncomputable def AlgebraicGeometry.sigmaSpec {ι : Type u} (R : ιCommRingCat) :
                          ( fun (i : ι) => Spec (R i)) Spec { carrier := (i : ι) → (R i), commRing := Pi.commRing }

                          The canonical map ∐ Spec Rᵢ ⟶ Spec (Π Rᵢ). This is an isomorphism when the product is finite.

                          Equations
                          Instances For
                            @[simp]
                            theorem AlgebraicGeometry.ι_sigmaSpec {ι : Type u} (R : ιCommRingCat) (i : ι) :
                            @[simp]
                            theorem AlgebraicGeometry.ι_sigmaSpec_assoc {ι : Type u} (R : ιCommRingCat) (i : ι) {Z : Scheme} (h : Spec { carrier := (i : ι) → (R i), commRing := Pi.commRing } Z) :
                            instance AlgebraicGeometry.instIsOpenImmersionMapOfHomForallEvalRingHom {ι : Type u} (i : ι) (R : ιType (max u_1 u)) [(i : ι) → CommRing (R i)] :
                            instance AlgebraicGeometry.instIsAffineSigmaObjScheme {σ : Type v} (g : σScheme) [Finite σ] [∀ (i : σ), IsAffine (g i)] :
                            theorem AlgebraicGeometry.IsAffineOpen.iSup_of_disjoint {σ : Type v} {X : Scheme} [Finite σ] {U : σX.Opens} (hU : ∀ (i : σ), IsAffineOpen (U i)) (hU' : Pairwise (Function.onFun Disjoint U)) :
                            theorem AlgebraicGeometry.IsAffineOpen.biSup_of_disjoint {σ : Type v} {X : Scheme} {s : Set σ} (hs : s.Finite) {U : σX.Opens} (hU : is, IsAffineOpen (U i)) (hU' : s.Pairwise (Function.onFun Disjoint U)) :
                            IsAffineOpen (⨆ is, U i)