@@ -22,6 +22,14 @@ In this file we prove that tensoring with a flat module is an exact functor.
22
22
- `Module.Flat.iff_tensorRight_preservesFiniteLimits`: an `R`-module `M` is flat if and only if
23
23
right tensoring with `M` preserves finite limits, i.e. the functor `M ⊗ -` is left exact.
24
24
25
+ - `Module.Flat.iff_lTensor_preserves_shortComplex_exact`: an `R`-module `M` is flat if and only if
26
+ for every short exact sequence `0 ⟶ A ⟶ B ⟶ C ⟶ 0`, `0 ⟶ M ⊗ A ⟶ M ⊗ B ⟶ M ⊗ C ⟶ 0` is also
27
+ a short exact sequence.
28
+
29
+ - `Module.Flat.iff_rTensor_preserves_shortComplex_exact`: an `R`-module `M` is flat if and only if
30
+ for every short exact sequence `0 ⟶ A ⟶ B ⟶ C ⟶ 0`, `0 ⟶ A ⊗ M ⟶ B ⊗ M ⟶ C ⊗ M ⟶ 0` is also
31
+ a short exact sequence.
32
+
25
33
- `Module.Flat.higherTorIsoZero`: if an `R`-module `M` is flat, then `Torⁿ(M, N) ≅ 0` for all `N`
26
34
and all `n ≥ 1`.
27
35
@@ -86,10 +94,10 @@ noncomputable instance [flat : Flat R M] {X Y : ModuleCat.{u} R} (f : X ⟶ Y) :
86
94
have mono0 : Mono ι :=
87
95
{ right_cancellation := fun {Z} g h H => by
88
96
let c' : Limits.Cone (Limits.parallelPair f 0 ) :=
89
- ⟨Z, ⟨fun | .zero => h ≫ ι | .one => 0 , by rintro _ _ (⟨j⟩|⟨j⟩) <;> simpa [ι] using H ⟩⟩
97
+ ⟨Z, ⟨fun | .zero => h ≫ ι | .one => 0 , by rintro _ _ (⟨j⟩|⟨j⟩) <;> simp [ι, H] ⟩⟩
90
98
91
99
rw [hc.uniq c' g, hc.uniq c' h] <;>
92
- rintro (⟨j⟩|⟨j⟩) <;> simpa [ι] using H }
100
+ rintro (⟨j⟩|⟨j⟩) <;> try simp [ι, H] <;> try simpa [ι, c' ] using H }
93
101
let s : ShortComplex (ModuleCat R) := .mk ι f $ by simp [ι]
94
102
have exact0 : s.Exact:= by
95
103
refine ShortComplex.exact_of_f_is_kernel _ $
@@ -101,13 +109,10 @@ noncomputable instance [flat : Flat R M] {X Y : ModuleCat.{u} R} (f : X ⟶ Y) :
101
109
· rfl
102
110
· rfl
103
111
104
- let s' := s.map (tensorLeft M)
105
- let f' := M ◁ f; let ι' := M ◁ ι
106
- have exact1 : s'.Exact := by
107
- apply lTensor_shortComplex_exact; assumption
112
+ let s' := s.map (tensorLeft M); let f' := M ◁ f; let ι' := M ◁ ι
113
+ have exact1 : s'.Exact := by apply lTensor_shortComplex_exact; assumption
108
114
109
115
have mono1 : Mono ι' := by
110
-
111
116
rw [ModuleCat.mono_iff_injective] at mono0 ⊢
112
117
exact lTensor_preserves_injective_linearMap _ mono0
113
118
@@ -163,23 +168,27 @@ section Tor
163
168
164
169
open scoped ZeroObject
165
170
171
+ /-
172
+ For a flat module `M`, higher tor groups vanish.
173
+ -/
174
+ noncomputable def higherTorIsoZero [Flat R M] (n : ℕ) (N : ModuleCat.{u} R) :
175
+ ((Tor _ (n + 1 )).obj M).obj N ≅ 0 :=
176
+ let pN := ProjectiveResolution.of N
177
+ pN.isoLeftDerivedObj (tensorLeft M) (n + 1 ) ≪≫
178
+ (Limits.IsZero.isoZero $ HomologicalComplex.exactAt_iff_isZero_homology _ _ |>.1 $
179
+ lTensor_shortComplex_exact M (pN.complex.sc (n + 1 )) (pN.complex_exactAt_succ n))
180
+
181
+
166
182
/--
167
183
For a flat module `M`, higher tor groups vanish.
168
184
-/
169
- noncomputable def higherTorIsoZero [flat : Flat R M] (n : ℕ) (N : ModuleCat.{u} R) :
170
- ((Tor' _ (n + 1 )).obj N).obj M ≅ 0 := by
171
- dsimp [Tor', Functor.flip]
185
+ noncomputable def higherTor'IsoZero [Flat R M] (n : ℕ) (N : ModuleCat.{u} R) :
186
+ ((Tor' _ (n + 1 )).obj N).obj M ≅ 0 :=
172
187
let pN := ProjectiveResolution.of N
173
- refine' pN.isoLeftDerivedObj (tensorRight M) (n + 1 ) ≪≫ ?_
174
- refine Limits.IsZero.isoZero ?_
175
- dsimp only [HomologicalComplex.homologyFunctor_obj]
176
- rw [← HomologicalComplex.exactAt_iff_isZero_homology, HomologicalComplex.exactAt_iff,
177
- ← exact_iff_shortComplex_exact, ModuleCat.exact_iff, Eq.comm, ← LinearMap.exact_iff]
178
- refine iff_rTensor_exact |>.1 flat ?_
179
- rw [LinearMap.exact_iff, Eq.comm, ← ModuleCat.exact_iff]
180
- have := pN.complex_exactAt_succ n
181
- rw [HomologicalComplex.exactAt_iff, ← exact_iff_shortComplex_exact] at this
182
- exact this
188
+ pN.isoLeftDerivedObj (tensorRight M) (n + 1 ) ≪≫
189
+ (Limits.IsZero.isoZero $ HomologicalComplex.exactAt_iff_isZero_homology _ _ |>.1 $
190
+ rTensor_shortComplex_exact M (pN.complex.sc (n + 1 )) (pN.complex_exactAt_succ n))
191
+
183
192
184
193
end Tor
185
194
0 commit comments