| author | haftmann | 
| Mon, 07 Apr 2008 15:37:31 +0200 | |
| changeset 26564 | 631ce7f6bdc6 | 
| parent 26304 | 02fbd0e7954a | 
| child 27104 | 791607529f6d | 
| permissions | -rw-r--r-- | 
| 24999 | 1 | (* ID: $Id$ | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Type of indices *}
 | |
| 6 | ||
| 7 | theory Code_Index | |
| 25691 | 8 | imports ATP_Linkup | 
| 24999 | 9 | begin | 
| 10 | ||
| 11 | text {*
 | |
| 25767 | 12 |   Indices are isomorphic to HOL @{typ nat} but
 | 
| 24999 | 13 | mapped to target-language builtin integers | 
| 14 | *} | |
| 15 | ||
| 16 | subsection {* Datatype of indices *}
 | |
| 17 | ||
| 26140 | 18 | typedef index = "UNIV \<Colon> nat set" | 
| 19 | morphisms nat_of_index index_of_nat by rule | |
| 24999 | 20 | |
| 26140 | 21 | lemma index_of_nat_nat_of_index [simp]: | 
| 22 | "index_of_nat (nat_of_index k) = k" | |
| 23 | by (rule nat_of_index_inverse) | |
| 24999 | 24 | |
| 26140 | 25 | lemma nat_of_index_index_of_nat [simp]: | 
| 26 | "nat_of_index (index_of_nat n) = n" | |
| 27 | by (rule index_of_nat_inverse) | |
| 28 | (unfold index_def, rule UNIV_I) | |
| 24999 | 29 | |
| 30 | lemma index: | |
| 25767 | 31 | "(\<And>n\<Colon>index. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (index_of_nat n))" | 
| 24999 | 32 | proof | 
| 25767 | 33 | fix n :: nat | 
| 34 | assume "\<And>n\<Colon>index. PROP P n" | |
| 35 | then show "PROP P (index_of_nat n)" . | |
| 24999 | 36 | next | 
| 25767 | 37 | fix n :: index | 
| 38 | assume "\<And>n\<Colon>nat. PROP P (index_of_nat n)" | |
| 39 | then have "PROP P (index_of_nat (nat_of_index n))" . | |
| 40 | then show "PROP P n" by simp | |
| 24999 | 41 | qed | 
| 42 | ||
| 26140 | 43 | lemma index_case: | 
| 44 | assumes "\<And>n. k = index_of_nat n \<Longrightarrow> P" | |
| 45 | shows P | |
| 46 | by (rule assms [of "nat_of_index k"]) simp | |
| 47 | ||
| 26304 | 48 | lemma index_induct_raw: | 
| 26140 | 49 | assumes "\<And>n. P (index_of_nat n)" | 
| 50 | shows "P k" | |
| 51 | proof - | |
| 52 | from assms have "P (index_of_nat (nat_of_index k))" . | |
| 53 | then show ?thesis by simp | |
| 54 | qed | |
| 55 | ||
| 56 | lemma nat_of_index_inject [simp]: | |
| 57 | "nat_of_index k = nat_of_index l \<longleftrightarrow> k = l" | |
| 58 | by (rule nat_of_index_inject) | |
| 59 | ||
| 60 | lemma index_of_nat_inject [simp]: | |
| 61 | "index_of_nat n = index_of_nat m \<longleftrightarrow> n = m" | |
| 62 | by (auto intro!: index_of_nat_inject simp add: index_def) | |
| 63 | ||
| 64 | instantiation index :: zero | |
| 65 | begin | |
| 66 | ||
| 67 | definition [simp, code func del]: | |
| 68 | "0 = index_of_nat 0" | |
| 69 | ||
| 70 | instance .. | |
| 71 | ||
| 72 | end | |
| 73 | ||
| 74 | definition [simp]: | |
| 75 | "Suc_index k = index_of_nat (Suc (nat_of_index k))" | |
| 76 | ||
| 77 | lemma index_induct: "P 0 \<Longrightarrow> (\<And>k. P k \<Longrightarrow> P (Suc_index k)) \<Longrightarrow> P k" | |
| 78 | proof - | |
| 79 | assume "P 0" then have init: "P (index_of_nat 0)" by simp | |
| 80 | assume "\<And>k. P k \<Longrightarrow> P (Suc_index k)" | |
| 81 | then have "\<And>n. P (index_of_nat n) \<Longrightarrow> P (Suc_index (index_of_nat (n)))" . | |
| 82 | then have step: "\<And>n. P (index_of_nat n) \<Longrightarrow> P (index_of_nat (Suc n))" by simp | |
| 83 | from init step have "P (index_of_nat (nat_of_index k))" | |
| 84 | by (induct "nat_of_index k") simp_all | |
| 85 | then show "P k" by simp | |
| 86 | qed | |
| 87 | ||
| 88 | lemma Suc_not_Zero_index: "Suc_index k \<noteq> 0" | |
| 89 | by simp | |
| 90 | ||
| 91 | lemma Zero_not_Suc_index: "0 \<noteq> Suc_index k" | |
| 92 | by simp | |
| 93 | ||
| 94 | lemma Suc_Suc_index_eq: "Suc_index k = Suc_index l \<longleftrightarrow> k = l" | |
| 95 | by simp | |
| 96 | ||
| 97 | rep_datatype index | |
| 98 | distinct Suc_not_Zero_index Zero_not_Suc_index | |
| 99 | inject Suc_Suc_index_eq | |
| 100 | induction index_induct | |
| 101 | ||
| 102 | lemmas [code func del] = index.recs index.cases | |
| 103 | ||
| 104 | declare index_case [case_names nat, cases type: index] | |
| 105 | declare index_induct [case_names nat, induct type: index] | |
| 106 | ||
| 107 | lemma [code func]: | |
| 108 | "index_size = nat_of_index" | |
| 109 | proof (rule ext) | |
| 110 | fix k | |
| 111 | have "index_size k = nat_size (nat_of_index k)" | |
| 112 | by (induct k rule: index.induct) (simp_all del: zero_index_def Suc_index_def, simp_all) | |
| 113 | also have "nat_size (nat_of_index k) = nat_of_index k" by (induct "nat_of_index k") simp_all | |
| 114 | finally show "index_size k = nat_of_index k" . | |
| 115 | qed | |
| 116 | ||
| 117 | lemma [code func]: | |
| 118 | "size = nat_of_index" | |
| 119 | proof (rule ext) | |
| 120 | fix k | |
| 121 | show "size k = nat_of_index k" | |
| 122 | by (induct k) (simp_all del: zero_index_def Suc_index_def, simp_all) | |
| 123 | qed | |
| 124 | ||
| 125 | lemma [code func]: | |
| 126 | "k = l \<longleftrightarrow> nat_of_index k = nat_of_index l" | |
| 127 | by (cases k, cases l) simp | |
| 24999 | 128 | |
| 129 | ||
| 25767 | 130 | subsection {* Indices as datatype of ints *}
 | 
| 131 | ||
| 132 | instantiation index :: number | |
| 133 | begin | |
| 24999 | 134 | |
| 25767 | 135 | definition | 
| 136 | "number_of = index_of_nat o nat" | |
| 137 | ||
| 138 | instance .. | |
| 139 | ||
| 140 | end | |
| 24999 | 141 | |
| 26264 | 142 | lemma nat_of_index_number [simp]: | 
| 143 | "nat_of_index (number_of k) = number_of k" | |
| 144 | by (simp add: number_of_index_def nat_number_of_def number_of_is_id) | |
| 145 | ||
| 24999 | 146 | code_datatype "number_of \<Colon> int \<Rightarrow> index" | 
| 147 | ||
| 148 | ||
| 149 | subsection {* Basic arithmetic *}
 | |
| 150 | ||
| 25767 | 151 | instantiation index :: "{minus, ordered_semidom, Divides.div, linorder}"
 | 
| 152 | begin | |
| 24999 | 153 | |
| 154 | lemma zero_index_code [code inline, code func]: | |
| 155 | "(0\<Colon>index) = Numeral0" | |
| 25767 | 156 | by (simp add: number_of_index_def Pls_def) | 
| 25967 | 157 | lemma [code post]: "Numeral0 = (0\<Colon>index)" | 
| 158 | using zero_index_code .. | |
| 25767 | 159 | |
| 160 | definition [simp, code func del]: | |
| 161 | "(1\<Colon>index) = index_of_nat 1" | |
| 24999 | 162 | |
| 163 | lemma one_index_code [code inline, code func]: | |
| 164 | "(1\<Colon>index) = Numeral1" | |
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
26009diff
changeset | 165 | by (simp add: number_of_index_def Pls_def Bit1_def) | 
| 25967 | 166 | lemma [code post]: "Numeral1 = (1\<Colon>index)" | 
| 167 | using one_index_code .. | |
| 25767 | 168 | |
| 169 | definition [simp, code func del]: | |
| 170 | "n + m = index_of_nat (nat_of_index n + nat_of_index m)" | |
| 171 | ||
| 172 | lemma plus_index_code [code func]: | |
| 173 | "index_of_nat n + index_of_nat m = index_of_nat (n + m)" | |
| 174 | by simp | |
| 175 | ||
| 176 | definition [simp, code func del]: | |
| 177 | "n - m = index_of_nat (nat_of_index n - nat_of_index m)" | |
| 178 | ||
| 179 | definition [simp, code func del]: | |
| 180 | "n * m = index_of_nat (nat_of_index n * nat_of_index m)" | |
| 181 | ||
| 182 | lemma times_index_code [code func]: | |
| 183 | "index_of_nat n * index_of_nat m = index_of_nat (n * m)" | |
| 24999 | 184 | by simp | 
| 185 | ||
| 25767 | 186 | definition [simp, code func del]: | 
| 187 | "n div m = index_of_nat (nat_of_index n div nat_of_index m)" | |
| 24999 | 188 | |
| 25767 | 189 | definition [simp, code func del]: | 
| 190 | "n mod m = index_of_nat (nat_of_index n mod nat_of_index m)" | |
| 24999 | 191 | |
| 25767 | 192 | lemma div_index_code [code func]: | 
| 193 | "index_of_nat n div index_of_nat m = index_of_nat (n div m)" | |
| 194 | by simp | |
| 25335 | 195 | |
| 25767 | 196 | lemma mod_index_code [code func]: | 
| 197 | "index_of_nat n mod index_of_nat m = index_of_nat (n mod m)" | |
| 198 | by simp | |
| 24999 | 199 | |
| 25767 | 200 | definition [simp, code func del]: | 
| 201 | "n \<le> m \<longleftrightarrow> nat_of_index n \<le> nat_of_index m" | |
| 24999 | 202 | |
| 25767 | 203 | definition [simp, code func del]: | 
| 204 | "n < m \<longleftrightarrow> nat_of_index n < nat_of_index m" | |
| 24999 | 205 | |
| 25767 | 206 | lemma less_eq_index_code [code func]: | 
| 207 | "index_of_nat n \<le> index_of_nat m \<longleftrightarrow> n \<le> m" | |
| 208 | by simp | |
| 24999 | 209 | |
| 25767 | 210 | lemma less_index_code [code func]: | 
| 211 | "index_of_nat n < index_of_nat m \<longleftrightarrow> n < m" | |
| 212 | by simp | |
| 24999 | 213 | |
| 25767 | 214 | instance by default (auto simp add: left_distrib index) | 
| 215 | ||
| 216 | end | |
| 24999 | 217 | |
| 26140 | 218 | lemma Suc_index_minus_one: "Suc_index n - 1 = n" by simp | 
| 219 | ||
| 25928 | 220 | lemma index_of_nat_code [code]: | 
| 25918 | 221 | "index_of_nat = of_nat" | 
| 222 | proof | |
| 223 | fix n :: nat | |
| 224 | have "of_nat n = index_of_nat n" | |
| 225 | by (induct n) simp_all | |
| 226 | then show "index_of_nat n = of_nat n" | |
| 227 | by (rule sym) | |
| 228 | qed | |
| 229 | ||
| 25928 | 230 | lemma index_not_eq_zero: "i \<noteq> index_of_nat 0 \<longleftrightarrow> i \<ge> 1" | 
| 231 | by (cases i) auto | |
| 232 | ||
| 233 | definition | |
| 234 | nat_of_index_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" | |
| 235 | where | |
| 236 | "nat_of_index_aux i n = nat_of_index i + n" | |
| 237 | ||
| 238 | lemma nat_of_index_aux_code [code]: | |
| 239 | "nat_of_index_aux i n = (if i = 0 then n else nat_of_index_aux (i - 1) (Suc n))" | |
| 240 | by (auto simp add: nat_of_index_aux_def index_not_eq_zero) | |
| 241 | ||
| 242 | lemma nat_of_index_code [code]: | |
| 243 | "nat_of_index i = nat_of_index_aux i 0" | |
| 244 | by (simp add: nat_of_index_aux_def) | |
| 25918 | 245 | |
| 24999 | 246 | |
| 247 | subsection {* ML interface *}
 | |
| 248 | ||
| 249 | ML {*
 | |
| 250 | structure Index = | |
| 251 | struct | |
| 252 | ||
| 25928 | 253 | fun mk k = HOLogic.mk_number @{typ index} k;
 | 
| 24999 | 254 | |
| 255 | end; | |
| 256 | *} | |
| 257 | ||
| 258 | ||
| 26009 | 259 | subsection {* Specialized @{term "op - \<Colon> index \<Rightarrow> index \<Rightarrow> index"},
 | 
| 260 |   @{term "op div \<Colon> index \<Rightarrow> index \<Rightarrow> index"} and @{term "op mod \<Colon> index \<Rightarrow> index \<Rightarrow> index"}
 | |
| 261 | operations *} | |
| 262 | ||
| 263 | definition | |
| 264 | minus_index_aux :: "index \<Rightarrow> index \<Rightarrow> index" | |
| 265 | where | |
| 266 | [code func del]: "minus_index_aux = op -" | |
| 267 | ||
| 268 | lemma [code func]: "op - = minus_index_aux" | |
| 269 | using minus_index_aux_def .. | |
| 270 | ||
| 271 | definition | |
| 272 | div_mod_index :: "index \<Rightarrow> index \<Rightarrow> index \<times> index" | |
| 273 | where | |
| 274 | [code func del]: "div_mod_index n m = (n div m, n mod m)" | |
| 275 | ||
| 276 | lemma [code func]: | |
| 277 | "div_mod_index n m = (if m = 0 then (0, n) else (n div m, n mod m))" | |
| 278 | unfolding div_mod_index_def by auto | |
| 279 | ||
| 280 | lemma [code func]: | |
| 281 | "n div m = fst (div_mod_index n m)" | |
| 282 | unfolding div_mod_index_def by simp | |
| 283 | ||
| 284 | lemma [code func]: | |
| 285 | "n mod m = snd (div_mod_index n m)" | |
| 286 | unfolding div_mod_index_def by simp | |
| 287 | ||
| 288 | ||
| 24999 | 289 | subsection {* Code serialization *}
 | 
| 290 | ||
| 25767 | 291 | text {* Implementation of indices by bounded integers *}
 | 
| 292 | ||
| 24999 | 293 | code_type index | 
| 294 | (SML "int") | |
| 295 | (OCaml "int") | |
| 25967 | 296 | (Haskell "Int") | 
| 24999 | 297 | |
| 298 | code_instance index :: eq | |
| 299 | (Haskell -) | |
| 300 | ||
| 301 | setup {*
 | |
| 25928 | 302 |   fold (Numeral.add_code @{const_name number_index_inst.number_of_index}
 | 
| 303 | false false) ["SML", "OCaml", "Haskell"] | |
| 24999 | 304 | *} | 
| 305 | ||
| 25918 | 306 | code_reserved SML Int int | 
| 307 | code_reserved OCaml Pervasives int | |
| 24999 | 308 | |
| 309 | code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index" | |
| 25928 | 310 | (SML "Int.+/ ((_),/ (_))") | 
| 25967 | 311 | (OCaml "Pervasives.( + )") | 
| 24999 | 312 | (Haskell infixl 6 "+") | 
| 313 | ||
| 26009 | 314 | code_const "minus_index_aux \<Colon> index \<Rightarrow> index \<Rightarrow> index" | 
| 25918 | 315 | (SML "Int.max/ (_/ -/ _,/ 0 : int)") | 
| 316 | (OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ") | |
| 317 | (Haskell "max/ (_/ -/ _)/ (0 :: Int)") | |
| 24999 | 318 | |
| 319 | code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index" | |
| 25928 | 320 | (SML "Int.*/ ((_),/ (_))") | 
| 25967 | 321 | (OCaml "Pervasives.( * )") | 
| 24999 | 322 | (Haskell infixl 7 "*") | 
| 323 | ||
| 26009 | 324 | code_const div_mod_index | 
| 325 | (SML "(fn n => fn m =>/ (n div m, n mod m))") | |
| 326 | (OCaml "(fun n -> fun m ->/ (n '/ m, n mod m))") | |
| 327 | (Haskell "divMod") | |
| 25928 | 328 | |
| 24999 | 329 | code_const "op = \<Colon> index \<Rightarrow> index \<Rightarrow> bool" | 
| 330 | (SML "!((_ : Int.int) = _)") | |
| 25967 | 331 | (OCaml "!((_ : int) = _)") | 
| 24999 | 332 | (Haskell infixl 4 "==") | 
| 333 | ||
| 334 | code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool" | |
| 25928 | 335 | (SML "Int.<=/ ((_),/ (_))") | 
| 25967 | 336 | (OCaml "!((_ : int) <= _)") | 
| 24999 | 337 | (Haskell infix 4 "<=") | 
| 338 | ||
| 339 | code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool" | |
| 25928 | 340 | (SML "Int.</ ((_),/ (_))") | 
| 25967 | 341 | (OCaml "!((_ : int) < _)") | 
| 24999 | 342 | (Haskell infix 4 "<") | 
| 343 | ||
| 344 | end |