| author | haftmann | 
| Wed, 17 Dec 2008 12:10:40 +0100 | |
| changeset 29135 | 20b42397e293 | 
| parent 27982 | 2aaa4a5569a6 | 
| child 29823 | 0ab754d13ccd | 
| permissions | -rw-r--r-- | 
| 17024 | 1 | (* Title: HOL/Extraction/Pigeonhole.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Stefan Berghofer, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 6 | header {* The pigeonhole principle *}
 | |
| 7 | ||
| 22737 | 8 | theory Pigeonhole | 
| 25418 | 9 | imports Util Efficient_Nat | 
| 22737 | 10 | begin | 
| 17024 | 11 | |
| 12 | text {*
 | |
| 13 | We formalize two proofs of the pigeonhole principle, which lead | |
| 14 | to extracted programs of quite different complexity. The original | |
| 15 | formalization of these proofs in {\sc Nuprl} is due to
 | |
| 16 | Aleksey Nogin \cite{Nogin-ENTCS-2000}.
 | |
| 17 | ||
| 18 | This proof yields a polynomial program. | |
| 19 | *} | |
| 20 | ||
| 21 | theorem pigeonhole: | |
| 22 | "\<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" | |
| 23 | proof (induct n) | |
| 24 | case 0 | |
| 25 | hence "Suc 0 \<le> Suc 0 \<and> 0 < Suc 0 \<and> f (Suc 0) = f 0" by simp | |
| 17604 | 26 | thus ?case by iprover | 
| 17024 | 27 | next | 
| 28 | case (Suc n) | |
| 29 |   {
 | |
| 30 | fix k | |
| 31 | have | |
| 32 | "k \<le> Suc (Suc n) \<Longrightarrow> | |
| 33 | (\<And>i j. Suc k \<le> i \<Longrightarrow> i \<le> Suc (Suc n) \<Longrightarrow> j < i \<Longrightarrow> f i \<noteq> f j) \<Longrightarrow> | |
| 34 | (\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j)" | |
| 35 | proof (induct k) | |
| 36 | case 0 | |
| 37 | let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i" | |
| 38 | have "\<not> (\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j)" | |
| 39 | proof | |
| 40 | assume "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" | |
| 41 | then obtain i j where i: "i \<le> Suc n" and j: "j < i" | |
| 17604 | 42 | and f: "?f i = ?f j" by iprover | 
| 17024 | 43 | from j have i_nz: "Suc 0 \<le> i" by simp | 
| 44 | from i have iSSn: "i \<le> Suc (Suc n)" by simp | |
| 45 | have S0SSn: "Suc 0 \<le> Suc (Suc n)" by simp | |
| 46 | show False | |
| 47 | proof cases | |
| 48 | assume fi: "f i = Suc n" | |
| 49 | show False | |
| 50 | proof cases | |
| 51 | assume fj: "f j = Suc n" | |
| 52 | from i_nz and iSSn and j have "f i \<noteq> f j" by (rule 0) | |
| 53 | moreover from fi have "f i = f j" | |
| 54 | by (simp add: fj [symmetric]) | |
| 55 | ultimately show ?thesis .. | |
| 56 | next | |
| 57 | from i and j have "j < Suc (Suc n)" by simp | |
| 58 | with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f j" | |
| 59 | by (rule 0) | |
| 60 | moreover assume "f j \<noteq> Suc n" | |
| 61 | with fi and f have "f (Suc (Suc n)) = f j" by simp | |
| 62 | ultimately show False .. | |
| 63 | qed | |
| 64 | next | |
| 65 | assume fi: "f i \<noteq> Suc n" | |
| 66 | show False | |
| 67 | proof cases | |
| 68 | from i have "i < Suc (Suc n)" by simp | |
| 69 | with S0SSn and le_refl have "f (Suc (Suc n)) \<noteq> f i" | |
| 70 | by (rule 0) | |
| 71 | moreover assume "f j = Suc n" | |
| 72 | with fi and f have "f (Suc (Suc n)) = f i" by simp | |
| 73 | ultimately show False .. | |
| 74 | next | |
| 75 | from i_nz and iSSn and j | |
| 76 | have "f i \<noteq> f j" by (rule 0) | |
| 77 | moreover assume "f j \<noteq> Suc n" | |
| 78 | with fi and f have "f i = f j" by simp | |
| 79 | ultimately show False .. | |
| 80 | qed | |
| 81 | qed | |
| 82 | qed | |
| 83 | moreover have "\<And>i. i \<le> Suc n \<Longrightarrow> ?f i \<le> n" | |
| 84 | proof - | |
| 85 | fix i assume "i \<le> Suc n" | |
| 86 | hence i: "i < Suc (Suc n)" by simp | |
| 87 | have "f (Suc (Suc n)) \<noteq> f i" | |
| 88 | by (rule 0) (simp_all add: i) | |
| 89 | moreover have "f (Suc (Suc n)) \<le> Suc n" | |
| 90 | by (rule Suc) simp | |
| 91 | moreover from i have "i \<le> Suc (Suc n)" by simp | |
| 92 | hence "f i \<le> Suc n" by (rule Suc) | |
| 93 | ultimately show "?thesis i" | |
| 94 | by simp | |
| 95 | qed | |
| 96 | hence "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" | |
| 97 | by (rule Suc) | |
| 98 | ultimately show ?case .. | |
| 99 | next | |
| 100 | case (Suc k) | |
| 25418 | 101 | from search [OF nat_eq_dec] show ?case | 
| 17024 | 102 | proof | 
| 103 | assume "\<exists>j<Suc k. f (Suc k) = f j" | |
| 17604 | 104 | thus ?case by (iprover intro: le_refl) | 
| 17024 | 105 | next | 
| 106 | assume nex: "\<not> (\<exists>j<Suc k. f (Suc k) = f j)" | |
| 107 | have "\<exists>i j. i \<le> k \<and> j < i \<and> f i = f j" | |
| 108 | proof (rule Suc) | |
| 109 | from Suc show "k \<le> Suc (Suc n)" by simp | |
| 110 | fix i j assume k: "Suc k \<le> i" and i: "i \<le> Suc (Suc n)" | |
| 111 | and j: "j < i" | |
| 112 | show "f i \<noteq> f j" | |
| 113 | proof cases | |
| 114 | assume eq: "i = Suc k" | |
| 115 | show ?thesis | |
| 116 | proof | |
| 117 | assume "f i = f j" | |
| 118 | hence "f (Suc k) = f j" by (simp add: eq) | |
| 17604 | 119 | with nex and j and eq show False by iprover | 
| 17024 | 120 | qed | 
| 121 | next | |
| 122 | assume "i \<noteq> Suc k" | |
| 123 | with k have "Suc (Suc k) \<le> i" by simp | |
| 124 | thus ?thesis using i and j by (rule Suc) | |
| 125 | qed | |
| 126 | qed | |
| 17604 | 127 | thus ?thesis by (iprover intro: le_SucI) | 
| 17024 | 128 | qed | 
| 129 | qed | |
| 130 | } | |
| 131 | note r = this | |
| 132 | show ?case by (rule r) simp_all | |
| 133 | qed | |
| 134 | ||
| 135 | text {*
 | |
| 136 | The following proof, although quite elegant from a mathematical point of view, | |
| 137 | leads to an exponential program: | |
| 138 | *} | |
| 139 | ||
| 140 | theorem pigeonhole_slow: | |
| 141 | "\<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" | |
| 142 | proof (induct n) | |
| 143 | case 0 | |
| 144 | have "Suc 0 \<le> Suc 0" .. | |
| 145 | moreover have "0 < Suc 0" .. | |
| 146 | moreover from 0 have "f (Suc 0) = f 0" by simp | |
| 17604 | 147 | ultimately show ?case by iprover | 
| 17024 | 148 | next | 
| 149 | case (Suc n) | |
| 25418 | 150 | from search [OF nat_eq_dec] show ?case | 
| 17024 | 151 | proof | 
| 152 | assume "\<exists>j < Suc (Suc n). f (Suc (Suc n)) = f j" | |
| 17604 | 153 | thus ?case by (iprover intro: le_refl) | 
| 17024 | 154 | next | 
| 155 | assume "\<not> (\<exists>j < Suc (Suc n). f (Suc (Suc n)) = f j)" | |
| 17604 | 156 | hence nex: "\<forall>j < Suc (Suc n). f (Suc (Suc n)) \<noteq> f j" by iprover | 
| 17024 | 157 | let ?f = "\<lambda>i. if f i = Suc n then f (Suc (Suc n)) else f i" | 
| 158 | have "\<And>i. i \<le> Suc n \<Longrightarrow> ?f i \<le> n" | |
| 159 | proof - | |
| 160 | fix i assume i: "i \<le> Suc n" | |
| 161 | show "?thesis i" | |
| 162 | proof (cases "f i = Suc n") | |
| 163 | case True | |
| 164 | from i and nex have "f (Suc (Suc n)) \<noteq> f i" by simp | |
| 165 | with True have "f (Suc (Suc n)) \<noteq> Suc n" by simp | |
| 166 | moreover from Suc have "f (Suc (Suc n)) \<le> Suc n" by simp | |
| 167 | ultimately have "f (Suc (Suc n)) \<le> n" by simp | |
| 168 | with True show ?thesis by simp | |
| 169 | next | |
| 170 | case False | |
| 171 | from Suc and i have "f i \<le> Suc n" by simp | |
| 172 | with False show ?thesis by simp | |
| 173 | qed | |
| 174 | qed | |
| 175 | hence "\<exists>i j. i \<le> Suc n \<and> j < i \<and> ?f i = ?f j" by (rule Suc) | |
| 176 | then obtain i j where i: "i \<le> Suc n" and ji: "j < i" and f: "?f i = ?f j" | |
| 17604 | 177 | by iprover | 
| 17024 | 178 | have "f i = f j" | 
| 179 | proof (cases "f i = Suc n") | |
| 180 | case True | |
| 181 | show ?thesis | |
| 182 | proof (cases "f j = Suc n") | |
| 183 | assume "f j = Suc n" | |
| 184 | with True show ?thesis by simp | |
| 185 | next | |
| 186 | assume "f j \<noteq> Suc n" | |
| 187 | moreover from i ji nex have "f (Suc (Suc n)) \<noteq> f j" by simp | |
| 188 | ultimately show ?thesis using True f by simp | |
| 189 | qed | |
| 190 | next | |
| 191 | case False | |
| 192 | show ?thesis | |
| 193 | proof (cases "f j = Suc n") | |
| 194 | assume "f j = Suc n" | |
| 195 | moreover from i nex have "f (Suc (Suc n)) \<noteq> f i" by simp | |
| 196 | ultimately show ?thesis using False f by simp | |
| 197 | next | |
| 198 | assume "f j \<noteq> Suc n" | |
| 199 | with False f show ?thesis by simp | |
| 200 | qed | |
| 201 | qed | |
| 202 | moreover from i have "i \<le> Suc (Suc n)" by simp | |
| 17604 | 203 | ultimately show ?thesis using ji by iprover | 
| 17024 | 204 | qed | 
| 205 | qed | |
| 206 | ||
| 207 | extract pigeonhole pigeonhole_slow | |
| 208 | ||
| 209 | text {*
 | |
| 210 | The programs extracted from the above proofs look as follows: | |
| 211 | @{thm [display] pigeonhole_def}
 | |
| 212 | @{thm [display] pigeonhole_slow_def}
 | |
| 213 | The program for searching for an element in an array is | |
| 214 | @{thm [display,eta_contract=false] search_def}
 | |
| 215 | The correctness statement for @{term "pigeonhole"} is
 | |
| 216 | @{thm [display] pigeonhole_correctness [no_vars]}
 | |
| 217 | ||
| 218 | In order to analyze the speed of the above programs, | |
| 219 | we generate ML code from them. | |
| 220 | *} | |
| 221 | ||
| 27982 | 222 | instantiation nat :: default | 
| 223 | begin | |
| 224 | ||
| 225 | definition "default = (0::nat)" | |
| 226 | ||
| 227 | instance .. | |
| 228 | ||
| 229 | end | |
| 230 | ||
| 231 | instantiation * :: (default, default) default | |
| 232 | begin | |
| 233 | ||
| 234 | definition "default = (default, default)" | |
| 235 | ||
| 236 | instance .. | |
| 237 | ||
| 238 | end | |
| 239 | ||
| 240 | consts_code | |
| 241 |   "default :: nat" ("{* 0::nat *}")
 | |
| 242 |   "default :: nat \<times> nat" ("{* (0::nat, 0::nat) *}")
 | |
| 243 | ||
| 20837 | 244 | definition | 
| 23810 | 245 | "test n u = pigeonhole n (\<lambda>m. m - 1)" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21127diff
changeset | 246 | definition | 
| 23810 | 247 | "test' n u = pigeonhole_slow n (\<lambda>m. m - 1)" | 
| 22507 | 248 | definition | 
| 249 | "test'' u = pigeonhole 8 (op ! [0, 1, 2, 3, 4, 5, 6, 3, 7, 8])" | |
| 20837 | 250 | |
| 27436 | 251 | code_module PH | 
| 22507 | 252 | contains | 
| 253 | test = test | |
| 254 | test' = test' | |
| 255 | test'' = test'' | |
| 256 | ||
| 27436 | 257 | ML "timeit (PH.test 10)" | 
| 258 | ML "timeit (@{code test} 10)" 
 | |
| 22507 | 259 | |
| 27436 | 260 | ML "timeit (PH.test' 10)" | 
| 261 | ML "timeit (@{code test'} 10)"
 | |
| 22507 | 262 | |
| 27436 | 263 | ML "timeit (PH.test 20)" | 
| 264 | ML "timeit (@{code test} 20)"
 | |
| 22507 | 265 | |
| 27436 | 266 | ML "timeit (PH.test' 20)" | 
| 267 | ML "timeit (@{code test'} 20)"
 | |
| 22507 | 268 | |
| 27436 | 269 | ML "timeit (PH.test 25)" | 
| 270 | ML "timeit (@{code test} 25)"
 | |
| 22507 | 271 | |
| 27436 | 272 | ML "timeit (PH.test' 25)" | 
| 273 | ML "timeit (@{code test'} 25)"
 | |
| 22507 | 274 | |
| 27436 | 275 | ML "timeit (PH.test 500)" | 
| 276 | ML "timeit (@{code test} 500)"
 | |
| 22507 | 277 | |
| 27436 | 278 | ML "timeit PH.test''" | 
| 279 | ML "timeit @{code test''}"
 | |
| 20837 | 280 | |
| 17024 | 281 | end |