Skip to content

Commit 65d6572

Browse files
erdOnejoelriou
andcommitted
refactor(AlgebraicGeometry): Rework Morphisms/Basic to phase away from TFAEs (#14430)
Previously the contents of `Morphism/Baisic` were all implementation detail that needed to be copied over to each new morphism class. In this PR we clean up the file and promote it into a proper interface of the API. We also phase away from TFAEs in favor of easier to use iff lemmas. We introduce the following two interfaces: ## `IsLocalAtTarget` - `AlgebraicGeometry.IsLocalAtTarget`: We say that `IsLocalAtTarget P` for `P : MorphismProperty Scheme` if 1. `P` respects isomorphisms. 2. If `P` holds for `f : X ⟶ Y`, then `P` holds for `f ∣_ U` for any open `U` of `Y`. 3. If `P` holds for `f ∣_ U` for an open cover `U` of `Y`, then `P` holds for `f`. For a morphism property `P` local at the target and `f : X ⟶ Y`, we provide these API lemmas: - `AlgebraicGeometry.IsLocalAtTarget.of_isPullback`: `P` is preserved under pullback along open immersions. - `AlgebraicGeometry.IsLocalAtTarget.restrict`: `P f → P (f ∣_ U)` for an open `U` of `Y`. - `AlgebraicGeometry.IsLocalAtTarget.iff_of_iSup_eq_top`: `P f ↔ ∀ i, P (f ∣_ U i)` for a family `U i` of open sets covering `Y`. - `AlgebraicGeometry.IsLocalAtTarget.iff_of_openCover`: `P f ↔ ∀ i, P (𝒰.pullbackHom f i)` for `𝒰 : Y.openCover`. ## `HasAffineProperty` - `AlgebraicGeometry.HasAffineProperty`: `HasAffineProperty P Q` is a type class asserting that `P` is local at the target, and over affine schemes, it is equivalent to `Q : AffineTargetMorphismProperty`. For `HasAffineProperty P Q` and `f : X ⟶ Y`, we provide these API lemmas: - `AlgebraicGeometry.HasAffineProperty.of_isPullback`: `P` is preserved under pullback along open immersions from affine schemes. - `AlgebraicGeometry.HasAffineProperty.restrict`: `P f → Q (f ∣_ U)` for affine `U` of `Y`. - `AlgebraicGeometry.HasAffineProperty.iff_of_iSup_eq_top`: `P f ↔ ∀ i, Q (f ∣_ U i)` for a family `U i` of affine open sets covering `Y`. - `AlgebraicGeometry.HasAffineProperty.iff_of_openCover`: `P f ↔ ∀ i, P (𝒰.pullbackHom f i)` for affine open covers `𝒰` of `Y`. - `AlgebraicGeometry.HasAffineProperty.stableUnderBaseChange_mk`: If `Q` is stable under affine base change, then `P` is stable under arbitrary base change. Co-authored-by: Andrew Yang <[email protected]> Co-authored-by: Joël Riou <[email protected]>
1 parent db141d8 commit 65d6572

File tree

13 files changed

+666
-881
lines changed

13 files changed

+666
-881
lines changed

Mathlib/Algebra/Category/Ring/Basic.lean

+3
Original file line numberDiff line numberDiff line change
@@ -501,6 +501,9 @@ lemma coe_id {X : CommRingCat} : (𝟙 X : X → X) = id := rfl
501501

502502
lemma coe_comp {X Y Z : CommRingCat} {f : X ⟶ Y} {g : Y ⟶ Z} : (f ≫ g : X → Z) = g ∘ f := rfl
503503

504+
/-- Specialization of `ConcreteCategory.id_apply` because `simp` can't see through the defeq. -/
505+
@[simp] lemma id_apply (R : CommRingCat) (x : R) : 𝟙 R x = x := rfl
506+
504507
@[simp] lemma forget_map {X Y : CommRingCat} (f : X ⟶ Y) :
505508
(forget CommRingCat).map f = (f : X → Y) := rfl
506509

Mathlib/AlgebraicGeometry/AffineScheme.lean

+13-1
Original file line numberDiff line numberDiff line change
@@ -225,10 +225,13 @@ instance Scheme.isAffine_affineBasisCover (X : Scheme) (i : X.affineBasisCover.J
225225
isAffine_Spec _
226226
#align algebraic_geometry.Scheme.affine_basis_cover_is_affine AlgebraicGeometry.Scheme.isAffine_affineBasisCover
227227

228-
instance Scheme.isAffine_affineOpenCover (X : Scheme) (𝒰 : X.AffineOpenCover) (i: 𝒰.J) :
228+
instance Scheme.isAffine_affineOpenCover (X : Scheme) (𝒰 : X.AffineOpenCover) (i : 𝒰.J) :
229229
IsAffine (𝒰.openCover.obj i) :=
230230
inferInstanceAs (IsAffine (Spec (𝒰.obj i)))
231231

232+
instance {X} [IsAffine X] (i) : IsAffine ((Scheme.openCoverOfIsIso (𝟙 X)).obj i) := by
233+
dsimp; infer_instance
234+
232235
theorem isBasis_affine_open (X : Scheme) : Opens.IsBasis X.affineOpens := by
233236
rw [Opens.isBasis_iff_nbhd]
234237
rintro U x (hU : x ∈ (U : Set X))
@@ -325,6 +328,12 @@ theorem image_of_isOpenImmersion (f : X ⟶ Y) [H : IsOpenImmersion f] :
325328
exact Set.image_eq_range _ _
326329
#align algebraic_geometry.is_affine_open.image_is_open_immersion AlgebraicGeometry.IsAffineOpen.image_of_isOpenImmersion
327330

331+
theorem preimage_of_isIso {U : Opens Y.carrier} (hU : IsAffineOpen U) (f : X ⟶ Y) [IsIso f] :
332+
IsAffineOpen (f ⁻¹ᵁ U) :=
333+
haveI : IsAffine _ := hU
334+
isAffine_of_isIso (f ∣_ U)
335+
#align algebraic_geometry.is_affine_open.map_is_iso AlgebraicGeometry.IsAffineOpen.preimage_of_isIso
336+
328337
theorem _root_.AlgebraicGeometry.Scheme.Hom.isAffineOpen_iff_of_isOpenImmersion
329338
(f : AlgebraicGeometry.Scheme.Hom X Y) [H : IsOpenImmersion f] {U : Opens X} :
330339
IsAffineOpen (f ''ᵁ U) ↔ IsAffineOpen U := by
@@ -435,6 +444,9 @@ theorem basicOpen :
435444
exact Opens.ext (PrimeSpectrum.localization_away_comap_range (Localization.Away f) f).symm
436445
#align algebraic_geometry.is_affine_open.basic_open_is_affine AlgebraicGeometry.IsAffineOpen.basicOpen
437446

447+
instance [IsAffine X] (r : Γ(X, ⊤)) : IsAffine (X ∣_ᵤ X.basicOpen r) :=
448+
(isAffineOpen_top X).basicOpen _
449+
438450
theorem ιOpens_basicOpen_preimage (r : Γ(X, ⊤)) :
439451
IsAffineOpen (Scheme.ιOpens (X.basicOpen r) ⁻¹ᵁ U) := by
440452
apply (Scheme.ιOpens (X.basicOpen r)).isAffineOpen_iff_of_isOpenImmersion.mp

Mathlib/AlgebraicGeometry/Cover/Open.lean

+44
Original file line numberDiff line numberDiff line change
@@ -192,6 +192,15 @@ def OpenCover.pullbackCover {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ X
192192
· rw [← TopCat.epi_iff_surjective]; infer_instance
193193
#align algebraic_geometry.Scheme.open_cover.pullback_cover AlgebraicGeometry.Scheme.OpenCover.pullbackCover
194194

195+
/-- The family of morphisms from the pullback cover to the original cover. -/
196+
def OpenCover.pullbackHom {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ X) (i) :
197+
(𝒰.pullbackCover f).obj i ⟶ 𝒰.obj i :=
198+
pullback.snd f (𝒰.map i)
199+
200+
@[reassoc (attr := simp)]
201+
lemma OpenCover.pullbackHom_map {X W : Scheme.{u}} (𝒰 : X.OpenCover) (f : W ⟶ X) (i) :
202+
𝒰.pullbackHom f i ≫ 𝒰.map i = (𝒰.pullbackCover f).map i ≫ f := pullback.condition.symm
203+
195204
/-- Given an open cover on `X`, we may pull them back along a morphism `f : W ⟶ X` to obtain
196205
an open cover of `W`. This is similar to `Scheme.OpenCover.pullbackCover`, but here we
197206
take `pullback (𝒰.map x) f` instead of `pullback f (𝒰.map x)`. -/
@@ -337,6 +346,41 @@ def OpenCover.affineRefinement {X : Scheme.{u}} (𝓤 : X.OpenCover) : X.AffineO
337346
f := (𝓤.bind fun j => (𝓤.obj j).affineCover).f
338347
covers := (𝓤.bind fun j => (𝓤.obj j).affineCover).covers
339348

349+
/-- The pullback of the affine refinement is the pullback of the affine cover. -/
350+
def OpenCover.pullbackCoverAffineRefinementObjIso (f : X ⟶ Y) (𝒰 : Y.OpenCover) (i) :
351+
(𝒰.affineRefinement.openCover.pullbackCover f).obj i ≅
352+
((𝒰.obj i.1).affineCover.pullbackCover (𝒰.pullbackHom f i.1)).obj i.2 :=
353+
pullbackSymmetry _ _ ≪≫ (pullbackRightPullbackFstIso _ _ _).symm ≪≫
354+
pullbackSymmetry _ _ ≪≫ asIso (pullback.map _ _ _ _ (pullbackSymmetry _ _).hom (𝟙 _) (𝟙 _)
355+
(by simp [pullbackHom]) (by simp))
356+
357+
@[reassoc]
358+
lemma OpenCover.pullbackCoverAffineRefinementObjIso_inv_map (f : X ⟶ Y) (𝒰 : Y.OpenCover) (i) :
359+
(𝒰.pullbackCoverAffineRefinementObjIso f i).inv ≫
360+
(𝒰.affineRefinement.openCover.pullbackCover f).map i =
361+
((𝒰.obj i.1).affineCover.pullbackCover (𝒰.pullbackHom f i.1)).map i.2
362+
(𝒰.pullbackCover f).map i.1 := by
363+
simp only [pullbackCover_obj, AffineOpenCover.openCover_obj, AffineOpenCover.openCover_map,
364+
pullbackCoverAffineRefinementObjIso, Iso.trans_inv, asIso_inv, Iso.symm_inv, Category.assoc,
365+
pullbackCover_map, pullbackSymmetry_inv_comp_fst, IsIso.inv_comp_eq, limit.lift_π_assoc, id_eq,
366+
PullbackCone.mk_pt, cospan_left, PullbackCone.mk_π_app, pullbackSymmetry_hom_comp_fst]
367+
convert pullbackSymmetry_inv_comp_snd_assoc
368+
((𝒰.obj i.1).affineCover.map i.2) (pullback.fst _ _) _ using 2
369+
exact pullbackRightPullbackFstIso_hom_snd _ _ _
370+
371+
@[reassoc]
372+
lemma OpenCover.pullbackCoverAffineRefinementObjIso_inv_pullbackHom
373+
(f : X ⟶ Y) (𝒰 : Y.OpenCover) (i) :
374+
(𝒰.pullbackCoverAffineRefinementObjIso f i).inv ≫
375+
𝒰.affineRefinement.openCover.pullbackHom f i =
376+
(𝒰.obj i.1).affineCover.pullbackHom (𝒰.pullbackHom f i.1) i.2 := by
377+
simp only [pullbackCover_obj, pullbackHom, AffineOpenCover.openCover_obj,
378+
AffineOpenCover.openCover_map, pullbackCoverAffineRefinementObjIso, Iso.trans_inv, asIso_inv,
379+
Iso.symm_inv, Category.assoc, pullbackSymmetry_inv_comp_snd, IsIso.inv_comp_eq, limit.lift_π,
380+
id_eq, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.comp_id]
381+
convert pullbackSymmetry_inv_comp_fst ((𝒰.obj i.1).affineCover.map i.2) (pullback.fst _ _)
382+
exact pullbackRightPullbackFstIso_hom_fst _ _ _
383+
340384
section category
341385

342386
/--

0 commit comments

Comments
 (0)