@@ -380,6 +380,8 @@ include hf in
380380Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S`
381381with 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]
385387theorem 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+
391413include hf in
392414/--
393415Given a formally smooth `R`-algebra `P` and a surjective algebra homomorphism `f : P →ₐ[R] S`
0 commit comments