@@ -9,6 +9,7 @@ import category_theory.limits.preserves.shapes.pullbacks
9
9
import topology.sheaves.functors
10
10
import algebraic_geometry.Scheme
11
11
import category_theory.limits.shapes.strict_initial
12
+ import algebra.category.CommRing.instances
12
13
13
14
/-!
14
15
# Open immersions of structured spaces
@@ -974,7 +975,7 @@ begin
974
975
apply_with limits.comp_preserves_limit { instances := ff },
975
976
apply_instance,
976
977
apply preserves_limit_of_iso_diagram _ (diagram_iso_cospan.{u} _).symm,
977
- dsimp,
978
+ dsimp [SheafedSpace.forget_to_PresheafedSpace, -subtype.val_eq_coe] ,
978
979
apply_instance,
979
980
end
980
981
@@ -1062,4 +1063,148 @@ end pullback
1062
1063
1063
1064
end LocallyRingedSpace.is_open_immersion
1064
1065
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
+
1065
1210
end algebraic_geometry
0 commit comments