Ballot problem #
This file proves Theorem 30 from the 100 Theorems List.
The ballot problem asks, if in an election, candidate A receives p
votes whereas candidate B
receives q
votes where p > q
, what is the probability that candidate A is strictly ahead
throughout the count. The probability of this is (p - q) / (p + q)
.
Main definitions #
counted_sequence
: given natural numbersp
andq
,counted_sequence p q
is the set of all lists containingp
of1
s andq
of-1
s representing the votes of candidate A and B respectively.stays_positive
: is the set of lists of integers which suffix has positive sum. In particular, the intersection of this set withcounted_sequence
is the set of lists where candidate A is strictly ahead.
Main result #
ballot
: the ballot problem.
counted_sequence p q
is the set of lists of integers for which every element is +1
or -1
,
there are p
lots of +1
and q
lots of -1
.
This represents vote sequences where candidate +1
receives p
votes and candidate -1
receives
q
votes.
theorem
ballot.mem_counted_sequence_iff_perm
{p q : ℕ}
{l : list ℤ} :
l ∈ ballot.counted_sequence p q ↔ l ~ list.replicate p 1 ++ list.replicate q (-1)
An alternative definition of counted_sequence
that uses list.perm
.
@[simp]
@[simp]
theorem
ballot.length_of_mem_counted_sequence
{p q : ℕ}
{l : list ℤ}
(hl : l ∈ ballot.counted_sequence p q) :
theorem
ballot.counted_ne_nil_left
{p q : ℕ}
(hp : p ≠ 0)
{l : list ℤ}
(hl : l ∈ ballot.counted_sequence p q) :
theorem
ballot.counted_ne_nil_right
{p q : ℕ}
(hq : q ≠ 0)
{l : list ℤ}
(hl : l ∈ ballot.counted_sequence p q) :
theorem
ballot.counted_succ_succ
(p q : ℕ) :
ballot.counted_sequence (p + 1) (q + 1) = list.cons 1 '' ballot.counted_sequence p (q + 1) ∪ list.cons (-1) '' ballot.counted_sequence (p + 1) q
theorem
ballot.disjoint_bits
(p q : ℕ) :
disjoint (list.cons 1 '' ballot.counted_sequence p (q + 1)) (list.cons (-1) '' ballot.counted_sequence (p + 1) q)
theorem
ballot.count_counted_sequence
(p q : ℕ) :
⇑measure_theory.measure.count (ballot.counted_sequence p q) = ↑((p + q).choose p)
theorem
ballot.ballot_same
(p : ℕ) :
⇑(probability_theory.cond_count (ballot.counted_sequence (p + 1) (p + 1))) ballot.stays_positive = 0
theorem
ballot.ballot_pos
(p q : ℕ) :
⇑(probability_theory.cond_count (ballot.counted_sequence (p + 1) (q + 1) ∩ {l : list ℤ | l.head = 1})) ballot.stays_positive = ⇑(probability_theory.cond_count (ballot.counted_sequence p (q + 1))) ballot.stays_positive
theorem
ballot.ballot_neg
(p q : ℕ)
(qp : q < p) :
⇑(probability_theory.cond_count (ballot.counted_sequence (p + 1) (q + 1) ∩ {l : list ℤ | l.head = 1}ᶜ)) ballot.stays_positive = ⇑(probability_theory.cond_count (ballot.counted_sequence (p + 1) q)) ballot.stays_positive
theorem
ballot.ballot_problem'
(q p : ℕ) :
q < p → (⇑(probability_theory.cond_count (ballot.counted_sequence p q)) ballot.stays_positive).to_real = (↑p - ↑q) / (↑p + ↑q)
theorem
ballot.ballot_problem
(q p : ℕ) :
q < p → ⇑(probability_theory.cond_count (ballot.counted_sequence p q)) ballot.stays_positive = (↑p - ↑q) / (↑p + ↑q)
The ballot problem.