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