author | nipkow |
Thu, 26 Jun 2025 17:30:33 +0200 | |
changeset 82772 | 59b937edcff8 |
parent 76987 | 4c275405faae |
permissions | -rw-r--r-- |
39157
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
wenzelm
parents:
37678
diff
changeset
|
1 |
(* Title: HOL/Proofs/Extraction/Pigeonhole.thy |
17024 | 2 |
Author: Stefan Berghofer, TU Muenchen |
3 |
*) |
|
4 |
||
61986 | 5 |
section \<open>The pigeonhole principle\<close> |
17024 | 6 |
|
22737 | 7 |
theory Pigeonhole |
67320 | 8 |
imports Util "HOL-Library.Realizers" "HOL-Library.Code_Target_Numeral" |
22737 | 9 |
begin |
17024 | 10 |
|
61986 | 11 |
text \<open> |
63361 | 12 |
We formalize two proofs of the pigeonhole principle, which lead |
13 |
to extracted programs of quite different complexity. The original |
|
14 |
formalization of these proofs in {\sc Nuprl} is due to |
|
76987 | 15 |
Aleksey Nogin \<^cite>\<open>"Nogin-ENTCS-2000"\<close>. |
17024 | 16 |
|
63361 | 17 |
This proof yields a polynomial program. |
61986 | 18 |
\<close> |
17024 | 19 |
|
20 |
theorem pigeonhole: |
|
21 |
"\<And>f. (\<And>i. i \<le> Suc n \<Longrightarrow> f i \<le> n) \<Longrightarrow> \<exists>i j. i \<le> Suc n \<and> j < i \<and> f i = f j" |
|
22 |
proof (induct n) |
|
23 |
case 0 |
|
63361 | 24 |
then have "Suc 0 \<le> Suc 0 \<and> 0 < Suc 0 \<and> f (Suc 0) = f 0" by simp |
25 |
then show ?case by iprover |
|
17024 | 26 |
next |
27 |
case (Suc n) |
|
63361 | 28 |
have r: |
29 |
"k \<le> Suc (Suc n) \<Longrightarrow> |
|
30 |
(\<And>i j. Suc k \<le> i \<Longrightarrow> i \<le> Suc (Suc n) \<Longrightarrow> j < i \<Longrightarrow> f i \<noteq> f j) \<Longrightarrow> |
|
31 |
(\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j)" for k |
|
32 |
proof (induct k) |
|
33 |
case 0 |
|
34 |
let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i" |
|
35 |
have "\<not> (\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j)" |
|
36 |
proof |
|
37 |
assume "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" |
|
38 |
then obtain i j where i: "i \<le> Suc n" and j: "j < i" and f: "?f i = ?f j" |
|
39 |
by iprover |
|
40 |
from j have i_nz: "Suc 0 \<le> i" by simp |
|
41 |
from i have iSSn: "i \<le> Suc (Suc n)" by simp |
|
42 |
have S0SSn: "Suc 0 \<le> Suc (Suc n)" by simp |
|
43 |
show False |
|
44 |
proof cases |
|
45 |
assume fi: "f i = Suc n" |
|
46 |
show False |
|
47 |
proof cases |
|
48 |
assume fj: "f j = Suc n" |
|
49 |
from i_nz and iSSn and j have "f i \<noteq> f j" by (rule 0) |
|
50 |
moreover from fi have "f i = f j" |
|
51 |
by (simp add: fj [symmetric]) |
|
52 |
ultimately show ?thesis .. |
|
53 |
next |
|
54 |
from i and j have "j < Suc (Suc n)" by simp |
|
55 |
with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f j" |
|
56 |
by (rule 0) |
|
57 |
moreover assume "f j \<noteq> Suc n" |
|
58 |
with fi and f have "f (Suc (Suc n)) = f j" by simp |
|
59 |
ultimately show False .. |
|
60 |
qed |
|
61 |
next |
|
62 |
assume fi: "f i \<noteq> Suc n" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
63 |
show False |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
64 |
proof cases |
63361 | 65 |
from i have "i < Suc (Suc n)" by simp |
66 |
with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f i" |
|
67 |
by (rule 0) |
|
68 |
moreover assume "f j = Suc n" |
|
69 |
with fi and f have "f (Suc (Suc n)) = f i" by simp |
|
70 |
ultimately show False .. |
|
71 |
next |
|
72 |
from i_nz and iSSn and j |
|
73 |
have "f i \<noteq> f j" by (rule 0) |
|
74 |
moreover assume "f j \<noteq> Suc n" |
|
75 |
with fi and f have "f i = f j" by simp |
|
76 |
ultimately show False .. |
|
77 |
qed |
|
78 |
qed |
|
79 |
qed |
|
80 |
moreover have "?f i \<le> n" if "i \<le> Suc n" for i |
|
81 |
proof - |
|
82 |
from that have i: "i < Suc (Suc n)" by simp |
|
83 |
have "f (Suc (Suc n)) \<noteq> f i" |
|
84 |
by (rule 0) (simp_all add: i) |
|
85 |
moreover have "f (Suc (Suc n)) \<le> Suc n" |
|
86 |
by (rule Suc) simp |
|
87 |
moreover from i have "i \<le> Suc (Suc n)" by simp |
|
88 |
then have "f i \<le> Suc n" by (rule Suc) |
|
89 |
ultimately show ?thesis |
|
90 |
by simp |
|
91 |
qed |
|
92 |
then have "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" |
|
93 |
by (rule Suc) |
|
94 |
ultimately show ?case .. |
|
95 |
next |
|
96 |
case (Suc k) |
|
97 |
from search [OF nat_eq_dec] show ?case |
|
98 |
proof |
|
99 |
assume "\<exists>j<Suc k. f (Suc k) = f j" |
|
100 |
then show ?case by (iprover intro: le_refl) |
|
101 |
next |
|
102 |
assume nex: "\<not> (\<exists>j<Suc k. f (Suc k) = f j)" |
|
103 |
have "\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j" |
|
104 |
proof (rule Suc) |
|
105 |
from Suc show "k \<le> Suc (Suc n)" by simp |
|
106 |
fix i j assume k: "Suc k \<le> i" and i: "i \<le> Suc (Suc n)" |
|
107 |
and j: "j < i" |
|
108 |
show "f i \<noteq> f j" |
|
109 |
proof cases |
|
110 |
assume eq: "i = Suc k" |
|
111 |
show ?thesis |
|
112 |
proof |
|
113 |
assume "f i = f j" |
|
114 |
then have "f (Suc k) = f j" by (simp add: eq) |
|
115 |
with nex and j and eq show False by iprover |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
116 |
qed |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
117 |
next |
63361 | 118 |
assume "i \<noteq> Suc k" |
119 |
with k have "Suc (Suc k) \<le> i" by simp |
|
120 |
then show ?thesis using i and j by (rule Suc) |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
121 |
qed |
17024 | 122 |
qed |
63361 | 123 |
then show ?thesis by (iprover intro: le_SucI) |
17024 | 124 |
qed |
63361 | 125 |
qed |
17024 | 126 |
show ?case by (rule r) simp_all |
127 |
qed |
|
128 |
||
61986 | 129 |
text \<open> |
63361 | 130 |
The following proof, although quite elegant from a mathematical point of view, |
131 |
leads to an exponential program: |
|
61986 | 132 |
\<close> |
17024 | 133 |
|
134 |
theorem pigeonhole_slow: |
|
135 |
"\<And>f. (\<And>i. i \<le> Suc n \<Longrightarrow> f i \<le> n) \<Longrightarrow> \<exists>i j. i \<le> Suc n \<and> j < i \<and> f i = f j" |
|
136 |
proof (induct n) |
|
137 |
case 0 |
|
138 |
have "Suc 0 \<le> Suc 0" .. |
|
139 |
moreover have "0 < Suc 0" .. |
|
140 |
moreover from 0 have "f (Suc 0) = f 0" by simp |
|
17604 | 141 |
ultimately show ?case by iprover |
17024 | 142 |
next |
143 |
case (Suc n) |
|
25418 | 144 |
from search [OF nat_eq_dec] show ?case |
17024 | 145 |
proof |
146 |
assume "\<exists>j < Suc (Suc n). f (Suc (Suc n)) = f j" |
|
63361 | 147 |
then show ?case by (iprover intro: le_refl) |
17024 | 148 |
next |
149 |
assume "\<not> (\<exists>j < Suc (Suc n). f (Suc (Suc n)) = f j)" |
|
63361 | 150 |
then have nex: "\<forall>j < Suc (Suc n). f (Suc (Suc n)) \<noteq> f j" by iprover |
17024 | 151 |
let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i" |
152 |
have "\<And>i. i \<le> Suc n \<Longrightarrow> ?f i \<le> n" |
|
153 |
proof - |
|
154 |
fix i assume i: "i \<le> Suc n" |
|
155 |
show "?thesis i" |
|
156 |
proof (cases "f i = Suc n") |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
157 |
case True |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
158 |
from i and nex have "f (Suc (Suc n)) \<noteq> f i" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
159 |
with True have "f (Suc (Suc n)) \<noteq> Suc n" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
160 |
moreover from Suc have "f (Suc (Suc n)) \<le> Suc n" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
161 |
ultimately have "f (Suc (Suc n)) \<le> n" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
162 |
with True show ?thesis by simp |
17024 | 163 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
164 |
case False |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
165 |
from Suc and i have "f i \<le> Suc n" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
166 |
with False show ?thesis by simp |
17024 | 167 |
qed |
168 |
qed |
|
63361 | 169 |
then have "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" by (rule Suc) |
17024 | 170 |
then obtain i j where i: "i \<le> Suc n" and ji: "j < i" and f: "?f i = ?f j" |
17604 | 171 |
by iprover |
17024 | 172 |
have "f i = f j" |
173 |
proof (cases "f i = Suc n") |
|
174 |
case True |
|
175 |
show ?thesis |
|
176 |
proof (cases "f j = Suc n") |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
177 |
assume "f j = Suc n" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
178 |
with True show ?thesis by simp |
17024 | 179 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
180 |
assume "f j \<noteq> Suc n" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
181 |
moreover from i ji nex have "f (Suc (Suc n)) \<noteq> f j" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
182 |
ultimately show ?thesis using True f by simp |
17024 | 183 |
qed |
184 |
next |
|
185 |
case False |
|
186 |
show ?thesis |
|
187 |
proof (cases "f j = Suc n") |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
188 |
assume "f j = Suc n" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
189 |
moreover from i nex have "f (Suc (Suc n)) \<noteq> f i" by simp |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
190 |
ultimately show ?thesis using False f by simp |
17024 | 191 |
next |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
192 |
assume "f j \<noteq> Suc n" |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
29823
diff
changeset
|
193 |
with False f show ?thesis by simp |
17024 | 194 |
qed |
195 |
qed |
|
196 |
moreover from i have "i \<le> Suc (Suc n)" by simp |
|
17604 | 197 |
ultimately show ?thesis using ji by iprover |
17024 | 198 |
qed |
199 |
qed |
|
200 |
||
201 |
extract pigeonhole pigeonhole_slow |
|
202 |
||
61986 | 203 |
text \<open> |
63361 | 204 |
The programs extracted from the above proofs look as follows: |
205 |
@{thm [display] pigeonhole_def} |
|
206 |
@{thm [display] pigeonhole_slow_def} |
|
207 |
The program for searching for an element in an array is |
|
208 |
@{thm [display,eta_contract=false] search_def} |
|
69604 | 209 |
The correctness statement for \<^term>\<open>pigeonhole\<close> is |
63361 | 210 |
@{thm [display] pigeonhole_correctness [no_vars]} |
17024 | 211 |
|
63361 | 212 |
In order to analyze the speed of the above programs, |
213 |
we generate ML code from them. |
|
61986 | 214 |
\<close> |
17024 | 215 |
|
27982 | 216 |
instantiation nat :: default |
217 |
begin |
|
218 |
||
219 |
definition "default = (0::nat)" |
|
220 |
||
221 |
instance .. |
|
222 |
||
223 |
end |
|
224 |
||
37678
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
haftmann
parents:
37287
diff
changeset
|
225 |
instantiation prod :: (default, default) default |
27982 | 226 |
begin |
227 |
||
228 |
definition "default = (default, default)" |
|
229 |
||
230 |
instance .. |
|
231 |
||
232 |
end |
|
233 |
||
63361 | 234 |
definition "test n u = pigeonhole (nat_of_integer n) (\<lambda>m. m - 1)" |
235 |
definition "test' n u = pigeonhole_slow (nat_of_integer n) (\<lambda>m. m - 1)" |
|
236 |
definition "test'' u = pigeonhole 8 (List.nth [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])" |
|
20837 | 237 |
|
63361 | 238 |
ML_val "timeit (@{code test} 10)" |
51272 | 239 |
ML_val "timeit (@{code test'} 10)" |
240 |
ML_val "timeit (@{code test} 20)" |
|
241 |
ML_val "timeit (@{code test'} 20)" |
|
242 |
ML_val "timeit (@{code test} 25)" |
|
243 |
ML_val "timeit (@{code test'} 25)" |
|
244 |
ML_val "timeit (@{code test} 500)" |
|
245 |
ML_val "timeit @{code test''}" |
|
37287 | 246 |
|
17024 | 247 |
end |