Skip to content

Commit bc6f44c

Browse files
committed
chore : fix camelcase naming of PDF, CDF and PMF (#9681)
Replaces occurences of Pdf, Cdf and Pmf within camelCase situations by PDF, CDF and PMF respectively.
1 parent be8dd31 commit bc6f44c

File tree

7 files changed

+539
-539
lines changed

7 files changed

+539
-539
lines changed

Mathlib/Probability/Cdf.lean

Lines changed: 11 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -15,11 +15,11 @@ Two probability measures are equal if and only if they have the same cdf.
1515
## Main definitions
1616
1717
* `ProbabilityTheory.cdf μ`: cumulative distribution function of `μ : Measure ℝ`, defined as the
18-
conditional cdf (`ProbabilityTheory.condCdf`) of the product measure
18+
conditional cdf (`ProbabilityTheory.condCDF`) of the product measure
1919
`(Measure.dirac Unit.unit).prod μ` evaluated at `Unit.unit`.
2020
2121
The definition could be replaced by the more elementary `cdf μ x = (μ (Iic x)).toReal`, but using
22-
`condCdf` gives us access to its API, from which most properties of the cdf follow directly.
22+
`condCDF` gives us access to its API, from which most properties of the cdf follow directly.
2323
2424
## Main statements
2525
@@ -30,11 +30,11 @@ The definition could be replaced by the more elementary `cdf μ x = (μ (Iic x))
3030
3131
## TODO
3232
33-
The definition could be extended to a finite measure by rescaling `condCdf`, but it would be nice
33+
The definition could be extended to a finite measure by rescaling `condCDF`, but it would be nice
3434
to have more structure on Stieltjes functions first. Right now, if `f` is a Stieltjes function,
3535
`2 • f` makes no sense. We could define Stieltjes functions as a submodule.
3636
37-
The definition could be extended to `ℝⁿ`, either by extending the definition of `condCdf`, or by
37+
The definition could be extended to `ℝⁿ`, either by extending the definition of `condCDF`, or by
3838
using another construction here.
3939
-/
4040

@@ -49,29 +49,29 @@ for probability measures. In that case, it satisfies `cdf μ x = (μ (Iic x)).to
4949
`ProbabilityTheory.cdf_eq_toReal`). -/
5050
noncomputable
5151
def cdf (μ : Measure ℝ) : StieltjesFunction :=
52-
condCdf ((Measure.dirac Unit.unit).prod μ) Unit.unit
52+
condCDF ((Measure.dirac Unit.unit).prod μ) Unit.unit
5353

5454
section ExplicitMeasureArg
5555
variable (μ : Measure ℝ)
5656

5757
/-- The cdf is non-negative. -/
58-
lemma cdf_nonneg (x : ℝ) : 0 ≤ cdf μ x := condCdf_nonneg _ _ _
58+
lemma cdf_nonneg (x : ℝ) : 0 ≤ cdf μ x := condCDF_nonneg _ _ _
5959

6060
/-- The cdf is lower or equal to 1. -/
61-
lemma cdf_le_one (x : ℝ) : cdf μ x ≤ 1 := condCdf_le_one _ _ _
61+
lemma cdf_le_one (x : ℝ) : cdf μ x ≤ 1 := condCDF_le_one _ _ _
6262

6363
/-- The cdf is monotone. -/
64-
lemma monotone_cdf : Monotone (cdf μ) := (condCdf _ _).mono
64+
lemma monotone_cdf : Monotone (cdf μ) := (condCDF _ _).mono
6565

6666
/-- The cdf tends to 0 at -∞. -/
67-
lemma tendsto_cdf_atBot : Tendsto (cdf μ) atBot (𝓝 0) := tendsto_condCdf_atBot _ _
67+
lemma tendsto_cdf_atBot : Tendsto (cdf μ) atBot (𝓝 0) := tendsto_condCDF_atBot _ _
6868

6969
/-- The cdf tends to 1 at +∞. -/
70-
lemma tendsto_cdf_atTop : Tendsto (cdf μ) atTop (𝓝 1) := tendsto_condCdf_atTop _ _
70+
lemma tendsto_cdf_atTop : Tendsto (cdf μ) atTop (𝓝 1) := tendsto_condCDF_atTop _ _
7171

7272
lemma ofReal_cdf [IsProbabilityMeasure μ] (x : ℝ) : ENNReal.ofReal (cdf μ x) = μ (Iic x) := by
7373
have := IsProbabilityMeasure.toIsFiniteMeasure (Measure.prod (Measure.dirac ()) μ)
74-
have h := lintegral_condCdf ((Measure.dirac Unit.unit).prod μ) x
74+
have h := lintegral_condCDF ((Measure.dirac Unit.unit).prod μ) x
7575
simpa only [MeasureTheory.Measure.fst_prod, Measure.prod_prod, measure_univ, one_mul,
7676
lintegral_dirac] using h
7777

Mathlib/Probability/Distributions/Exponential.lean

Lines changed: 47 additions & 47 deletions
Original file line numberDiff line numberDiff line change
@@ -14,17 +14,17 @@ import Mathlib.Probability.Distributions.Gamma
1414
Define the Exponential measure over the reals.
1515
1616
## Main definitions
17-
* `exponentialPdfReal`: the function `r x ↦ r * exp (-(r * x)` for `0 ≤ x`
17+
* `exponentialPDFReal`: the function `r x ↦ r * exp (-(r * x)` for `0 ≤ x`
1818
or `0` else, which is the probability density function of a exponential distribution with
1919
rate `r` (when `hr : 0 < r`).
20-
* `exponentialPdf`: `ℝ≥0∞`-valued pdf,
21-
`exponentialPdf r = ENNReal.ofReal (exponentialPdfReal r)`.
20+
* `exponentialPDF`: `ℝ≥0∞`-valued pdf,
21+
`exponentialPDF r = ENNReal.ofReal (exponentialPDFReal r)`.
2222
* `expMeasure`: an exponential measure on `ℝ`, parametrized by its rate `r`.
23-
* `exponentialCdfReal`: the Cdf given by the definition of CDF in `ProbabilityTheory.Cdf` applied to
23+
* `exponentialCDFReal`: the CDF given by the definition of CDF in `ProbabilityTheory.CDF` applied to
2424
the exponential measure.
2525
2626
## Main results
27-
* `exponentialCdfReal_eq`: Proof that the `exponentialCdfReal` given by the definition equals the
27+
* `exponentialCDFReal_eq`: Proof that the `exponentialCDFReal` given by the definition equals the
2828
known function given as `r x ↦ 1 - exp (- (r * x))` for `0 ≤ x` or `0` else.
2929
-/
3030

@@ -34,59 +34,59 @@ open MeasureTheory Real Set Filter Topology
3434

3535
namespace ProbabilityTheory
3636

37-
section ExponentialPdf
37+
section ExponentialPDF
3838

3939
/-- The pdf of the exponential distribution depending on its rate -/
4040
noncomputable
41-
def exponentialPdfReal (r x : ℝ) : ℝ :=
42-
gammaPdfReal 1 r x
41+
def exponentialPDFReal (r x : ℝ) : ℝ :=
42+
gammaPDFReal 1 r x
4343

4444
/-- The pdf of the exponential distribution, as a function valued in `ℝ≥0∞` -/
4545
noncomputable
46-
def exponentialPdf (r x : ℝ) : ℝ≥0∞ :=
47-
ENNReal.ofReal (exponentialPdfReal r x)
46+
def exponentialPDF (r x : ℝ) : ℝ≥0∞ :=
47+
ENNReal.ofReal (exponentialPDFReal r x)
4848

49-
lemma exponentialPdf_eq (r x : ℝ) :
50-
exponentialPdf r x = ENNReal.ofReal (if 0 ≤ x then r * exp (-(r * x)) else 0) := by
51-
rw [exponentialPdf, exponentialPdfReal, gammaPdfReal]
49+
lemma exponentialPDF_eq (r x : ℝ) :
50+
exponentialPDF r x = ENNReal.ofReal (if 0 ≤ x then r * exp (-(r * x)) else 0) := by
51+
rw [exponentialPDF, exponentialPDFReal, gammaPDFReal]
5252
simp only [rpow_one, Gamma_one, div_one, sub_self, rpow_zero, mul_one]
5353

54-
lemma exponentialPdf_of_neg {r x : ℝ} (hx : x < 0) : exponentialPdf r x = 0 := gammaPdf_of_neg hx
54+
lemma exponentialPDF_of_neg {r x : ℝ} (hx : x < 0) : exponentialPDF r x = 0 := gammaPDF_of_neg hx
5555

56-
lemma exponentialPdf_of_nonneg {r x : ℝ} (hx : 0 ≤ x) :
57-
exponentialPdf r x = ENNReal.ofReal (r * rexp (-(r * x))) := by
58-
simp only [exponentialPdf_eq, if_pos hx]
56+
lemma exponentialPDF_of_nonneg {r x : ℝ} (hx : 0 ≤ x) :
57+
exponentialPDF r x = ENNReal.ofReal (r * rexp (-(r * x))) := by
58+
simp only [exponentialPDF_eq, if_pos hx]
5959

6060
/-- The Lebesgue integral of the exponential pdf over nonpositive reals equals 0-/
61-
lemma lintegral_exponentialPdf_of_nonpos {x r : ℝ} (hx : x ≤ 0) :
62-
∫⁻ y in Iio x, exponentialPdf r y = 0 := lintegral_gammaPdf_of_nonpos hx
61+
lemma lintegral_exponentialPDF_of_nonpos {x r : ℝ} (hx : x ≤ 0) :
62+
∫⁻ y in Iio x, exponentialPDF r y = 0 := lintegral_gammaPDF_of_nonpos hx
6363

6464
/-- The exponential pdf is measurable. -/
6565
@[measurability]
66-
lemma measurable_exponentialPdfReal (r : ℝ) : Measurable (exponentialPdfReal r) :=
67-
measurable_gammaPdfReal 1 r
66+
lemma measurable_exponentialPDFReal (r : ℝ) : Measurable (exponentialPDFReal r) :=
67+
measurable_gammaPDFReal 1 r
6868

6969
-- The exponential pdf is strongly measurable -/
7070
@[measurability]
71-
lemma stronglyMeasurable_exponentialPdfReal (r : ℝ) :
72-
StronglyMeasurable (exponentialPdfReal r) := stronglyMeasurable_gammaPdfReal 1 r
71+
lemma stronglyMeasurable_exponentialPDFReal (r : ℝ) :
72+
StronglyMeasurable (exponentialPDFReal r) := stronglyMeasurable_gammaPDFReal 1 r
7373

7474
/-- The exponential pdf is positive for all positive reals -/
75-
lemma exponentialPdfReal_pos {x r : ℝ} (hr : 0 < r) (hx : 0 < x) :
76-
0 < exponentialPdfReal r x := gammaPdfReal_pos zero_lt_one hr hx
75+
lemma exponentialPDFReal_pos {x r : ℝ} (hr : 0 < r) (hx : 0 < x) :
76+
0 < exponentialPDFReal r x := gammaPDFReal_pos zero_lt_one hr hx
7777

7878
/-- The exponential pdf is nonnegative-/
79-
lemma exponentialPdfReal_nonneg {r : ℝ} (hr : 0 < r) (x : ℝ) :
80-
0exponentialPdfReal r x := gammaPdfReal_nonneg zero_lt_one hr x
79+
lemma exponentialPDFReal_nonneg {r : ℝ} (hr : 0 < r) (x : ℝ) :
80+
0exponentialPDFReal r x := gammaPDFReal_nonneg zero_lt_one hr x
8181

8282
open Measure
8383

8484
/-- The pdf of the exponential distribution integrates to 1 -/
8585
@[simp]
86-
lemma lintegral_exponentialPdf_eq_one {r : ℝ} (hr : 0 < r) : ∫⁻ x, exponentialPdf r x = 1 :=
87-
lintegral_gammaPdf_eq_one zero_lt_one hr
86+
lemma lintegral_exponentialPDF_eq_one {r : ℝ} (hr : 0 < r) : ∫⁻ x, exponentialPDF r x = 1 :=
87+
lintegral_gammaPDF_eq_one zero_lt_one hr
8888

89-
end ExponentialPdf
89+
end ExponentialPDF
9090

9191
open MeasureTheory
9292

@@ -97,20 +97,20 @@ def expMeasure (r : ℝ) : Measure ℝ := gammaMeasure 1 r
9797
lemma isProbabilityMeasureExponential {r : ℝ} (hr : 0 < r) :
9898
IsProbabilityMeasure (expMeasure r) := isProbabilityMeasureGamma zero_lt_one hr
9999

100-
section ExponentialCdf
100+
section ExponentialCDF
101101

102102
/-- CDF of the exponential distribution -/
103103
noncomputable
104-
def exponentialCdfReal (r : ℝ) : StieltjesFunction :=
104+
def exponentialCDFReal (r : ℝ) : StieltjesFunction :=
105105
cdf (expMeasure r)
106106

107-
lemma exponentialCdfReal_eq_integral {r : ℝ} (hr : 0 < r) (x : ℝ) :
108-
exponentialCdfReal r x = ∫ x in Iic x, exponentialPdfReal r x :=
109-
gammaCdfReal_eq_integral zero_lt_one hr x
107+
lemma exponentialCDFReal_eq_integral {r : ℝ} (hr : 0 < r) (x : ℝ) :
108+
exponentialCDFReal r x = ∫ x in Iic x, exponentialPDFReal r x :=
109+
gammaCDFReal_eq_integral zero_lt_one hr x
110110

111-
lemma exponentialCdfReal_eq_lintegral {r : ℝ} (hr : 0 < r) (x : ℝ) :
112-
exponentialCdfReal r x = ENNReal.toReal (∫⁻ x in Iic x, exponentialPdf r x) :=
113-
gammaCdfReal_eq_lintegral zero_lt_one hr x
111+
lemma exponentialCDFReal_eq_lintegral {r : ℝ} (hr : 0 < r) (x : ℝ) :
112+
exponentialCDFReal r x = ENNReal.toReal (∫⁻ x in Iic x, exponentialPDF r x) :=
113+
gammaCDFReal_eq_lintegral zero_lt_one hr x
114114

115115
open Topology
116116

@@ -126,18 +126,18 @@ lemma exp_neg_integrableOn_Ioc {b x : ℝ} (hb : 0 < b) :
126126
simp only [neg_mul_eq_neg_mul]
127127
exact (exp_neg_integrableOn_Ioi _ hb).mono_set Ioc_subset_Ioi_self
128128

129-
lemma lintegral_exponentialPdf_eq_antiDeriv {r : ℝ} (hr : 0 < r) (x : ℝ) :
130-
∫⁻ y in Iic x, exponentialPdf r y
129+
lemma lintegral_exponentialPDF_eq_antiDeriv {r : ℝ} (hr : 0 < r) (x : ℝ) :
130+
∫⁻ y in Iic x, exponentialPDF r y
131131
= ENNReal.ofReal (if 0 ≤ x then 1 - exp (-(r * x)) else 0) := by
132132
split_ifs with h
133133
case neg =>
134-
simp only [exponentialPdf_eq]
134+
simp only [exponentialPDF_eq]
135135
rw [set_lintegral_congr_fun measurableSet_Iic, lintegral_zero, ENNReal.ofReal_zero]
136136
exact ae_of_all _ fun a (_ : a ≤ _) ↦ by rw [if_neg (by linarith), ENNReal.ofReal_eq_zero]
137137
case pos =>
138-
rw [lintegral_Iic_eq_lintegral_Iio_add_Icc _ h, lintegral_exponentialPdf_of_nonpos (le_refl 0),
138+
rw [lintegral_Iic_eq_lintegral_Iio_add_Icc _ h, lintegral_exponentialPDF_of_nonpos (le_refl 0),
139139
zero_add]
140-
simp only [exponentialPdf_eq]
140+
simp only [exponentialPDF_eq]
141141
rw [set_lintegral_congr_fun measurableSet_Icc (ae_of_all _
142142
(by intro a ⟨(hle : _ ≤ a), _⟩; rw [if_pos hle]))]
143143
rw [← ENNReal.toReal_eq_toReal _ ENNReal.ofReal_ne_top, ← integral_eq_lintegral_of_nonneg_ae
@@ -164,13 +164,13 @@ lemma lintegral_exponentialPdf_eq_antiDeriv {r : ℝ} (hr : 0 < r) (x : ℝ) :
164164
exact Integrable.const_mul (exp_neg_integrableOn_Ioc hr) _
165165

166166
/-- The CDF of the exponential distribution equals ``1 - exp (-(r * x))``-/
167-
lemma exponentialCdfReal_eq {r : ℝ} (hr : 0 < r) (x : ℝ) :
168-
exponentialCdfReal r x = if 0 ≤ x then 1 - exp (-(r * x)) else 0 := by
169-
rw [exponentialCdfReal_eq_lintegral hr, lintegral_exponentialPdf_eq_antiDeriv hr x,
167+
lemma exponentialCDFReal_eq {r : ℝ} (hr : 0 < r) (x : ℝ) :
168+
exponentialCDFReal r x = if 0 ≤ x then 1 - exp (-(r * x)) else 0 := by
169+
rw [exponentialCDFReal_eq_lintegral hr, lintegral_exponentialPDF_eq_antiDeriv hr x,
170170
ENNReal.toReal_ofReal_eq_iff]
171171
split_ifs with h
172172
· simp only [sub_nonneg, exp_le_one_iff, Left.neg_nonpos_iff, gt_iff_lt, ge_iff_le]
173173
exact mul_nonneg hr.le h
174174
· exact le_rfl
175175

176-
end ExponentialCdf
176+
end ExponentialCDF

0 commit comments

Comments
 (0)