Documentation

Mathlib.AlgebraicGeometry.OpenImmersion

Open immersions of schemes #

@[reducible, inline]

A morphism of Schemes is an open immersion if it is an open immersion as a morphism of LocallyRingedSpaces

Equations
Instances For

    To show that a locally ringed space is a scheme, it suffices to show that it has a jointly surjective family of open immersions from affine schemes.

    Equations
    Instances For
      @[deprecated AlgebraicGeometry.IsOpenImmersion.isOpen_range]

      Alias of AlgebraicGeometry.IsOpenImmersion.isOpen_range.

      The image of an open immersion as an open set.

      Equations
      • f.opensRange = { carrier := Set.range f.val.base, is_open' := }
      Instances For
        @[reducible, inline]

        The functor opens X ⥤ opens Y associated with an open immersion f : X ⟶ Y.

        Equations
        Instances For

          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 the image (as an open set) of U under an open immersion f.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              theorem AlgebraicGeometry.Scheme.Hom.image_le_image_of_le {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] {U : TopologicalSpace.Opens X.toPresheafedSpace} {V : TopologicalSpace.Opens X.toPresheafedSpace} (e : U V) :
              f.opensFunctor.obj U f.opensFunctor.obj V
              @[simp]
              theorem AlgebraicGeometry.Scheme.Hom.opensFunctor_map_homOfLE {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] {U : TopologicalSpace.Opens X.toPresheafedSpace} {V : TopologicalSpace.Opens X.toPresheafedSpace} (e : U V) :
              @[simp]
              theorem AlgebraicGeometry.Scheme.Hom.preimage_image_eq {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens X.toPresheafedSpace) :
              (TopologicalSpace.Opens.map f.val.base).obj (f.opensFunctor.obj U) = U
              theorem AlgebraicGeometry.Scheme.Hom.image_preimage_eq_opensRange_inter {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens Y.toPresheafedSpace) :
              f.opensFunctor.obj ((TopologicalSpace.Opens.map f.val.base).obj U) = f.opensRange U
              def AlgebraicGeometry.Scheme.Hom.invApp {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens X.toPresheafedSpace) :
              X.presheaf.obj { unop := U } Y.presheaf.obj { unop := f.opensFunctor.obj U }

              The isomorphism Γ(X, U) ⟶ Γ(Y, f(U)) induced by an open immersion f : X ⟶ Y.

              Equations
              Instances For
                @[simp]
                theorem AlgebraicGeometry.Scheme.Hom.invApp_naturality_assoc {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] {U : TopologicalSpace.Opens X.toPresheafedSpace} {V : TopologicalSpace.Opens X.toPresheafedSpace} (i : { unop := U } { unop := V }) {Z : CommRingCat} (h : Y.presheaf.obj { unop := f.opensFunctor.obj V } Z) :
                CategoryTheory.CategoryStruct.comp (X.presheaf.map i) (CategoryTheory.CategoryStruct.comp (f.invApp V) h) = CategoryTheory.CategoryStruct.comp (f.invApp U) (CategoryTheory.CategoryStruct.comp (Y.presheaf.map (f.opensFunctor.op.map i)) h)
                @[simp]
                theorem AlgebraicGeometry.Scheme.Hom.invApp_naturality {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] {U : TopologicalSpace.Opens X.toPresheafedSpace} {V : TopologicalSpace.Opens X.toPresheafedSpace} (i : { unop := U } { unop := V }) :
                CategoryTheory.CategoryStruct.comp (X.presheaf.map i) (f.invApp V) = CategoryTheory.CategoryStruct.comp (f.invApp U) (Y.presheaf.map (f.opensFunctor.op.map i))
                theorem AlgebraicGeometry.Scheme.Hom.inv_invApp {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens X.toPresheafedSpace) :
                CategoryTheory.inv (f.invApp U) = CategoryTheory.CategoryStruct.comp (f.app (f.opensFunctor.obj U)) (X.presheaf.map (CategoryTheory.eqToHom ).op)
                @[simp]
                theorem AlgebraicGeometry.Scheme.Hom.app_invApp_assoc {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens Y.toPresheafedSpace) {Z : CommRingCat} (h : Y.presheaf.obj { unop := f.opensFunctor.obj ((TopologicalSpace.Opens.map f.val.base).obj U) } Z) :
                @[simp]
                theorem AlgebraicGeometry.Scheme.Hom.app_invApp'_assoc {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens Y.toPresheafedSpace) (hU : U f.opensRange) {Z : CommRingCat} (h : Y.presheaf.obj { unop := f.opensFunctor.obj ((TopologicalSpace.Opens.map f.val.base).obj U) } Z) :
                theorem AlgebraicGeometry.Scheme.Hom.app_invApp' {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens Y.toPresheafedSpace) (hU : U f.opensRange) :
                CategoryTheory.CategoryStruct.comp (f.app U) (f.invApp ((TopologicalSpace.Opens.map f.val.base).obj U)) = Y.presheaf.map (CategoryTheory.eqToHom ).op

                A variant of app_invApp that gives an eqToHom instead of homOfLE.

                @[simp]
                theorem AlgebraicGeometry.Scheme.Hom.invApp_app_assoc {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens X.toPresheafedSpace) {Z : CommRingCat} (h : X.presheaf.obj { unop := (TopologicalSpace.Opens.map f.val.base).obj (f.opensFunctor.obj U) } Z) :
                @[simp]
                theorem AlgebraicGeometry.Scheme.Hom.invApp_app {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X.Hom Y) [H : AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens X.toPresheafedSpace) :
                CategoryTheory.CategoryStruct.comp (f.invApp U) (f.app (f.opensFunctor.obj U)) = X.presheaf.map (CategoryTheory.eqToHom ).op
                theorem AlgebraicGeometry.Scheme.Hom.appLE_invApp_apply {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X Y) [AlgebraicGeometry.IsOpenImmersion f] {U : TopologicalSpace.Opens Y.toPresheafedSpace} {V : TopologicalSpace.Opens X.toPresheafedSpace} (e : V (TopologicalSpace.Opens.map f.val.base).obj U) (x : (CategoryTheory.forget CommRingCat).obj (Y.presheaf.obj { unop := U })) :

                The open sets of an open subscheme corresponds to the open sets containing in the image.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  theorem AlgebraicGeometry.Scheme.exists_affine_mem_range_and_range_subset {X : AlgebraicGeometry.Scheme} {x : X.toPresheafedSpace} {U : TopologicalSpace.Opens X.toPresheafedSpace} (hxU : x U) :
                  ∃ (R : CommRingCat) (f : AlgebraicGeometry.Spec R X), AlgebraicGeometry.IsOpenImmersion f x Set.range f.val.base Set.range f.val.base U

                  If X ⟶ Y is an open immersion, and Y is a scheme, then so is X.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    theorem AlgebraicGeometry.Scheme.restrict_carrier {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) :
                    (X.restrict h).toPresheafedSpace = U
                    @[simp]
                    theorem AlgebraicGeometry.Scheme.restrict_presheaf_obj {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X✝.toPresheafedSpace} (h : OpenEmbedding f) (X : (TopologicalSpace.Opens U)ᵒᵖ) :
                    (X✝.restrict h).presheaf.obj X = X✝.presheaf.obj { unop := .functor.obj X.unop }

                    The restriction of a Scheme along an open embedding.

                    Equations
                    Instances For
                      theorem AlgebraicGeometry.Scheme.restrict_toPresheafedSpace {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) :
                      (X.restrict h).toPresheafedSpace = X.restrict h
                      @[simp]
                      theorem AlgebraicGeometry.Scheme.ofRestrict_val_base {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) :
                      (X.ofRestrict h).val.base = f
                      theorem AlgebraicGeometry.Scheme.ofRestrict_val_c_app {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) (V : (TopologicalSpace.Opens X.toPresheafedSpace)ᵒᵖ) :
                      (X.ofRestrict h).val.c.app V = X.presheaf.map (.adjunction.counit.app V.unop).op
                      def AlgebraicGeometry.Scheme.ofRestrict {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) :
                      X.restrict h X

                      The canonical map from the restriction to the subspace.

                      Equations
                      • X.ofRestrict h = X.ofRestrict h
                      Instances For
                        @[simp]
                        theorem AlgebraicGeometry.Scheme.ofRestrict_app {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) (V : TopologicalSpace.Opens X.toPresheafedSpace) :
                        AlgebraicGeometry.Scheme.Hom.app (X.ofRestrict h) V = X.presheaf.map (.adjunction.counit.app V).op
                        instance AlgebraicGeometry.IsOpenImmersion.ofRestrict {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) :
                        Equations
                        • =
                        @[simp]
                        theorem AlgebraicGeometry.Scheme.ofRestrict_appLE {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) (V : TopologicalSpace.Opens X.toPresheafedSpace) (W : TopologicalSpace.Opens (X.restrict h).toPresheafedSpace) (e : W (TopologicalSpace.Opens.map (X.ofRestrict h).val.base).obj V) :
                        AlgebraicGeometry.Scheme.Hom.appLE (X.ofRestrict h) V W e = X.presheaf.map (CategoryTheory.homOfLE ).op
                        @[simp]
                        theorem AlgebraicGeometry.Scheme.restrict_presheaf_map {U : TopCat} (X : AlgebraicGeometry.Scheme) {f : U TopCat.of X.toPresheafedSpace} (h : OpenEmbedding f) (V : (TopologicalSpace.Opens (X.restrict h).toPresheafedSpace)ᵒᵖ) (W : (TopologicalSpace.Opens (X.restrict h).toPresheafedSpace)ᵒᵖ) (i : V W) :
                        (X.restrict h).presheaf.map i = X.presheaf.map (CategoryTheory.homOfLE ).op

                        An open immersion induces an isomorphism from the domain onto the image

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          theorem AlgebraicGeometry.IsOpenImmersion.range_pullback_fst_of_right {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} {Z : AlgebraicGeometry.Scheme} (f : X Z) (g : Y Z) [H : AlgebraicGeometry.IsOpenImmersion f] :
                          Set.range CategoryTheory.Limits.pullback.fst.val.base = ((TopologicalSpace.Opens.map g.val.base).obj { carrier := Set.range f.val.base, is_open' := }).carrier

                          The universal property of open immersions: For an open immersion f : X ⟶ Z, given any morphism of schemes g : Y ⟶ Z whose topological image is contained in the image of f, we can lift this morphism to a unique Y ⟶ X that commutes with these maps.

                          Equations
                          Instances For

                            Two open immersions with equal range are isomorphic.

                            Equations
                            • One or more equations did not get rendered due to their size.
                            Instances For
                              noncomputable def AlgebraicGeometry.IsOpenImmersion.ΓIso {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X Y) [AlgebraicGeometry.IsOpenImmersion f] (U : TopologicalSpace.Opens Y.toPresheafedSpace) :
                              X.presheaf.obj { unop := (TopologicalSpace.Opens.map f.val.base).obj U } Y.presheaf.obj { unop := AlgebraicGeometry.Scheme.Hom.opensRange f U }

                              If f is an open immersion X ⟶ Y, the global sections of X are naturally isomorphic to the sections of Y over the image of f.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                theorem AlgebraicGeometry.Scheme.image_basicOpen {X : AlgebraicGeometry.Scheme} {Y : AlgebraicGeometry.Scheme} (f : X Y) [H : AlgebraicGeometry.IsOpenImmersion f] {U : TopologicalSpace.Opens X.toPresheafedSpace} (r : (X.presheaf.obj { unop := U })) :