src/HOL/Library/Efficient_Nat.thy
author haftmann
Thu, 26 Aug 2010 10:16:22 +0200
changeset 39004 f9cd27cbe8a4
parent 38206 3bf1fffcdd48
child 39007 567b94f8bb6e
permissions -rw-r--r--
code_include Scala: qualify module nmae
haftmann@23854
     1
(*  Title:      HOL/Library/Efficient_Nat.thy
haftmann@25931
     2
    Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
haftmann@23854
     3
*)
haftmann@23854
     4
haftmann@25931
     5
header {* Implementation of natural numbers by target-language integers *}
haftmann@23854
     6
haftmann@23854
     7
theory Efficient_Nat
haftmann@31203
     8
imports Code_Integer Main
haftmann@23854
     9
begin
haftmann@23854
    10
haftmann@23854
    11
text {*
haftmann@25931
    12
  When generating code for functions on natural numbers, the
haftmann@25931
    13
  canonical representation using @{term "0::nat"} and
haftmann@37363
    14
  @{term Suc} is unsuitable for computations involving large
haftmann@25931
    15
  numbers.  The efficiency of the generated code can be improved
haftmann@25931
    16
  drastically by implementing natural numbers by target-language
haftmann@25931
    17
  integers.  To do this, just include this theory.
haftmann@23854
    18
*}
haftmann@23854
    19
haftmann@25931
    20
subsection {* Basic arithmetic *}
haftmann@23854
    21
haftmann@23854
    22
text {*
haftmann@25931
    23
  Most standard arithmetic functions on natural numbers are implemented
haftmann@25931
    24
  using their counterparts on the integers:
haftmann@23854
    25
*}
haftmann@23854
    26
haftmann@25931
    27
code_datatype number_nat_inst.number_of_nat
haftmann@23854
    28
haftmann@32061
    29
lemma zero_nat_code [code, code_unfold_post]:
haftmann@25931
    30
  "0 = (Numeral0 :: nat)"
haftmann@25931
    31
  by simp
haftmann@23854
    32
haftmann@32061
    33
lemma one_nat_code [code, code_unfold_post]:
haftmann@25931
    34
  "1 = (Numeral1 :: nat)"
haftmann@25931
    35
  by simp
haftmann@23854
    36
haftmann@25931
    37
lemma Suc_code [code]:
haftmann@25931
    38
  "Suc n = n + 1"
haftmann@25931
    39
  by simp
haftmann@23854
    40
haftmann@25931
    41
lemma plus_nat_code [code]:
haftmann@25931
    42
  "n + m = nat (of_nat n + of_nat m)"
haftmann@25931
    43
  by simp
haftmann@23854
    44
haftmann@25931
    45
lemma minus_nat_code [code]:
haftmann@25931
    46
  "n - m = nat (of_nat n - of_nat m)"
haftmann@25931
    47
  by simp
haftmann@23854
    48
haftmann@25931
    49
lemma times_nat_code [code]:
haftmann@25931
    50
  "n * m = nat (of_nat n * of_nat m)"
haftmann@25931
    51
  unfolding of_nat_mult [symmetric] by simp
haftmann@23854
    52
haftmann@38195
    53
lemma divmod_nat_code [code]:
haftmann@38195
    54
  "divmod_nat n m = prod_fun nat nat (pdivmod (of_nat n) (of_nat m))"
haftmann@38195
    55
  by (simp add: prod_fun_def split_def pdivmod_def nat_div_distrib nat_mod_distrib divmod_nat_div_mod)
haftmann@23854
    56
haftmann@25931
    57
lemma eq_nat_code [code]:
haftmann@28351
    58
  "eq_class.eq n m \<longleftrightarrow> eq_class.eq (of_nat n \<Colon> int) (of_nat m)"
haftmann@28351
    59
  by (simp add: eq)
haftmann@28351
    60
haftmann@28351
    61
lemma eq_nat_refl [code nbe]:
haftmann@28351
    62
  "eq_class.eq (n::nat) n \<longleftrightarrow> True"
haftmann@28351
    63
  by (rule HOL.eq_refl)
haftmann@25931
    64
haftmann@25931
    65
lemma less_eq_nat_code [code]:
haftmann@25931
    66
  "n \<le> m \<longleftrightarrow> (of_nat n \<Colon> int) \<le> of_nat m"
haftmann@25931
    67
  by simp
haftmann@25931
    68
haftmann@25931
    69
lemma less_nat_code [code]:
haftmann@25931
    70
  "n < m \<longleftrightarrow> (of_nat n \<Colon> int) < of_nat m"
haftmann@25931
    71
  by simp
haftmann@25931
    72
haftmann@25931
    73
subsection {* Case analysis *}
haftmann@24423
    74
haftmann@23854
    75
text {*
haftmann@23854
    76
  Case analysis on natural numbers is rephrased using a conditional
haftmann@23854
    77
  expression:
haftmann@23854
    78
*}
haftmann@23854
    79
haftmann@31998
    80
lemma [code, code_unfold]:
haftmann@25615
    81
  "nat_case = (\<lambda>f g n. if n = 0 then f else g (n - 1))"
haftmann@25931
    82
  by (auto simp add: expand_fun_eq dest!: gr0_implies_Suc)
haftmann@25615
    83
haftmann@23854
    84
haftmann@23854
    85
subsection {* Preprocessors *}
haftmann@23854
    86
haftmann@23854
    87
text {*
haftmann@23854
    88
  In contrast to @{term "Suc n"}, the term @{term "n + (1::nat)"} is no longer
haftmann@23854
    89
  a constructor term. Therefore, all occurrences of this term in a position
haftmann@23854
    90
  where a pattern is expected (i.e.\ on the left-hand side of a recursion
haftmann@23854
    91
  equation or in the arguments of an inductive relation in an introduction
haftmann@23854
    92
  rule) must be eliminated.
haftmann@23854
    93
  This can be accomplished by applying the following transformation rules:
haftmann@23854
    94
*}
haftmann@23854
    95
haftmann@29874
    96
lemma Suc_if_eq: "(\<And>n. f (Suc n) \<equiv> h n) \<Longrightarrow> f 0 \<equiv> g \<Longrightarrow>
haftmann@29874
    97
  f n \<equiv> if n = 0 then g else h (n - 1)"
haftmann@31954
    98
  by (rule eq_reflection) (cases n, simp_all)
haftmann@29874
    99
haftmann@25931
   100
lemma Suc_clause: "(\<And>n. P n (Suc n)) \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> P (n - 1) n"
haftmann@29869
   101
  by (cases n) simp_all
haftmann@23854
   102
haftmann@23854
   103
text {*
haftmann@23854
   104
  The rules above are built into a preprocessor that is plugged into
haftmann@23854
   105
  the code generator. Since the preprocessor for introduction rules
haftmann@23854
   106
  does not know anything about modes, some of the modes that worked
haftmann@23854
   107
  for the canonical representation of natural numbers may no longer work.
haftmann@23854
   108
*}
haftmann@23854
   109
haftmann@23854
   110
(*<*)
haftmann@27609
   111
setup {*
haftmann@27609
   112
let
haftmann@23854
   113
haftmann@31954
   114
fun remove_suc thy thms =
haftmann@23854
   115
  let
haftmann@23854
   116
    val vname = Name.variant (map fst
haftmann@29874
   117
      (fold (Term.add_var_names o Thm.full_prop_of) thms [])) "n";
haftmann@23854
   118
    val cv = cterm_of thy (Var ((vname, 0), HOLogic.natT));
haftmann@23854
   119
    fun lhs_of th = snd (Thm.dest_comb
haftmann@31954
   120
      (fst (Thm.dest_comb (cprop_of th))));
haftmann@31954
   121
    fun rhs_of th = snd (Thm.dest_comb (cprop_of th));
haftmann@23854
   122
    fun find_vars ct = (case term_of ct of
haftmann@29869
   123
        (Const (@{const_name Suc}, _) $ Var _) => [(cv, snd (Thm.dest_comb ct))]
haftmann@23854
   124
      | _ $ _ =>
haftmann@23854
   125
        let val (ct1, ct2) = Thm.dest_comb ct
haftmann@23854
   126
        in 
haftmann@23854
   127
          map (apfst (fn ct => Thm.capply ct ct2)) (find_vars ct1) @
haftmann@23854
   128
          map (apfst (Thm.capply ct1)) (find_vars ct2)
haftmann@23854
   129
        end
haftmann@23854
   130
      | _ => []);
haftmann@23854
   131
    val eqs = maps
haftmann@23854
   132
      (fn th => map (pair th) (find_vars (lhs_of th))) thms;
haftmann@23854
   133
    fun mk_thms (th, (ct, cv')) =
haftmann@23854
   134
      let
haftmann@23854
   135
        val th' =
haftmann@23854
   136
          Thm.implies_elim
haftmann@23854
   137
           (Conv.fconv_rule (Thm.beta_conversion true)
haftmann@23854
   138
             (Drule.instantiate'
haftmann@23854
   139
               [SOME (ctyp_of_term ct)] [SOME (Thm.cabs cv ct),
haftmann@23854
   140
                 SOME (Thm.cabs cv' (rhs_of th)), NONE, SOME cv']
haftmann@31954
   141
               @{thm Suc_if_eq})) (Thm.forall_intr cv' th)
haftmann@23854
   142
      in
haftmann@23854
   143
        case map_filter (fn th'' =>
haftmann@23854
   144
            SOME (th'', singleton
wenzelm@36603
   145
              (Variable.trade (K (fn [th'''] => [th''' RS th']))
wenzelm@36603
   146
                (Variable.global_thm_context th'')) th'')
haftmann@23854
   147
          handle THM _ => NONE) thms of
haftmann@23854
   148
            [] => NONE
haftmann@23854
   149
          | thps =>
haftmann@23854
   150
              let val (ths1, ths2) = split_list thps
haftmann@23854
   151
              in SOME (subtract Thm.eq_thm (th :: ths1) thms @ ths2) end
haftmann@23854
   152
      end
haftmann@29874
   153
  in get_first mk_thms eqs end;
haftmann@29874
   154
haftmann@34893
   155
fun eqn_suc_base_preproc thy thms =
haftmann@29874
   156
  let
haftmann@31954
   157
    val dest = fst o Logic.dest_equals o prop_of;
haftmann@29874
   158
    val contains_suc = exists_Const (fn (c, _) => c = @{const_name Suc});
haftmann@29874
   159
  in
haftmann@29874
   160
    if forall (can dest) thms andalso exists (contains_suc o dest) thms
haftmann@32348
   161
      then thms |> perhaps_loop (remove_suc thy) |> (Option.map o map) Drule.zero_var_indexes
haftmann@29874
   162
       else NONE
haftmann@23854
   163
  end;
haftmann@23854
   164
haftmann@34893
   165
val eqn_suc_preproc = Code_Preproc.simple_functrans eqn_suc_base_preproc;
haftmann@23854
   166
haftmann@23854
   167
fun remove_suc_clause thy thms =
haftmann@23854
   168
  let
haftmann@23854
   169
    val vname = Name.variant (map fst
wenzelm@29258
   170
      (fold (Term.add_var_names o Thm.full_prop_of) thms [])) "x";
haftmann@24222
   171
    fun find_var (t as Const (@{const_name Suc}, _) $ (v as Var _)) = SOME (t, v)
haftmann@23854
   172
      | find_var (t $ u) = (case find_var t of NONE => find_var u | x => x)
haftmann@23854
   173
      | find_var _ = NONE;
haftmann@23854
   174
    fun find_thm th =
wenzelm@35625
   175
      let val th' = Conv.fconv_rule Object_Logic.atomize th
haftmann@23854
   176
      in Option.map (pair (th, th')) (find_var (prop_of th')) end
haftmann@23854
   177
  in
haftmann@23854
   178
    case get_first find_thm thms of
haftmann@23854
   179
      NONE => thms
haftmann@23854
   180
    | SOME ((th, th'), (Sucv, v)) =>
haftmann@23854
   181
        let
haftmann@23854
   182
          val cert = cterm_of (Thm.theory_of_thm th);
wenzelm@35625
   183
          val th'' = Object_Logic.rulify (Thm.implies_elim
haftmann@23854
   184
            (Conv.fconv_rule (Thm.beta_conversion true)
haftmann@23854
   185
              (Drule.instantiate' []
haftmann@23854
   186
                [SOME (cert (lambda v (Abs ("x", HOLogic.natT,
haftmann@23854
   187
                   abstract_over (Sucv,
haftmann@23854
   188
                     HOLogic.dest_Trueprop (prop_of th')))))),
haftmann@24222
   189
                 SOME (cert v)] @{thm Suc_clause}))
haftmann@23854
   190
            (Thm.forall_intr (cert v) th'))
haftmann@23854
   191
        in
haftmann@23854
   192
          remove_suc_clause thy (map (fn th''' =>
haftmann@23854
   193
            if (op = o pairself prop_of) (th''', th) then th'' else th''') thms)
haftmann@23854
   194
        end
haftmann@23854
   195
  end;
haftmann@23854
   196
haftmann@23854
   197
fun clause_suc_preproc thy ths =
haftmann@23854
   198
  let
haftmann@23854
   199
    val dest = fst o HOLogic.dest_mem o HOLogic.dest_Trueprop
haftmann@23854
   200
  in
haftmann@23854
   201
    if forall (can (dest o concl_of)) ths andalso
wenzelm@29287
   202
      exists (fn th => exists (exists_Const (fn (c, _) => c = @{const_name Suc}))
wenzelm@29287
   203
        (map_filter (try dest) (concl_of th :: prems_of th))) ths
haftmann@23854
   204
    then remove_suc_clause thy ths else ths
haftmann@23854
   205
  end;
haftmann@27609
   206
in
haftmann@27609
   207
haftmann@34893
   208
  Code_Preproc.add_functrans ("eqn_Suc", eqn_suc_preproc)
haftmann@23854
   209
  #> Codegen.add_preprocessor clause_suc_preproc
haftmann@27609
   210
haftmann@27609
   211
end;
haftmann@23854
   212
*}
haftmann@23854
   213
(*>*)
haftmann@23854
   214
haftmann@27609
   215
haftmann@25931
   216
subsection {* Target language setup *}
haftmann@25931
   217
haftmann@25931
   218
text {*
haftmann@25967
   219
  For ML, we map @{typ nat} to target language integers, where we
haftmann@34899
   220
  ensure that values are always non-negative.
haftmann@25931
   221
*}
haftmann@25931
   222
haftmann@25931
   223
code_type nat
haftmann@27496
   224
  (SML "IntInf.int")
haftmann@25931
   225
  (OCaml "Big'_int.big'_int")
haftmann@38195
   226
  (Eval "int")
haftmann@25931
   227
haftmann@25931
   228
types_code
haftmann@25931
   229
  nat ("int")
haftmann@25931
   230
attach (term_of) {*
haftmann@25931
   231
val term_of_nat = HOLogic.mk_number HOLogic.natT;
haftmann@25931
   232
*}
haftmann@25931
   233
attach (test) {*
haftmann@25931
   234
fun gen_nat i =
haftmann@25931
   235
  let val n = random_range 0 i
haftmann@25931
   236
  in (n, fn () => term_of_nat n) end;
haftmann@25931
   237
*}
haftmann@25931
   238
haftmann@25931
   239
text {*
haftmann@34899
   240
  For Haskell ans Scala we define our own @{typ nat} type.  The reason
haftmann@34899
   241
  is that we have to distinguish type class instances for @{typ nat}
haftmann@34899
   242
  and @{typ int}.
haftmann@25967
   243
*}
haftmann@25967
   244
haftmann@25967
   245
code_include Haskell "Nat" {*
haftmann@37034
   246
newtype Nat = Nat Integer deriving (Eq, Show, Read);
haftmann@25967
   247
haftmann@25967
   248
instance Num Nat where {
haftmann@25967
   249
  fromInteger k = Nat (if k >= 0 then k else 0);
haftmann@25967
   250
  Nat n + Nat m = Nat (n + m);
haftmann@25967
   251
  Nat n - Nat m = fromInteger (n - m);
haftmann@25967
   252
  Nat n * Nat m = Nat (n * m);
haftmann@25967
   253
  abs n = n;
haftmann@25967
   254
  signum _ = 1;
haftmann@25967
   255
  negate n = error "negate Nat";
haftmann@25967
   256
};
haftmann@25967
   257
haftmann@25967
   258
instance Ord Nat where {
haftmann@25967
   259
  Nat n <= Nat m = n <= m;
haftmann@25967
   260
  Nat n < Nat m = n < m;
haftmann@25967
   261
};
haftmann@25967
   262
haftmann@25967
   263
instance Real Nat where {
haftmann@25967
   264
  toRational (Nat n) = toRational n;
haftmann@25967
   265
};
haftmann@25967
   266
haftmann@25967
   267
instance Enum Nat where {
haftmann@25967
   268
  toEnum k = fromInteger (toEnum k);
haftmann@25967
   269
  fromEnum (Nat n) = fromEnum n;
haftmann@25967
   270
};
haftmann@25967
   271
haftmann@25967
   272
instance Integral Nat where {
haftmann@25967
   273
  toInteger (Nat n) = n;
haftmann@25967
   274
  divMod n m = quotRem n m;
haftmann@38195
   275
  quotRem (Nat n) (Nat m)
haftmann@38195
   276
    | (m == 0) = (0, Nat n)
haftmann@38195
   277
    | otherwise = (Nat k, Nat l) where (k, l) = quotRem n m;
haftmann@25967
   278
};
haftmann@25967
   279
*}
haftmann@25967
   280
haftmann@25967
   281
code_reserved Haskell Nat
haftmann@25967
   282
haftmann@34899
   283
code_include Scala "Nat" {*
haftmann@34931
   284
import scala.Math
haftmann@34931
   285
haftmann@34899
   286
object Nat {
haftmann@34899
   287
haftmann@34899
   288
  def apply(numeral: BigInt): Nat = new Nat(numeral max 0)
haftmann@34899
   289
  def apply(numeral: Int): Nat = Nat(BigInt(numeral))
haftmann@34899
   290
  def apply(numeral: String): Nat = Nat(BigInt(numeral))
haftmann@34899
   291
haftmann@34899
   292
}
haftmann@34899
   293
haftmann@34899
   294
class Nat private(private val value: BigInt) {
haftmann@34899
   295
haftmann@34899
   296
  override def hashCode(): Int = this.value.hashCode()
haftmann@34899
   297
haftmann@34899
   298
  override def equals(that: Any): Boolean = that match {
haftmann@34899
   299
    case that: Nat => this equals that
haftmann@34899
   300
    case _ => false
haftmann@34899
   301
  }
haftmann@34899
   302
haftmann@34899
   303
  override def toString(): String = this.value.toString
haftmann@34899
   304
haftmann@34899
   305
  def equals(that: Nat): Boolean = this.value == that.value
haftmann@34899
   306
haftmann@34899
   307
  def as_BigInt: BigInt = this.value
haftmann@37865
   308
  def as_Int: Int = if (this.value >= Int.MinValue && this.value <= Int.MaxValue)
haftmann@34931
   309
      this.value.intValue
haftmann@38206
   310
    else error("Int value out of range: " + this.value.toString)
haftmann@34899
   311
haftmann@34899
   312
  def +(that: Nat): Nat = new Nat(this.value + that.value)
haftmann@37222
   313
  def -(that: Nat): Nat = Nat(this.value - that.value)
haftmann@34899
   314
  def *(that: Nat): Nat = new Nat(this.value * that.value)
haftmann@34899
   315
haftmann@34899
   316
  def /%(that: Nat): (Nat, Nat) = if (that.value == 0) (new Nat(0), this)
haftmann@34899
   317
    else {
haftmann@34899
   318
      val (k, l) = this.value /% that.value
haftmann@34899
   319
      (new Nat(k), new Nat(l))
haftmann@34899
   320
    }
haftmann@34899
   321
haftmann@34899
   322
  def <=(that: Nat): Boolean = this.value <= that.value
haftmann@34899
   323
haftmann@34899
   324
  def <(that: Nat): Boolean = this.value < that.value
haftmann@34899
   325
haftmann@34899
   326
}
haftmann@34899
   327
*}
haftmann@34899
   328
haftmann@34899
   329
code_reserved Scala Nat
haftmann@34899
   330
haftmann@25967
   331
code_type nat
haftmann@29730
   332
  (Haskell "Nat.Nat")
haftmann@39004
   333
  (Scala "Nat.Nat")
haftmann@25967
   334
haftmann@25967
   335
code_instance nat :: eq
haftmann@25967
   336
  (Haskell -)
haftmann@25967
   337
haftmann@25967
   338
text {*
haftmann@25931
   339
  Natural numerals.
haftmann@25931
   340
*}
haftmann@25931
   341
haftmann@32061
   342
lemma [code_unfold_post]:
haftmann@25931
   343
  "nat (number_of i) = number_nat_inst.number_of_nat i"
haftmann@25931
   344
  -- {* this interacts as desired with @{thm nat_number_of_def} *}
haftmann@25931
   345
  by (simp add: number_nat_inst.number_of_nat)
haftmann@25931
   346
haftmann@25931
   347
setup {*
haftmann@25931
   348
  fold (Numeral.add_code @{const_name number_nat_inst.number_of_nat}
haftmann@38195
   349
    false Code_Printer.literal_positive_numeral) ["SML", "OCaml", "Haskell", "Scala"]
haftmann@25931
   350
*}
haftmann@25931
   351
haftmann@25931
   352
text {*
haftmann@25931
   353
  Since natural numbers are implemented
haftmann@25967
   354
  using integers in ML, the coercion function @{const "of_nat"} of type
haftmann@25931
   355
  @{typ "nat \<Rightarrow> int"} is simply implemented by the identity function.
haftmann@37363
   356
  For the @{const nat} function for converting an integer to a natural
haftmann@25931
   357
  number, we give a specific implementation using an ML function that
haftmann@25931
   358
  returns its input value, provided that it is non-negative, and otherwise
haftmann@25931
   359
  returns @{text "0"}.
haftmann@25931
   360
*}
haftmann@25931
   361
haftmann@32065
   362
definition int :: "nat \<Rightarrow> int" where
haftmann@28562
   363
  [code del]: "int = of_nat"
haftmann@25931
   364
haftmann@28562
   365
lemma int_code' [code]:
haftmann@25931
   366
  "int (number_of l) = (if neg (number_of l \<Colon> int) then 0 else number_of l)"
haftmann@25931
   367
  unfolding int_nat_number_of [folded int_def] ..
haftmann@25931
   368
haftmann@28562
   369
lemma nat_code' [code]:
haftmann@25931
   370
  "nat (number_of l) = (if neg (number_of l \<Colon> int) then 0 else number_of l)"
huffman@28969
   371
  unfolding nat_number_of_def number_of_is_id neg_def by simp
haftmann@25931
   372
haftmann@32061
   373
lemma of_nat_int [code_unfold_post]:
haftmann@25931
   374
  "of_nat = int" by (simp add: int_def)
haftmann@25931
   375
haftmann@32065
   376
lemma of_nat_aux_int [code_unfold]:
haftmann@32065
   377
  "of_nat_aux (\<lambda>i. i + 1) k 0 = int k"
haftmann@32065
   378
  by (simp add: int_def Nat.of_nat_code)
haftmann@32065
   379
haftmann@25931
   380
code_const int
haftmann@25931
   381
  (SML "_")
haftmann@25931
   382
  (OCaml "_")
haftmann@25931
   383
haftmann@25931
   384
consts_code
haftmann@25931
   385
  int ("(_)")
haftmann@25931
   386
  nat ("\<module>nat")
haftmann@25931
   387
attach {*
haftmann@25931
   388
fun nat i = if i < 0 then 0 else i;
haftmann@25931
   389
*}
haftmann@25931
   390
haftmann@25967
   391
code_const nat
haftmann@25967
   392
  (SML "IntInf.max/ (/0,/ _)")
haftmann@25967
   393
  (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int")
haftmann@38195
   394
  (Eval "Integer.max/ _/ 0")
haftmann@25967
   395
haftmann@35689
   396
text {* For Haskell and Scala, things are slightly different again. *}
haftmann@25967
   397
haftmann@25967
   398
code_const int and nat
haftmann@25967
   399
  (Haskell "toInteger" and "fromInteger")
haftmann@39004
   400
  (Scala "!_.as'_BigInt" and "Nat.Nat")
haftmann@25931
   401
haftmann@38195
   402
text {* Conversion from and to code numerals. *}
haftmann@25931
   403
haftmann@31205
   404
code_const Code_Numeral.of_nat
haftmann@25967
   405
  (SML "IntInf.toInt")
haftmann@31377
   406
  (OCaml "_")
haftmann@38195
   407
  (Haskell "!(fromInteger/ ./ toInteger)")
haftmann@39004
   408
  (Scala "!Natural.Nat(_.as'_BigInt)")
haftmann@38195
   409
  (Eval "_")
haftmann@25967
   410
haftmann@31205
   411
code_const Code_Numeral.nat_of
haftmann@25931
   412
  (SML "IntInf.fromInt")
haftmann@31377
   413
  (OCaml "_")
haftmann@38195
   414
  (Haskell "!(fromInteger/ ./ toInteger)")
haftmann@39004
   415
  (Scala "!Nat.Nat(_.as'_BigInt)")
haftmann@38195
   416
  (Eval "_")
haftmann@25931
   417
haftmann@25931
   418
text {* Using target language arithmetic operations whenever appropriate *}
haftmann@25931
   419
haftmann@25931
   420
code_const "op + \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
haftmann@25931
   421
  (SML "IntInf.+ ((_), (_))")
haftmann@25931
   422
  (OCaml "Big'_int.add'_big'_int")
haftmann@25931
   423
  (Haskell infixl 6 "+")
haftmann@34899
   424
  (Scala infixl 7 "+")
haftmann@38195
   425
  (Eval infixl 8 "+")
haftmann@34899
   426
haftmann@34899
   427
code_const "op - \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
haftmann@34899
   428
  (Haskell infixl 6 "-")
haftmann@34899
   429
  (Scala infixl 7 "-")
haftmann@25931
   430
haftmann@25931
   431
code_const "op * \<Colon> nat \<Rightarrow> nat \<Rightarrow> nat"
haftmann@25931
   432
  (SML "IntInf.* ((_), (_))")
haftmann@25931
   433
  (OCaml "Big'_int.mult'_big'_int")
haftmann@25931
   434
  (Haskell infixl 7 "*")
haftmann@34899
   435
  (Scala infixl 8 "*")
haftmann@38195
   436
  (Eval infixl 9 "*")
haftmann@25931
   437
haftmann@38195
   438
code_const divmod_nat
haftmann@26009
   439
  (SML "IntInf.divMod/ ((_),/ (_))")
haftmann@26009
   440
  (OCaml "Big'_int.quomod'_big'_int")
haftmann@26009
   441
  (Haskell "divMod")
haftmann@34899
   442
  (Scala infixl 8 "/%")
haftmann@38195
   443
  (Eval "Integer.div'_mod")
haftmann@25931
   444
haftmann@28346
   445
code_const "eq_class.eq \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
haftmann@25931
   446
  (SML "!((_ : IntInf.int) = _)")
haftmann@25931
   447
  (OCaml "Big'_int.eq'_big'_int")
haftmann@25931
   448
  (Haskell infixl 4 "==")
haftmann@34899
   449
  (Scala infixl 5 "==")
haftmann@38195
   450
  (Eval infixl 6 "=")
haftmann@25931
   451
haftmann@25931
   452
code_const "op \<le> \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
haftmann@25931
   453
  (SML "IntInf.<= ((_), (_))")
haftmann@25931
   454
  (OCaml "Big'_int.le'_big'_int")
haftmann@25931
   455
  (Haskell infix 4 "<=")
haftmann@34899
   456
  (Scala infixl 4 "<=")
haftmann@38195
   457
  (Eval infixl 6 "<=")
haftmann@25931
   458
haftmann@25931
   459
code_const "op < \<Colon> nat \<Rightarrow> nat \<Rightarrow> bool"
haftmann@25931
   460
  (SML "IntInf.< ((_), (_))")
haftmann@25931
   461
  (OCaml "Big'_int.lt'_big'_int")
haftmann@25931
   462
  (Haskell infix 4 "<")
haftmann@34899
   463
  (Scala infixl 4 "<")
haftmann@38195
   464
  (Eval infixl 6 "<")
haftmann@25931
   465
haftmann@25931
   466
consts_code
haftmann@28522
   467
  "0::nat"                     ("0")
haftmann@28522
   468
  "1::nat"                     ("1")
haftmann@25931
   469
  Suc                          ("(_ +/ 1)")
haftmann@25931
   470
  "op + \<Colon>  nat \<Rightarrow> nat \<Rightarrow> nat"   ("(_ +/ _)")
haftmann@25931
   471
  "op * \<Colon>  nat \<Rightarrow> nat \<Rightarrow> nat"   ("(_ */ _)")
haftmann@25931
   472
  "op \<le> \<Colon>  nat \<Rightarrow> nat \<Rightarrow> bool"  ("(_ <=/ _)")
haftmann@25931
   473
  "op < \<Colon>  nat \<Rightarrow> nat \<Rightarrow> bool"  ("(_ </ _)")
haftmann@25931
   474
haftmann@25931
   475
haftmann@28228
   476
text {* Evaluation *}
haftmann@28228
   477
haftmann@28562
   478
lemma [code, code del]:
haftmann@32657
   479
  "(Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term) = Code_Evaluation.term_of" ..
haftmann@28228
   480
haftmann@32657
   481
code_const "Code_Evaluation.term_of \<Colon> nat \<Rightarrow> term"
haftmann@28228
   482
  (SML "HOLogic.mk'_number/ HOLogic.natT")
haftmann@28228
   483
haftmann@28228
   484
haftmann@25931
   485
text {* Module names *}
haftmann@23854
   486
haftmann@23854
   487
code_modulename SML
haftmann@33364
   488
  Efficient_Nat Arith
haftmann@23854
   489
haftmann@23854
   490
code_modulename OCaml
haftmann@33364
   491
  Efficient_Nat Arith
haftmann@23854
   492
haftmann@23854
   493
code_modulename Haskell
haftmann@33364
   494
  Efficient_Nat Arith
haftmann@23854
   495
wenzelm@36176
   496
hide_const int
haftmann@23854
   497
haftmann@23854
   498
end