Documentation

Mathlib.RingTheory.AdicCompletion.Basic

Completion of a module with respect to an ideal. #

In this file we define the notions of Hausdorff, precomplete, and complete for an R-module M with respect to an ideal I:

Main definitions #

class IsHausdorff {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :

A module M is Hausdorff with respect to an ideal I if ⋂ I^n M = 0.

Instances
    theorem IsHausdorff.haus' {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] [self : IsHausdorff I M] (x : M) :
    (∀ (n : ), x 0 [SMOD I ^ n ])x = 0
    class IsPrecomplete {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :

    A module M is precomplete with respect to an ideal I if every Cauchy sequence converges.

    Instances
      theorem IsPrecomplete.prec' {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] [self : IsPrecomplete I M] (f : M) :
      (∀ {m n : }, m nf m f n [SMOD I ^ m ])∃ (L : M), ∀ (n : ), f n L [SMOD I ^ n ]
      class IsAdicComplete {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] extends IsHausdorff , IsPrecomplete :

      A module M is I-adically complete if it is Hausdorff and precomplete.

        Instances
          theorem IsHausdorff.haus {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] :
          IsHausdorff I M∀ (x : M), (∀ (n : ), x 0 [SMOD I ^ n ])x = 0
          theorem isHausdorff_iff {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] :
          IsHausdorff I M ∀ (x : M), (∀ (n : ), x 0 [SMOD I ^ n ])x = 0
          theorem IsPrecomplete.prec {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] :
          IsPrecomplete I M∀ {f : M}, (∀ {m n : }, m nf m f n [SMOD I ^ m ])∃ (L : M), ∀ (n : ), f n L [SMOD I ^ n ]
          theorem isPrecomplete_iff {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] :
          IsPrecomplete I M ∀ (f : M), (∀ {m n : }, m nf m f n [SMOD I ^ m ])∃ (L : M), ∀ (n : ), f n L [SMOD I ^ n ]
          @[reducible, inline]
          noncomputable abbrev Hausdorffification {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
          Type u_2

          The Hausdorffification of a module with respect to an ideal.

          Equations
          Instances For
            noncomputable def AdicCompletion.transitionMap {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {m : } {n : } (hmn : m n) :
            M I ^ n →ₗ[R] M I ^ m

            The canonical linear map M ⧸ (I ^ n • ⊤) →ₗ[R] M ⧸ (I ^ m • ⊤) for m ≤ n used to define AdicCompletion.

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

              The completion of a module with respect to an ideal. This is not necessarily Hausdorff. In fact, this is only complete if the ideal is finitely generated.

              Equations
              Instances For
                noncomputable instance IsHausdorff.bot {R : Type u_1} [CommRing R] (M : Type u_2) [AddCommGroup M] [Module R M] :
                Equations
                • =
                theorem IsHausdorff.subsingleton {R : Type u_1} [CommRing R] {M : Type u_2} [AddCommGroup M] [Module R M] (h : IsHausdorff M) :
                @[instance 100]
                noncomputable instance IsHausdorff.of_subsingleton {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] [Subsingleton M] :
                Equations
                • =
                theorem IsHausdorff.iInf_pow_smul {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (h : IsHausdorff I M) :
                ⨅ (n : ), I ^ n =
                noncomputable def Hausdorffification.of {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :

                The canonical linear map to the Hausdorffification.

                Equations
                Instances For
                  theorem Hausdorffification.induction_on {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {C : Hausdorffification I MProp} (x : Hausdorffification I M) (ih : ∀ (x : M), C ((Hausdorffification.of I M) x)) :
                  C x
                  noncomputable instance Hausdorffification.instIsHausdorff {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                  Equations
                  • =
                  noncomputable def Hausdorffification.lift {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] [h : IsHausdorff I N] (f : M →ₗ[R] N) :

                  Universal property of Hausdorffification: any linear map to a Hausdorff module extends to a unique map from the Hausdorffification.

                  Equations
                  Instances For
                    theorem Hausdorffification.lift_of {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] [h : IsHausdorff I N] (f : M →ₗ[R] N) (x : M) :
                    theorem Hausdorffification.lift_comp_of {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] [h : IsHausdorff I N] (f : M →ₗ[R] N) :
                    theorem Hausdorffification.lift_eq {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] [h : IsHausdorff I N] (f : M →ₗ[R] N) (g : Hausdorffification I M →ₗ[R] N) (hg : g ∘ₗ Hausdorffification.of I M = f) :

                    Uniqueness of lift.

                    noncomputable instance IsPrecomplete.bot {R : Type u_1} [CommRing R] (M : Type u_2) [AddCommGroup M] [Module R M] :
                    Equations
                    • =
                    noncomputable instance IsPrecomplete.top {R : Type u_1} [CommRing R] (M : Type u_2) [AddCommGroup M] [Module R M] :
                    Equations
                    • =
                    @[instance 100]
                    noncomputable instance IsPrecomplete.of_subsingleton {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] [Subsingleton M] :
                    Equations
                    • =
                    noncomputable def AdicCompletion.submodule {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                    Submodule R ((n : ) → M I ^ n )

                    AdicCompletion is the submodule of compatible families in ∀ n : ℕ, M ⧸ (I ^ n • ⊤).

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      noncomputable instance AdicCompletion.instZero {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instAdd {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instNeg {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instSub {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instSMulNat {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instSMulInt {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instAddCommGroup {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instSMul {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable instance AdicCompletion.instModule {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      Equations
                      noncomputable def AdicCompletion.incl {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                      AdicCompletion I M →ₗ[R] (n : ) → M I ^ n

                      The canonical inclusion from the completion to the product.

                      Equations
                      Instances For
                        @[simp]
                        theorem AdicCompletion.incl_apply {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (x : AdicCompletion I M) (n : ) :
                        (AdicCompletion.incl I M) x n = x n
                        noncomputable def AdicCompletion.of {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :

                        The canonical linear map to the completion.

                        Equations
                        Instances For
                          @[simp]
                          theorem AdicCompletion.of_apply {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (x : M) (n : ) :
                          ((AdicCompletion.of I M) x) n = (I ^ n ).mkQ x
                          noncomputable def AdicCompletion.eval {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :

                          Linearly evaluating a sequence in the completion at a given input.

                          Equations
                          Instances For
                            @[simp]
                            theorem AdicCompletion.coe_eval {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :
                            (AdicCompletion.eval I M n) = fun (f : AdicCompletion I M) => f n
                            theorem AdicCompletion.eval_apply {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) (f : AdicCompletion I M) :
                            (AdicCompletion.eval I M n) f = f n
                            theorem AdicCompletion.eval_of {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) (x : M) :
                            (AdicCompletion.eval I M n) ((AdicCompletion.of I M) x) = (I ^ n ).mkQ x
                            @[simp]
                            theorem AdicCompletion.eval_comp_of {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :
                            AdicCompletion.eval I M n ∘ₗ AdicCompletion.of I M = (I ^ n ).mkQ
                            @[simp]
                            theorem AdicCompletion.range_eval {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :
                            @[simp]
                            theorem AdicCompletion.val_zero {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :
                            0 n = 0
                            @[simp]
                            theorem AdicCompletion.val_add {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (n : ) (f : AdicCompletion I M) (g : AdicCompletion I M) :
                            (f + g) n = f n + g n
                            @[simp]
                            theorem AdicCompletion.val_sub {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (n : ) (f : AdicCompletion I M) (g : AdicCompletion I M) :
                            (f - g) n = f n - g n
                            @[simp]
                            theorem AdicCompletion.val_sum {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {α : Type u_4} (s : Finset α) (f : αAdicCompletion I M) (n : ) :
                            (s.sum f) n = as, (f a) n
                            theorem AdicCompletion.val_smul {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (n : ) (r : R) (f : AdicCompletion I M) :
                            (r f) n = r f n
                            theorem AdicCompletion.ext_iff {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {x : AdicCompletion I M} {y : AdicCompletion I M} :
                            x = y ∀ (n : ), x n = y n
                            theorem AdicCompletion.ext {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {x : AdicCompletion I M} {y : AdicCompletion I M} (h : ∀ (n : ), x n = y n) :
                            x = y
                            noncomputable instance AdicCompletion.instIsHausdorff {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                            Equations
                            • =
                            @[simp]
                            theorem AdicCompletion.transitionMap_mk {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {m : } {n : } (hmn : m n) (x : M) :
                            @[simp]
                            theorem AdicCompletion.transitionMap_eq {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :
                            AdicCompletion.transitionMap I M = LinearMap.id
                            @[simp]
                            theorem AdicCompletion.transitionMap_comp {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {m : } {n : } {k : } (hmn : m n) (hnk : n k) :
                            @[simp]
                            theorem AdicCompletion.transitionMap_comp_apply {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {m : } {n : } {k : } (hmn : m n) (hnk : n k) (x : M I ^ k ) :
                            @[simp]
                            theorem AdicCompletion.transitionMap_comp_eval_apply {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {m : } {n : } (hmn : m n) (x : AdicCompletion I M) :
                            (AdicCompletion.transitionMap I M hmn) (x n) = x m
                            @[simp]
                            theorem AdicCompletion.transitionMap_comp_eval {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {m : } {n : } (hmn : m n) :
                            noncomputable def AdicCompletion.IsAdicCauchy {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (f : M) :

                            A sequence ℕ → M is an I-adic Cauchy sequence if for every m ≤ n, f m ≡ f n modulo I ^ m • ⊤.

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

                              The type of I-adic Cauchy sequences.

                              Equations
                              Instances For
                                noncomputable def AdicCompletion.AdicCauchySequence.submodule {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                Submodule R (M)

                                The type of I-adic cauchy sequences is a submodule of the product ℕ → M.

                                Equations
                                Instances For
                                  noncomputable instance AdicCompletion.AdicCauchySequence.instZero {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                  Equations
                                  noncomputable instance AdicCompletion.AdicCauchySequence.instAdd {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                  Equations
                                  noncomputable instance AdicCompletion.AdicCauchySequence.instNeg {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                  Equations
                                  noncomputable instance AdicCompletion.AdicCauchySequence.instSub {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                  Equations
                                  Equations
                                  Equations
                                  noncomputable instance AdicCompletion.AdicCauchySequence.instSMul {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                  Equations
                                  noncomputable instance AdicCompletion.AdicCauchySequence.instModule {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :
                                  Equations
                                  @[simp]
                                  theorem AdicCompletion.AdicCauchySequence.zero_apply {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (n : ) :
                                  0 n = 0
                                  @[simp]
                                  theorem AdicCompletion.AdicCauchySequence.add_apply {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (n : ) (f : AdicCompletion.AdicCauchySequence I M) (g : AdicCompletion.AdicCauchySequence I M) :
                                  (f + g) n = f n + g n
                                  @[simp]
                                  theorem AdicCompletion.AdicCauchySequence.sub_apply {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (n : ) (f : AdicCompletion.AdicCauchySequence I M) (g : AdicCompletion.AdicCauchySequence I M) :
                                  (f - g) n = f n - g n
                                  @[simp]
                                  theorem AdicCompletion.AdicCauchySequence.smul_apply {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] (n : ) (r : R) (f : AdicCompletion.AdicCauchySequence I M) :
                                  (r f) n = r f n
                                  theorem AdicCompletion.AdicCauchySequence.ext_iff {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {x : AdicCompletion.AdicCauchySequence I M} {y : AdicCompletion.AdicCauchySequence I M} :
                                  x = y ∀ (n : ), x n = y n
                                  theorem AdicCompletion.AdicCauchySequence.ext {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {x : AdicCompletion.AdicCauchySequence I M} {y : AdicCompletion.AdicCauchySequence I M} (h : ∀ (n : ), x n = y n) :
                                  x = y
                                  theorem AdicCompletion.AdicCauchySequence.mk_eq_mk {R : Type u_1} [CommRing R] {I : Ideal R} {M : Type u_2} [AddCommGroup M] [Module R M] {m : } {n : } (hmn : m n) (f : AdicCompletion.AdicCauchySequence I M) :

                                  The defining property of an adic cauchy sequence unwrapped.

                                  theorem AdicCompletion.isAdicCauchy_iff {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (f : M) :
                                  AdicCompletion.IsAdicCauchy I M f ∀ (n : ), f n f (n + 1) [SMOD I ^ n ]

                                  The I-adic cauchy condition can be checked on successive n.

                                  noncomputable def AdicCompletion.AdicCauchySequence.mk {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (f : M) (h : ∀ (n : ), f n f (n + 1) [SMOD I ^ n ]) :

                                  Construct I-adic cauchy sequence from sequence satisfying the successive cauchy condition.

                                  Equations
                                  Instances For
                                    @[simp]
                                    theorem AdicCompletion.AdicCauchySequence.mk_coe {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (f : M) (h : ∀ (n : ), f n f (n + 1) [SMOD I ^ n ]) :
                                    ∀ (a : ), (AdicCompletion.AdicCauchySequence.mk I M f h) a = f a
                                    noncomputable def AdicCompletion.mk {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] :

                                    The canonical linear map from cauchy sequences to the completion.

                                    Equations
                                    Instances For
                                      @[simp]
                                      theorem AdicCompletion.mk_apply_coe {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (f : AdicCompletion.AdicCauchySequence I M) (n : ) :
                                      ((AdicCompletion.mk I M) f) n = (I ^ n ).mkQ (f n)
                                      theorem AdicCompletion.mk_zero_of {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] (f : AdicCompletion.AdicCauchySequence I M) (h : ∃ (k : ), nk, mn, ln, f m I ^ l ) :

                                      Criterion for checking that an adic cauchy sequence is mapped to zero in the adic completion.

                                      Every element in the adic completion is represented by a Cauchy sequence.

                                      theorem AdicCompletion.induction_on {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] {p : AdicCompletion I MProp} (x : AdicCompletion I M) (h : ∀ (f : AdicCompletion.AdicCauchySequence I M), p ((AdicCompletion.mk I M) f)) :
                                      p x

                                      To show a statement about an element of adicCompletion I M, it suffices to check it on Cauchy sequences.

                                      noncomputable def AdicCompletion.lift {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] (f : (n : ) → M →ₗ[R] N I ^ n ) (h : ∀ {m n : } (hle : m n), AdicCompletion.transitionMap I N hle ∘ₗ f n = f m) :

                                      Lift a compatible family of linear maps M →ₗ[R] N ⧸ (I ^ n • ⊤ : Submodule R N) to the I-adic completion of M.

                                      Equations
                                      • AdicCompletion.lift I f h = { toFun := fun (x : M) => fun (n : ) => (f n) x, , map_add' := , map_smul' := }
                                      Instances For
                                        @[simp]
                                        theorem AdicCompletion.eval_lift {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] (f : (n : ) → M →ₗ[R] N I ^ n ) (h : ∀ {m n : } (hle : m n), AdicCompletion.transitionMap I N hle ∘ₗ f n = f m) (n : ) :
                                        AdicCompletion.eval I N n ∘ₗ AdicCompletion.lift I f = f n
                                        @[simp]
                                        theorem AdicCompletion.eval_lift_apply {R : Type u_1} [CommRing R] (I : Ideal R) {M : Type u_2} [AddCommGroup M] [Module R M] {N : Type u_3} [AddCommGroup N] [Module R N] (f : (n : ) → M →ₗ[R] N I ^ n ) (h : ∀ {m n : } (hle : m n), AdicCompletion.transitionMap I N hle ∘ₗ f n = f m) (n : ) (x : M) :
                                        ((AdicCompletion.lift I f ) x) n = (f n) x
                                        noncomputable instance IsAdicComplete.bot {R : Type u_1} [CommRing R] (M : Type u_2) [AddCommGroup M] [Module R M] :
                                        Equations
                                        • =
                                        @[instance 100]
                                        noncomputable instance IsAdicComplete.of_subsingleton {R : Type u_1} [CommRing R] (I : Ideal R) (M : Type u_2) [AddCommGroup M] [Module R M] [Subsingleton M] :
                                        Equations
                                        • =
                                        theorem IsAdicComplete.le_jacobson_bot {R : Type u_1} [CommRing R] (I : Ideal R) [IsAdicComplete I R] :
                                        I .jacobson