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