src/HOL/Library/Code_Index.thy
author haftmann
Wed, 02 Jan 2008 15:14:23 +0100
changeset 25767 852bce03412a
parent 25691 8f8d83af100a
child 25918 82dd239e0f65
permissions -rw-r--r--
index now a copy of nat rather than int
haftmann@24999
     1
(*  ID:         $Id$
haftmann@24999
     2
    Author:     Florian Haftmann, TU Muenchen
haftmann@24999
     3
*)
haftmann@24999
     4
haftmann@24999
     5
header {* Type of indices *}
haftmann@24999
     6
haftmann@24999
     7
theory Code_Index
haftmann@25691
     8
imports ATP_Linkup
haftmann@24999
     9
begin
haftmann@24999
    10
haftmann@24999
    11
text {*
haftmann@25767
    12
  Indices are isomorphic to HOL @{typ nat} but
haftmann@24999
    13
  mapped to target-language builtin integers
haftmann@24999
    14
*}
haftmann@24999
    15
haftmann@24999
    16
subsection {* Datatype of indices *}
haftmann@24999
    17
haftmann@25767
    18
datatype index = index_of_nat nat
haftmann@24999
    19
haftmann@24999
    20
lemmas [code func del] = index.recs index.cases
haftmann@24999
    21
haftmann@25767
    22
primrec
haftmann@25767
    23
  nat_of_index :: "index \<Rightarrow> nat"
haftmann@24999
    24
where
haftmann@25767
    25
  "nat_of_index (index_of_nat k) = k"
haftmann@25767
    26
lemmas [code func del] = nat_of_index.simps
haftmann@24999
    27
haftmann@24999
    28
lemma index_id [simp]:
haftmann@25767
    29
  "index_of_nat (nat_of_index n) = n"
haftmann@25767
    30
  by (cases n) simp_all
haftmann@25767
    31
haftmann@25767
    32
lemma nat_of_index_inject [simp]:
haftmann@25767
    33
  "nat_of_index n = nat_of_index m \<longleftrightarrow> n = m"
haftmann@25767
    34
  by (cases n) auto
haftmann@24999
    35
haftmann@24999
    36
lemma index:
haftmann@25767
    37
  "(\<And>n\<Colon>index. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (index_of_nat n))"
haftmann@24999
    38
proof
haftmann@25767
    39
  fix n :: nat
haftmann@25767
    40
  assume "\<And>n\<Colon>index. PROP P n"
haftmann@25767
    41
  then show "PROP P (index_of_nat n)" .
haftmann@24999
    42
next
haftmann@25767
    43
  fix n :: index
haftmann@25767
    44
  assume "\<And>n\<Colon>nat. PROP P (index_of_nat n)"
haftmann@25767
    45
  then have "PROP P (index_of_nat (nat_of_index n))" .
haftmann@25767
    46
  then show "PROP P n" by simp
haftmann@24999
    47
qed
haftmann@24999
    48
haftmann@25767
    49
lemma [code func]: "size (n\<Colon>index) = 0"
haftmann@25767
    50
  by (cases n) simp_all
haftmann@24999
    51
haftmann@24999
    52
haftmann@25767
    53
subsection {* Indices as datatype of ints *}
haftmann@24999
    54
haftmann@25767
    55
instantiation index :: number
haftmann@25767
    56
begin
haftmann@25767
    57
haftmann@25767
    58
definition
haftmann@25767
    59
  "number_of = index_of_nat o nat"
haftmann@25767
    60
haftmann@25767
    61
instance ..
haftmann@25767
    62
haftmann@25767
    63
end
haftmann@24999
    64
haftmann@24999
    65
code_datatype "number_of \<Colon> int \<Rightarrow> index"
haftmann@24999
    66
haftmann@24999
    67
haftmann@24999
    68
subsection {* Basic arithmetic *}
haftmann@24999
    69
haftmann@25767
    70
instantiation index :: "{minus, ordered_semidom, Divides.div, linorder}"
haftmann@25767
    71
begin
haftmann@24999
    72
haftmann@25767
    73
definition [simp, code func del]:
haftmann@25767
    74
  "(0\<Colon>index) = index_of_nat 0"
haftmann@24999
    75
haftmann@24999
    76
lemma zero_index_code [code inline, code func]:
haftmann@24999
    77
  "(0\<Colon>index) = Numeral0"
haftmann@25767
    78
  by (simp add: number_of_index_def Pls_def)
haftmann@25767
    79
haftmann@25767
    80
definition [simp, code func del]:
haftmann@25767
    81
  "(1\<Colon>index) = index_of_nat 1"
haftmann@24999
    82
haftmann@24999
    83
lemma one_index_code [code inline, code func]:
haftmann@24999
    84
  "(1\<Colon>index) = Numeral1"
haftmann@25767
    85
  by (simp add: number_of_index_def Pls_def Bit_def)
haftmann@25767
    86
haftmann@25767
    87
definition [simp, code func del]:
haftmann@25767
    88
  "n + m = index_of_nat (nat_of_index n + nat_of_index m)"
haftmann@25767
    89
haftmann@25767
    90
lemma plus_index_code [code func]:
haftmann@25767
    91
  "index_of_nat n + index_of_nat m = index_of_nat (n + m)"
haftmann@24999
    92
  by simp
haftmann@24999
    93
haftmann@25767
    94
definition [simp, code func del]:
haftmann@25767
    95
  "n - m = index_of_nat (nat_of_index n - nat_of_index m)"
haftmann@24999
    96
haftmann@25767
    97
definition [simp, code func del]:
haftmann@25767
    98
  "n * m = index_of_nat (nat_of_index n * nat_of_index m)"
haftmann@24999
    99
haftmann@25767
   100
lemma times_index_code [code func]:
haftmann@25767
   101
  "index_of_nat n * index_of_nat m = index_of_nat (n * m)"
haftmann@25767
   102
  by simp
haftmann@25335
   103
haftmann@25767
   104
definition [simp, code func del]:
haftmann@25767
   105
  "n div m = index_of_nat (nat_of_index n div nat_of_index m)"
haftmann@24999
   106
haftmann@25767
   107
definition [simp, code func del]:
haftmann@25767
   108
  "n mod m = index_of_nat (nat_of_index n mod nat_of_index m)"
haftmann@24999
   109
haftmann@25767
   110
lemma div_index_code [code func]:
haftmann@25767
   111
  "index_of_nat n div index_of_nat m = index_of_nat (n div m)"
haftmann@25767
   112
  by simp
haftmann@24999
   113
haftmann@25767
   114
lemma mod_index_code [code func]:
haftmann@25767
   115
  "index_of_nat n mod index_of_nat m = index_of_nat (n mod m)"
haftmann@25767
   116
  by simp
haftmann@24999
   117
haftmann@25767
   118
definition [simp, code func del]:
haftmann@25767
   119
  "n \<le> m \<longleftrightarrow> nat_of_index n \<le> nat_of_index m"
haftmann@24999
   120
haftmann@25767
   121
definition [simp, code func del]:
haftmann@25767
   122
  "n < m \<longleftrightarrow> nat_of_index n < nat_of_index m"
haftmann@24999
   123
haftmann@25767
   124
lemma less_eq_index_code [code func]:
haftmann@25767
   125
  "index_of_nat n \<le> index_of_nat m \<longleftrightarrow> n \<le> m"
haftmann@25767
   126
  by simp
haftmann@24999
   127
haftmann@25767
   128
lemma less_index_code [code func]:
haftmann@25767
   129
  "index_of_nat n < index_of_nat m \<longleftrightarrow> n < m"
haftmann@25767
   130
  by simp
haftmann@24999
   131
haftmann@25767
   132
instance by default (auto simp add: left_distrib index)
haftmann@25767
   133
haftmann@25767
   134
end
haftmann@24999
   135
haftmann@24999
   136
haftmann@24999
   137
subsection {* ML interface *}
haftmann@24999
   138
haftmann@24999
   139
ML {*
haftmann@24999
   140
structure Index =
haftmann@24999
   141
struct
haftmann@24999
   142
haftmann@25767
   143
fun mk k = @{term index_of_nat} $ HOLogic.mk_number @{typ index} k;
haftmann@24999
   144
haftmann@24999
   145
end;
haftmann@24999
   146
*}
haftmann@24999
   147
haftmann@24999
   148
haftmann@24999
   149
subsection {* Code serialization *}
haftmann@24999
   150
haftmann@25767
   151
text {* Pecularity for operations with potentially negative result *}
haftmann@25767
   152
haftmann@25767
   153
definition
haftmann@25767
   154
  minus_index' :: "index \<Rightarrow> index \<Rightarrow> index"
haftmann@25767
   155
where
haftmann@25767
   156
  [code func del]: "minus_index' = op -"
haftmann@25767
   157
haftmann@25767
   158
lemma minus_index_code [code func]:
haftmann@25767
   159
  "n - m = (let q = minus_index' n m
haftmann@25767
   160
    in if q < 0 then 0 else q)"
haftmann@25767
   161
  by (simp add: minus_index'_def Let_def)
haftmann@25767
   162
haftmann@25767
   163
text {* Implementation of indices by bounded integers *}
haftmann@25767
   164
haftmann@24999
   165
code_type index
haftmann@24999
   166
  (SML "int")
haftmann@24999
   167
  (OCaml "int")
haftmann@24999
   168
  (Haskell "Integer")
haftmann@24999
   169
haftmann@24999
   170
code_instance index :: eq
haftmann@24999
   171
  (Haskell -)
haftmann@24999
   172
haftmann@24999
   173
setup {*
haftmann@24999
   174
  fold (fn target => CodeTarget.add_pretty_numeral target true
haftmann@24999
   175
    @{const_name number_index_inst.number_of_index}
haftmann@24999
   176
    @{const_name Numeral.B0} @{const_name Numeral.B1}
haftmann@24999
   177
    @{const_name Numeral.Pls} @{const_name Numeral.Min}
haftmann@24999
   178
    @{const_name Numeral.Bit}
haftmann@24999
   179
  ) ["SML", "OCaml", "Haskell"]
haftmann@24999
   180
*}
haftmann@24999
   181
haftmann@24999
   182
code_reserved SML int
haftmann@24999
   183
code_reserved OCaml int
haftmann@24999
   184
haftmann@24999
   185
code_const "op + \<Colon> index \<Rightarrow> index \<Rightarrow> index"
haftmann@24999
   186
  (SML "Int.+ ((_), (_))")
haftmann@24999
   187
  (OCaml "Pervasives.+")
haftmann@24999
   188
  (Haskell infixl 6 "+")
haftmann@24999
   189
haftmann@25767
   190
code_const "minus_index' \<Colon> index \<Rightarrow> index \<Rightarrow> index"
haftmann@24999
   191
  (SML "Int.- ((_), (_))")
haftmann@24999
   192
  (OCaml "Pervasives.-")
haftmann@24999
   193
  (Haskell infixl 6 "-")
haftmann@24999
   194
haftmann@24999
   195
code_const "op * \<Colon> index \<Rightarrow> index \<Rightarrow> index"
haftmann@24999
   196
  (SML "Int.* ((_), (_))")
haftmann@24999
   197
  (OCaml "Pervasives.*")
haftmann@24999
   198
  (Haskell infixl 7 "*")
haftmann@24999
   199
haftmann@24999
   200
code_const "op = \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
haftmann@24999
   201
  (SML "!((_ : Int.int) = _)")
haftmann@24999
   202
  (OCaml "!((_ : Pervasives.int) = _)")
haftmann@24999
   203
  (Haskell infixl 4 "==")
haftmann@24999
   204
haftmann@24999
   205
code_const "op \<le> \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
haftmann@24999
   206
  (SML "Int.<= ((_), (_))")
haftmann@24999
   207
  (OCaml "!((_ : Pervasives.int) <= _)")
haftmann@24999
   208
  (Haskell infix 4 "<=")
haftmann@24999
   209
haftmann@24999
   210
code_const "op < \<Colon> index \<Rightarrow> index \<Rightarrow> bool"
haftmann@24999
   211
  (SML "Int.< ((_), (_))")
haftmann@24999
   212
  (OCaml "!((_ : Pervasives.int) < _)")
haftmann@24999
   213
  (Haskell infix 4 "<")
haftmann@24999
   214
haftmann@25767
   215
code_const "op div \<Colon> index \<Rightarrow> index \<Rightarrow> index"
haftmann@25767
   216
  (SML "IntInf.div ((_), (_))")
haftmann@25767
   217
  (OCaml "Big'_int.div'_big'_int")
haftmann@25767
   218
  (Haskell "div")
haftmann@25767
   219
haftmann@25767
   220
code_const "op mod \<Colon> index \<Rightarrow> index \<Rightarrow> index"
haftmann@25767
   221
  (SML "IntInf.mod ((_), (_))")
haftmann@25767
   222
  (OCaml "Big'_int.mod'_big'_int")
haftmann@25767
   223
  (Haskell "mod")
haftmann@25767
   224
haftmann@24999
   225
code_reserved SML Int
haftmann@24999
   226
code_reserved OCaml Pervasives
haftmann@24999
   227
haftmann@24999
   228
end