src/HOL/Code_Numeral.thy
author huffman
Mon, 20 Feb 2012 12:37:17 +0100
changeset 47418 d1dcb91a512e
parent 46899 9f113cdf3d66
child 47506 cde737f9c911
permissions -rw-r--r--
use qualified constant names instead of suffixes (from Florian Haftmann)
haftmann@29752
     1
(* Author: Florian Haftmann, TU Muenchen *)
haftmann@24999
     2
haftmann@31205
     3
header {* Type of target language numerals *}
haftmann@24999
     4
haftmann@31205
     5
theory Code_Numeral
haftmann@33301
     6
imports Nat_Numeral Nat_Transfer Divides
haftmann@24999
     7
begin
haftmann@24999
     8
haftmann@24999
     9
text {*
haftmann@31205
    10
  Code numerals are isomorphic to HOL @{typ nat} but
haftmann@31205
    11
  mapped to target-language builtin numerals.
haftmann@24999
    12
*}
haftmann@24999
    13
haftmann@31205
    14
subsection {* Datatype of target language numerals *}
haftmann@24999
    15
haftmann@31205
    16
typedef (open) code_numeral = "UNIV \<Colon> nat set"
wenzelm@46567
    17
  morphisms nat_of of_nat ..
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@31205
    30
lemma code_numeral:
haftmann@31205
    31
  "(\<And>n\<Colon>code_numeral. PROP P n) \<equiv> (\<And>n\<Colon>nat. PROP P (of_nat n))"
haftmann@24999
    32
proof
haftmann@25767
    33
  fix n :: nat
haftmann@31205
    34
  assume "\<And>n\<Colon>code_numeral. PROP P n"
haftmann@29752
    35
  then show "PROP P (of_nat n)" .
haftmann@24999
    36
next
haftmann@31205
    37
  fix n :: code_numeral
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@31205
    43
lemma code_numeral_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
haftmann@31205
    48
lemma code_numeral_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@31205
    64
instantiation code_numeral :: 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
huffman@47418
    74
definition Suc where [simp]:
huffman@47418
    75
  "Suc k = of_nat (Nat.Suc (nat_of k))"
haftmann@26140
    76
huffman@47418
    77
rep_datatype "0 \<Colon> code_numeral" Suc
haftmann@26140
    78
proof -
haftmann@31205
    79
  fix P :: "code_numeral \<Rightarrow> bool"
haftmann@31205
    80
  fix k :: code_numeral
haftmann@29752
    81
  assume "P 0" then have init: "P (of_nat 0)" by simp
huffman@47418
    82
  assume "\<And>k. P k \<Longrightarrow> P (Suc k)"
huffman@47418
    83
    then have "\<And>n. P (of_nat n) \<Longrightarrow> P (Suc (of_nat n))" .
huffman@47418
    84
    then have step: "\<And>n. P (of_nat n) \<Longrightarrow> P (of_nat (Nat.Suc n))" by simp
haftmann@29752
    85
  from init step have "P (of_nat (nat_of k))"
berghofe@34915
    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@31205
    90
declare code_numeral_case [case_names nat, cases type: code_numeral]
haftmann@31205
    91
declare code_numeral.induct [case_names nat, induct type: code_numeral]
haftmann@26140
    92
haftmann@31205
    93
lemma code_numeral_decr [termination_simp]:
huffman@47418
    94
  "k \<noteq> of_nat 0 \<Longrightarrow> nat_of k - Nat.Suc 0 < nat_of k"
haftmann@30245
    95
  by (cases k) simp
haftmann@30245
    96
haftmann@30245
    97
lemma [simp, code]:
haftmann@31205
    98
  "code_numeral_size = nat_of"
haftmann@26140
    99
proof (rule ext)
haftmann@26140
   100
  fix k
haftmann@31205
   101
  have "code_numeral_size k = nat_size (nat_of k)"
huffman@47418
   102
    by (induct k rule: code_numeral.induct) (simp_all del: zero_code_numeral_def Suc_def, simp_all)
berghofe@34915
   103
  also have "nat_size (nat_of k) = nat_of k" by (induct ("nat_of k")) simp_all
haftmann@31205
   104
  finally show "code_numeral_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"
huffman@47418
   112
  by (induct k) (simp_all del: zero_code_numeral_def Suc_def, simp_all)
haftmann@26140
   113
qed
haftmann@26140
   114
haftmann@31205
   115
lemmas [code del] = code_numeral.recs code_numeral.cases
haftmann@30245
   116
haftmann@28562
   117
lemma [code]:
haftmann@39086
   118
  "HOL.equal k l \<longleftrightarrow> HOL.equal (nat_of k) (nat_of l)"
haftmann@39086
   119
  by (cases k, cases l) (simp add: equal)
haftmann@24999
   120
haftmann@28351
   121
lemma [code nbe]:
haftmann@39086
   122
  "HOL.equal (k::code_numeral) k \<longleftrightarrow> True"
haftmann@39086
   123
  by (rule equal_refl)
haftmann@28351
   124
haftmann@24999
   125
haftmann@38195
   126
subsection {* Code numerals as datatype of ints *}
haftmann@24999
   127
haftmann@31205
   128
instantiation code_numeral :: 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@31205
   140
  by (simp add: number_of_code_numeral_def nat_number_of_def number_of_is_id)
haftmann@26264
   141
haftmann@31205
   142
code_datatype "number_of \<Colon> int \<Rightarrow> code_numeral"
haftmann@24999
   143
haftmann@24999
   144
haftmann@24999
   145
subsection {* Basic arithmetic *}
haftmann@24999
   146
haftmann@35028
   147
instantiation code_numeral :: "{minus, linordered_semidom, semiring_div, linorder}"
haftmann@25767
   148
begin
haftmann@24999
   149
haftmann@28708
   150
definition [simp, code del]:
haftmann@31205
   151
  "(1\<Colon>code_numeral) = 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@33335
   175
qed (auto simp add: code_numeral left_distrib intro: mult_commute)
haftmann@28708
   176
haftmann@28708
   177
end
haftmann@28708
   178
haftmann@46899
   179
lemma zero_code_numeral_code [code]:
haftmann@31205
   180
  "(0\<Colon>code_numeral) = Numeral0"
haftmann@31205
   181
  by (simp add: number_of_code_numeral_def Pls_def)
haftmann@46899
   182
haftmann@46899
   183
lemma [code_abbrev]: "Numeral0 = (0\<Colon>code_numeral)"
haftmann@31205
   184
  using zero_code_numeral_code ..
haftmann@25767
   185
haftmann@46899
   186
lemma one_code_numeral_code [code]:
haftmann@31205
   187
  "(1\<Colon>code_numeral) = Numeral1"
haftmann@31205
   188
  by (simp add: number_of_code_numeral_def Pls_def Bit1_def)
haftmann@46899
   189
haftmann@46899
   190
lemma [code_abbrev]: "Numeral1 = (1\<Colon>code_numeral)"
haftmann@31205
   191
  using one_code_numeral_code ..
haftmann@25767
   192
haftmann@31205
   193
lemma plus_code_numeral_code [code nbe]:
haftmann@29752
   194
  "of_nat n + of_nat m = of_nat (n + m)"
haftmann@24999
   195
  by simp
haftmann@24999
   196
huffman@47418
   197
definition subtract :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral" where
huffman@47418
   198
  [simp]: "subtract = minus"
haftmann@24999
   199
huffman@47418
   200
lemma subtract_code [code nbe]:
huffman@47418
   201
  "subtract (of_nat n) (of_nat m) = of_nat (n - m)"
haftmann@28708
   202
  by simp
haftmann@24999
   203
haftmann@31205
   204
lemma minus_code_numeral_code [code]:
huffman@47418
   205
  "minus = subtract"
haftmann@28708
   206
  by simp
haftmann@28708
   207
haftmann@31205
   208
lemma times_code_numeral_code [code nbe]:
haftmann@29752
   209
  "of_nat n * of_nat m = of_nat (n * m)"
haftmann@25767
   210
  by simp
haftmann@25335
   211
haftmann@31205
   212
lemma less_eq_code_numeral_code [code nbe]:
haftmann@29752
   213
  "of_nat n \<le> of_nat m \<longleftrightarrow> n \<le> m"
haftmann@25767
   214
  by simp
haftmann@24999
   215
haftmann@31205
   216
lemma less_code_numeral_code [code nbe]:
haftmann@29752
   217
  "of_nat n < of_nat m \<longleftrightarrow> n < m"
haftmann@25767
   218
  by simp
haftmann@24999
   219
haftmann@31259
   220
lemma code_numeral_zero_minus_one:
haftmann@31259
   221
  "(0::code_numeral) - 1 = 0"
haftmann@31259
   222
  by simp
haftmann@31259
   223
haftmann@31259
   224
lemma Suc_code_numeral_minus_one:
huffman@47418
   225
  "Suc n - 1 = n"
haftmann@31259
   226
  by simp
haftmann@26140
   227
haftmann@29752
   228
lemma of_nat_code [code]:
haftmann@29752
   229
  "of_nat = Nat.of_nat"
haftmann@25918
   230
proof
haftmann@25918
   231
  fix n :: nat
haftmann@29752
   232
  have "Nat.of_nat n = of_nat n"
haftmann@25918
   233
    by (induct n) simp_all
haftmann@29752
   234
  then show "of_nat n = Nat.of_nat n"
haftmann@25918
   235
    by (rule sym)
haftmann@25918
   236
qed
haftmann@25918
   237
haftmann@31205
   238
lemma code_numeral_not_eq_zero: "i \<noteq> of_nat 0 \<longleftrightarrow> i \<ge> 1"
haftmann@25928
   239
  by (cases i) auto
haftmann@25928
   240
haftmann@31205
   241
definition nat_of_aux :: "code_numeral \<Rightarrow> nat \<Rightarrow> nat" where
haftmann@29752
   242
  "nat_of_aux i n = nat_of i + n"
haftmann@25928
   243
haftmann@29752
   244
lemma nat_of_aux_code [code]:
huffman@47418
   245
  "nat_of_aux i n = (if i = 0 then n else nat_of_aux (i - 1) (Nat.Suc n))"
haftmann@31205
   246
  by (auto simp add: nat_of_aux_def code_numeral_not_eq_zero)
haftmann@25928
   247
haftmann@29752
   248
lemma nat_of_code [code]:
haftmann@29752
   249
  "nat_of i = nat_of_aux i 0"
haftmann@29752
   250
  by (simp add: nat_of_aux_def)
haftmann@25918
   251
huffman@47418
   252
definition div_mod :: "code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<times> code_numeral" where
huffman@47418
   253
  [code del]: "div_mod n m = (n div m, n mod m)"
haftmann@26009
   254
haftmann@28562
   255
lemma [code]:
huffman@47418
   256
  "div_mod n m = (if m = 0 then (0, n) else (n div m, n mod m))"
huffman@47418
   257
  unfolding div_mod_def by auto
haftmann@26009
   258
haftmann@28562
   259
lemma [code]:
huffman@47418
   260
  "n div m = fst (div_mod n m)"
huffman@47418
   261
  unfolding div_mod_def by simp
haftmann@26009
   262
haftmann@28562
   263
lemma [code]:
huffman@47418
   264
  "n mod m = snd (div_mod n m)"
huffman@47418
   265
  unfolding div_mod_def by simp
haftmann@26009
   266
haftmann@31205
   267
definition int_of :: "code_numeral \<Rightarrow> int" where
haftmann@31192
   268
  "int_of = Nat.of_nat o nat_of"
haftmann@26009
   269
haftmann@31192
   270
lemma int_of_code [code]:
haftmann@31192
   271
  "int_of k = (if k = 0 then 0
haftmann@31192
   272
    else (if k mod 2 = 0 then 2 * int_of (k div 2) else 2 * int_of (k div 2) + 1))"
haftmann@33335
   273
proof -
haftmann@33335
   274
  have "(nat_of k div 2) * 2 + nat_of k mod 2 = nat_of k" 
haftmann@33335
   275
    by (rule mod_div_equality)
haftmann@33335
   276
  then have "int ((nat_of k div 2) * 2 + nat_of k mod 2) = int (nat_of k)" 
haftmann@33335
   277
    by simp
haftmann@33335
   278
  then have "int (nat_of k) = int (nat_of k div 2) * 2 + int (nat_of k mod 2)" 
huffman@45692
   279
    unfolding of_nat_mult of_nat_add by simp
haftmann@33335
   280
  then show ?thesis by (auto simp add: int_of_def mult_ac)
haftmann@33335
   281
qed
haftmann@28708
   282
haftmann@28708
   283
huffman@47418
   284
text {* Lazy Evaluation of an indexed function *}
bulwahn@36049
   285
huffman@47418
   286
function iterate_upto :: "(code_numeral \<Rightarrow> 'a) \<Rightarrow> code_numeral \<Rightarrow> code_numeral \<Rightarrow> 'a Predicate.pred"
bulwahn@36049
   287
where
huffman@47418
   288
  "iterate_upto f n m =
huffman@47418
   289
    Predicate.Seq (%u. if n > m then Predicate.Empty
huffman@47418
   290
     else Predicate.Insert (f n) (iterate_upto f (n + 1) m))"
bulwahn@36049
   291
by pat_completeness auto
bulwahn@36049
   292
bulwahn@36049
   293
termination by (relation "measure (%(f, n, m). Code_Numeral.nat_of (m + 1 - n))") auto
bulwahn@36049
   294
huffman@47418
   295
hide_const (open) of_nat nat_of Suc  subtract int_of iterate_upto
huffman@47418
   296
haftmann@28708
   297
haftmann@28228
   298
subsection {* Code generator setup *}
haftmann@24999
   299
haftmann@38195
   300
text {* Implementation of code numerals by bounded integers *}
haftmann@25767
   301
haftmann@31205
   302
code_type code_numeral
haftmann@24999
   303
  (SML "int")
haftmann@31377
   304
  (OCaml "Big'_int.big'_int")
haftmann@38185
   305
  (Haskell "Integer")
haftmann@38195
   306
  (Scala "BigInt")
haftmann@24999
   307
haftmann@39086
   308
code_instance code_numeral :: equal
haftmann@24999
   309
  (Haskell -)
haftmann@24999
   310
haftmann@24999
   311
setup {*
haftmann@38195
   312
  Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
haftmann@38195
   313
    false Code_Printer.literal_naive_numeral "SML"
haftmann@38185
   314
  #> fold (Numeral.add_code @{const_name number_code_numeral_inst.number_of_code_numeral}
haftmann@38195
   315
    false Code_Printer.literal_numeral) ["OCaml", "Haskell", "Scala"]
haftmann@24999
   316
*}
haftmann@24999
   317
haftmann@25918
   318
code_reserved SML Int int
haftmann@38195
   319
code_reserved Eval Integer
haftmann@24999
   320
huffman@47418
   321
code_const "plus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
haftmann@25928
   322
  (SML "Int.+/ ((_),/ (_))")
haftmann@31377
   323
  (OCaml "Big'_int.add'_big'_int")
haftmann@24999
   324
  (Haskell infixl 6 "+")
haftmann@34886
   325
  (Scala infixl 7 "+")
haftmann@38195
   326
  (Eval infixl 8 "+")
haftmann@24999
   327
huffman@47418
   328
code_const "Code_Numeral.subtract \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
haftmann@25918
   329
  (SML "Int.max/ (_/ -/ _,/ 0 : int)")
haftmann@31377
   330
  (OCaml "Big'_int.max'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)/ Big'_int.zero'_big'_int")
haftmann@38185
   331
  (Haskell "max/ (_/ -/ _)/ (0 :: Integer)")
haftmann@34886
   332
  (Scala "!(_/ -/ _).max(0)")
haftmann@38195
   333
  (Eval "Integer.max/ (_/ -/ _)/ 0")
haftmann@24999
   334
huffman@47418
   335
code_const "times \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
haftmann@25928
   336
  (SML "Int.*/ ((_),/ (_))")
haftmann@31377
   337
  (OCaml "Big'_int.mult'_big'_int")
haftmann@24999
   338
  (Haskell infixl 7 "*")
haftmann@34886
   339
  (Scala infixl 8 "*")
haftmann@38195
   340
  (Eval infixl 8 "*")
haftmann@24999
   341
huffman@47418
   342
code_const Code_Numeral.div_mod
haftmann@38195
   343
  (SML "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Int.div (n, m), Int.mod (n, m)))")
haftmann@34898
   344
  (OCaml "Big'_int.quomod'_big'_int/ (Big'_int.abs'_big'_int _)/ (Big'_int.abs'_big'_int _)")
haftmann@26009
   345
  (Haskell "divMod")
haftmann@38195
   346
  (Scala "!((k: BigInt) => (l: BigInt) =>/ if (l == 0)/ (BigInt(0), k) else/ (k.abs '/% l.abs))")
haftmann@40060
   347
  (Eval "!(fn n => fn m =>/ if m = 0/ then (0, n) else/ (Integer.div'_mod n m))")
haftmann@25928
   348
haftmann@39086
   349
code_const "HOL.equal \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
haftmann@24999
   350
  (SML "!((_ : Int.int) = _)")
haftmann@31377
   351
  (OCaml "Big'_int.eq'_big'_int")
haftmann@39499
   352
  (Haskell infix 4 "==")
haftmann@34886
   353
  (Scala infixl 5 "==")
haftmann@38195
   354
  (Eval "!((_ : int) = _)")
haftmann@24999
   355
huffman@47418
   356
code_const "less_eq \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
haftmann@25928
   357
  (SML "Int.<=/ ((_),/ (_))")
haftmann@31377
   358
  (OCaml "Big'_int.le'_big'_int")
haftmann@24999
   359
  (Haskell infix 4 "<=")
haftmann@34898
   360
  (Scala infixl 4 "<=")
haftmann@38195
   361
  (Eval infixl 6 "<=")
haftmann@24999
   362
huffman@47418
   363
code_const "less \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> bool"
haftmann@25928
   364
  (SML "Int.</ ((_),/ (_))")
haftmann@31377
   365
  (OCaml "Big'_int.lt'_big'_int")
haftmann@24999
   366
  (Haskell infix 4 "<")
haftmann@34898
   367
  (Scala infixl 4 "<")
haftmann@38195
   368
  (Eval infixl 6 "<")
haftmann@24999
   369
huffman@47418
   370
code_modulename SML
huffman@47418
   371
  Code_Numeral Arith
huffman@47418
   372
huffman@47418
   373
code_modulename OCaml
huffman@47418
   374
  Code_Numeral Arith
huffman@47418
   375
huffman@47418
   376
code_modulename Haskell
huffman@47418
   377
  Code_Numeral Arith
huffman@47418
   378
haftmann@24999
   379
end