Skip to content

Commit cddb28f

Browse files
robin-carlierRobin CarliererdOnejcommelin
committed
feat(RingTheory/Flat): free and projective modules over a CommRing are flat (#7567)
Prove that flat modules are stable under retracts, isomorphisms and direct sums. Use this to deduce that free and projective modules are flat. Co-authored-by: Robin Carlier <robin.carlier@ens-lyon.fr> Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent 1ff3bf3 commit cddb28f

File tree

1 file changed

+133
-4
lines changed

1 file changed

+133
-4
lines changed

Mathlib/RingTheory/Flat.lean

Lines changed: 133 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johan Commelin
55
-/
66
import Mathlib.RingTheory.Noetherian
7+
import Mathlib.Algebra.DirectSum.Module
8+
import Mathlib.Algebra.DirectSum.Finsupp
9+
import Mathlib.LinearAlgebra.DirectSum.TensorProduct
10+
import Mathlib.LinearAlgebra.FreeModule.Basic
11+
import Mathlib.Algebra.Module.Projective
712

813
#align_import ring_theory.flat from "leanprover-community/mathlib"@"62c0a4ef1441edb463095ea02a06e87f3dfe135c"
914

@@ -23,6 +28,14 @@ This result is not yet formalised.
2328
2429
* `Module.Flat`: the predicate asserting that an `R`-module `M` is flat.
2530
31+
## Main theorems
32+
33+
* `Module.Flat.of_retract`: retracts of flat modules are flat
34+
* `Module.Flat.of_linearEquiv`: modules linearly equivalent to a flat modules are flat
35+
* `Module.Flat.directSum`: arbitrary direct sums of flat modules are flat
36+
* `Module.Flat.of_free`: free modules are flat
37+
* `Module.Flat.of_projective`: projective modules are flat
38+
2639
## TODO
2740
2841
* Show that tensoring with a flat module preserves injective morphisms.
@@ -46,13 +59,13 @@ This result is not yet formalised.
4659
-/
4760

4861

49-
universe u v
62+
universe u v w
5063

5164
namespace Module
5265

53-
open Function (Injective)
66+
open Function (Injective Surjective)
5467

55-
open LinearMap (lsmul)
68+
open LinearMap (lsmul rTensor lTensor)
5669

5770
open TensorProduct
5871

@@ -64,7 +77,9 @@ class Flat (R : Type u) (M : Type v) [CommRing R] [AddCommGroup M] [Module R M]
6477

6578
namespace Flat
6679

67-
open TensorProduct LinearMap Submodule
80+
variable (R : Type u) [CommRing R]
81+
82+
open LinearMap Submodule
6883

6984
instance self (R : Type u) [CommRing R] : Flat R R :=
7085
by
@@ -76,6 +91,120 @@ instance self (R : Type u) [CommRing R] : Flat R R :=
7691
lift.tmul, Submodule.subtype_apply, Algebra.id.smul_eq_mul, lsmul_apply]⟩
7792
#align module.flat.self Module.Flat.self
7893

94+
variable (M : Type v) [AddCommGroup M] [Module R M]
95+
96+
lemma iff_rTensor_injective :
97+
Flat R M ↔ (∀ ⦃I : Ideal R⦄ (_ : I.FG), Injective (rTensor M I.subtype)) := by
98+
have aux : ∀ (I : Ideal R), ((TensorProduct.lid R M).comp (rTensor M I.subtype)) =
99+
(TensorProduct.lift ((lsmul R M).comp I.subtype))
100+
· intro I; apply TensorProduct.ext'; intro x y; simp
101+
constructor
102+
· intro F I hI
103+
erw [← Equiv.comp_injective _ (TensorProduct.lid R M).toEquiv]
104+
have h₁ := F.out hI
105+
rw [← aux] at h₁
106+
refine h₁
107+
· intro h₁
108+
constructor
109+
intro I hI
110+
rw [← aux]
111+
simp [h₁ hI]
112+
113+
variable (N : Type w) [AddCommGroup N] [Module R N]
114+
115+
/-- A retract of a flat `R`-module is flat. -/
116+
lemma of_retract [f : Flat R M] (i : N →ₗ[R] M) (r : M →ₗ[R] N) (h : r.comp i = LinearMap.id) :
117+
Flat R N := by
118+
rw [iff_rTensor_injective] at *
119+
intro I hI
120+
have h₁ : Function.Injective (lTensor R i)
121+
· apply Function.RightInverse.injective (g := (lTensor R r))
122+
intro x
123+
rw [← LinearMap.comp_apply, ← lTensor_comp, h]
124+
simp
125+
rw [← Function.Injective.of_comp_iff h₁ (rTensor N I.subtype), ← LinearMap.coe_comp]
126+
rw [LinearMap.lTensor_comp_rTensor, ←LinearMap.rTensor_comp_lTensor]
127+
rw [LinearMap.coe_comp, Function.Injective.of_comp_iff (f hI)]
128+
apply Function.RightInverse.injective (g := lTensor _ r)
129+
intro x
130+
rw [← LinearMap.comp_apply, ← lTensor_comp, h]
131+
simp
132+
133+
/-- A `R`-module linearly equivalent to a flat `R`-module is flat. -/
134+
lemma of_linearEquiv [f : Flat R M] (e : N ≃ₗ[R] M) : Flat R N := by
135+
have h : e.symm.toLinearMap.comp e.toLinearMap = LinearMap.id := by simp
136+
exact of_retract _ _ _ e.toLinearMap e.symm.toLinearMap h
137+
138+
end Flat
139+
140+
namespace Flat
141+
142+
open DirectSum LinearMap Submodule
143+
144+
variable (R : Type u) [CommRing R]
145+
146+
/-- A direct sum of flat `R`-modules is flat. -/
147+
instance directSum (ι : Type v) (M : ι → Type w) [(i : ι) → AddCommGroup (M i)]
148+
[(i : ι) → Module R (M i)] [F : (i : ι) → (Flat R (M i))] : Flat R (⨁ i, M i) := by
149+
classical
150+
rw [iff_rTensor_injective]
151+
intro I hI
152+
rw [← Equiv.comp_injective _ (TensorProduct.lid R (⨁ i, M i)).toEquiv]
153+
set η₁ := TensorProduct.lid R (⨁ i, M i)
154+
set η := (fun i ↦ TensorProduct.lid R (M i))
155+
set φ := (fun i ↦ rTensor (M i) I.subtype)
156+
set π := (fun i ↦ component R ι (fun j ↦ M j) i)
157+
set ψ := (TensorProduct.directSumRight R {x // x ∈ I} (fun i ↦ M i)).symm.toLinearMap with psi_def
158+
set ρ := rTensor (⨁ i, M i) I.subtype
159+
set τ := (fun i ↦ component R ι (fun j ↦ ({x // x ∈ I} ⊗[R] (M j))) i)
160+
rw [← Equiv.injective_comp (TensorProduct.directSumRight _ _ _).symm.toEquiv]
161+
rw [LinearEquiv.coe_toEquiv, ← LinearEquiv.coe_coe, ← LinearMap.coe_comp]
162+
rw [LinearEquiv.coe_toEquiv, ← LinearEquiv.coe_coe, ← LinearMap.coe_comp]
163+
rw [← psi_def, injective_iff_map_eq_zero ((η₁.comp ρ).comp ψ)]
164+
have h₁ : ∀ (i : ι), (π i).comp ((η₁.comp ρ).comp ψ) = (η i).comp ((φ i).comp (τ i))
165+
· intro i
166+
apply DirectSum.linearMap_ext
167+
intro j
168+
apply TensorProduct.ext'
169+
intro a m
170+
simp only [coe_comp, LinearEquiv.coe_coe, Function.comp_apply, directSumRight_symm_lof_tmul,
171+
rTensor_tmul, coeSubtype, lid_tmul, map_smul]
172+
rw [DirectSum.component.of, DirectSum.component.of]
173+
by_cases h₂ : j = i
174+
· subst j; simp
175+
· simp [h₂]
176+
intro a ha; rw [DirectSum.ext_iff R]; intro i
177+
have f := LinearMap.congr_arg (f:= (π i)) ha
178+
erw [LinearMap.congr_fun (h₁ i) a] at f
179+
rw [(map_zero (π i) : (π i) 0 = (0 : M i))] at f
180+
have h₂ := (F i)
181+
rw [iff_rTensor_injective] at h₂
182+
have h₃ := h₂ hI
183+
simp only [coe_comp, LinearEquiv.coe_coe, Function.comp_apply, AddEquivClass.map_eq_zero_iff,
184+
h₃, LinearMap.map_eq_zero_iff] at f
185+
simp [f]
186+
187+
/-- Free `R`-modules over discrete types are flat. -/
188+
instance finsupp (ι : Type v) : Flat R (ι →₀ R) :=
189+
let _ := Classical.decEq ι
190+
of_linearEquiv R _ _ (finsuppLEquivDirectSum R R ι)
191+
192+
variable (M : Type v) [AddCommGroup M] [Module R M]
193+
194+
instance of_free [Free R M] : Flat R M := of_linearEquiv R _ _ (Free.repr R M)
195+
196+
/-- A projective module with a discrete type of generator is flat -/
197+
lemma of_projective_surjective (ι : Type w) [Projective R M] (p : (ι →₀ R) →ₗ[R] M)
198+
(hp : Surjective p) : Flat R M := by
199+
have h := Module.projective_lifting_property p (LinearMap.id) hp
200+
cases h with
201+
| _ e he => exact of_retract R _ _ _ _ he
202+
203+
instance of_projective [h : Projective R M] : Flat R M := by
204+
rw [Module.projective_def'] at h
205+
cases h with
206+
| _ e he => exact of_retract R _ _ _ _ he
207+
79208
end Flat
80209

81210
end Module

0 commit comments

Comments
 (0)