Documentation

Mathlib.RingTheory.DividedPowers.Basic

Divided powers #

Let A be a commutative (semi)ring and I be an ideal of A. A divided power structure on I is the datum of operations a n ↦ dpow a n satisfying relations that model the intuitive formula dpow n a = a ^ n / n ! and collected by the structure DividedPowers. The list of axioms is embedded in the structure: To avoid coercions, we rather consider DividedPowers.dpow : ℕ → A → A, extended by 0.

References #

Discussion #

structure DividedPowers {A : Type u_1} [CommSemiring A] (I : Ideal A) :
Type u_1

The divided power structure on an ideal I of a commutative ring A

Instances For
    noncomputable def dividedPowersBot (A : Type u_1) [CommSemiring A] :

    The canonical DividedPowers structure on the zero ideal

    Equations
      Instances For
        theorem dividedPowersBot_dpow_eq {A : Type u_1} [CommSemiring A] [DecidableEq A] (n : ) (a : A) :
        (dividedPowersBot A).dpow n a = if a = 0 n = 0 then 1 else 0
        Equations
          instance instCoeFunDividedPowersForallNatForall {A : Type u_1} [CommSemiring A] (I : Ideal A) :
          CoeFun (DividedPowers I) fun (x : DividedPowers I) => AA

          The coercion from the divided powers structures to functions

          Equations
            theorem DividedPowers.ext {A : Type u_1} [CommSemiring A] {I : Ideal A} (hI hI' : DividedPowers I) (h_eq : ∀ (n : ) {x : A}, x IhI.dpow n x = hI'.dpow n x) :
            hI = hI'
            theorem DividedPowers.ext_iff {A : Type u_1} [CommSemiring A] {I : Ideal A} {hI hI' : DividedPowers I} :
            hI = hI' ∀ (n : ) {x : A}, x IhI.dpow n x = hI'.dpow n x
            theorem DividedPowers.dpow_add' {A : Type u_1} [CommSemiring A] {I : Ideal A} {a b : A} (hI : DividedPowers I) {n : } (ha : a I) (hb : b I) :
            hI.dpow n (a + b) = kFinset.range (n + 1), hI.dpow k a * hI.dpow (n - k) b

            Variant of DividedPowers.dpow_add with a sum on range (n + 1)

            def DividedPowers.exp {A : Type u_1} [CommSemiring A] {I : Ideal A} (hI : DividedPowers I) (a : A) :

            The exponential series of an element in the context of divided powers, Σ (dpow n a) X ^ n

            Equations
              Instances For
                theorem DividedPowers.exp_add' {A : Type u_1} [CommSemiring A] {a b : A} (dp : AA) (dp_add : ∀ (n : ), dp n (a + b) = kFinset.antidiagonal n, dp k.1 a * dp k.2 b) :
                (PowerSeries.mk fun (n : ) => dp n (a + b)) = (PowerSeries.mk fun (n : ) => dp n a) * PowerSeries.mk fun (n : ) => dp n b

                A more general of DividedPowers.exp_add

                theorem DividedPowers.exp_add {A : Type u_1} [CommSemiring A] {I : Ideal A} {a b : A} (hI : DividedPowers I) (ha : a I) (hb : b I) :
                hI.exp (a + b) = hI.exp a * hI.exp b
                theorem DividedPowers.dpow_smul {A : Type u_1} [CommSemiring A] {I : Ideal A} {a b : A} (hI : DividedPowers I) {n : } (ha : a I) :
                hI.dpow n (b a) = b ^ n hI.dpow n a
                theorem DividedPowers.dpow_mul_right {A : Type u_1} [CommSemiring A] {I : Ideal A} {a b : A} (hI : DividedPowers I) {n : } (ha : a I) :
                hI.dpow n (a * b) = hI.dpow n a * b ^ n
                theorem DividedPowers.dpow_smul_right {A : Type u_1} [CommSemiring A] {I : Ideal A} {a b : A} (hI : DividedPowers I) {n : } (ha : a I) :
                hI.dpow n (a b) = hI.dpow n a b ^ n
                theorem DividedPowers.factorial_mul_dpow_eq_pow {A : Type u_1} [CommSemiring A] {I : Ideal A} {a : A} (hI : DividedPowers I) {n : } (ha : a I) :
                n.factorial * hI.dpow n a = a ^ n
                theorem DividedPowers.dpow_eval_zero {A : Type u_1} [CommSemiring A] {I : Ideal A} (hI : DividedPowers I) {n : } (hn : n 0) :
                hI.dpow n 0 = 0
                theorem DividedPowers.nilpotent_of_mem_dpIdeal {A : Type u_1} [CommSemiring A] {I : Ideal A} {a : A} {n : } (hn : n 0) (hnI : ∀ {y : A}, y In y = 0) (hI : DividedPowers I) (ha : a I) :
                a ^ n = 0

                If an element of a divided power ideal is killed by multiplication by some nonzero integer n, then its nth power is zero.

                Proposition 1.2.7 of [Berthelot-1974], part (i).

                theorem DividedPowers.coincide_on_smul {A : Type u_1} [CommSemiring A] {I : Ideal A} {a : A} (hI : DividedPowers I) {J : Ideal A} (hJ : DividedPowers J) {n : } (ha : a I J) :
                hI.dpow n a = hJ.dpow n a

                If J is another ideal of A with divided powers, then the divided powers of I and J coincide on I • J

                [Berthelot-1974], 1.6.1 (ii)

                theorem DividedPowers.prod_dpow {A : Type u_1} [CommSemiring A] {I : Ideal A} {a : A} (hI : DividedPowers I) {ι : Type u_2} {s : Finset ι} {n : ι} (ha : a I) :
                is, hI.dpow (n i) a = (Nat.multinomial s n) * hI.dpow (s.sum n) a

                A product of divided powers is a multinomial coefficient times the divided power

                [Roby-1965], formula (III')

                theorem DividedPowers.dpow_sum' {A : Type u_1} [CommSemiring A] {M : Type u_2} [AddCommMonoid M] {I : AddSubmonoid M} (dpow : MA) (dpow_zero : ∀ {x : M}, x Idpow 0 x = 1) (dpow_add : ∀ {n : } {x y : M}, x Iy Idpow n (x + y) = kFinset.antidiagonal n, dpow k.1 x * dpow k.2 y) (dpow_eval_zero : ∀ {n : }, n 0dpow n 0 = 0) {ι : Type u_3} [DecidableEq ι] {s : Finset ι} {x : ιM} (hx : is, x i I) {n : } :
                dpow n (s.sum x) = ks.sym n, is, dpow (Multiset.count i k) (x i)

                Lemma towards dpow_sum when we only have partial information on a divided power ideal

                theorem DividedPowers.dpow_sum {A : Type u_1} [CommSemiring A] {I : Ideal A} (hI : DividedPowers I) {ι : Type u_2} [DecidableEq ι] {s : Finset ι} {x : ιA} (hx : is, x i I) {n : } :
                hI.dpow n (s.sum x) = ks.sym n, is, hI.dpow (Multiset.count i k) (x i)

                A “multinomial” theorem for divided powers — without multinomial coefficients

                def DividedPowers.ofRingEquiv {A : Type u_1} {B : Type u_2} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} {e : A ≃+* B} (h : Ideal.map e I = J) (hI : DividedPowers I) :

                Transfer divided powers under an equivalence

                Equations
                  Instances For
                    @[simp]
                    theorem DividedPowers.ofRingEquiv_dpow {A : Type u_1} {B : Type u_2} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} {e : A ≃+* B} (h : Ideal.map e I = J) (hI : DividedPowers I) {n : } {b : B} :
                    (ofRingEquiv h hI).dpow n b = e (hI.dpow n (e.symm b))
                    theorem DividedPowers.ofRingEquiv_dpow_apply {A : Type u_1} {B : Type u_2} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} {e : A ≃+* B} (h : Ideal.map e I = J) (hI : DividedPowers I) {n : } {a : A} :
                    (ofRingEquiv h hI).dpow n (e a) = e (hI.dpow n a)
                    def DividedPowers.equiv {A : Type u_1} {B : Type u_2} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} {e : A ≃+* B} (h : Ideal.map e I = J) :

                    Transfer divided powers under an equivalence (Equiv version)

                    Equations
                      Instances For
                        theorem DividedPowers.equiv_apply {A : Type u_1} {B : Type u_2} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} {e : A ≃+* B} (h : Ideal.map e I = J) (hI : DividedPowers I) (n : ) (b : B) :
                        ((equiv h) hI).dpow n b = e (hI.dpow n (e.symm b))
                        theorem DividedPowers.equiv_apply' {A : Type u_1} {B : Type u_2} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} {e : A ≃+* B} (h : Ideal.map e I = J) (hI : DividedPowers I) {n : } {a : A} :
                        ((equiv h) hI).dpow n (e a) = e (hI.dpow n a)

                        Variant of DividedPowers.equiv_apply