Documentation

Mathlib.AlgebraicGeometry.Scheme

The category of schemes #

A scheme is a locally ringed space such that every point is contained in some open set where there is an isomorphism of presheaves between the restriction to that open set, and the structure sheaf of Spec R, for some commutative ring R.

A morphism of schemes is just a morphism of the underlying locally ringed spaces.

We define Scheme as an X : LocallyRingedSpace, along with a proof that every point has an open neighbourhood U so that the restriction of X to U is isomorphic, as a locally ringed space, to Spec.toLocallyRingedSpace.obj (op R) for some R : CommRingCat.

Instances For
    @[reducible, inline]

    The type of open sets of a scheme.

    Equations
    Instances For

      A morphism between schemes is a morphism between the underlying locally ringed spaces.

      Instances For
        @[reducible, inline]

        Cast a morphism of schemes into morphisms of local ringed spaces.

        Equations
        Instances For

          Schemes are a full subcategory of locally ringed spaces.

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

          Pretty printer defined by notation3 command.

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

            f ⁻¹ᵁ U is notation for (Opens.map f.base).obj U, the preimage of an open set U under f.

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

              Γ(X, U) is notation for X.presheaf.obj (op U).

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

                Pretty printer defined by notation3 command.

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

                  The structure sheaf of a scheme.

                  Equations
                  Instances For
                    @[reducible, inline]

                    Given a morphism of schemes f : X ⟶ Y, and open U ⊆ Y, this is the induced map Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U).

                    Equations
                    Instances For
                      @[reducible, inline]

                      Given a morphism of schemes f : X ⟶ Y, this is the induced map Γ(Y, ⊤) ⟶ Γ(X, ⊤).

                      Equations
                      Instances For

                        Given a morphism of schemes f : X ⟶ Y, and open sets U ⊆ Y, V ⊆ f ⁻¹' U, this is the induced map Γ(Y, U) ⟶ Γ(X, V).

                        Equations
                        Instances For
                          @[simp]
                          @[simp]
                          theorem AlgebraicGeometry.Scheme.Hom.appLE_congr {X Y : Scheme} (f : X.Hom Y) {U U' : Y.Opens} {V V' : X.Opens} (e : V (TopologicalSpace.Opens.map f.base).obj U) (e₁ : U = U') (e₂ : V = V') (P : {R S : CommRingCat} → (R S) → Prop) :
                          P (f.appLE U V e) P (f.appLE U' V' )

                          A morphism of schemes f : X ⟶ Y induces a local ring homomorphism from Y.presheaf.stalk (f x) to X.presheaf.stalk x for any x : X.

                          Equations
                          Instances For
                            theorem AlgebraicGeometry.Scheme.Hom.ext {X Y : Scheme} {f g : X Y} (h_base : f.base = g.base) (h_app : ∀ (U : Y.Opens), CategoryTheory.CategoryStruct.comp (app f U) (X.presheaf.map (CategoryTheory.eqToHom ).op) = app g U) :
                            f = g
                            theorem AlgebraicGeometry.Scheme.Hom.ext' {X Y : Scheme} {f g : X Y} (h : toLRSHom f = toLRSHom g) :
                            f = g

                            An alternative ext lemma for scheme morphisms.

                            theorem AlgebraicGeometry.Scheme.Hom.preimage_iSup {X Y : Scheme} (f : X.Hom Y) {ι : Sort u_1} (U : ιY.Opens) :
                            theorem AlgebraicGeometry.Scheme.Hom.preimage_iSup_eq_top {X Y : Scheme} (f : X.Hom Y) {ι : Sort u_1} {U : ιY.Opens} (hU : iSup U = ) :
                            ⨆ (i : ι), (TopologicalSpace.Opens.map f.base).obj (U i) =

                            The forgetful functor from Scheme to LocallyRingedSpace.

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

                              The forget functor Scheme ⥤ LocallyRingedSpace is fully faithful.

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

                                An isomorphism of schemes induces a homeomorphism of the underlying topological spaces.

                                Equations
                                Instances For

                                  Alias of AlgebraicGeometry.Scheme.homeoOfIso.


                                  An isomorphism of schemes induces a homeomorphism of the underlying topological spaces.

                                  Equations
                                  Instances For

                                    An isomorphism of schemes induces a homeomorphism of the underlying topological spaces.

                                    Equations
                                    Instances For
                                      @[deprecated AlgebraicGeometry.Scheme.inv_appTop (since := "2024-11-23")]

                                      Alias of AlgebraicGeometry.Scheme.inv_appTop.

                                      The spectrum of a commutative ring, as a scheme.

                                      Equations
                                      Instances For

                                        The induced map of a ring homomorphism on the ring spectra, as a morphism of schemes.

                                        Equations
                                        Instances For

                                          The spectrum, as a contravariant functor from commutative rings to schemes.

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

                                            The empty scheme.

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

                                              The counit (SpecΓIdentity.inv.op) of the adjunction ΓSpec as an isomorphism. This is almost never needed in practical use cases. Use ΓSpecIso instead.

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

                                                The subset of the underlying space where the given section does not vanish.

                                                Equations
                                                Instances For
                                                  @[simp]

                                                  A variant of mem_basicOpen for bundled x : U.

                                                  A variant of mem_basicOpen without the x ∈ U assumption.

                                                  @[simp]
                                                  theorem AlgebraicGeometry.Scheme.basicOpen_mul (X : Scheme) {U : X.Opens} (f g : (X.presheaf.obj (Opposite.op U))) :
                                                  theorem AlgebraicGeometry.Scheme.basicOpen_pow (X : Scheme) {U : X.Opens} (f : (X.presheaf.obj (Opposite.op U))) {n : } (h : 0 < n) :
                                                  X.basicOpen (f ^ n) = X.basicOpen f

                                                  The zero locus of a set of sections s over an open set U is the closed set consisting of the complement of U and of all points of U, where all elements of f vanish.

                                                  Equations
                                                  Instances For
                                                    theorem AlgebraicGeometry.Scheme.zeroLocus_def (X : Scheme) {U : X.Opens} (s : Set (X.presheaf.obj (Opposite.op U))) :
                                                    X.zeroLocus s = fs, (X.basicOpen f).carrier
                                                    @[simp]
                                                    theorem AlgebraicGeometry.Scheme.mem_zeroLocus_iff (X : Scheme) {U : X.Opens} (s : Set (X.presheaf.obj (Opposite.op U))) (x : X.toPresheafedSpace) :
                                                    x X.zeroLocus s fs, xX.basicOpen f
                                                    theorem AlgebraicGeometry.Scheme.zeroLocus_mono (X : Scheme) {U : X.Opens} {s t : Set (X.presheaf.obj (Opposite.op U))} (h : s t) :
                                                    theorem AlgebraicGeometry.germ_eq_zero_of_pow_mul_eq_zero {X : Scheme} {U : TopologicalSpace.Opens X.toPresheafedSpace} (x : U) {f s : (X.presheaf.obj (Opposite.op U))} (hx : x X.basicOpen s) {n : } (hf : s ^ n * f = 0) :

                                                    If x = y, the stalk maps are isomorphic.

                                                    Equations
                                                    Instances For