Skip to content

Commit 6d4ab88

Browse files
committed
feat: relative differentials as a presheaf of modules (#14014)
In this PR, we define the type `M.Derivation φ` of derivations with values in a presheaf of `R`-modules `M` relative to a morphism of `φ : S ⟶ F.op ⋙ R` of presheaves of commutative rings over categories `C` and `D` that are related by a functor `F : C ⥤ D`. In the particular case `F` is the identity functor, we construct the universal derivation. Co-authored-by: Joël Riou <[email protected]>
1 parent c4f1cc8 commit 6d4ab88

File tree

3 files changed

+235
-2
lines changed

3 files changed

+235
-2
lines changed

Mathlib.lean

+1
Original file line numberDiff line numberDiff line change
@@ -83,6 +83,7 @@ import Mathlib.Algebra.Category.ModuleCat.Biproducts
8383
import Mathlib.Algebra.Category.ModuleCat.ChangeOfRings
8484
import Mathlib.Algebra.Category.ModuleCat.Colimits
8585
import Mathlib.Algebra.Category.ModuleCat.Differentials.Basic
86+
import Mathlib.Algebra.Category.ModuleCat.Differentials.Presheaf
8687
import Mathlib.Algebra.Category.ModuleCat.EpiMono
8788
import Mathlib.Algebra.Category.ModuleCat.FilteredColimits
8889
import Mathlib.Algebra.Category.ModuleCat.Free
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,231 @@
1+
/-
2+
Copyright (c) 2024 Joël Riou. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Joël Riou
5+
-/
6+
import Mathlib.Algebra.Category.ModuleCat.Presheaf
7+
import Mathlib.Algebra.Category.ModuleCat.Differentials.Basic
8+
9+
/-!
10+
# The presheaf of differentials of a presheaf of modules
11+
12+
In this file, we define the type `M.Derivation φ` of derivations
13+
with values in a presheaf of `R`-modules `M` relative to
14+
a morphism of `φ : S ⟶ F.op ⋙ R` of presheaves of commutative rings
15+
over categories `C` and `D` that are related by a functor `F : C ⥤ D`.
16+
We formalize the notion of universal derivation.
17+
18+
Geometrically, if `f : X ⟶ S` is a morphisms of schemes (or more generally
19+
a morphism of commutative ringed spaces), we would like to apply
20+
these definitions in the case where `F` is the pullback functors from
21+
open subsets of `S` to open subsets of `X` and `φ` is the
22+
morphism $O_S ⟶ f_* O_X$.
23+
24+
In order to prove that there exists a universal derivation, the target
25+
of which shall be called the presheaf of relative differentials of `φ`,
26+
we first study the case where `F` is the identity functor.
27+
In this case where we have a morphism of presheaves of commutative
28+
rings `φ' : S' ⟶ R`, we construct a derivation
29+
`DifferentialsConstruction.derivation'` which is universal.
30+
Then, the general case (TODO) shall be obtained by observing that
31+
derivations for `S ⟶ F.op ⋙ R` identify to derivations
32+
for `S' ⟶ R` where `S'` is the pullback by `F` of the presheaf of
33+
commutative rings `S` (the data is the same: it suffices
34+
to show that the two vanishing conditions `d_app` are equivalent).
35+
36+
-/
37+
38+
universe v u v₁ v₂ u₁ u₂
39+
40+
open CategoryTheory
41+
42+
variable {C : Type u₁} [Category.{v₁} C] {D : Type u₂} [Category.{v₂} D]
43+
44+
namespace PresheafOfModules
45+
46+
variable {S : Cᵒᵖ ⥤ CommRingCat.{u}} {F : C ⥤ D} {S' R : Dᵒᵖ ⥤ CommRingCat.{u}}
47+
(M N : PresheafOfModules.{v} (R ⋙ forget₂ _ _))
48+
(φ : S ⟶ F.op ⋙ R) (φ' : S' ⟶ R)
49+
50+
/-- Given a morphism of presheaves of commutative rings `φ : S ⟶ F.op ⋙ R`,
51+
this is the type of relative `φ`-derivation of a presheaf of `R`-modules `M`. -/
52+
@[ext]
53+
structure Derivation where
54+
/-- the underlying additive map `R.obj X →+ M.obj X` of a derivation -/
55+
d {X : Dᵒᵖ} : R.obj X →+ M.obj X
56+
d_mul {X : Dᵒᵖ} (a b : R.obj X) : d (a * b) = a • d b + b • d a := by aesop_cat
57+
d_map {X Y : Dᵒᵖ} (f : X ⟶ Y) (x : R.obj X) :
58+
d (R.map f x) = M.map f (d x) := by aesop_cat
59+
d_app {X : Cᵒᵖ} (a : S.obj X) : d (φ.app X a) = 0 := by aesop_cat
60+
61+
namespace Derivation
62+
63+
-- Note: `d_app` cannot be a simp lemma because `dsimp` would
64+
-- simplify the composition of functors `R ⋙ forget₂ _ _`
65+
attribute [simp] d_mul d_map
66+
67+
variable {M N φ}
68+
69+
lemma congr_d {d d' : M.Derivation φ} (h : d = d') {X : Dᵒᵖ} (b : R.obj X) :
70+
d.d b = d'.d b := by rw [h]
71+
72+
variable (d : M.Derivation φ)
73+
74+
@[simp] lemma d_one (X : Dᵒᵖ) : d.d (X := X) 1 = 0 := by
75+
simpa using d.d_mul (X := X) 1 1
76+
77+
/-- The postcomposition of a derivation by a morphism of presheaves of modules. -/
78+
@[simps! d_apply]
79+
def postcomp (f : M ⟶ N) : N.Derivation φ where
80+
d := (f.app _).toAddMonoidHom.comp d.d
81+
d_map _ _ := by simp [naturality_apply]
82+
d_app {X} a := by
83+
dsimp
84+
erw [d_app, map_zero]
85+
86+
/-- The universal property that a derivation `d : M.Derivation φ` must
87+
satisfy so that the presheaf of modules `M` can be considered as the presheaf of
88+
(relative) differentials of a presheaf of commutative rings `φ : S ⟶ F.op ⋙ R`. -/
89+
structure Universal where
90+
/-- An absolyte derivation of `M'` descends as a morphism `M ⟶ M'`. -/
91+
desc {M' : PresheafOfModules (R ⋙ forget₂ CommRingCat RingCat)}
92+
(d' : M'.Derivation φ) : M ⟶ M'
93+
fac {M' : PresheafOfModules (R ⋙ forget₂ CommRingCat RingCat)}
94+
(d' : M'.Derivation φ) : d.postcomp (desc d') = d' := by aesop_cat
95+
postcomp_injective {M' : PresheafOfModules (R ⋙ forget₂ CommRingCat RingCat)}
96+
(φ φ' : M ⟶ M') (h : d.postcomp φ = d.postcomp φ') : φ = φ' := by aesop_cat
97+
98+
attribute [simp] Universal.fac
99+
100+
instance : Subsingleton d.Universal where
101+
allEq h₁ h₂ := by
102+
suffices ∀ {M' : PresheafOfModules (R ⋙ forget₂ CommRingCat RingCat)}
103+
(d' : M'.Derivation φ), h₁.desc d' = h₂.desc d' by
104+
cases h₁
105+
cases h₂
106+
simp only [Universal.mk.injEq]
107+
ext : 2
108+
apply this
109+
intro M' d'
110+
apply h₁.postcomp_injective
111+
simp
112+
113+
end Derivation
114+
115+
/-- The property that there exists a universal derivation for
116+
a morphism of presheaves of commutative rings `S ⟶ F.op ⋙ R`. -/
117+
class HasDifferentials : Prop where
118+
exists_universal_derivation : ∃ (M : PresheafOfModules.{u} (R ⋙ forget₂ _ _))
119+
(d : M.Derivation φ), Nonempty d.Universal
120+
121+
/-- Given a morphism of presheaves of commutative rings `φ : S ⟶ R`,
122+
this is the type of relative `φ`-derivation of a presheaf of `R`-modules `M`. -/
123+
abbrev Derivation' : Type _ := M.Derivation (F := 𝟭 D) φ'
124+
125+
namespace Derivation'
126+
127+
variable {M φ'}
128+
129+
@[simp]
130+
nonrec lemma d_app (d : M.Derivation' φ') {X : Dᵒᵖ} (a : S'.obj X) :
131+
d.d (φ'.app X a) = 0 :=
132+
d.d_app _
133+
134+
/-- The derivation relative to the morphism of commutative rings `φ'.app X` induced by
135+
a derivation relative to a morphism of presheaves of commutative rings. -/
136+
noncomputable def app (d : M.Derivation' φ') (X : Dᵒᵖ) : (M.obj X).Derivation (φ'.app X) :=
137+
ModuleCat.Derivation.mk (fun b ↦ d.d b)
138+
139+
@[simp]
140+
lemma app_apply (d : M.Derivation' φ') {X : Dᵒᵖ} (b : R.obj X) :
141+
(d.app X).d b = d.d b := rfl
142+
143+
section
144+
145+
variable (d : ∀ (X : Dᵒᵖ), (M.obj X).Derivation (φ'.app X))
146+
(d_map : ∀ ⦃X Y : Dᵒᵖ⦄ (f : X ⟶ Y) (x : R.obj X),
147+
(d Y).d ((R.map f) x) = (M.map f) ((d X).d x))
148+
149+
/-- Given a morphism of presheaves of commutative rings `φ'`, this is the
150+
in derivation `M.Derivation' φ'` that is given by a compatible family of derivations
151+
with values in the modules `M.obj X` for all `X`. -/
152+
def mk : M.Derivation' φ' where
153+
d {X} := AddMonoidHom.mk' (d X).d (by simp)
154+
155+
@[simp]
156+
lemma mk_app (X : Dᵒᵖ) : (mk d d_map).app X = d X := rfl
157+
158+
/-- Constructor for `Derivation.Universal` in the case `F` is the identity functor. -/
159+
def Universal.mk {d : M.Derivation' φ'}
160+
(desc : ∀ {M' : PresheafOfModules (R ⋙ forget₂ _ _)}
161+
(_ : M'.Derivation' φ'), M ⟶ M')
162+
(fac : ∀ {M' : PresheafOfModules (R ⋙ forget₂ _ _)}
163+
(d' : M'.Derivation' φ'), d.postcomp (desc d') = d')
164+
(postcomp_injective : ∀ {M' : PresheafOfModules (R ⋙ forget₂ _ _)}
165+
(α β : M ⟶ M'), d.postcomp α = d.postcomp β → α = β) : d.Universal where
166+
desc := desc
167+
fac := fac
168+
postcomp_injective := postcomp_injective
169+
170+
end
171+
172+
end Derivation'
173+
174+
namespace DifferentialsConstruction
175+
176+
/-- Auxiliary definition for `relativeDifferentials'`. -/
177+
noncomputable def relativeDifferentials'BundledCore :
178+
BundledCorePresheafOfModules.{u} (R ⋙ forget₂ _ _) where
179+
obj X := CommRingCat.KaehlerDifferential (φ'.app X)
180+
map f := CommRingCat.KaehlerDifferential.map (φ'.naturality f)
181+
182+
/-- The presheaf of relative differentials of a morphism of presheaves of
183+
commutative rings. -/
184+
noncomputable def relativeDifferentials' :
185+
PresheafOfModules.{u} (R ⋙ forget₂ _ _) :=
186+
(relativeDifferentials'BundledCore φ').toPresheafOfModules
187+
188+
@[simp]
189+
lemma relativeDifferentials'_obj (X : Dᵒᵖ) :
190+
(relativeDifferentials' φ').obj X =
191+
CommRingCat.KaehlerDifferential (φ'.app X) := rfl
192+
193+
-- Note: this cannot be a simp lemma because `dsimp` would
194+
-- simplify the composition of functors `R ⋙ forget₂ _ _`
195+
lemma relativeDifferentials'_map_apply {X Y : Dᵒᵖ} (f : X ⟶ Y)
196+
(x : CommRingCat.KaehlerDifferential (φ'.app X)) :
197+
(relativeDifferentials' φ').map f x =
198+
CommRingCat.KaehlerDifferential.map (φ'.naturality f) x := rfl
199+
200+
lemma relativeDifferentials'_map_d {X Y : Dᵒᵖ} (f : X ⟶ Y)
201+
(x : R.obj X) :
202+
(relativeDifferentials' φ').map f (CommRingCat.KaehlerDifferential.d x) =
203+
CommRingCat.KaehlerDifferential.d (R.map f x) := by
204+
rw [relativeDifferentials'_map_apply, CommRingCat.KaehlerDifferential.map_d]
205+
206+
/-- The universal derivation. -/
207+
noncomputable def derivation' : (relativeDifferentials' φ').Derivation' φ' :=
208+
Derivation'.mk (fun X ↦ CommRingCat.KaehlerDifferential.D (φ'.app X)) (fun X Y f x ↦ by
209+
rw [relativeDifferentials'_map_apply, CommRingCat.KaehlerDifferential.map_d])
210+
211+
/-- The derivation `derivation' φ'` is universal. -/
212+
noncomputable def isUniversal' : (derivation' φ').Universal :=
213+
Derivation'.Universal.mk
214+
(fun {M'} d' ↦ Hom.mk'' (fun X ↦ (d'.app X).desc) (fun X Y f ↦
215+
CommRingCat.KaehlerDifferential.ext (fun b ↦ by
216+
dsimp [ModuleCat.ofHom]
217+
erw [restrictionApp_apply, restrictionApp_apply]
218+
simp only [relativeDifferentials'_map_d, ModuleCat.Derivation.desc_d,
219+
d'.app_apply, d'.d_map])))
220+
(fun {M'} d' ↦ by
221+
ext X b
222+
apply ModuleCat.Derivation.desc_d)
223+
(fun {M} α β h ↦ by
224+
ext1 X
225+
exact CommRingCat.KaehlerDifferential.ext (Derivation.congr_d h))
226+
227+
instance : HasDifferentials (F := 𝟭 D) φ' := ⟨_, _, ⟨isUniversal' φ'⟩⟩
228+
229+
end DifferentialsConstruction
230+
231+
end PresheafOfModules

Mathlib/Algebra/Category/ModuleCat/Presheaf.lean

+3-2
Original file line numberDiff line numberDiff line change
@@ -384,12 +384,13 @@ structure BundledCorePresheafOfModules where
384384
map {X Y : Cᵒᵖ} (f : X ⟶ Y) : obj X ⟶ (ModuleCat.restrictScalars (R.map f)).obj (obj Y)
385385
/-- `map` is compatible with the identities -/
386386
map_id (X : Cᵒᵖ) :
387-
map (𝟙 X) = (ModuleCat.restrictScalarsId' (R.map (𝟙 X)) (R.map_id X)).inv.app (obj X)
387+
map (𝟙 X) = (ModuleCat.restrictScalarsId' (R.map (𝟙 X)) (R.map_id X)).inv.app (obj X) := by
388+
aesop
388389
/-- `map` is compatible with the composition -/
389390
map_comp {X Y Z : Cᵒᵖ} (f : X ⟶ Y) (g : Y ⟶ Z) :
390391
map (f ≫ g) = map f ≫ (ModuleCat.restrictScalars (R.map f)).map (map g) ≫
391392
(ModuleCat.restrictScalarsComp' (R.map f) (R.map g) (R.map (f ≫ g))
392-
(R.map_comp f g)).inv.app (obj Z)
393+
(R.map_comp f g)).inv.app (obj Z) := by aesop
393394

394395
namespace BundledCorePresheafOfModules
395396

0 commit comments

Comments
 (0)