Skip to content

Commit

Permalink
feat(AlgebraicGeometry): proper morphisms of schemes (#17863)
Browse files Browse the repository at this point in the history
We define proper morphisms of schemes and show standard stability properties.

Partly from the valuative criterion project.

Co-authored by: Andrew Yang
  • Loading branch information
chrisflav committed Oct 17, 2024
1 parent 389e097 commit 23a1a4b
Show file tree
Hide file tree
Showing 5 changed files with 98 additions and 1 deletion.
1 change: 1 addition & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -868,6 +868,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.FiniteType
import Mathlib.AlgebraicGeometry.Morphisms.IsIso
import Mathlib.AlgebraicGeometry.Morphisms.OpenImmersion
import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion
import Mathlib.AlgebraicGeometry.Morphisms.Proper
import Mathlib.AlgebraicGeometry.Morphisms.QuasiCompact
import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated
import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties
Expand Down
14 changes: 14 additions & 0 deletions Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Amelia Livingston, Christian Merten, Jonas van der Schaaf
import Mathlib.AlgebraicGeometry.Morphisms.Preimmersion
import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated
import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties
import Mathlib.AlgebraicGeometry.Morphisms.FiniteType
import Mathlib.AlgebraicGeometry.ResidueField
import Mathlib.RingTheory.RingHom.Surjective

Expand Down Expand Up @@ -269,4 +270,17 @@ lemma IsClosedImmersion.stableUnderBaseChange :
exact ⟨inferInstance, RingHom.surjective_stableUnderBaseChange.pullback_fst_app_top _
RingHom.surjective_respectsIso f _ hsurj⟩

/-- Closed immersions are locally of finite type. -/
instance (priority := 900) {X Y : Scheme.{u}} (f : X ⟶ Y) [h : IsClosedImmersion f] :
LocallyOfFiniteType f := by
wlog hY : IsAffine Y
· rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := @LocallyOfFiniteType) _
(iSup_affineOpens_eq_top Y)]
intro U
have H : IsClosedImmersion (f ∣_ U) := IsLocalAtTarget.restrict h U
exact this _ U.2
obtain ⟨_, hf⟩ := h.isAffine_surjective_of_isAffine
rw [HasRingHomProperty.iff_of_isAffine (P := @LocallyOfFiniteType)]
exact RingHom.FiniteType.of_surjective (Scheme.Hom.app f ⊤) hf

end AlgebraicGeometry
3 changes: 3 additions & 0 deletions Mathlib/AlgebraicGeometry/Morphisms/FiniteType.lean
Original file line number Diff line number Diff line change
Expand Up @@ -59,6 +59,9 @@ theorem locallyOfFiniteType_of_comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z)
[LocallyOfFiniteType (f ≫ g)] : LocallyOfFiniteType f :=
HasRingHomProperty.of_comp (fun _ _ ↦ RingHom.FiniteType.of_comp_finiteType) ‹_›

instance : MorphismProperty.IsMultiplicative @LocallyOfFiniteType where
id_mem _ := inferInstance

open scoped TensorProduct in
lemma locallyOfFiniteType_stableUnderBaseChange :
MorphismProperty.StableUnderBaseChange @LocallyOfFiniteType :=
Expand Down
69 changes: 69 additions & 0 deletions Mathlib/AlgebraicGeometry/Morphisms/Proper.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,69 @@
/-
Copyright (c) 2024 Christian Merten, Andrew Yang. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Christian Merten, Andrew Yang
-/
import Mathlib.AlgebraicGeometry.Morphisms.Separated
import Mathlib.AlgebraicGeometry.Morphisms.UniversallyClosed
import Mathlib.AlgebraicGeometry.Morphisms.FiniteType

/-!
# Proper morphisms
A morphism of schemes is proper if it is separated, universally closed and (locally) of finite type.
Note that we don't require quasi-compact, since this is implied by universally closed (TODO).
-/


noncomputable section

open CategoryTheory

universe u

namespace AlgebraicGeometry

variable {X Y : Scheme.{u}} (f : X ⟶ Y)

/-- A morphism is proper if it is separated, universally closed and locally of finite type. -/
@[mk_iff]
class IsProper extends IsSeparated f, UniversallyClosed f, LocallyOfFiniteType f : Prop where

lemma isProper_eq : @IsProper =
(@IsSeparated ⊓ @UniversallyClosed : MorphismProperty Scheme) ⊓ @LocallyOfFiniteType := by
ext X Y f
rw [isProper_iff, ← and_assoc]
rfl

namespace IsProper

instance : MorphismProperty.RespectsIso @IsProper := by
rw [isProper_eq]
infer_instance

instance stableUnderComposition : MorphismProperty.IsStableUnderComposition @IsProper := by
rw [isProper_eq]
infer_instance

instance : MorphismProperty.IsMultiplicative @IsProper := by
rw [isProper_eq]
infer_instance

instance (priority := 900) [IsClosedImmersion f] : IsProper f where

lemma stableUnderBaseChange : MorphismProperty.StableUnderBaseChange @IsProper := by
rw [isProper_eq]
exact MorphismProperty.StableUnderBaseChange.inf
(MorphismProperty.StableUnderBaseChange.inf
IsSeparated.stableUnderBaseChange universallyClosed_stableUnderBaseChange)
locallyOfFiniteType_stableUnderBaseChange

instance : IsLocalAtTarget @IsProper := by
rw [isProper_eq]
infer_instance

end IsProper

end AlgebraicGeometry
12 changes: 11 additions & 1 deletion Mathlib/AlgebraicGeometry/Morphisms/UniversallyClosed.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2022 Andrew Yang. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Andrew Yang
-/
import Mathlib.AlgebraicGeometry.Morphisms.Constructors
import Mathlib.AlgebraicGeometry.Morphisms.ClosedImmersion
import Mathlib.Topology.LocalAtTarget

/-!
Expand Down Expand Up @@ -40,6 +40,13 @@ class UniversallyClosed (f : X ⟶ Y) : Prop where
theorem universallyClosed_eq : @UniversallyClosed = universally (topologically @IsClosedMap) := by
ext X Y f; rw [universallyClosed_iff]

instance (priority := 900) [IsClosedImmersion f] : UniversallyClosed f := by
rw [universallyClosed_eq]
intro X' Y' i₁ i₂ f' hf
have hf' : IsClosedImmersion f' :=
IsClosedImmersion.stableUnderBaseChange hf.flip inferInstance
exact hf'.base_closed.isClosedMap

theorem universallyClosed_respectsIso : RespectsIso @UniversallyClosed :=
universallyClosed_eq.symm ▸ universally_respectsIso (topologically @IsClosedMap)

Expand All @@ -59,6 +66,9 @@ instance universallyClosedTypeComp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z)
[hf : UniversallyClosed f] [hg : UniversallyClosed g] : UniversallyClosed (f ≫ g) :=
comp_mem _ _ _ hf hg

instance : MorphismProperty.IsMultiplicative @UniversallyClosed where
id_mem _ := inferInstance

instance universallyClosed_fst {X Y Z : Scheme} (f : X ⟶ Z) (g : Y ⟶ Z) [hg : UniversallyClosed g] :
UniversallyClosed (pullback.fst f g) :=
universallyClosed_stableUnderBaseChange.fst f g hg
Expand Down

0 comments on commit 23a1a4b

Please sign in to comment.