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