| 24999 |      1 | (*  ID:         $Id$
 | 
|  |      2 |     Author:     Florian Haftmann, TU Muenchen
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
|  |      5 | header {* Type of indices *}
 | 
|  |      6 | 
 | 
|  |      7 | theory Code_Index
 | 
|  |      8 | imports PreList
 | 
|  |      9 | begin
 | 
|  |     10 | 
 | 
|  |     11 | text {*
 | 
|  |     12 |   Indices are isomorphic to HOL @{typ int} but
 | 
|  |     13 |   mapped to target-language builtin integers
 | 
|  |     14 | *}
 | 
|  |     15 | 
 | 
|  |     16 | subsection {* Datatype of indices *}
 | 
|  |     17 | 
 | 
|  |     18 | datatype index = index_of_int int
 | 
|  |     19 | 
 | 
|  |     20 | lemmas [code func del] = index.recs index.cases
 | 
|  |     21 | 
 | 
|  |     22 | fun
 | 
|  |     23 |   int_of_index :: "index \<Rightarrow> int"
 | 
|  |     24 | where
 | 
|  |     25 |   "int_of_index (index_of_int k) = k"
 | 
|  |     26 | lemmas [code func del] = int_of_index.simps
 | 
|  |     27 | 
 | 
|  |     28 | lemma index_id [simp]:
 | 
|  |     29 |   "index_of_int (int_of_index k) = k"
 | 
|  |     30 |   by (cases k) simp_all
 | 
|  |     31 | 
 | 
|  |     32 | lemma index:
 | 
|  |     33 |   "(\<And>k\<Colon>index. PROP P k) \<equiv> (\<And>k\<Colon>int. PROP P (index_of_int k))"
 | 
|  |     34 | proof
 | 
|  |     35 |   fix k :: int
 | 
|  |     36 |   assume "\<And>k\<Colon>index. PROP P k"
 | 
|  |     37 |   then show "PROP P (index_of_int k)" .
 | 
|  |     38 | next
 | 
|  |     39 |   fix k :: index
 | 
|  |     40 |   assume "\<And>k\<Colon>int. PROP P (index_of_int k)"
 | 
|  |     41 |   then have "PROP P (index_of_int (int_of_index k))" .
 | 
|  |     42 |   then show "PROP P k" by simp
 | 
|  |     43 | qed
 | 
|  |     44 | 
 | 
|  |     45 | lemma [code func]: "size (k\<Colon>index) = 0"
 | 
|  |     46 |   by (cases k) simp_all
 | 
|  |     47 | 
 | 
|  |     48 | 
 | 
|  |     49 | subsection {* Built-in integers as datatype on numerals *}
 | 
|  |     50 | 
 | 
|  |     51 | instance index :: number
 | 
|  |     52 |   "number_of \<equiv> index_of_int" ..
 | 
|  |     53 | 
 | 
|  |     54 | code_datatype "number_of \<Colon> int \<Rightarrow> index"
 | 
|  |     55 | 
 | 
|  |     56 | lemma number_of_index_id [simp]:
 | 
|  |     57 |   "number_of (int_of_index k) = k"
 | 
|  |     58 |   unfolding number_of_index_def by simp
 | 
|  |     59 | 
 | 
|  |     60 | lemma number_of_index_shift:
 | 
|  |     61 |   "number_of k = index_of_int (number_of k)"
 | 
|  |     62 |   by (simp add: number_of_is_id number_of_index_def)
 | 
|  |     63 | 
 | 
| 25335 |     64 | lemma int_of_index_number_of [simp]:
 | 
|  |     65 |   "int_of_index (number_of k) = number_of k"
 | 
|  |     66 |   unfolding number_of_index_def number_of_is_id by simp
 | 
|  |     67 | 
 | 
| 24999 |     68 | 
 | 
|  |     69 | subsection {* Basic arithmetic *}
 | 
|  |     70 | 
 | 
|  |     71 | instance index :: zero
 | 
|  |     72 |   [simp]: "0 \<equiv> index_of_int 0" ..
 | 
|  |     73 | lemmas [code func del] = zero_index_def
 | 
|  |     74 | 
 | 
|  |     75 | instance index :: one
 | 
|  |     76 |   [simp]: "1 \<equiv> index_of_int 1" ..
 | 
|  |     77 | lemmas [code func del] = one_index_def
 | 
|  |     78 | 
 | 
|  |     79 | instance index :: plus
 | 
|  |     80 |   [simp]: "k + l \<equiv> index_of_int (int_of_index k + int_of_index l)" ..
 | 
|  |     81 | lemmas [code func del] = plus_index_def
 | 
|  |     82 | lemma plus_index_code [code func]:
 | 
|  |     83 |   "index_of_int k + index_of_int l = index_of_int (k + l)"
 | 
|  |     84 |   unfolding plus_index_def by simp
 | 
|  |     85 | 
 | 
|  |     86 | instance index :: minus
 | 
|  |     87 |   [simp]: "- k \<equiv> index_of_int (- int_of_index k)"
 | 
|  |     88 |   [simp]: "k - l \<equiv> index_of_int (int_of_index k - int_of_index l)" ..
 | 
|  |     89 | lemmas [code func del] = uminus_index_def minus_index_def
 | 
|  |     90 | lemma uminus_index_code [code func]:
 | 
|  |     91 |   "- index_of_int k \<equiv> index_of_int (- k)"
 | 
|  |     92 |   unfolding uminus_index_def by simp
 | 
|  |     93 | lemma minus_index_code [code func]:
 | 
|  |     94 |   "index_of_int k - index_of_int l = index_of_int (k - l)"
 | 
|  |     95 |   unfolding minus_index_def by simp
 | 
|  |     96 | 
 | 
|  |     97 | instance index :: times
 | 
|  |     98 |   [simp]: "k * l \<equiv> index_of_int (int_of_index k * int_of_index l)" ..
 | 
|  |     99 | lemmas [code func del] = times_index_def
 | 
|  |    100 | lemma times_index_code [code func]:
 | 
|  |    101 |   "index_of_int k * index_of_int l = index_of_int (k * l)"
 | 
|  |    102 |   unfolding times_index_def by simp
 | 
|  |    103 | 
 | 
|  |    104 | instance index :: ord
 | 
|  |    105 |   [simp]: "k \<le> l \<equiv> int_of_index k \<le> int_of_index l"
 | 
|  |    106 |   [simp]: "k < l \<equiv> int_of_index k < int_of_index l" ..
 | 
|  |    107 | lemmas [code func del] = less_eq_index_def less_index_def
 | 
|  |    108 | lemma less_eq_index_code [code func]:
 | 
|  |    109 |   "index_of_int k \<le> index_of_int l \<longleftrightarrow> k \<le> l"
 | 
|  |    110 |   unfolding less_eq_index_def by simp
 | 
|  |    111 | lemma less_index_code [code func]:
 | 
|  |    112 |   "index_of_int k < index_of_int l \<longleftrightarrow> k < l"
 | 
|  |    113 |   unfolding less_index_def by simp
 | 
|  |    114 | 
 | 
| 25335 |    115 | instance index :: "Divides.div"
 | 
|  |    116 |   [simp]: "k div l \<equiv> index_of_int (int_of_index k div int_of_index l)"
 | 
|  |    117 |   [simp]: "k mod l \<equiv> index_of_int (int_of_index k mod int_of_index l)" ..
 | 
|  |    118 | 
 | 
| 24999 |    119 | instance index :: ring_1
 | 
|  |    120 |   by default (auto simp add: left_distrib right_distrib)
 | 
|  |    121 | 
 | 
|  |    122 | lemma of_nat_index: "of_nat n = index_of_int (of_nat n)"
 | 
|  |    123 | proof (induct n)
 | 
|  |    124 |   case 0 show ?case by simp
 | 
|  |    125 | next
 | 
|  |    126 |   case (Suc n)
 | 
|  |    127 |   then have "int_of_index (index_of_int (int n))
 | 
|  |    128 |     = int_of_index (of_nat n)" by simp
 | 
|  |    129 |   then have "int n = int_of_index (of_nat n)" by simp
 | 
|  |    130 |   then show ?case by simp
 | 
|  |    131 | qed
 | 
|  |    132 | 
 | 
|  |    133 | instance index :: number_ring
 | 
|  |    134 |   by default
 | 
|  |    135 |     (simp_all add: left_distrib number_of_index_def of_int_of_nat of_nat_index)
 | 
|  |    136 | 
 | 
|  |    137 | lemma zero_index_code [code inline, code func]:
 | 
|  |    138 |   "(0\<Colon>index) = Numeral0"
 | 
|  |    139 |   by simp
 | 
|  |    140 | 
 | 
|  |    141 | lemma one_index_code [code inline, code func]:
 | 
|  |    142 |   "(1\<Colon>index) = Numeral1"
 | 
|  |    143 |   by simp
 | 
|  |    144 | 
 | 
|  |    145 | instance index :: abs
 | 
| 25502 |    146 |   "\<bar>k\<Colon>index\<bar> \<equiv> if k < 0 then -k else k" ..
 | 
| 24999 |    147 | 
 | 
|  |    148 | lemma index_of_int [code func]:
 | 
|  |    149 |   "index_of_int k = (if k = 0 then 0
 | 
|  |    150 |     else if k = -1 then -1
 | 
|  |    151 |     else let (l, m) = divAlg (k, 2) in 2 * index_of_int l +
 | 
|  |    152 |       (if m = 0 then 0 else 1))"
 | 
|  |    153 |   by (simp add: number_of_index_shift Let_def split_def divAlg_mod_div) arith
 | 
|  |    154 | 
 | 
| 25335 |    155 | lemma int_of_index [code func]:
 | 
|  |    156 |   "int_of_index k = (if k = 0 then 0
 | 
|  |    157 |     else if k = -1 then -1
 | 
|  |    158 |     else let l = k div 2; m = k mod 2 in 2 * int_of_index l +
 | 
|  |    159 |       (if m = 0 then 0 else 1))"
 | 
|  |    160 |   by (auto simp add: number_of_index_shift Let_def split_def) arith
 | 
|  |    161 | 
 | 
| 24999 |    162 | 
 | 
|  |    163 | subsection {* Conversion to and from @{typ nat} *}
 | 
|  |    164 | 
 | 
|  |    165 | definition
 | 
|  |    166 |   nat_of_index :: "index \<Rightarrow> nat"
 | 
|  |    167 | where
 | 
|  |    168 |   [code func del]: "nat_of_index = nat o int_of_index"
 | 
|  |    169 | 
 | 
|  |    170 | definition
 | 
|  |    171 |   nat_of_index_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" where
 | 
|  |    172 |   [code func del]: "nat_of_index_aux i n = nat_of_index i + n"
 | 
|  |    173 | 
 | 
|  |    174 | lemma nat_of_index_aux_code [code]:
 | 
|  |    175 |   "nat_of_index_aux i n = (if i \<le> 0 then n else nat_of_index_aux (i - 1) (Suc n))"
 | 
|  |    176 |   by (auto simp add: nat_of_index_aux_def nat_of_index_def)
 | 
|  |    177 | 
 | 
|  |    178 | lemma nat_of_index_code [code]:
 | 
|  |    179 |   "nat_of_index i = nat_of_index_aux i 0"
 | 
|  |    180 |   by (simp add: nat_of_index_aux_def)
 | 
|  |    181 | 
 | 
|  |    182 | definition
 | 
|  |    183 |   index_of_nat :: "nat \<Rightarrow> index"
 | 
|  |    184 | where
 | 
|  |    185 |   [code func del]: "index_of_nat = index_of_int o of_nat"
 | 
|  |    186 | 
 | 
|  |    187 | lemma index_of_nat [code func]:
 | 
|  |    188 |   "index_of_nat 0 = 0"
 | 
|  |    189 |   "index_of_nat (Suc n) = index_of_nat n + 1"
 | 
|  |    190 |   unfolding index_of_nat_def by simp_all
 | 
|  |    191 | 
 | 
|  |    192 | lemma index_nat_id [simp]:
 | 
|  |    193 |   "nat_of_index (index_of_nat n) = n"
 | 
|  |    194 |   "index_of_nat (nat_of_index i) = (if i \<le> 0 then 0 else i)"
 | 
|  |    195 |   unfolding index_of_nat_def nat_of_index_def by simp_all
 | 
|  |    196 | 
 | 
|  |    197 | 
 | 
|  |    198 | subsection {* ML interface *}
 | 
|  |    199 | 
 | 
|  |    200 | ML {*
 | 
|  |    201 | structure Index =
 | 
|  |    202 | struct
 | 
|  |    203 | 
 | 
|  |    204 | fun mk k = @{term index_of_int} $ HOLogic.mk_number @{typ index} k;
 | 
|  |    205 | 
 | 
|  |    206 | end;
 | 
|  |    207 | *}
 | 
|  |    208 | 
 | 
|  |    209 | 
 | 
|  |    210 | subsection {* Code serialization *}
 | 
|  |    211 | 
 | 
|  |    212 | code_type index
 | 
|  |    213 |   (SML "int")
 | 
|  |    214 |   (OCaml "int")
 | 
|  |    215 |   (Haskell "Integer")
 | 
|  |    216 | 
 | 
|  |    217 | code_instance index :: eq
 | 
|  |    218 |   (Haskell -)
 | 
|  |    219 | 
 | 
|  |    220 | setup {*
 | 
|  |    221 |   fold (fn target => CodeTarget.add_pretty_numeral target true
 | 
|  |    222 |     @{const_name number_index_inst.number_of_index}
 | 
|  |    223 |     @{const_name Numeral.B0} @{const_name Numeral.B1}
 | 
|  |    224 |     @{const_name Numeral.Pls} @{const_name Numeral.Min}
 | 
|  |    225 |     @{const_name Numeral.Bit}
 | 
|  |    226 |   ) ["SML", "OCaml", "Haskell"]
 | 
|  |    227 | *}
 | 
|  |    228 | 
 | 
|  |    229 | code_reserved SML int
 | 
|  |    230 | code_reserved OCaml int
 | 
|  |    231 | 
 | 
|  |    232 | code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"
 | 
|  |    233 |   (SML "Int.+ ((_), (_))")
 | 
|  |    234 |   (OCaml "Pervasives.+")
 | 
|  |    235 |   (Haskell infixl 6 "+")
 | 
|  |    236 | 
 | 
|  |    237 | code_const "uminus \<Colon> index \<Rightarrow> index"
 | 
|  |    238 |   (SML "Int.~")
 | 
|  |    239 |   (OCaml "Pervasives.~-")
 | 
|  |    240 |   (Haskell "negate")
 | 
|  |    241 | 
 | 
|  |    242 | code_const "op - \<Colon> index \<Rightarrow> index \<Rightarrow> index"
 | 
|  |    243 |   (SML "Int.- ((_), (_))")
 | 
|  |    244 |   (OCaml "Pervasives.-")
 | 
|  |    245 |   (Haskell infixl 6 "-")
 | 
|  |    246 | 
 | 
|  |    247 | code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"
 | 
|  |    248 |   (SML "Int.* ((_), (_))")
 | 
|  |    249 |   (OCaml "Pervasives.*")
 | 
|  |    250 |   (Haskell infixl 7 "*")
 | 
|  |    251 | 
 | 
|  |    252 | code_const "op = \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
 | 
|  |    253 |   (SML "!((_ : Int.int) = _)")
 | 
|  |    254 |   (OCaml "!((_ : Pervasives.int) = _)")
 | 
|  |    255 |   (Haskell infixl 4 "==")
 | 
|  |    256 | 
 | 
|  |    257 | code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
 | 
|  |    258 |   (SML "Int.<= ((_), (_))")
 | 
|  |    259 |   (OCaml "!((_ : Pervasives.int) <= _)")
 | 
|  |    260 |   (Haskell infix 4 "<=")
 | 
|  |    261 | 
 | 
|  |    262 | code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
 | 
|  |    263 |   (SML "Int.< ((_), (_))")
 | 
|  |    264 |   (OCaml "!((_ : Pervasives.int) < _)")
 | 
|  |    265 |   (Haskell infix 4 "<")
 | 
|  |    266 | 
 | 
|  |    267 | code_reserved SML Int
 | 
|  |    268 | code_reserved OCaml Pervasives
 | 
|  |    269 | 
 | 
|  |    270 | end
 |