Documentation

Mathlib.Probability.Kernel.Disintegration.MeasurableStieltjes

Measurable parametric Stieltjes functions #

We provide tools to build a measurable function α → StieltjesFunction with limits 0 at -∞ and 1 at +∞ for all a : α from a measurable function f : α → ℚ → ℝ. These measurable parametric Stieltjes functions are cumulative distribution functions (CDF) of transition kernels. The reason for going through instead of defining directly a Stieltjes function is that since is countable, building a measurable function is easier and we can obtain properties of the form ∀ᵐ (a : α) ∂μ, ∀ (q : ℚ), ... (for some measure μ on α) by proving the weaker ∀ (q : ℚ), ∀ᵐ (a : α) ∂μ, ....

This construction will be possible if f a : ℚ → ℝ satisfies a package of properties for all a: monotonicity, limits at +-∞ and a continuity property. We define IsRatStieltjesPoint f a to state that this is the case at a and define the property IsMeasurableRatCDF f that f is measurable and IsRatStieltjesPoint f a for all a. The function α → StieltjesFunction obtained by extending f by continuity from the right is then called IsMeasurableRatCDF.stieltjesFunction.

In applications, we will often only have IsRatStieltjesPoint f a almost surely with respect to some measure. In order to turn that almost everywhere property into an everywhere property we define toRatCDF (f : α → ℚ → ℝ) := fun a q ↦ if IsRatStieltjesPoint f a then f a q else defaultRatCDF q, which satisfies the property IsMeasurableRatCDF (toRatCDF f).

Finally, we define stieltjesOfMeasurableRat, composition of toRatCDF and IsMeasurableRatCDF.stieltjesFunction.

Main definitions #

structure ProbabilityTheory.IsRatStieltjesPoint {α : Type u_1} (f : α) (a : α) :

a : α is a Stieltjes point for f : α → ℚ → ℝ if f a is monotone with limit 0 at -∞ and 1 at +∞ and satisfies a continuity property.

Instances For
    theorem ProbabilityTheory.IsRatStieltjesPoint.mono {α : Type u_1} {f : α} {a : α} (self : ProbabilityTheory.IsRatStieltjesPoint f a) :
    Monotone (f a)
    theorem ProbabilityTheory.IsRatStieltjesPoint.tendsto_atTop_one {α : Type u_1} {f : α} {a : α} (self : ProbabilityTheory.IsRatStieltjesPoint f a) :
    Filter.Tendsto (f a) Filter.atTop (nhds 1)
    theorem ProbabilityTheory.IsRatStieltjesPoint.tendsto_atBot_zero {α : Type u_1} {f : α} {a : α} (self : ProbabilityTheory.IsRatStieltjesPoint f a) :
    Filter.Tendsto (f a) Filter.atBot (nhds 0)
    theorem ProbabilityTheory.IsRatStieltjesPoint.iInf_rat_gt_eq {α : Type u_1} {f : α} {a : α} (self : ProbabilityTheory.IsRatStieltjesPoint f a) (t : ) :
    ⨅ (r : (Set.Ioi t)), f a r = f a t
    theorem ProbabilityTheory.IsRatStieltjesPoint.ite {α : Type u_1} {f : α} {g : α} {a : α} (p : αProp) [DecidablePred p] (hf : p aProbabilityTheory.IsRatStieltjesPoint f a) (hg : ¬p aProbabilityTheory.IsRatStieltjesPoint g a) :
    ProbabilityTheory.IsRatStieltjesPoint (fun (a : α) => if p a then f a else g a) a
    structure ProbabilityTheory.IsMeasurableRatCDF {α : Type u_1} [MeasurableSpace α] (f : α) :

    A function f : α → ℚ → ℝ is a (kernel) rational cumulative distribution function if it is measurable in the first argument and if f a satisfies a list of properties for all a : α: monotonicity between 0 at -∞ and 1 at +∞ and a form of continuity.

    A function with these properties can be extended to a measurable function α → StieltjesFunction. See ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunction.

    Instances For
      theorem ProbabilityTheory.IsMeasurableRatCDF.nonneg {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (q : ) :
      0 f a q
      theorem ProbabilityTheory.IsMeasurableRatCDF.le_one {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (q : ) :
      f a q 1
      theorem ProbabilityTheory.IsMeasurableRatCDF.iInf_rat_gt_eq {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (q : ) :
      ⨅ (r : (Set.Ioi q)), f a r = f a q

      A function with the property IsMeasurableRatCDF. Used in a piecewise construction to convert a function which only satisfies the properties defining IsMeasurableRatCDF on some set into a true IsMeasurableRatCDF.

      Equations
      Instances For
        noncomputable def ProbabilityTheory.toRatCDF {α : Type u_1} (f : α) :
        α

        Turn a function f : α → ℚ → ℝ into another with the property IsRatStieltjesPoint f a everywhere. At a that does not satisfy that property, f a is replaced by an arbitrary suitable function. Mainly useful when f satisfies the property IsRatStieltjesPoint f a almost everywhere with respect to some measure.

        Equations
        Instances For
          theorem ProbabilityTheory.toRatCDF_unit_prod {α : Type u_1} {f : α} (a : α) :
          @[irreducible]
          noncomputable def ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunctionAux {α : Type u_4} (f : α) :
          α

          Auxiliary definition for IsMeasurableRatCDF.stieltjesFunction: turn f : α → ℚ → ℝ into a function α → ℝ → ℝ by assigning to f a x the infimum of f a q over q : ℚ with x < q.

          Equations
          Instances For
            theorem ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunctionAux_def {α : Type u_4} (f : α) (a : α) (x : ) :
            ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunctionAux f a x = ⨅ (q : { q' : // x < q' }), f a q
            theorem ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunctionAux_def' {α : Type u_1} (f : α) (a : α) :
            ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunctionAux f a = fun (t : ) => ⨅ (r : { r' : // t < r' }), f a r

            Extend a function f : α → ℚ → ℝ with property IsMeasurableRatCDF from to , to a function α → StieltjesFunction.

            Equations
            Instances For
              theorem ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunction_eq {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (r : ) :
              (hf.stieltjesFunction a) r = f a r
              theorem ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunction_nonneg {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (r : ) :
              0 (hf.stieltjesFunction a) r
              theorem ProbabilityTheory.IsMeasurableRatCDF.stieltjesFunction_le_one {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (x : ) :
              (hf.stieltjesFunction a) x 1
              theorem ProbabilityTheory.IsMeasurableRatCDF.tendsto_stieltjesFunction_atBot {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) :
              Filter.Tendsto ((hf.stieltjesFunction a)) Filter.atBot (nhds 0)
              theorem ProbabilityTheory.IsMeasurableRatCDF.tendsto_stieltjesFunction_atTop {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) :
              Filter.Tendsto ((hf.stieltjesFunction a)) Filter.atTop (nhds 1)
              theorem ProbabilityTheory.IsMeasurableRatCDF.measurable_stieltjesFunction {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (x : ) :
              Measurable fun (a : α) => (hf.stieltjesFunction a) x
              theorem ProbabilityTheory.IsMeasurableRatCDF.measure_stieltjesFunction_Iic {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) (x : ) :
              (hf.stieltjesFunction a).measure (Set.Iic x) = ENNReal.ofReal ((hf.stieltjesFunction a) x)
              theorem ProbabilityTheory.IsMeasurableRatCDF.measure_stieltjesFunction_univ {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) (a : α) :
              (hf.stieltjesFunction a).measure Set.univ = 1
              theorem ProbabilityTheory.IsMeasurableRatCDF.measurable_measure_stieltjesFunction {α : Type u_1} [MeasurableSpace α] {f : α} (hf : ProbabilityTheory.IsMeasurableRatCDF f) :
              Measurable fun (a : α) => (hf.stieltjesFunction a).measure
              noncomputable def ProbabilityTheory.stieltjesOfMeasurableRat {α : Type u_1} [MeasurableSpace α] (f : α) (hf : Measurable f) :

              Turn a measurable function f : α → ℚ → ℝ into a measurable function α → StieltjesFunction. Composition of toRatCDF and IsMeasurableRatCDF.stieltjesFunction.

              Equations
              Instances For
                theorem ProbabilityTheory.measure_stieltjesOfMeasurableRat_univ {α : Type u_1} [MeasurableSpace α] {f : α} (hf : Measurable f) (a : α) :
                (ProbabilityTheory.stieltjesOfMeasurableRat f hf a).measure Set.univ = 1