author | wenzelm |
Fri, 25 Oct 2024 13:43:12 +0200 | |
changeset 81258 | 74647c464cbd |
parent 68188 | 2af1f142f855 |
permissions | -rw-r--r-- |
81258 | 1 |
(* Title: HOL/ex/Birthday_Paradox.thy |
2 |
Author: Lukas Bulwahn, TU Muenchen, 2007 |
|
40632 | 3 |
*) |
4 |
||
61343 | 5 |
section \<open>A Formulation of the Birthday Paradox\<close> |
40632 | 6 |
|
43238
04c886a1d1a5
renaming the formalisation of the birthday problem to a proper English name
bulwahn
parents:
40632
diff
changeset
|
7 |
theory Birthday_Paradox |
81258 | 8 |
imports "HOL-Library.FuncSet" |
40632 | 9 |
begin |
10 |
||
61343 | 11 |
section \<open>Cardinality\<close> |
40632 | 12 |
|
13 |
lemma card_product_dependent: |
|
14 |
assumes "finite S" |
|
81258 | 15 |
and "\<forall>x \<in> S. finite (T x)" |
40632 | 16 |
shows "card {(x, y). x \<in> S \<and> y \<in> T x} = (\<Sum>x \<in> S. card (T x))" |
81258 | 17 |
using card_SigmaI[OF assms, symmetric] |
18 |
by (auto intro!: arg_cong[where f=card] simp add: Sigma_def) |
|
40632 | 19 |
|
20 |
lemma card_extensional_funcset_inj_on: |
|
21 |
assumes "finite S" "finite T" "card S \<le> card T" |
|
22 |
shows "card {f \<in> extensional_funcset S T. inj_on f S} = fact (card T) div (fact (card T - card S))" |
|
81258 | 23 |
using assms |
40632 | 24 |
proof (induct S arbitrary: T rule: finite_induct) |
25 |
case empty |
|
81258 | 26 |
from this show ?case |
27 |
by (simp add: Collect_conv_if PiE_empty_domain) |
|
40632 | 28 |
next |
29 |
case (insert x S) |
|
81258 | 30 |
have finite_delete: "finite {f : extensional_funcset S (T - {x}). inj_on f S}" for x |
31 |
proof - |
|
32 |
from \<open>finite T\<close> have "finite (T - {x})" |
|
33 |
by auto |
|
34 |
from \<open>finite S\<close> this have *: "finite (extensional_funcset S (T - {x}))" |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
43238
diff
changeset
|
35 |
by (rule finite_PiE) |
81258 | 36 |
have "{f : extensional_funcset S (T - {x}). inj_on f S} \<subseteq> (extensional_funcset S (T - {x}))" |
37 |
by auto |
|
38 |
with * show ?thesis |
|
40632 | 39 |
by (auto intro: finite_subset) |
81258 | 40 |
qed |
41 |
from insert have hyps: "\<forall>y \<in> T. card ({g. g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S}) = |
|
42 |
fact (card T - 1) div fact ((card T - 1) - card S)"(is "\<forall> _ \<in> T. _ = ?k") |
|
43 |
by auto |
|
61343 | 44 |
from extensional_funcset_extend_domain_inj_on_eq[OF \<open>x \<notin> S\<close>] |
67613 | 45 |
have "card {f. f \<in> extensional_funcset (insert x S) T \<and> inj_on f (insert x S)} = |
81258 | 46 |
card ((\<lambda>(y, g). g(x := y)) ` {(y, g). y \<in> T \<and> g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S})" |
40632 | 47 |
by metis |
81258 | 48 |
also from extensional_funcset_extend_domain_inj_onI[OF \<open>x \<notin> S\<close>, of T] |
49 |
have "\<dots> = card {(y, g). y \<in> T \<and> g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S}" |
|
40632 | 50 |
by (simp add: card_image) |
51 |
also have "card {(y, g). y \<in> T \<and> g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S} = |
|
81258 | 52 |
card {(y, g). y \<in> T \<and> g \<in> {f \<in> extensional_funcset S (T - {y}). inj_on f S}}" |
53 |
by auto |
|
54 |
also from \<open>finite T\<close> finite_delete |
|
55 |
have "\<dots> = (\<Sum>y \<in> T. card {g. g \<in> extensional_funcset S (T - {y}) \<and> inj_on g S})" |
|
40632 | 56 |
by (subst card_product_dependent) auto |
81258 | 57 |
also from hyps have "\<dots> = (card T) * ?k" |
40632 | 58 |
by auto |
81258 | 59 |
also have "\<dots> = card T * fact (card T - 1) div fact (card T - card (insert x S))" |
40632 | 60 |
using insert unfolding div_mult1_eq[of "card T" "fact (card T - 1)"] |
61 |
by (simp add: fact_mod) |
|
81258 | 62 |
also have "\<dots> = fact (card T) div fact (card T - card (insert x S))" |
59730
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents:
59669
diff
changeset
|
63 |
using insert by (simp add: fact_reduce[of "card T"]) |
40632 | 64 |
finally show ?case . |
65 |
qed |
|
66 |
||
67 |
lemma card_extensional_funcset_not_inj_on: |
|
68 |
assumes "finite S" "finite T" "card S \<le> card T" |
|
81258 | 69 |
shows "card {f \<in> extensional_funcset S T. \<not> inj_on f S} = |
70 |
(card T) ^ (card S) - (fact (card T)) div (fact (card T - card S))" |
|
40632 | 71 |
proof - |
81258 | 72 |
have subset: "{f \<in> extensional_funcset S T. inj_on f S} \<subseteq> extensional_funcset S T" |
73 |
by auto |
|
74 |
from finite_subset[OF subset] assms |
|
75 |
have finite: "finite {f : extensional_funcset S T. inj_on f S}" |
|
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
43238
diff
changeset
|
76 |
by (auto intro!: finite_PiE) |
81258 | 77 |
have "{f \<in> extensional_funcset S T. \<not> inj_on f S} = |
78 |
extensional_funcset S T - {f \<in> extensional_funcset S T. inj_on f S}" by auto |
|
40632 | 79 |
from assms this finite subset show ?thesis |
64272 | 80 |
by (simp add: card_Diff_subset card_PiE card_extensional_funcset_inj_on prod_constant) |
40632 | 81 |
qed |
82 |
||
64272 | 83 |
lemma prod_upto_nat_unfold: |
84 |
"prod f {m..(n::nat)} = (if n < m then 1 else (if n = 0 then f 0 else f n * prod f {m..(n - 1)}))" |
|
40632 | 85 |
by auto (auto simp add: gr0_conv_Suc atLeastAtMostSuc_conv) |
86 |
||
81258 | 87 |
|
61343 | 88 |
section \<open>Birthday paradox\<close> |
40632 | 89 |
|
43238
04c886a1d1a5
renaming the formalisation of the birthday problem to a proper English name
bulwahn
parents:
40632
diff
changeset
|
90 |
lemma birthday_paradox: |
40632 | 91 |
assumes "card S = 23" "card T = 365" |
92 |
shows "2 * card {f \<in> extensional_funcset S T. \<not> inj_on f S} \<ge> card (extensional_funcset S T)" |
|
93 |
proof - |
|
81258 | 94 |
from \<open>card S = 23\<close> \<open>card T = 365\<close> have "finite S" "finite T" "card S \<le> card T" |
95 |
by (auto intro: card_ge_0_finite) |
|
40632 | 96 |
from assms show ?thesis |
61343 | 97 |
using card_PiE[OF \<open>finite S\<close>, of "\<lambda>i. T"] \<open>finite S\<close> |
81258 | 98 |
card_extensional_funcset_not_inj_on[OF \<open>finite S\<close> \<open>finite T\<close> \<open>card S \<le> card T\<close>] |
64272 | 99 |
by (simp add: fact_div_fact prod_upto_nat_unfold prod_constant) |
40632 | 100 |
qed |
101 |
||
102 |
end |