Documentation

Mathlib.RingTheory.AdicCompletion.AsTensorProduct

Adic completion as tensor product #

In this file we examine properties of the natural map

AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] AdicCompletion I M.

We show (in the AdicCompletion namespace):

As a corollary we obtain

TODO #

noncomputable def AdicCompletion.ofTensorProduct {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :

The natural AdicCompletion I R-linear map from AdicCompletion I R ⊗[R] M to the adic completion of M.

Equations
Instances For
    @[simp]
    theorem AdicCompletion.ofTensorProduct_tmul {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (r : AdicCompletion I R) (x : M) :
    noncomputable def AdicCompletion.ofTensorProductEquivOfPiFintype {R : Type u_1} [CommRing R] (I : Ideal R) (ι : Type u_4) [Fintype ι] [DecidableEq ι] :

    ofTensorProduct as an equiv in the case of M = R^ι where ι is finite.

    Equations
    Instances For

      If M is a finite R-module, then the canonical map AdicCompletion I R ⊗[R] M →ₗ AdicCompletion I M is surjective.

      Noetherian case #

      Suppose R is Noetherian. Then we show that the canonical map AdicCompletion I R ⊗[R] M →ₗ[AdicCompletion I R] AdicCompletion I M is an isomorphism for every finite R-module M.

      The strategy is the following: Choose a surjection f : (ι → R) →ₗ[R] M and consider the following commutative diagram:

       AdicCompletion I R ⊗[R] ker f -→ AdicCompletion I R ⊗[R] (ι → R) -→ AdicCompletion I R ⊗[R] M -→ 0
                     |                             |                                 |                  |
                     ↓                             ↓                                 ↓                  ↓
          AdicCompletion I (ker f) ------→ AdicCompletion I (ι → R) -------→ AdicCompletion I M ------→ 0
      

      The vertical maps are given by ofTensorProduct. By the previous section we know that the second vertical map is an isomorphism. Since R is Noetherian, ker f is finitely-generated, so again by the previous section the first vertical map is surjective.

      Moreover, both rows are exact by right-exactness of the tensor product and exactness of adic completions over Noetherian rings. Hence we conclude by the 5-lemma.

      If R is a Noetherian ring and M is a finite R-module, then the natural map given by AdicCompletion.ofTensorProduct is an isomorphism.

      ofTensorProduct packaged as linear equiv if M is a finite R-module and R is Noetherian.

      Equations
      Instances For

        Adic completion of a Noetherian ring R is flat over R.

        Equations
        • =