Skip to content

Commit 1c56cac

Browse files
committed
feat(RingTheory): Jacobian criterion for smoothness of local algebras (#20326)
Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com>
1 parent 571927b commit 1c56cac

File tree

3 files changed

+60
-0
lines changed

3 files changed

+60
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4666,6 +4666,7 @@ import Mathlib.RingTheory.SimpleRing.Field
46664666
import Mathlib.RingTheory.SimpleRing.Matrix
46674667
import Mathlib.RingTheory.Smooth.Basic
46684668
import Mathlib.RingTheory.Smooth.Kaehler
4669+
import Mathlib.RingTheory.Smooth.Local
46694670
import Mathlib.RingTheory.Smooth.Pi
46704671
import Mathlib.RingTheory.Smooth.StandardSmooth
46714672
import Mathlib.RingTheory.Support

Mathlib/RingTheory/Smooth/Kaehler.lean

Lines changed: 22 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -380,6 +380,8 @@ include hf in
380380
Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S`
381381
with kernel `I` (typically a presentation `R[X] → S`),
382382
`S` is formally smooth iff the `P`-linear map `I/I² → S ⊗[P] Ω[P⁄R]` is split injective.
383+
Also see `Algebra.Extension.formallySmooth_iff_split_injection`
384+
for the version in terms of `Extension`.
383385
-/
384386
@[stacks 031I]
385387
theorem Algebra.FormallySmooth.iff_split_injection :
@@ -388,6 +390,26 @@ theorem Algebra.FormallySmooth.iff_split_injection :
388390
simp only [nonempty_subtype] at this
389391
rw [this, ← Algebra.FormallySmooth.iff_split_surjection _ hf]
390392

393+
/--
394+
Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S`
395+
with kernel `I` (typically a presentation `R[X] → S`),
396+
`S` is formally smooth iff the `P`-linear map `I/I² → S ⊗[P] Ω[P⁄R]` is split injective.
397+
-/
398+
@[stacks 031I]
399+
theorem Algebra.Extension.formallySmooth_iff_split_injection
400+
(P : Algebra.Extension.{u} R S) [FormallySmooth R P.Ring] :
401+
Algebra.FormallySmooth R S ↔ ∃ l, l ∘ₗ P.cotangentComplex = LinearMap.id := by
402+
refine (Algebra.FormallySmooth.iff_split_injection P.algebraMap_surjective).trans ?_
403+
let e : P.ker.Cotangent ≃ₗ[P.Ring] P.Cotangent :=
404+
{ __ := AddEquiv.refl _, map_smul' r m := by ext1; simp; rfl }
405+
constructor
406+
· intro ⟨l, hl⟩
407+
exact ⟨(e.comp l).extendScalarsOfSurjective P.algebraMap_surjective,
408+
LinearMap.ext (DFunLike.congr_fun hl : _)⟩
409+
· intro ⟨l, hl⟩
410+
exact ⟨e.symm.toLinearMap ∘ₗ l.restrictScalars P.Ring,
411+
LinearMap.ext (DFunLike.congr_fun hl : _)⟩
412+
391413
include hf in
392414
/--
393415
Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S`
Lines changed: 37 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,37 @@
1+
/-
2+
Copyright (c) 2024 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 Mathlib.RingTheory.LocalRing.Module
7+
import Mathlib.RingTheory.Smooth.Kaehler
8+
import Mathlib.RingTheory.TensorProduct.Free
9+
10+
/-!
11+
# Formally smooth local algebras
12+
-/
13+
14+
open TensorProduct IsLocalRing KaehlerDifferential
15+
16+
/--
17+
The **Jacobian criterion** for smoothness of local algebras.
18+
Suppose `S` is a local `R`-algebra, and `0 → I → P → S → 0` is a presentation such that
19+
`P` is formally-smooth over `R`, `Ω[P⁄R]` is finite free over `P`,
20+
(typically satisfied when `P` is the localization of a polynomial ring of finite type)
21+
and `I` is finitely generated.
22+
Then `S` is formally smooth iff `k ⊗ₛ I/I² → k ⊗ₚ Ω[P/R]` is injective,
23+
where `k` is the residue field of `S`.
24+
-/
25+
theorem Algebra.FormallySmooth.iff_injective_lTensor_residueField {R S} [CommRing R]
26+
[CommRing S] [IsLocalRing S] [Algebra R S] (P : Algebra.Extension R S)
27+
[FormallySmooth R P.Ring]
28+
[Module.Free P.Ring (Ω[P.Ring⁄R])] [Module.Finite P.Ring (Ω[P.Ring⁄R])]
29+
(h' : P.ker.FG) :
30+
Algebra.FormallySmooth R S ↔
31+
Function.Injective (P.cotangentComplex.lTensor (ResidueField S)) := by
32+
have : Module.Finite P.Ring P.Cotangent :=
33+
have : Module.Finite P.Ring P.ker := ⟨(Submodule.fg_top _).mpr h'⟩
34+
.of_surjective _ Extension.Cotangent.mk_surjective
35+
have : Module.Finite S P.Cotangent := Module.Finite.of_restrictScalars_finite P.Ring _ _
36+
rw [← IsLocalRing.split_injective_iff_lTensor_residueField_injective,
37+
P.formallySmooth_iff_split_injection]

0 commit comments

Comments
 (0)