@@ -84,6 +84,15 @@ def is_fraction {U : opens (prime_spectrum.Top R)} (f : Π x : U, localizations
84
84
∃ (r s : R), ∀ x : U,
85
85
¬ (s ∈ x.1 .as_ideal) ∧ f x * (localization.of _).to_map s = (localization.of _).to_map r
86
86
87
+ lemma is_fraction.eq_mk' {U : opens (prime_spectrum.Top R)} {f : Π x : U, localizations R x}
88
+ (hf : is_fraction f) :
89
+ ∃ (r s : R) , ∀ x : U, ∃ (hs : s ∉ x.1 .as_ideal), f x = (localization.of _).mk' r ⟨s, hs⟩ :=
90
+ begin
91
+ rcases hf with ⟨r, s, h⟩,
92
+ refine ⟨r, s, λ x, ⟨(h x).1 , ((localization_map.mk'_eq_iff_eq_mul _).mpr _).symm⟩⟩,
93
+ exact (h x).2 .symm,
94
+ end
95
+
87
96
variables (R)
88
97
89
98
/--
@@ -357,7 +366,7 @@ by rw [mul_comm, const_mul_cancel]
357
366
358
367
/-- The canonical ring homomorphism interpreting an element of `R` as
359
368
a section of the structure sheaf. -/
360
- def to_open (U : opens (prime_spectrum.Top R)) :
369
+ @[simps] def to_open (U : opens (prime_spectrum.Top R)) :
361
370
CommRing.of R ⟶ (structure_sheaf R).presheaf.obj (op U) :=
362
371
{ to_fun := λ f, ⟨λ x, (localization.of _).to_map f,
363
372
λ x, ⟨U, x.2 , 𝟙 _, f, 1 , λ y, ⟨(ideal.ne_top_iff_one _).1 y.1 .2 .1 ,
@@ -371,10 +380,6 @@ def to_open (U : opens (prime_spectrum.Top R)) :
371
380
to_open R U ≫ (structure_sheaf R).presheaf.map i.op = to_open R V :=
372
381
rfl
373
382
374
- @[simp] lemma to_open_apply (U : opens (prime_spectrum.Top R)) (f : R) (x : U) :
375
- (to_open R U f).1 x = (localization.of _).to_map f :=
376
- rfl
377
-
378
383
lemma to_open_eq_const (U : opens (prime_spectrum.Top R)) (f : R) : to_open R U f =
379
384
const R f 1 U (λ x _, (ideal.ne_top_iff_one _).1 x.2 .1 ) :=
380
385
subtype.eq $ funext $ λ x, eq.symm $ (localization.of _).mk'_one f
@@ -487,7 +492,7 @@ ring_hom.ext_iff.1 (to_stalk_comp_stalk_to_fiber_ring_hom R x) _
487
492
488
493
/-- The ring isomorphism between the stalk of the structure sheaf of `R` at a point `p`
489
494
corresponding to a prime ideal in `R` and the localization of `R` at `p`. -/
490
- def stalk_iso (x : prime_spectrum.Top R) :
495
+ @[simps] def stalk_iso (x : prime_spectrum.Top R) :
491
496
(structure_sheaf R).presheaf.stalk x ≅ CommRing.of (localization.at_prime x.as_ideal) :=
492
497
{ hom := stalk_to_fiber_ring_hom R x,
493
498
inv := localization_to_stalk R x,
@@ -777,4 +782,161 @@ def basic_open_iso (f : R) : (structure_sheaf R).presheaf.obj (op (basic_open f)
777
782
CommRing.of (localization.away f) :=
778
783
(as_iso (to_basic_open R f)).symm
779
784
785
+ section comap
786
+
787
+ variables {R} {S : Type u} [comm_ring S] {P : Type u} [comm_ring P]
788
+
789
+ /--
790
+ Given a ring homomorphism `f : R →+* S`, an open set `U` of the prime spectrum of `R` and an open
791
+ set `V` of the prime spectrum of `S`, such that `V ⊆ (comap f) ⁻¹' U`, we can push a section `s`
792
+ on `U` to a section on `V`, by composing with `localization.local_ring_hom _ _ f` from the left and
793
+ `comap f` from the right. Explicitly, if `s` evaluates on `comap f p` to `a / b`, its image on `V`
794
+ evaluates on `p` to `f(a) / f(b)`.
795
+
796
+ At the moment, we work with arbitrary dependent functions `s : Π x : U, localizations R x`. Below,
797
+ we prove the predicate `is_locally_fraction` is preserved by this map, hence it can be extended to
798
+ a morphism between the structure sheaves of `R` and `S`.
799
+ -/
800
+ def structure_sheaf.comap_fun (f : R →+* S) (U : opens (prime_spectrum.Top R))
801
+ (V : opens (prime_spectrum.Top S)) (hUV : V.1 ⊆ (comap f) ⁻¹' U.1 )
802
+ (s : Π x : U, localizations R x) (y : V) : localizations S y :=
803
+ localization.local_ring_hom (comap f y.1 ).as_ideal _ f rfl (s ⟨(comap f y.1 ), hUV y.2 ⟩ : _)
804
+
805
+ lemma structure_sheaf.comap_fun_is_locally_fraction (f : R →+* S)
806
+ (U : opens (prime_spectrum.Top R)) (V : opens (prime_spectrum.Top S))
807
+ (hUV : V.1 ⊆ (comap f) ⁻¹' U.1 ) (s : Π x : U, localizations R x)
808
+ (hs : (is_locally_fraction R).to_prelocal_predicate.pred s) :
809
+ (is_locally_fraction S).to_prelocal_predicate.pred (structure_sheaf.comap_fun f U V hUV s) :=
810
+ begin
811
+ rintro ⟨p, hpV⟩,
812
+ -- Since `s` is locally fraction, we can find a neighborhood `W` of `comap f p` in `U`, such
813
+ -- that `s = a / b` on `W`, for some ring elements `a, b : R`.
814
+ rcases hs ⟨comap f p, hUV hpV⟩ with ⟨W, m, iWU, a, b, h_frac⟩,
815
+ -- We claim that we can write our new section as the fraction `f a / f b` on the neighborhood
816
+ -- `(comap f) ⁻¹ W ⊓ V` of `p`.
817
+ refine ⟨opens.comap (comap_continuous f) W ⊓ V, ⟨m, hpV⟩, opens.inf_le_right _ _, f a, f b, _⟩,
818
+ rintro ⟨q, ⟨hqW, hqV⟩⟩,
819
+ specialize h_frac ⟨prime_spectrum.comap f q, hqW⟩,
820
+ refine ⟨h_frac.1 , _⟩,
821
+ dsimp only [structure_sheaf.comap_fun],
822
+ erw [← localization.local_ring_hom_to_map ((comap f q).as_ideal), ← ring_hom.map_mul,
823
+ h_frac.2 , localization.local_ring_hom_to_map],
824
+ refl,
825
+ end
826
+
827
+ /--
828
+ For a ring homomorphism `f : R →+* S` and open sets `U` and `V` of the prime spectra of `R` and
829
+ `S` such that `V ⊆ (comap f) ⁻¹ U`, the induced ring homomorphism from the structure sheaf of `R`
830
+ at `U` to the structure sheaf of `S` at `V`.
831
+
832
+ Explicitly, this map is given as follows: For a point `p : V`, if the section `s` evaluates on `p`
833
+ to the fraction `a / b`, its image on `V` evaluates on `p` to the fraction `f(a) / f(b)`.
834
+ -/
835
+ def structure_sheaf.comap (f : R →+* S) (U : opens (prime_spectrum.Top R))
836
+ (V : opens (prime_spectrum.Top S)) (hUV : V.1 ⊆ (comap f) ⁻¹' U.1 ) :
837
+ (structure_sheaf R).presheaf.obj (op U) →+* (structure_sheaf S).presheaf.obj (op V) :=
838
+ { to_fun := λ s, ⟨structure_sheaf.comap_fun f U V hUV s.1 ,
839
+ structure_sheaf.comap_fun_is_locally_fraction f U V hUV s.1 s.2 ⟩,
840
+ map_one' := subtype.ext $ funext $ λ p, by
841
+ { rw [subtype.coe_mk, subtype.val_eq_coe, structure_sheaf.comap_fun,
842
+ (sections_subring R (op U)).coe_one, pi.one_apply, ring_hom.map_one], refl },
843
+ map_zero' := subtype.ext $ funext $ λ p, by
844
+ { rw [subtype.coe_mk, subtype.val_eq_coe, structure_sheaf.comap_fun,
845
+ (sections_subring R (op U)).coe_zero, pi.zero_apply, ring_hom.map_zero], refl },
846
+ map_add' := λ s t, subtype.ext $ funext $ λ p, by
847
+ { rw [subtype.coe_mk, subtype.val_eq_coe, structure_sheaf.comap_fun,
848
+ (sections_subring R (op U)).coe_add, pi.add_apply, ring_hom.map_add], refl },
849
+ map_mul' := λ s t, subtype.ext $ funext $ λ p, by
850
+ { rw [subtype.coe_mk, subtype.val_eq_coe, structure_sheaf.comap_fun,
851
+ (sections_subring R (op U)).coe_mul, pi.mul_apply, ring_hom.map_mul], refl }
852
+ }
853
+
854
+ @[simp]
855
+ lemma structure_sheaf.comap_apply (f : R →+* S) (U : opens (prime_spectrum.Top R))
856
+ (V : opens (prime_spectrum.Top S)) (hUV : V.1 ⊆ (comap f) ⁻¹' U.1 )
857
+ (s : (structure_sheaf R).presheaf.obj (op U)) (p : V) :
858
+ (structure_sheaf.comap f U V hUV s).1 p =
859
+ localization.local_ring_hom (comap f p.1 ).as_ideal _ f rfl (s.1 ⟨(comap f p.1 ), hUV p.2 ⟩ : _) :=
860
+ rfl
861
+
862
+ lemma structure_sheaf.comap_const (f : R →+* S) (U : opens (prime_spectrum.Top R))
863
+ (V : opens (prime_spectrum.Top S)) (hUV : V.1 ⊆ (comap f) ⁻¹' U.1 )
864
+ (a b : R) (hb : ∀ x : prime_spectrum R, x ∈ U → b ∈ x.as_ideal.prime_compl) :
865
+ structure_sheaf.comap f U V hUV (const R a b U hb) =
866
+ const S (f a) (f b) V (λ p hpV, hb (comap f p) (hUV hpV)) :=
867
+ subtype.eq $ funext $ λ p,
868
+ begin
869
+ rw [structure_sheaf.comap_apply, const_apply, const_apply],
870
+ erw localization.local_ring_hom_mk',
871
+ refl,
872
+ end
873
+
874
+ /--
875
+ For an inclusion `i : V ⟶ U` between open sets of the prime spectrum of `R`, the comap of the
876
+ identity from OO_X(U) to OO_X(V) equals as the restriction map of the structure sheaf.
877
+
878
+ This is a generalization of the fact that, for fixed `U`, the comap of the identity from OO_X(U)
879
+ to OO_X(U) is the identity.
880
+ -/
881
+ lemma structure_sheaf.comap_id_eq_map (U V : opens (prime_spectrum.Top R)) (iVU : V ⟶ U) :
882
+ structure_sheaf.comap (ring_hom.id R) U V
883
+ (λ p hpV, le_of_hom iVU $ by rwa prime_spectrum.comap_id) =
884
+ (structure_sheaf R).presheaf.map iVU.op :=
885
+ ring_hom.ext $ λ s, subtype.eq $ funext $ λ p,
886
+ begin
887
+ rw structure_sheaf.comap_apply,
888
+ -- Unfortunately, we cannot use `localization.local_ring_hom_id` here, because
889
+ -- `comap (ring_hom.id R) p` is not *definitionally* equal to `p`. Instead, we use that we can
890
+ -- write `s` as a fraction `a/b` in a small neighborhood around `p`. Since
891
+ -- `comap (ring_hom.id R) p` equals `p`, it is also contained in the same neighborhood, hence
892
+ -- `s` equals `a/b` there too.
893
+ obtain ⟨W, hpW, iWU, h⟩ := s.2 (iVU p),
894
+ obtain ⟨a, b, h'⟩ := h.eq_mk',
895
+ obtain ⟨hb₁, s_eq₁⟩ := h' ⟨p, hpW⟩,
896
+ obtain ⟨hb₂, s_eq₂⟩ := h' ⟨comap (ring_hom.id _) p.1 , by rwa prime_spectrum.comap_id⟩,
897
+ dsimp only at s_eq₁ s_eq₂,
898
+ erw [s_eq₂, localization.local_ring_hom_mk', ← s_eq₁, ← res_apply],
899
+ end
900
+
901
+ /--
902
+ The comap of the identity is the identity. In this variant of the lemma, two open subsets `U` and
903
+ `V` are given as arguments, together with a proof that `U = V`. This is be useful when `U` and `V`
904
+ are not definitionally equal.
905
+ -/
906
+ lemma structure_sheaf.comap_id (U V : opens (prime_spectrum.Top R)) (hUV : U = V) :
907
+ structure_sheaf.comap (ring_hom.id R) U V
908
+ (λ p hpV, by rwa [hUV, prime_spectrum.comap_id]) =
909
+ eq_to_hom (show (structure_sheaf R).presheaf.obj (op U) = _, by rw hUV) :=
910
+ by erw [structure_sheaf.comap_id_eq_map U V (eq_to_hom hUV.symm), eq_to_hom_op, eq_to_hom_map]
911
+
912
+ @[simp] lemma structure_sheaf.comap_id' (U : opens (prime_spectrum.Top R)) :
913
+ structure_sheaf.comap (ring_hom.id R) U U (λ p hpU, by rwa prime_spectrum.comap_id) =
914
+ ring_hom.id _ :=
915
+ by { rw structure_sheaf.comap_id U U rfl, refl }
916
+
917
+ lemma structure_sheaf.comap_comp (f : R →+* S) (g : S →+* P) (U : opens (prime_spectrum.Top R))
918
+ (V : opens (prime_spectrum.Top S)) (W : opens (prime_spectrum.Top P))
919
+ (hUV : ∀ p ∈ V, comap f p ∈ U) (hVW : ∀ p ∈ W, comap g p ∈ V) :
920
+ structure_sheaf.comap (g.comp f) U W (λ p hpW, hUV (comap g p) (hVW p hpW)) =
921
+ (structure_sheaf.comap g V W hVW).comp (structure_sheaf.comap f U V hUV) :=
922
+ ring_hom.ext $ λ s, subtype.eq $ funext $ λ p,
923
+ begin
924
+ rw structure_sheaf.comap_apply,
925
+ erw localization.local_ring_hom_comp _ (comap g p.1 ).as_ideal,
926
+ -- refl works here, because `comap (g.comp f) p` is defeq to `comap f (comap g p)`
927
+ refl,
928
+ end
929
+
930
+ @[elementwise, reassoc] lemma to_open_comp_comap (f : R →+* S) :
931
+ to_open R ⊤ ≫ structure_sheaf.comap f ⊤ ⊤ (λ p hpV, trivial) =
932
+ @category_theory.category_struct.comp _ _ (CommRing.of R) (CommRing.of S) _ f (to_open S ⊤) :=
933
+ ring_hom.ext $ λ s, subtype.eq $ funext $ λ p,
934
+ begin
935
+ simp_rw [comp_apply, structure_sheaf.comap_apply, subtype.val_eq_coe, to_open_apply_coe],
936
+ erw localization.local_ring_hom_to_map,
937
+ refl,
938
+ end
939
+
940
+ end comap
941
+
780
942
end algebraic_geometry
0 commit comments