Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 36ba1ac

Browse files
erdOnejcommelin
andcommitted
feat(algebraic_geometry): Define open_covers of Schemes. (#10931)
Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent c374a3b commit 36ba1ac

File tree

4 files changed

+183
-2
lines changed

4 files changed

+183
-2
lines changed
Lines changed: 21 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,21 @@
1+
/-
2+
Copyright (c) 2021 Andrew Yang. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Andrew Yang
5+
-/
6+
import algebra.category.CommRing.basic
7+
import ring_theory.localization
8+
9+
/-!
10+
# Ring-theoretic results in terms of categorical languages
11+
-/
12+
13+
open category_theory
14+
15+
instance localization_unit_is_iso (R : CommRing) :
16+
is_iso (CommRing.of_hom $ algebra_map R (localization.away (1 : R))) :=
17+
is_iso.of_iso (is_localization.at_one R (localization.away (1 : R))).to_ring_equiv.to_CommRing_iso
18+
19+
instance localization_unit_is_iso' (R : CommRing) :
20+
@is_iso CommRing _ R _ (CommRing.of_hom $ algebra_map R (localization.away (1 : R))) :=
21+
by { cases R, exact localization_unit_is_iso _ }

src/algebraic_geometry/Scheme.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -54,6 +54,16 @@ induced_category.category Scheme.to_LocallyRingedSpace
5454
/-- The structure sheaf of a Scheme. -/
5555
protected abbreviation sheaf (X : Scheme) := X.to_SheafedSpace.sheaf
5656

57+
/-- The forgetful functor from `Scheme` to `LocallyRingedSpace`. -/
58+
@[simps, derive[full, faithful]]
59+
def forget_to_LocallyRingedSpace : Scheme ⥤ LocallyRingedSpace :=
60+
induced_functor _
61+
62+
/-- The forgetful functor from `Scheme` to `Top`. -/
63+
@[simps]
64+
def forget_to_Top : Scheme ⥤ Top :=
65+
Scheme.forget_to_LocallyRingedSpace ⋙ LocallyRingedSpace.forget_to_Top
66+
5767
/--
5868
The spectrum of a commutative ring, as a scheme.
5969
-/

src/algebraic_geometry/locally_ringed_space.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -121,12 +121,17 @@ instance : category LocallyRingedSpace :=
121121
assoc' := by { intros, ext1, simp, }, }.
122122

123123
/-- The forgetful functor from `LocallyRingedSpace` to `SheafedSpace CommRing`. -/
124-
def forget_to_SheafedSpace : LocallyRingedSpace ⥤ SheafedSpace CommRing :=
124+
@[simps] def forget_to_SheafedSpace : LocallyRingedSpace ⥤ SheafedSpace CommRing :=
125125
{ obj := λ X, X.to_SheafedSpace,
126126
map := λ X Y f, f.1, }
127127

128128
instance : faithful forget_to_SheafedSpace := {}
129129

130+
/-- The forgetful functor from `LocallyRingedSpace` to `Top`. -/
131+
@[simps]
132+
def forget_to_Top : LocallyRingedSpace ⥤ Top :=
133+
forget_to_SheafedSpace ⋙ SheafedSpace.forget _
134+
130135
@[simp] lemma comp_val {X Y Z : LocallyRingedSpace} (f : X ⟶ Y) (g : Y ⟶ Z) :
131136
(f ≫ g).val = f.val ≫ g.val := rfl
132137

src/algebraic_geometry/open_immersion.lean

Lines changed: 146 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,6 +9,7 @@ import category_theory.limits.preserves.shapes.pullbacks
99
import topology.sheaves.functors
1010
import algebraic_geometry.Scheme
1111
import category_theory.limits.shapes.strict_initial
12+
import algebra.category.CommRing.instances
1213

1314
/-!
1415
# Open immersions of structured spaces
@@ -974,7 +975,7 @@ begin
974975
apply_with limits.comp_preserves_limit { instances := ff },
975976
apply_instance,
976977
apply preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{u} _).symm,
977-
dsimp,
978+
dsimp [SheafedSpace.forget_to_PresheafedSpace, -subtype.val_eq_coe],
978979
apply_instance,
979980
end
980981

@@ -1062,4 +1063,148 @@ end pullback
10621063

10631064
end LocallyRingedSpace.is_open_immersion
10641065

1066+
lemma is_open_immersion.open_range {X Y : Scheme} (f : X ⟶ Y) [H : is_open_immersion f] :
1067+
is_open (set.range f.1.base) := H.base_open.open_range
1068+
1069+
section open_cover
1070+
1071+
namespace Scheme
1072+
1073+
/-- An open cover of `X` consists of a family of open immersions into `X`,
1074+
and for each `x : X` an open immersion (indexed by `f x`) that covers `x`.
1075+
1076+
This is merely a coverage in the Zariski pretopology, and it would be optimal
1077+
if we could reuse the existing API about pretopologies, However, the definitions of sieves and
1078+
grothendieck topologies uses `Prop`s, so that the actual open sets and immersions are hard to
1079+
obtain. Also, since such a coverage in the pretopology usually contains a proper class of
1080+
immersions, it is quite hard to glue them, reason about finite covers, etc.
1081+
-/
1082+
-- TODO: provide API to and from a presieve.
1083+
structure open_cover (X : Scheme.{u}) :=
1084+
(J : Type v)
1085+
(obj : Π (j : J), Scheme)
1086+
(map : Π (j : J), obj j ⟶ X)
1087+
(f : X.carrier → J)
1088+
(covers : ∀ x, x ∈ set.range ((map (f x)).1.base))
1089+
(is_open : ∀ x, is_open_immersion (map x) . tactic.apply_instance)
1090+
1091+
attribute [instance] open_cover.is_open
1092+
1093+
variables {X Y Z : Scheme.{u}} (𝒰 : open_cover X) (f : X ⟶ Z) (g : Y ⟶ Z)
1094+
variables [∀ x, has_pullback (𝒰.map x ≫ f) g]
1095+
1096+
/-- The affine cover of a scheme. -/
1097+
def affine_cover (X : Scheme) : open_cover X :=
1098+
{ J := X.carrier,
1099+
obj := λ x, Spec.obj $ opposite.op (X.local_affine x).some_spec.some,
1100+
map := λ x, ((X.local_affine x).some_spec.some_spec.some.inv ≫
1101+
X.to_LocallyRingedSpace.of_restrict _ : _),
1102+
f := λ x, x,
1103+
is_open := λ x, begin
1104+
apply_with PresheafedSpace.is_open_immersion.comp { instances := ff },
1105+
apply_instance,
1106+
apply PresheafedSpace.is_open_immersion.of_restrict,
1107+
end,
1108+
covers :=
1109+
begin
1110+
intro x,
1111+
erw coe_comp,
1112+
rw [set.range_comp, set.range_iff_surjective.mpr, set.image_univ],
1113+
erw subtype.range_coe_subtype,
1114+
exact (X.local_affine x).some.2,
1115+
rw ← Top.epi_iff_surjective,
1116+
change epi ((SheafedSpace.forget _).map (LocallyRingedSpace.forget_to_SheafedSpace.map _)),
1117+
apply_instance
1118+
end }
1119+
1120+
instance : inhabited X.open_cover := ⟨X.affine_cover⟩
1121+
1122+
/-- Given an open cover `{ Uᵢ }` of `X`, and for each `Uᵢ` an open cover, we may combine these
1123+
open covers to form an open cover of `X`. -/
1124+
def open_cover.bind (f : Π (x : 𝒰.J), open_cover (𝒰.obj x)) : open_cover X :=
1125+
{ J := Σ (i : 𝒰.J), (f i).J,
1126+
obj := λ x, (f x.1).obj x.2,
1127+
map := λ x, (f x.1).map x.2 ≫ 𝒰.map x.1,
1128+
f := λ x, ⟨_, (f _).f (𝒰.covers x).some⟩,
1129+
covers := λ x,
1130+
begin
1131+
let y := (𝒰.covers x).some,
1132+
have hy : (𝒰.map (𝒰.f x)).val.base y = x := (𝒰.covers x).some_spec,
1133+
rcases (f (𝒰.f x)).covers y with ⟨z, hz⟩,
1134+
change x ∈ set.range (((f (𝒰.f x)).map ((f (𝒰.f x)).f y) ≫ 𝒰.map (𝒰.f x)).1.base),
1135+
use z,
1136+
erw comp_apply,
1137+
rw [hz, hy],
1138+
end }
1139+
1140+
local attribute [reducible] CommRing.of CommRing.of_hom
1141+
1142+
instance val_base_is_iso {X Y : Scheme} (f : X ⟶ Y) [is_iso f] : is_iso f.1.base :=
1143+
Scheme.forget_to_Top.map_is_iso f
1144+
1145+
instance basic_open_is_open_immersion {R : CommRing} (f : R) :
1146+
algebraic_geometry.is_open_immersion (Scheme.Spec.map (CommRing.of_hom
1147+
(algebra_map R (localization.away f))).op) :=
1148+
begin
1149+
apply_with SheafedSpace.is_open_immersion.of_stalk_iso { instances := ff },
1150+
any_goals { apply_instance },
1151+
any_goals { apply_instance },
1152+
exact (prime_spectrum.localization_away_open_embedding (localization.away f) f : _),
1153+
intro x,
1154+
exact Spec_map_localization_is_iso R (submonoid.powers f) x,
1155+
end
1156+
1157+
/-- The basic open sets form an affine open cover of `Spec R`. -/
1158+
def affine_basis_cover_of_affine (R : CommRing) : open_cover (Spec.obj (opposite.op R)) :=
1159+
{ J := R,
1160+
obj := λ r, Spec.obj (opposite.op $ CommRing.of $ localization.away r),
1161+
map := λ r, Spec.map (quiver.hom.op (algebra_map R (localization.away r) : _)),
1162+
f := λ x, 1,
1163+
covers := λ r,
1164+
begin
1165+
rw set.range_iff_surjective.mpr ((Top.epi_iff_surjective _).mp _),
1166+
{ exact trivial },
1167+
{ apply_instance }
1168+
end,
1169+
is_open := λ x, algebraic_geometry.Scheme.basic_open_is_open_immersion x }
1170+
1171+
/-- We may bind the basic open sets of an open affine cover to form a affine cover that is also
1172+
a basis. -/
1173+
def affine_basis_cover (X : Scheme) : open_cover X :=
1174+
X.affine_cover.bind (λ x, affine_basis_cover_of_affine _)
1175+
1176+
lemma affine_basis_cover_map_range (X : Scheme)
1177+
(x : X.carrier) (r : (X.local_affine x).some_spec.some) :
1178+
set.range (X.affine_basis_cover.map ⟨x, r⟩).1.base =
1179+
(X.affine_cover.map x).1.base '' (prime_spectrum.basic_open r).1 :=
1180+
begin
1181+
erw [coe_comp, set.range_comp],
1182+
congr,
1183+
exact (prime_spectrum.localization_away_comap_range (localization.away r) r : _)
1184+
end
1185+
1186+
lemma affine_basis_cover_is_basis (X : Scheme) :
1187+
topological_space.is_topological_basis
1188+
{ x : set X.carrier | ∃ a : X.affine_basis_cover.J, x =
1189+
set.range ((X.affine_basis_cover.map a).1.base) } :=
1190+
begin
1191+
apply topological_space.is_topological_basis_of_open_of_nhds,
1192+
{ rintros _ ⟨a, rfl⟩,
1193+
exact is_open_immersion.open_range (X.affine_basis_cover.map a) },
1194+
{ rintros a U haU hU,
1195+
rcases X.affine_cover.covers a with ⟨x, e⟩,
1196+
let U' := (X.affine_cover.map (X.affine_cover.f a)).1.base ⁻¹' U,
1197+
have hxU' : x ∈ U' := by { rw ← e at haU, exact haU },
1198+
rcases prime_spectrum.is_basis_basic_opens.exists_subset_of_mem_open hxU'
1199+
((X.affine_cover.map (X.affine_cover.f a)).1.base.continuous_to_fun.is_open_preimage _ hU)
1200+
with ⟨_,⟨_,⟨s,rfl⟩,rfl⟩,hxV,hVU⟩,
1201+
refine ⟨_,⟨⟨_,s⟩,rfl⟩,_,_⟩; erw affine_basis_cover_map_range,
1202+
{ exact ⟨x,hxV,e⟩ },
1203+
{ rw set.image_subset_iff, exact hVU } }
1204+
end
1205+
1206+
end Scheme
1207+
1208+
end open_cover
1209+
10651210
end algebraic_geometry

0 commit comments

Comments
 (0)