Documentation

Mathlib.Probability.UniformOn

Classical probability #

The classical formulation of probability states that the probability of an event occurring in a finite probability space is the ratio of that event to all possible events. This notion can be expressed with measure theory using the counting measure. In particular, given the sets s and t, we define the probability of t occurring in s to be |s|⁻¹ * |s ∩ t|. With this definition, we recover the probability over the entire sample space when s = Set.univ.

Classical probability is often used in combinatorics and we prove some useful lemmas in this file for that purpose.

Main definition #

Notes #

The original aim of this file is to provide a measure theoretic method of describing the probability an element of a set s satisfies some predicate P. Our current formulation still allow us to describe this by abusing the definitional equality of sets and predicates by simply writing uniformOn s P. We should avoid this however as none of the lemmas are written for predicates.

Given a set s, uniformOn s is the uniform measure on s, defined as the counting measure conditioned by s. One should think of uniformOn s t as the proportion of s that is contained in t.

This is a probability measure when s is finite and nonempty and is given by ProbabilityTheory.uniformOn_isProbabilityMeasure.

Equations
Instances For
    @[simp]

    See uniformOn_eq_zero for a version assuming MeasurableSingletonClass Ω instead of MeasurableSet s.

    theorem ProbabilityTheory.uniformOn_singleton {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] (ω : Ω) (t : Set Ω) [Decidable (ω t)] :
    (uniformOn {ω}) t = if ω t then 1 else 0
    theorem ProbabilityTheory.uniformOn_self {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) :
    (uniformOn s) s = 1
    theorem ProbabilityTheory.uniformOn_eq_one_of {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s t : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) (ht : s t) :
    (uniformOn s) t = 1
    theorem ProbabilityTheory.uniformOn_inter {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s t u : Set Ω} (hs : s.Finite) :
    (uniformOn s) (t u) = (uniformOn (s t)) u * (uniformOn s) t
    theorem ProbabilityTheory.uniformOn_inter' {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s t u : Set Ω} (hs : s.Finite) :
    (uniformOn s) (t u) = (uniformOn (s u)) t * (uniformOn s) u
    theorem ProbabilityTheory.uniformOn_union {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s t u : Set Ω} (hs : s.Finite) (htu : Disjoint t u) :
    (uniformOn s) (t u) = (uniformOn s) t + (uniformOn s) u
    theorem ProbabilityTheory.uniformOn_compl {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (t : Set Ω) (hs : s.Finite) (hs' : s.Nonempty) :
    (uniformOn s) t + (uniformOn s) t = 1
    theorem ProbabilityTheory.uniformOn_disjoint_union {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s t u : Set Ω} (hs : s.Finite) (ht : t.Finite) (hst : Disjoint s t) :
    (uniformOn s) u * (uniformOn (s t)) s + (uniformOn t) u * (uniformOn (s t)) t = (uniformOn (s t)) u
    theorem ProbabilityTheory.uniformOn_add_compl_eq {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (u t : Set Ω) (hs : s.Finite) :
    (uniformOn (s u)) t * (uniformOn s) u + (uniformOn (s u)) t * (uniformOn s) u = (uniformOn s) t

    A version of the law of total probability for counting probabilities.