src/HOL/Library/Code_Index.thy
changeset 24999 1dbe785ed529
child 25335 182a001a7ea4
     1.1 --- /dev/null	Thu Jan 01 00:00:00 1970 +0000
     1.2 +++ b/src/HOL/Library/Code_Index.thy	Fri Oct 12 10:26:18 2007 +0200
     1.3 @@ -0,0 +1,255 @@
     1.4 +(*  ID:         $Id$
     1.5 +    Author:     Florian Haftmann, TU Muenchen
     1.6 +*)
     1.7 +
     1.8 +header {* Type of indices *}
     1.9 +
    1.10 +theory Code_Index
    1.11 +imports PreList
    1.12 +begin
    1.13 +
    1.14 +text {*
    1.15 +  Indices are isomorphic to HOL @{typ int} but
    1.16 +  mapped to target-language builtin integers
    1.17 +*}
    1.18 +
    1.19 +subsection {* Datatype of indices *}
    1.20 +
    1.21 +datatype index = index_of_int int
    1.22 +
    1.23 +lemmas [code func del] = index.recs index.cases
    1.24 +
    1.25 +fun
    1.26 +  int_of_index :: "index \<Rightarrow> int"
    1.27 +where
    1.28 +  "int_of_index (index_of_int k) = k"
    1.29 +lemmas [code func del] = int_of_index.simps
    1.30 +
    1.31 +lemma index_id [simp]:
    1.32 +  "index_of_int (int_of_index k) = k"
    1.33 +  by (cases k) simp_all
    1.34 +
    1.35 +lemma index:
    1.36 +  "(\<And>k\<Colon>index. PROP P k) \<equiv> (\<And>k\<Colon>int. PROP P (index_of_int k))"
    1.37 +proof
    1.38 +  fix k :: int
    1.39 +  assume "\<And>k\<Colon>index. PROP P k"
    1.40 +  then show "PROP P (index_of_int k)" .
    1.41 +next
    1.42 +  fix k :: index
    1.43 +  assume "\<And>k\<Colon>int. PROP P (index_of_int k)"
    1.44 +  then have "PROP P (index_of_int (int_of_index k))" .
    1.45 +  then show "PROP P k" by simp
    1.46 +qed
    1.47 +
    1.48 +lemma [code func]: "size (k\<Colon>index) = 0"
    1.49 +  by (cases k) simp_all
    1.50 +
    1.51 +
    1.52 +subsection {* Built-in integers as datatype on numerals *}
    1.53 +
    1.54 +instance index :: number
    1.55 +  "number_of \<equiv> index_of_int" ..
    1.56 +
    1.57 +code_datatype "number_of \<Colon> int \<Rightarrow> index"
    1.58 +
    1.59 +lemma number_of_index_id [simp]:
    1.60 +  "number_of (int_of_index k) = k"
    1.61 +  unfolding number_of_index_def by simp
    1.62 +
    1.63 +lemma number_of_index_shift:
    1.64 +  "number_of k = index_of_int (number_of k)"
    1.65 +  by (simp add: number_of_is_id number_of_index_def)
    1.66 +
    1.67 +
    1.68 +subsection {* Basic arithmetic *}
    1.69 +
    1.70 +instance index :: zero
    1.71 +  [simp]: "0 \<equiv> index_of_int 0" ..
    1.72 +lemmas [code func del] = zero_index_def
    1.73 +
    1.74 +instance index :: one
    1.75 +  [simp]: "1 \<equiv> index_of_int 1" ..
    1.76 +lemmas [code func del] = one_index_def
    1.77 +
    1.78 +instance index :: plus
    1.79 +  [simp]: "k + l \<equiv> index_of_int (int_of_index k + int_of_index l)" ..
    1.80 +lemmas [code func del] = plus_index_def
    1.81 +lemma plus_index_code [code func]:
    1.82 +  "index_of_int k + index_of_int l = index_of_int (k + l)"
    1.83 +  unfolding plus_index_def by simp
    1.84 +
    1.85 +instance index :: minus
    1.86 +  [simp]: "- k \<equiv> index_of_int (- int_of_index k)"
    1.87 +  [simp]: "k - l \<equiv> index_of_int (int_of_index k - int_of_index l)" ..
    1.88 +lemmas [code func del] = uminus_index_def minus_index_def
    1.89 +lemma uminus_index_code [code func]:
    1.90 +  "- index_of_int k \<equiv> index_of_int (- k)"
    1.91 +  unfolding uminus_index_def by simp
    1.92 +lemma minus_index_code [code func]:
    1.93 +  "index_of_int k - index_of_int l = index_of_int (k - l)"
    1.94 +  unfolding minus_index_def by simp
    1.95 +
    1.96 +instance index :: times
    1.97 +  [simp]: "k * l \<equiv> index_of_int (int_of_index k * int_of_index l)" ..
    1.98 +lemmas [code func del] = times_index_def
    1.99 +lemma times_index_code [code func]:
   1.100 +  "index_of_int k * index_of_int l = index_of_int (k * l)"
   1.101 +  unfolding times_index_def by simp
   1.102 +
   1.103 +instance index :: ord
   1.104 +  [simp]: "k \<le> l \<equiv> int_of_index k \<le> int_of_index l"
   1.105 +  [simp]: "k < l \<equiv> int_of_index k < int_of_index l" ..
   1.106 +lemmas [code func del] = less_eq_index_def less_index_def
   1.107 +lemma less_eq_index_code [code func]:
   1.108 +  "index_of_int k \<le> index_of_int l \<longleftrightarrow> k \<le> l"
   1.109 +  unfolding less_eq_index_def by simp
   1.110 +lemma less_index_code [code func]:
   1.111 +  "index_of_int k < index_of_int l \<longleftrightarrow> k < l"
   1.112 +  unfolding less_index_def by simp
   1.113 +
   1.114 +instance index :: ring_1
   1.115 +  by default (auto simp add: left_distrib right_distrib)
   1.116 +
   1.117 +lemma of_nat_index: "of_nat n = index_of_int (of_nat n)"
   1.118 +proof (induct n)
   1.119 +  case 0 show ?case by simp
   1.120 +next
   1.121 +  case (Suc n)
   1.122 +  then have "int_of_index (index_of_int (int n))
   1.123 +    = int_of_index (of_nat n)" by simp
   1.124 +  then have "int n = int_of_index (of_nat n)" by simp
   1.125 +  then show ?case by simp
   1.126 +qed
   1.127 +
   1.128 +instance index :: number_ring
   1.129 +  by default
   1.130 +    (simp_all add: left_distrib number_of_index_def of_int_of_nat of_nat_index)
   1.131 +
   1.132 +lemma zero_index_code [code inline, code func]:
   1.133 +  "(0\<Colon>index) = Numeral0"
   1.134 +  by simp
   1.135 +
   1.136 +lemma one_index_code [code inline, code func]:
   1.137 +  "(1\<Colon>index) = Numeral1"
   1.138 +  by simp
   1.139 +
   1.140 +instance index :: abs
   1.141 +  "\<bar>k\<bar> \<equiv> if k < 0 then -k else k" ..
   1.142 +
   1.143 +lemma index_of_int [code func]:
   1.144 +  "index_of_int k = (if k = 0 then 0
   1.145 +    else if k = -1 then -1
   1.146 +    else let (l, m) = divAlg (k, 2) in 2 * index_of_int l +
   1.147 +      (if m = 0 then 0 else 1))"
   1.148 +  by (simp add: number_of_index_shift Let_def split_def divAlg_mod_div) arith
   1.149 +
   1.150 +
   1.151 +subsection {* Conversion to and from @{typ nat} *}
   1.152 +
   1.153 +definition
   1.154 +  nat_of_index :: "index \<Rightarrow> nat"
   1.155 +where
   1.156 +  [code func del]: "nat_of_index = nat o int_of_index"
   1.157 +
   1.158 +definition
   1.159 +  nat_of_index_aux :: "index \<Rightarrow> nat \<Rightarrow> nat" where
   1.160 +  [code func del]: "nat_of_index_aux i n = nat_of_index i + n"
   1.161 +
   1.162 +lemma nat_of_index_aux_code [code]:
   1.163 +  "nat_of_index_aux i n = (if i \<le> 0 then n else nat_of_index_aux (i - 1) (Suc n))"
   1.164 +  by (auto simp add: nat_of_index_aux_def nat_of_index_def)
   1.165 +
   1.166 +lemma nat_of_index_code [code]:
   1.167 +  "nat_of_index i = nat_of_index_aux i 0"
   1.168 +  by (simp add: nat_of_index_aux_def)
   1.169 +
   1.170 +definition
   1.171 +  index_of_nat :: "nat \<Rightarrow> index"
   1.172 +where
   1.173 +  [code func del]: "index_of_nat = index_of_int o of_nat"
   1.174 +
   1.175 +lemma index_of_nat [code func]:
   1.176 +  "index_of_nat 0 = 0"
   1.177 +  "index_of_nat (Suc n) = index_of_nat n + 1"
   1.178 +  unfolding index_of_nat_def by simp_all
   1.179 +
   1.180 +lemma index_nat_id [simp]:
   1.181 +  "nat_of_index (index_of_nat n) = n"
   1.182 +  "index_of_nat (nat_of_index i) = (if i \<le> 0 then 0 else i)"
   1.183 +  unfolding index_of_nat_def nat_of_index_def by simp_all
   1.184 +
   1.185 +
   1.186 +subsection {* ML interface *}
   1.187 +
   1.188 +ML {*
   1.189 +structure Index =
   1.190 +struct
   1.191 +
   1.192 +fun mk k = @{term index_of_int} $ HOLogic.mk_number @{typ index} k;
   1.193 +
   1.194 +end;
   1.195 +*}
   1.196 +
   1.197 +
   1.198 +subsection {* Code serialization *}
   1.199 +
   1.200 +code_type index
   1.201 +  (SML "int")
   1.202 +  (OCaml "int")
   1.203 +  (Haskell "Integer")
   1.204 +
   1.205 +code_instance index :: eq
   1.206 +  (Haskell -)
   1.207 +
   1.208 +setup {*
   1.209 +  fold (fn target => CodeTarget.add_pretty_numeral target true
   1.210 +    @{const_name number_index_inst.number_of_index}
   1.211 +    @{const_name Numeral.B0} @{const_name Numeral.B1}
   1.212 +    @{const_name Numeral.Pls} @{const_name Numeral.Min}
   1.213 +    @{const_name Numeral.Bit}
   1.214 +  ) ["SML", "OCaml", "Haskell"]
   1.215 +*}
   1.216 +
   1.217 +code_reserved SML int
   1.218 +code_reserved OCaml int
   1.219 +
   1.220 +code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"
   1.221 +  (SML "Int.+ ((_), (_))")
   1.222 +  (OCaml "Pervasives.+")
   1.223 +  (Haskell infixl 6 "+")
   1.224 +
   1.225 +code_const "uminus \<Colon> index \<Rightarrow> index"
   1.226 +  (SML "Int.~")
   1.227 +  (OCaml "Pervasives.~-")
   1.228 +  (Haskell "negate")
   1.229 +
   1.230 +code_const "op - \<Colon> index \<Rightarrow> index \<Rightarrow> index"
   1.231 +  (SML "Int.- ((_), (_))")
   1.232 +  (OCaml "Pervasives.-")
   1.233 +  (Haskell infixl 6 "-")
   1.234 +
   1.235 +code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"
   1.236 +  (SML "Int.* ((_), (_))")
   1.237 +  (OCaml "Pervasives.*")
   1.238 +  (Haskell infixl 7 "*")
   1.239 +
   1.240 +code_const "op = \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
   1.241 +  (SML "!((_ : Int.int) = _)")
   1.242 +  (OCaml "!((_ : Pervasives.int) = _)")
   1.243 +  (Haskell infixl 4 "==")
   1.244 +
   1.245 +code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
   1.246 +  (SML "Int.<= ((_), (_))")
   1.247 +  (OCaml "!((_ : Pervasives.int) <= _)")
   1.248 +  (Haskell infix 4 "<=")
   1.249 +
   1.250 +code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
   1.251 +  (SML "Int.< ((_), (_))")
   1.252 +  (OCaml "!((_ : Pervasives.int) < _)")
   1.253 +  (Haskell infix 4 "<")
   1.254 +
   1.255 +code_reserved SML Int
   1.256 +code_reserved OCaml Pervasives
   1.257 +
   1.258 +end