Documentation

Mathlib.Algebra.Group.ForwardDiff

Forward difference operators and Newton series #

We define the forward difference operator, sending f to the function x ↦ f (x + h) - f x for a given h (for any additive semigroup, taking values in an abelian group). The notation Δ_[h] is defined for this operator, scoped in namespace fwdDiff.

We prove two key formulae about this operator:

We also prove some auxiliary results about iterated forward differences of the function n ↦ n.choose k.

def fwdDiff {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f : MG) :
MG

Forward difference operator, fwdDiff h f n = f (n + h) - f n. The notation Δ_[h] for this operator is available in the fwdDiff namespace.

Equations
Instances For

    Forward difference operator, fwdDiff h f n = f (n + h) - f n. The notation Δ_[h] for this operator is available in the fwdDiff namespace.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      @[simp]
      theorem fwdDiff_add {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f g : MG) :
      fwdDiff h (f + g) = fwdDiff h f + fwdDiff h g
      @[simp]
      theorem fwdDiff_const {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (g : G) :
      (fwdDiff h fun (x : M) => g) = fun (x : M) => 0
      theorem fwdDiff_smul {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) {R : Type} [Ring R] [Module R G] (f : MR) (g : MG) :
      fwdDiff h (f g) = fwdDiff h f g + f fwdDiff h g + fwdDiff h f fwdDiff h g
      @[simp]
      theorem fwdDiff_const_smul {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) {R : Type u_3} [Semiring R] [Module R G] (r : R) (f : MG) :
      fwdDiff h (r f) = r fwdDiff h f
      @[simp]
      theorem fwdDiff_smul_const {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) {R : Type} [Ring R] [Module R G] (f : MR) (g : G) :
      (fwdDiff h fun (y : M) => f y g) = fwdDiff h f fun (x : M) => g

      Forward-difference and shift operators as linear endomorphisms #

      This section contains versions of the forward-difference operator and the shift operator bundled as -linear endomorphisms. These are useful for certain proofs; but they are slightly annoying to use, as the source and target types of the maps have to be specified each time, and various coercions need to be un-wound when the operators are applied, so we also provide the un-bundled version.

      def fwdDiff_aux.fwdDiffₗ (M : Type u_1) (G : Type u_2) [AddCommMonoid M] [AddCommGroup G] (h : M) :
      Module.End (MG)

      Linear-endomorphism version of the forward difference operator.

      Equations
      Instances For
        @[simp]
        theorem fwdDiff_aux.fwdDiffₗ_apply (M : Type u_1) (G : Type u_2) [AddCommMonoid M] [AddCommGroup G] (h : M) (f : MG) (a✝ : M) :
        (fwdDiff_aux.fwdDiffₗ M G h) f a✝ = fwdDiff h f a✝
        theorem fwdDiff_aux.coe_fwdDiffₗ {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) :
        theorem fwdDiff_aux.coe_fwdDiffₗ_pow {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (n : ) :
        def fwdDiff_aux.shiftₗ (M : Type u_1) (G : Type u_2) [AddCommMonoid M] [AddCommGroup G] (h : M) :
        Module.End (MG)

        Linear-endomorphism version of the shift-by-1 operator.

        Equations
        Instances For
          theorem fwdDiff_aux.shiftₗ_apply {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f : MG) (y : M) :
          (fwdDiff_aux.shiftₗ M G h) f y = f (y + h)
          theorem fwdDiff_aux.shiftₗ_pow_apply {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f : MG) (k : ) (y : M) :
          (fwdDiff_aux.shiftₗ M G h ^ k) f y = f (y + k h)
          @[simp]
          theorem fwdDiff_finset_sum {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) {α : Type u_3} (s : Finset α) (f : αMG) :
          fwdDiff h (∑ ks, f k) = ks, fwdDiff h (f k)
          @[simp]
          theorem fwdDiff_iter_add {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f g : MG) (n : ) :
          (fwdDiff h)^[n] (f + g) = (fwdDiff h)^[n] f + (fwdDiff h)^[n] g
          @[simp]
          theorem fwdDiff_iter_const_smul {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) {R : Type u_3} [Semiring R] [Module R G] (r : R) (f : MG) (n : ) :
          (fwdDiff h)^[n] (r f) = r (fwdDiff h)^[n] f
          @[simp]
          theorem fwdDiff_iter_finset_sum {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) {α : Type u_3} (s : Finset α) (f : αMG) (n : ) :
          (fwdDiff h)^[n] (∑ ks, f k) = ks, (fwdDiff h)^[n] (f k)
          theorem fwdDiff_iter_eq_sum_shift {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f : MG) (n : ) (y : M) :
          (fwdDiff h)^[n] f y = kFinset.range (n + 1), ((-1) ^ (n - k) * (n.choose k)) f (y + k h)

          Express the n-th forward difference of f at y in terms of the values f (y + k), for 0 ≤ k ≤ n.

          theorem shift_eq_sum_fwdDiff_iter {M : Type u_1} {G : Type u_2} [AddCommMonoid M] [AddCommGroup G] (h : M) (f : MG) (n : ) (y : M) :
          f (y + n h) = kFinset.range (n + 1), n.choose k (fwdDiff h)^[k] f y

          Gregory-Newton formula expressing f (y + n • h) in terms of the iterated forward differences of f at y.

          theorem fwdDiff_choose (j : ) :
          (fwdDiff 1 fun (x : ) => (x.choose (j + 1))) = fun (x : ) => (x.choose j)
          theorem fwdDiff_iter_choose (j k : ) :
          ((fwdDiff 1)^[k] fun (x : ) => (x.choose (k + j))) = fun (x : ) => (x.choose j)
          theorem fwdDiff_iter_choose_zero (m n : ) :
          (fwdDiff 1)^[n] (fun (x : ) => (x.choose m)) 0 = if n = m then 1 else 0