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
    @[deprecated ProbabilityTheory.uniformOn]

    Alias of ProbabilityTheory.uniformOn.


    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
      @[deprecated ProbabilityTheory.uniformOn_empty_meas]

      Alias of ProbabilityTheory.uniformOn_empty_meas.

      @[deprecated ProbabilityTheory.uniformOn_empty]

      Alias of ProbabilityTheory.uniformOn_empty.

      theorem ProbabilityTheory.finite_of_uniformOn_ne_zero {Ω : Type u_1} [MeasurableSpace Ω] {s : Set Ω} {t : Set Ω} (h : (ProbabilityTheory.uniformOn s) t 0) :
      s.Finite
      @[deprecated ProbabilityTheory.finite_of_uniformOn_ne_zero]
      theorem ProbabilityTheory.finite_of_condCount_ne_zero {Ω : Type u_1} [MeasurableSpace Ω] {s : Set Ω} {t : Set Ω} (h : (ProbabilityTheory.uniformOn s) t 0) :
      s.Finite

      Alias of ProbabilityTheory.finite_of_uniformOn_ne_zero.

      theorem ProbabilityTheory.uniformOn_univ {Ω : Type u_1} [MeasurableSpace Ω] [Fintype Ω] {s : Set Ω} :
      (ProbabilityTheory.uniformOn Set.univ) s = MeasureTheory.Measure.count s / (Fintype.card Ω)
      @[deprecated ProbabilityTheory.uniformOn_univ]
      theorem ProbabilityTheory.condCount_univ {Ω : Type u_1} [MeasurableSpace Ω] [Fintype Ω] {s : Set Ω} :
      (ProbabilityTheory.uniformOn Set.univ) s = MeasureTheory.Measure.count s / (Fintype.card Ω)

      Alias of ProbabilityTheory.uniformOn_univ.

      @[deprecated ProbabilityTheory.uniformOn_isProbabilityMeasure]

      Alias of ProbabilityTheory.uniformOn_isProbabilityMeasure.

      theorem ProbabilityTheory.uniformOn_singleton {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] (ω : Ω) (t : Set Ω) [Decidable (ω t)] :
      (ProbabilityTheory.uniformOn {ω}) t = if ω t then 1 else 0
      @[deprecated ProbabilityTheory.uniformOn_singleton]
      theorem ProbabilityTheory.condCount_singleton {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] (ω : Ω) (t : Set Ω) [Decidable (ω t)] :
      (ProbabilityTheory.uniformOn {ω}) t = if ω t then 1 else 0

      Alias of ProbabilityTheory.uniformOn_singleton.

      @[deprecated ProbabilityTheory.uniformOn_inter_self]

      Alias of ProbabilityTheory.uniformOn_inter_self.

      theorem ProbabilityTheory.uniformOn_self {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) :
      @[deprecated ProbabilityTheory.uniformOn_self]
      theorem ProbabilityTheory.condCount_self {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) :

      Alias of ProbabilityTheory.uniformOn_self.

      theorem ProbabilityTheory.uniformOn_eq_one_of {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} {t : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) (ht : s t) :
      @[deprecated ProbabilityTheory.uniformOn_eq_one_of]
      theorem ProbabilityTheory.condCount_eq_one_of {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} {t : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) (ht : s t) :

      Alias of ProbabilityTheory.uniformOn_eq_one_of.

      @[deprecated ProbabilityTheory.pred_true_of_uniformOn_eq_one]

      Alias of ProbabilityTheory.pred_true_of_uniformOn_eq_one.

      @[deprecated ProbabilityTheory.uniformOn_eq_zero_iff]
      theorem ProbabilityTheory.condCount_eq_zero_iff {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} {t : Set Ω} (hs : s.Finite) :

      Alias of ProbabilityTheory.uniformOn_eq_zero_iff.

      theorem ProbabilityTheory.uniformOn_of_univ {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) :
      @[deprecated ProbabilityTheory.uniformOn_of_univ]
      theorem ProbabilityTheory.condCount_of_univ {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (hs : s.Finite) (hs' : s.Nonempty) :

      Alias of ProbabilityTheory.uniformOn_of_univ.

      @[deprecated ProbabilityTheory.uniformOn_inter]

      Alias of ProbabilityTheory.uniformOn_inter.

      @[deprecated ProbabilityTheory.uniformOn_inter']

      Alias of ProbabilityTheory.uniformOn_inter'.

      @[deprecated ProbabilityTheory.uniformOn_union]
      theorem ProbabilityTheory.condCount_union {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} {t : Set Ω} {u : Set Ω} (hs : s.Finite) (htu : Disjoint t u) :

      Alias of ProbabilityTheory.uniformOn_union.

      theorem ProbabilityTheory.uniformOn_compl {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (t : Set Ω) (hs : s.Finite) (hs' : s.Nonempty) :
      @[deprecated ProbabilityTheory.uniformOn_compl]
      theorem ProbabilityTheory.condCount_compl {Ω : Type u_1} [MeasurableSpace Ω] [MeasurableSingletonClass Ω] {s : Set Ω} (t : Set Ω) (hs : s.Finite) (hs' : s.Nonempty) :

      Alias of ProbabilityTheory.uniformOn_compl.

      @[deprecated ProbabilityTheory.uniformOn_disjoint_union]

      Alias of ProbabilityTheory.uniformOn_disjoint_union.

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

      @[deprecated ProbabilityTheory.uniformOn_add_compl_eq]

      Alias of ProbabilityTheory.uniformOn_add_compl_eq.


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