| author | wenzelm | 
| Tue, 24 Mar 2009 16:11:09 +0100 | |
| changeset 30705 | e8ab35c6ade6 | 
| parent 30663 | 0b6aff7451b2 | 
| child 30926 | 3a30613aa469 | 
| permissions | -rw-r--r-- | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 24999 | 2 | |
| 3 | header {* Type of indices *}
 | |
| 4 | ||
| 5 | theory Code_Index | |
| 30663 
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
 haftmann parents: 
30245diff
changeset | 6 | imports Main | 
| 24999 | 7 | begin | 
| 8 | ||
| 9 | text {*
 | |
| 25767 | 10 |   Indices are isomorphic to HOL @{typ nat} but
 | 
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26304diff
changeset | 11 | mapped to target-language builtin integers. | 
| 24999 | 12 | *} | 
| 13 | ||
| 14 | subsection {* Datatype of indices *}
 | |
| 15 | ||
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 16 | typedef (open) index = "UNIV \<Colon> nat set" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 17 | morphisms nat_of of_nat by rule | 
| 24999 | 18 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 19 | lemma of_nat_nat_of [simp]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 20 | "of_nat (nat_of k) = k" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 21 | by (rule nat_of_inverse) | 
| 24999 | 22 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 23 | lemma nat_of_of_nat [simp]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 24 | "nat_of (of_nat n) = n" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 25 | by (rule of_nat_inverse) (rule UNIV_I) | 
| 24999 | 26 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 27 | lemma [measure_function]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 28 | "is_measure nat_of" by (rule is_measure_trivial) | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 29 | |
| 24999 | 30 | lemma index: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 31 | "(\<And>n\<Colon>index. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))" | 
| 24999 | 32 | proof | 
| 25767 | 33 | fix n :: nat | 
| 34 | assume "\<And>n\<Colon>index. PROP P n" | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 35 | then show "PROP P (of_nat n)" . | 
| 24999 | 36 | next | 
| 25767 | 37 | fix n :: index | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 38 | assume "\<And>n\<Colon>nat. PROP P (of_nat n)" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 39 | then have "PROP P (of_nat (nat_of n))" . | 
| 25767 | 40 | then show "PROP P n" by simp | 
| 24999 | 41 | qed | 
| 42 | ||
| 26140 | 43 | lemma index_case: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 44 | assumes "\<And>n. k = of_nat n \<Longrightarrow> P" | 
| 26140 | 45 | shows P | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 46 | by (rule assms [of "nat_of k"]) simp | 
| 26140 | 47 | |
| 26304 | 48 | lemma index_induct_raw: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 49 | assumes "\<And>n. P (of_nat n)" | 
| 26140 | 50 | shows "P k" | 
| 51 | proof - | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 52 | from assms have "P (of_nat (nat_of k))" . | 
| 26140 | 53 | then show ?thesis by simp | 
| 54 | qed | |
| 55 | ||
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 56 | lemma nat_of_inject [simp]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 57 | "nat_of k = nat_of l \<longleftrightarrow> k = l" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 58 | by (rule nat_of_inject) | 
| 26140 | 59 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 60 | lemma of_nat_inject [simp]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 61 | "of_nat n = of_nat m \<longleftrightarrow> n = m" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 62 | by (rule of_nat_inject) (rule UNIV_I)+ | 
| 26140 | 63 | |
| 64 | instantiation index :: zero | |
| 65 | begin | |
| 66 | ||
| 28562 | 67 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 68 | "0 = of_nat 0" | 
| 26140 | 69 | |
| 70 | instance .. | |
| 71 | ||
| 72 | end | |
| 73 | ||
| 74 | definition [simp]: | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 75 | "Suc_index k = of_nat (Suc (nat_of k))" | 
| 26140 | 76 | |
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26304diff
changeset | 77 | rep_datatype "0 \<Colon> index" Suc_index | 
| 26140 | 78 | proof - | 
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26304diff
changeset | 79 | fix P :: "index \<Rightarrow> bool" | 
| 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26304diff
changeset | 80 | fix k :: index | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 81 | assume "P 0" then have init: "P (of_nat 0)" by simp | 
| 26140 | 82 | assume "\<And>k. P k \<Longrightarrow> P (Suc_index k)" | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 83 | then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc_index (of_nat n))" . | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 84 | then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Suc n))" by simp | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 85 | from init step have "P (of_nat (nat_of k))" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 86 | by (induct "nat_of k") simp_all | 
| 26140 | 87 | then show "P k" by simp | 
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26304diff
changeset | 88 | qed simp_all | 
| 26140 | 89 | |
| 90 | declare index_case [case_names nat, cases type: index] | |
| 27104 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 haftmann parents: 
26304diff
changeset | 91 | declare index.induct [case_names nat, induct type: index] | 
| 26140 | 92 | |
| 30245 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 93 | lemma index_decr [termination_simp]: | 
| 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 94 | "k \<noteq> Code_Index.of_nat 0 \<Longrightarrow> Code_Index.nat_of k - Suc 0 < Code_Index.nat_of k" | 
| 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 95 | by (cases k) simp | 
| 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 96 | |
| 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 97 | lemma [simp, code]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 98 | "index_size = nat_of" | 
| 26140 | 99 | proof (rule ext) | 
| 100 | fix k | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 101 | have "index_size k = nat_size (nat_of k)" | 
| 26140 | 102 | by (induct k rule: index.induct) (simp_all del: zero_index_def Suc_index_def, simp_all) | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 103 | also have "nat_size (nat_of k) = nat_of k" by (induct "nat_of k") simp_all | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 104 | finally show "index_size k = nat_of k" . | 
| 26140 | 105 | qed | 
| 106 | ||
| 30245 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 107 | lemma [simp, code]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 108 | "size = nat_of" | 
| 26140 | 109 | proof (rule ext) | 
| 110 | fix k | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 111 | show "size k = nat_of k" | 
| 26140 | 112 | by (induct k) (simp_all del: zero_index_def Suc_index_def, simp_all) | 
| 113 | qed | |
| 114 | ||
| 30245 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 115 | lemmas [code del] = index.recs index.cases | 
| 
e67f42ac1157
consequent rewrite of index_size, size [index] to nat_of; support pseudo-primrec sepcifications with fun
 haftmann parents: 
29823diff
changeset | 116 | |
| 28562 | 117 | lemma [code]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 118 | "eq_class.eq k l \<longleftrightarrow> eq_class.eq (nat_of k) (nat_of l)" | 
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28228diff
changeset | 119 | by (cases k, cases l) (simp add: eq) | 
| 24999 | 120 | |
| 28351 | 121 | lemma [code nbe]: | 
| 122 | "eq_class.eq (k::index) k \<longleftrightarrow> True" | |
| 123 | by (rule HOL.eq_refl) | |
| 124 | ||
| 24999 | 125 | |
| 25767 | 126 | subsection {* Indices as datatype of ints *}
 | 
| 127 | ||
| 128 | instantiation index :: number | |
| 129 | begin | |
| 24999 | 130 | |
| 25767 | 131 | definition | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 132 | "number_of = of_nat o nat" | 
| 25767 | 133 | |
| 134 | instance .. | |
| 135 | ||
| 136 | end | |
| 24999 | 137 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 138 | lemma nat_of_number [simp]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 139 | "nat_of (number_of k) = number_of k" | 
| 26264 | 140 | by (simp add: number_of_index_def nat_number_of_def number_of_is_id) | 
| 141 | ||
| 24999 | 142 | code_datatype "number_of \<Colon> int \<Rightarrow> index" | 
| 143 | ||
| 144 | ||
| 145 | subsection {* Basic arithmetic *}
 | |
| 146 | ||
| 25767 | 147 | instantiation index :: "{minus, ordered_semidom, Divides.div, linorder}"
 | 
| 148 | begin | |
| 24999 | 149 | |
| 28562 | 150 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 151 | "(1\<Colon>index) = of_nat 1" | 
| 24999 | 152 | |
| 28562 | 153 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 154 | "n + m = of_nat (nat_of n + nat_of m)" | 
| 25767 | 155 | |
| 28562 | 156 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 157 | "n - m = of_nat (nat_of n - nat_of m)" | 
| 25767 | 158 | |
| 28562 | 159 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 160 | "n * m = of_nat (nat_of n * nat_of m)" | 
| 25767 | 161 | |
| 28562 | 162 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 163 | "n div m = of_nat (nat_of n div nat_of m)" | 
| 24999 | 164 | |
| 28562 | 165 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 166 | "n mod m = of_nat (nat_of n mod nat_of m)" | 
| 24999 | 167 | |
| 28562 | 168 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 169 | "n \<le> m \<longleftrightarrow> nat_of n \<le> nat_of m" | 
| 24999 | 170 | |
| 28562 | 171 | definition [simp, code del]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 172 | "n < m \<longleftrightarrow> nat_of n < nat_of m" | 
| 24999 | 173 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 174 | instance proof | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 175 | qed (auto simp add: left_distrib) | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 176 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 177 | end | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 178 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 179 | lemma zero_index_code [code inline, code]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 180 | "(0\<Colon>index) = Numeral0" | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 181 | by (simp add: number_of_index_def Pls_def) | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 182 | lemma [code post]: "Numeral0 = (0\<Colon>index)" | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 183 | using zero_index_code .. | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 184 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 185 | lemma one_index_code [code inline, code]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 186 | "(1\<Colon>index) = Numeral1" | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 187 | by (simp add: number_of_index_def Pls_def Bit1_def) | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 188 | lemma [code post]: "Numeral1 = (1\<Colon>index)" | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 189 | using one_index_code .. | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 190 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 191 | lemma plus_index_code [code nbe]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 192 | "of_nat n + of_nat m = of_nat (n + m)" | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 193 | by simp | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 194 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 195 | definition subtract_index :: "index \<Rightarrow> index \<Rightarrow> index" where | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 196 | [simp, code del]: "subtract_index = op -" | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 197 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 198 | lemma subtract_index_code [code nbe]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 199 | "subtract_index (of_nat n) (of_nat m) = of_nat (n - m)" | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 200 | by simp | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 201 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 202 | lemma minus_index_code [code]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 203 | "n - m = subtract_index n m" | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 204 | by simp | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 205 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 206 | lemma times_index_code [code nbe]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 207 | "of_nat n * of_nat m = of_nat (n * m)" | 
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 208 | by simp | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 209 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 210 | lemma less_eq_index_code [code nbe]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 211 | "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m" | 
| 25767 | 212 | by simp | 
| 24999 | 213 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 214 | lemma less_index_code [code nbe]: | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 215 | "of_nat n < of_nat m \<longleftrightarrow> n < m" | 
| 25767 | 216 | by simp | 
| 24999 | 217 | |
| 26140 | 218 | lemma Suc_index_minus_one: "Suc_index n - 1 = n" by simp | 
| 219 | ||
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 220 | lemma of_nat_code [code]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 221 | "of_nat = Nat.of_nat" | 
| 25918 | 222 | proof | 
| 223 | fix n :: nat | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 224 | have "Nat.of_nat n = of_nat n" | 
| 25918 | 225 | by (induct n) simp_all | 
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 226 | then show "of_nat n = Nat.of_nat n" | 
| 25918 | 227 | by (rule sym) | 
| 228 | qed | |
| 229 | ||
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 230 | lemma index_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1" | 
| 25928 | 231 | by (cases i) auto | 
| 232 | ||
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 233 | definition nat_of_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" where | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 234 | "nat_of_aux i n = nat_of i + n" | 
| 25928 | 235 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 236 | lemma nat_of_aux_code [code]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 237 | "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Suc n))" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 238 | by (auto simp add: nat_of_aux_def index_not_eq_zero) | 
| 25928 | 239 | |
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 240 | lemma nat_of_code [code]: | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 241 | "nat_of i = nat_of_aux i 0" | 
| 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 242 | by (simp add: nat_of_aux_def) | 
| 25918 | 243 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 244 | definition div_mod_index :: "index \<Rightarrow> index \<Rightarrow> index \<times> index" where | 
| 28562 | 245 | [code del]: "div_mod_index n m = (n div m, n mod m)" | 
| 26009 | 246 | |
| 28562 | 247 | lemma [code]: | 
| 26009 | 248 | "div_mod_index n m = (if m = 0 then (0, n) else (n div m, n mod m))" | 
| 249 | unfolding div_mod_index_def by auto | |
| 250 | ||
| 28562 | 251 | lemma [code]: | 
| 26009 | 252 | "n div m = fst (div_mod_index n m)" | 
| 253 | unfolding div_mod_index_def by simp | |
| 254 | ||
| 28562 | 255 | lemma [code]: | 
| 26009 | 256 | "n mod m = snd (div_mod_index n m)" | 
| 257 | unfolding div_mod_index_def by simp | |
| 258 | ||
| 29815 
9e94b7078fa5
mandatory prefix for index conversion operations
 haftmann parents: 
28708diff
changeset | 259 | hide (open) const of_nat nat_of | 
| 26009 | 260 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 261 | subsection {* ML interface *}
 | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 262 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 263 | ML {*
 | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 264 | structure Index = | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 265 | struct | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 266 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 267 | fun mk k = HOLogic.mk_number @{typ index} k;
 | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 268 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 269 | end; | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 270 | *} | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 271 | |
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 272 | |
| 28228 | 273 | subsection {* Code generator setup *}
 | 
| 24999 | 274 | |
| 25767 | 275 | text {* Implementation of indices by bounded integers *}
 | 
| 276 | ||
| 24999 | 277 | code_type index | 
| 278 | (SML "int") | |
| 279 | (OCaml "int") | |
| 25967 | 280 | (Haskell "Int") | 
| 24999 | 281 | |
| 282 | code_instance index :: eq | |
| 283 | (Haskell -) | |
| 284 | ||
| 285 | setup {*
 | |
| 25928 | 286 |   fold (Numeral.add_code @{const_name number_index_inst.number_of_index}
 | 
| 287 | false false) ["SML", "OCaml", "Haskell"] | |
| 24999 | 288 | *} | 
| 289 | ||
| 25918 | 290 | code_reserved SML Int int | 
| 291 | code_reserved OCaml Pervasives int | |
| 24999 | 292 | |
| 293 | code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index" | |
| 25928 | 294 | (SML "Int.+/ ((_),/ (_))") | 
| 25967 | 295 | (OCaml "Pervasives.( + )") | 
| 24999 | 296 | (Haskell infixl 6 "+") | 
| 297 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 298 | code_const "subtract_index \<Colon> index \<Rightarrow> index \<Rightarrow> index" | 
| 25918 | 299 | (SML "Int.max/ (_/ -/ _,/ 0 : int)") | 
| 300 | (OCaml "Pervasives.max/ (_/ -/ _)/ (0 : int) ") | |
| 301 | (Haskell "max/ (_/ -/ _)/ (0 :: Int)") | |
| 24999 | 302 | |
| 303 | code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index" | |
| 25928 | 304 | (SML "Int.*/ ((_),/ (_))") | 
| 25967 | 305 | (OCaml "Pervasives.( * )") | 
| 24999 | 306 | (Haskell infixl 7 "*") | 
| 307 | ||
| 26009 | 308 | code_const div_mod_index | 
| 29823 
0ab754d13ccd
session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
 haftmann parents: 
29815diff
changeset | 309 | (SML "(fn n => fn m =>/ if m = 0/ then (0, n) else/ (n div m, n mod m))") | 
| 
0ab754d13ccd
session Reflecion renamed to Decision_Procs, moved Dense_Linear_Order there
 haftmann parents: 
29815diff
changeset | 310 | (OCaml "(fun n -> fun m ->/ if m = 0/ then (0, n) else/ (n '/ m, n mod m))") | 
| 26009 | 311 | (Haskell "divMod") | 
| 25928 | 312 | |
| 28346 
b8390cd56b8f
discontinued special treatment of op = vs. eq_class.eq
 haftmann parents: 
28228diff
changeset | 313 | code_const "eq_class.eq \<Colon> index \<Rightarrow> index \<Rightarrow> bool" | 
| 24999 | 314 | (SML "!((_ : Int.int) = _)") | 
| 25967 | 315 | (OCaml "!((_ : int) = _)") | 
| 24999 | 316 | (Haskell infixl 4 "==") | 
| 317 | ||
| 318 | code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool" | |
| 25928 | 319 | (SML "Int.<=/ ((_),/ (_))") | 
| 25967 | 320 | (OCaml "!((_ : int) <= _)") | 
| 24999 | 321 | (Haskell infix 4 "<=") | 
| 322 | ||
| 323 | code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool" | |
| 25928 | 324 | (SML "Int.</ ((_),/ (_))") | 
| 25967 | 325 | (OCaml "!((_ : int) < _)") | 
| 24999 | 326 | (Haskell infix 4 "<") | 
| 327 | ||
| 28228 | 328 | text {* Evaluation *}
 | 
| 329 | ||
| 28562 | 330 | lemma [code, code del]: | 
| 28228 | 331 | "(Code_Eval.term_of \<Colon> index \<Rightarrow> term) = Code_Eval.term_of" .. | 
| 332 | ||
| 333 | code_const "Code_Eval.term_of \<Colon> index \<Rightarrow> term" | |
| 334 | (SML "HOLogic.mk'_number/ HOLogic.indexT/ (IntInf.fromInt/ _)") | |
| 335 | ||
| 24999 | 336 | end |