Documentation

Mathlib.MeasureTheory.Integral.IntervalIntegral.TrapezoidalRule

The trapezoidal rule #

This file contains a definition of integration on [[a, b]] via the trapezoidal rule, along with an error bound in terms of a bound on the second derivative of the integrand.

Main results #

References #

We follow the proof on (Wikipedia)[https://en.wikipedia.org/wiki/Trapezoidal_rule] for the error bound.

noncomputable def trapezoidal_integral (f : ) (N : ) (a b : ) :

Integration of f from a to b using the trapezoidal rule with N+1 total evaluations of f. (Note the off-by-one problem here: N counts the number of trapezoids, not the number of evaluations.)

Equations
Instances For
    noncomputable def trapezoidal_error (f : ) (N : ) (a b : ) :

    The absolute error of trapezoidal integration.

    Equations
    Instances For
      theorem trapezoidal_integral_symm (f : ) {N : } (N_nonzero : 0 < N) (a b : ) :

      Just like exact integration, the trapezoidal approximation retains the same magnitude but changes sign when the endpoints are swapped.

      theorem trapezoidal_error_symm (f : ) {N : } (N_nonzero : 0 < N) (a b : ) :

      The absolute error of the trapezoidal rule does not change when the endpoints are swapped.

      @[simp]
      theorem trapezoidal_integral_eq (f : ) (N : ) (a : ) :

      Just like exact integration, the trapezoidal integration from a to a is zero.

      @[simp]
      theorem trapezoidal_error_eq (f : ) (N : ) (a : ) :

      The error of the trapezoidal integration from a to a is zero.

      @[simp]
      theorem trapezoidal_integral_one (f : ) (a b : ) :
      trapezoidal_integral f 1 a b = (b - a) / 2 * (f a + f b)

      An exact formula for integration with a single trapezoid (the "midpoint rule").

      theorem sum_trapezoidal_integral_adjacent_intervals {f : } {N : } {a h : } (N_nonzero : 0 < N) :
      iFinset.range N, trapezoidal_integral f 1 (a + i * h) (a + (i + 1) * h) = trapezoidal_integral f N a (a + N * h)

      A basic trapezoidal equivalent to IntervalIntegral.sum_integral_adjacent_intervals. More general theorems are certainly possible, but many of them can be derived from repeated applications of this one.

      theorem trapezoidal_integral_ext {f : } {N : } {a h : } (N_nonzero : 0 < N) :
      trapezoidal_integral f N a (a + N * h) + trapezoidal_integral f 1 (a + N * h) (a + (N + 1) * h) = trapezoidal_integral f (N + 1) a (a + (N + 1) * h)

      A simplified version of the previous theorem, for use in proofs by induction and the like.

      theorem sum_trapezoidal_error_adjacent_intervals {f : } {N : } {a h : } (N_nonzero : 0 < N) (h_f_int : IntervalIntegrable f MeasureTheory.volume a (a + N * h)) :
      iFinset.range N, trapezoidal_error f 1 (a + i * h) (a + (i + 1) * h) = trapezoidal_error f N a (a + N * h)

      Since we have sum_[]_adjacent_intervals theorems for both exact and trapezoidal integration, it's natural to combine them into a similar formula for the error. This theorem is in particular used in the proof of the general error bound.

      theorem trapezoidal_error_le {f : } {a b : } (h_df : DifferentiableOn f (Set.uIcc a b)) (h_ddf : DifferentiableOn (derivWithin f (Set.uIcc a b)) (Set.uIcc a b)) (h_ddf_integrable : IntervalIntegrable (iteratedDerivWithin 2 f (Set.uIcc a b)) MeasureTheory.volume a b) {ζ : } (fpp_bound : ∀ (x : ), |iteratedDerivWithin 2 f (Set.uIcc a b) x| ζ) {N : } (N_nonzero : 0 < N) :
      |trapezoidal_error f N a b| |b - a| ^ 3 * ζ / (12 * N ^ 2)

      The standard error bound for trapezoidal integration on the general interval [[a, b]].

      theorem trapezoidal_error_le_of_c2 {f : } {a b : } (h_f_c2 : ContDiffOn 2 f (Set.uIcc a b)) {ζ : } (fpp_bound : ∀ (x : ), |iteratedDerivWithin 2 f (Set.uIcc a b) x| ζ) {N : } (N_nonzero : 0 < N) :
      |trapezoidal_error f N a b| |b - a| ^ 3 * ζ / (12 * N ^ 2)

      The error bound for trapezoidal integration in the slightly weaker, but very common, case where f is C^2.