Documentation

Mathlib.Algebra.Module.LocalizedModule.Submodule

Localization of Submodules #

Results about localizations of submodules and quotient modules are provided in this file.

Main results #

TODO #

def Submodule.localized₀ {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :

Let N be a localization of an R-module M at p. This is the localization of an R-submodule of M viewed as an R-submodule of N.

Equations
Instances For
    def Submodule.localized' {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :

    Let S be the localization of R at p and N be a localization of M at p. This is the localization of an R-submodule of M viewed as an S-submodule of N.

    Equations
    Instances For
      theorem Submodule.mem_localized₀ {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) (x : N) :
      x Submodule.localized₀ p f M' mM', ∃ (s : p), IsLocalizedModule.mk' f m s = x
      theorem Submodule.mem_localized' {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) (x : N) :
      x Submodule.localized' S p f M' mM', ∃ (s : p), IsLocalizedModule.mk' f m s = x
      theorem Submodule.restrictScalars_localized' {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :

      localized₀ is the same as localized' considered as a submodule over the base ring.

      @[reducible, inline]
      abbrev Submodule.localized {R : Type u_1} {M : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] (p : Submonoid R) (M' : Submodule R M) :

      The localization of an R-submodule of M at p viewed as an Rₚ-submodule of Mₚ.

      Equations
      Instances For
        @[simp]
        theorem Submodule.localized₀_bot {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] :
        @[simp]
        theorem Submodule.localized'_bot {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] :
        @[simp]
        theorem Submodule.localized₀_top {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] :
        @[simp]
        theorem Submodule.localized'_top {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] :
        @[simp]
        theorem Submodule.localized'_span {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (s : Set M) :
        def Submodule.toLocalized₀ {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :
        M' →ₗ[R] (Submodule.localized₀ p f M')

        The localization map of a submodule.

        Equations
        Instances For
          @[simp]
          theorem Submodule.toLocalized₀_apply_coe {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) (c : M') :
          ((Submodule.toLocalized₀ p f M') c) = f c
          def Submodule.toLocalized' {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :
          M' →ₗ[R] (Submodule.localized' S p f M')

          The localization map of a submodule.

          Equations
          Instances For
            @[simp]
            theorem Submodule.toLocalized'_apply_coe {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) (c : M') :
            ((Submodule.toLocalized' S p f M') c) = f c
            @[reducible, inline]
            abbrev Submodule.toLocalized {R : Type u_1} {M : Type u_3} [CommSemiring R] [AddCommMonoid M] [Module R M] (p : Submonoid R) (M' : Submodule R M) :
            M' →ₗ[R] (Submodule.localized p M')

            The localization map of a submodule.

            Equations
            Instances For
              instance Submodule.isLocalizedModule {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :
              Equations
              • =
              def Submodule.toLocalizedQuotient' {R : Type u_5} (S : Type u_6) {M : Type u_7} {N : Type u_8} [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :

              The localization map of a quotient module.

              Equations
              Instances For
                @[reducible, inline]
                abbrev Submodule.toLocalizedQuotient {R : Type u_5} {M : Type u_7} [CommRing R] [AddCommGroup M] [Module R M] (p : Submonoid R) (M' : Submodule R M) :

                The localization map of a quotient module.

                Equations
                Instances For
                  @[simp]
                  theorem Submodule.toLocalizedQuotient'_mk {R : Type u_5} (S : Type u_6) {M : Type u_7} {N : Type u_8} [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) (x : M) :
                  instance IsLocalizedModule.toLocalizedQuotient' {R : Type u_5} (S : Type u_6) {M : Type u_7} {N : Type u_8} [CommRing R] [CommRing S] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] (M' : Submodule R M) :
                  Equations
                  • =
                  theorem LinearMap.localized'_ker_eq_ker_localizedMap {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] {P : Type u_5} [AddCommGroup P] [Module R P] {Q : Type u_6} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] (g : M →ₗ[R] P) :
                  theorem LinearMap.ker_localizedMap_eq_localized'_ker {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] {P : Type u_5} [AddCommGroup P] [Module R P] {Q : Type u_6} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] (g : M →ₗ[R] P) :
                  noncomputable def LinearMap.toKerIsLocalized {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] {P : Type u_5} [AddCommGroup P] [Module R P] {Q : Type u_6} [AddCommGroup Q] [Module R Q] (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] (g : M →ₗ[R] P) :

                  The canonical map from the kernel of g to the kernel of g localized at a submonoid.

                  This is a localization map by LinearMap.toKerLocalized_isLocalizedModule.

                  Equations
                  Instances For
                    @[simp]
                    theorem LinearMap.toKerIsLocalized_apply_coe {R : Type u_1} {M : Type u_3} {N : Type u_4} [CommSemiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] {P : Type u_5} [AddCommGroup P] [Module R P] {Q : Type u_6} [AddCommGroup Q] [Module R Q] (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] (g : M →ₗ[R] P) (c : (LinearMap.ker g)) :
                    ((LinearMap.toKerIsLocalized p f f' g) c) = f c
                    theorem LinearMap.toKerLocalized_isLocalizedModule {R : Type u_1} (S : Type u_2) {M : Type u_3} {N : Type u_4} [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] [Algebra R S] [Module S N] [IsScalarTower R S N] (p : Submonoid R) [IsLocalization p S] (f : M →ₗ[R] N) [IsLocalizedModule p f] {P : Type u_5} [AddCommGroup P] [Module R P] {Q : Type u_6} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] (f' : P →ₗ[R] Q) [IsLocalizedModule p f'] (g : M →ₗ[R] P) :

                    The canonical map to the kernel of the localization of g is localizing. In other words, localization commutes with kernels.