src/HOL/Library/ML_Int.thy
author paulson
Thu, 27 Sep 2007 17:55:28 +0200
changeset 24742 73b8b42a36b6
parent 24716 483f0a64271f
permissions -rw-r--r--
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering theorems of Nat.thy are hidden by the Ordering.thy versions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     1
(*  ID:         $Id$
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     3
*)
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     4
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
     5
header {* Built-in integers for code generation *}
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     6
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     7
theory ML_Int
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
     8
imports PreList
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
     9
begin
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    10
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    11
subsection {* Datatype of built-in integers *}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    12
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    13
datatype ml_int = ml_int_of_int int
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    14
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    15
lemmas [code func del] = ml_int.recs ml_int.cases
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    16
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    17
fun
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    18
  int_of_ml_int :: "ml_int \<Rightarrow> int"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    19
where
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    20
  "int_of_ml_int (ml_int_of_int k) = k"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    21
lemmas [code func del] = int_of_ml_int.simps
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    22
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    23
lemma ml_int_id [simp]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    24
  "ml_int_of_int (int_of_ml_int k) = k"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    25
  by (cases k) simp_all
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    26
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    27
lemma ml_int:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    28
  "(\<And>k\<Colon>ml_int. PROP P k) \<equiv> (\<And>k\<Colon>int. PROP P (ml_int_of_int k))"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    29
proof
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    30
  fix k :: int
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    31
  assume "\<And>k\<Colon>ml_int. PROP P k"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    32
  then show "PROP P (ml_int_of_int k)" .
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    33
next
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    34
  fix k :: ml_int
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    35
  assume "\<And>k\<Colon>int. PROP P (ml_int_of_int k)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    36
  then have "PROP P (ml_int_of_int (int_of_ml_int k))" .
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    37
  then show "PROP P k" by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    38
qed
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    39
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    40
lemma [code func]: "size (k\<Colon>ml_int) = 0"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    41
  by (cases k) simp_all
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    42
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    43
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    44
subsection {* Built-in integers as datatype on numerals *}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    45
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    46
instance ml_int :: number
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    47
  "number_of \<equiv> ml_int_of_int" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    48
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    49
code_datatype "number_of \<Colon> int \<Rightarrow> ml_int"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    50
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    51
lemma number_of_ml_int_id [simp]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    52
  "number_of (int_of_ml_int k) = k"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    53
  unfolding number_of_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    54
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
    55
lemma number_of_ml_int_shift:
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
    56
  "number_of k = ml_int_of_int (number_of k)"
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
    57
  by (simp add: number_of_is_id number_of_ml_int_def)
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
    58
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    59
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    60
subsection {* Basic arithmetic *}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    61
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    62
instance ml_int :: zero
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    63
  [simp]: "0 \<equiv> ml_int_of_int 0" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    64
lemmas [code func del] = zero_ml_int_def
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    65
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    66
instance ml_int :: one
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    67
  [simp]: "1 \<equiv> ml_int_of_int 1" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    68
lemmas [code func del] = one_ml_int_def
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    69
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    70
instance ml_int :: plus
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    71
  [simp]: "k + l \<equiv> ml_int_of_int (int_of_ml_int k + int_of_ml_int l)" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    72
lemmas [code func del] = plus_ml_int_def
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    73
lemma plus_ml_int_code [code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    74
  "ml_int_of_int k + ml_int_of_int l = ml_int_of_int (k + l)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    75
  unfolding plus_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    76
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    77
instance ml_int :: minus
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    78
  [simp]: "- k \<equiv> ml_int_of_int (- int_of_ml_int k)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    79
  [simp]: "k - l \<equiv> ml_int_of_int (int_of_ml_int k - int_of_ml_int l)" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    80
lemmas [code func del] = uminus_ml_int_def minus_ml_int_def
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    81
lemma uminus_ml_int_code [code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    82
  "- ml_int_of_int k \<equiv> ml_int_of_int (- k)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    83
  unfolding uminus_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    84
lemma minus_ml_int_code [code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    85
  "ml_int_of_int k - ml_int_of_int l = ml_int_of_int (k - l)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    86
  unfolding minus_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    87
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    88
instance ml_int :: times
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    89
  [simp]: "k * l \<equiv> ml_int_of_int (int_of_ml_int k * int_of_ml_int l)" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    90
lemmas [code func del] = times_ml_int_def
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    91
lemma times_ml_int_code [code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    92
  "ml_int_of_int k * ml_int_of_int l = ml_int_of_int (k * l)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    93
  unfolding times_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    94
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    95
instance ml_int :: ord
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    96
  [simp]: "k \<le> l \<equiv> int_of_ml_int k \<le> int_of_ml_int l"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    97
  [simp]: "k < l \<equiv> int_of_ml_int k < int_of_ml_int l" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    98
lemmas [code func del] = less_eq_ml_int_def less_ml_int_def
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
    99
lemma less_eq_ml_int_code [code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   100
  "ml_int_of_int k \<le> ml_int_of_int l \<longleftrightarrow> k \<le> l"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   101
  unfolding less_eq_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   102
lemma less_ml_int_code [code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   103
  "ml_int_of_int k < ml_int_of_int l \<longleftrightarrow> k < l"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   104
  unfolding less_ml_int_def by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   105
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   106
instance ml_int :: ring_1
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   107
  by default (auto simp add: left_distrib right_distrib)
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   108
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   109
lemma of_nat_ml_int: "of_nat n = ml_int_of_int (of_nat n)"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   110
proof (induct n)
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   111
  case 0 show ?case by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   112
next
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   113
  case (Suc n)
24354
0fdabe28f0e6 remove int_of_nat; fix abs instance
huffman
parents: 24219
diff changeset
   114
  then have "int_of_ml_int (ml_int_of_int (int n))
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   115
    = int_of_ml_int (of_nat n)" by simp
24354
0fdabe28f0e6 remove int_of_nat; fix abs instance
huffman
parents: 24219
diff changeset
   116
  then have "int n = int_of_ml_int (of_nat n)" by simp
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   117
  then show ?case by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   118
qed
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   119
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   120
instance ml_int :: number_ring
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   121
  by default
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   122
    (simp_all add: left_distrib number_of_ml_int_def of_int_of_nat of_nat_ml_int)
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   123
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   124
lemma zero_ml_int_code [code inline, code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   125
  "(0\<Colon>ml_int) = Numeral0"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   126
  by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   127
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   128
lemma one_ml_int_code [code inline, code func]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   129
  "(1\<Colon>ml_int) = Numeral1"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   130
  by simp
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   131
24354
0fdabe28f0e6 remove int_of_nat; fix abs instance
huffman
parents: 24219
diff changeset
   132
instance ml_int :: abs
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   133
  "\<bar>k\<bar> \<equiv> if k < 0 then -k else k" ..
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   134
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   135
lemma ml_int_of_int [code func]:
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   136
  "ml_int_of_int k = (if k = 0 then 0
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   137
    else if k = -1 then -1
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   138
    else let (l, m) = divAlg (k, 2) in 2 * ml_int_of_int l +
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   139
      (if m = 0 then 0 else 1))"
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   140
  by (simp add: number_of_ml_int_shift Let_def split_def divAlg_mod_div) arith
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   141
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   142
24700
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   143
subsection {* Conversion to and from @{typ nat} *}
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   144
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   145
definition
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   146
  nat_of_ml_int :: "ml_int \<Rightarrow> nat"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   147
where
24700
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   148
  [code func del]: "nat_of_ml_int = nat o int_of_ml_int"
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   149
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   150
definition
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   151
  nat_of_ml_int_aux :: "ml_int \<Rightarrow> nat \<Rightarrow> nat" where
24700
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   152
  [code func del]: "nat_of_ml_int_aux i n = nat_of_ml_int i + n"
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   153
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   154
lemma nat_of_ml_int_aux_code [code]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   155
  "nat_of_ml_int_aux i n = (if i \<le> 0 then n else nat_of_ml_int_aux (i - 1) (Suc n))"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   156
  by (auto simp add: nat_of_ml_int_aux_def nat_of_ml_int_def)
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   157
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   158
lemma nat_of_ml_int_code [code]:
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   159
  "nat_of_ml_int i = nat_of_ml_int_aux i 0"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   160
  by (simp add: nat_of_ml_int_aux_def)
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   161
24700
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   162
definition
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   163
  ml_int_of_nat :: "nat \<Rightarrow> ml_int"
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   164
where
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   165
  [code func del]: "ml_int_of_nat = ml_int_of_int o of_nat"
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   166
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   167
lemma ml_int_of_nat [code func]:
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   168
  "ml_int_of_nat 0 = 0"
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   169
  "ml_int_of_nat (Suc n) = ml_int_of_nat n + 1"
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   170
  unfolding ml_int_of_nat_def by simp_all
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   171
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   172
lemma ml_int_nat_id [simp]:
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   173
  "nat_of_ml_int (ml_int_of_nat n) = n"
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   174
  "ml_int_of_nat (nat_of_ml_int i) = (if i \<le> 0 then 0 else i)"
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   175
  unfolding ml_int_of_nat_def nat_of_ml_int_def by simp_all
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   176
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   177
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   178
subsection {* ML interface *}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   179
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   180
ML {*
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   181
structure ML_Int =
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   182
struct
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   183
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   184
fun mk k = @{term ml_int_of_int} $ HOLogic.mk_number @{typ ml_int} k;
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   185
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   186
end;
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   187
*}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   188
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   189
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   190
subsection {* Code serialization *}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   191
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   192
code_type ml_int
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   193
  (SML "int")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   194
  (OCaml "int")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   195
  (Haskell "Integer")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   196
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   197
code_instance ml_int :: eq
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   198
  (Haskell -)
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   199
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   200
setup {*
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   201
  fold (fn target => CodeTarget.add_pretty_numeral target true
24700
291665d063e4 added conversions for natural numbers
haftmann
parents: 24423
diff changeset
   202
    @{const_name number_ml_int_inst.number_of_ml_int}
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   203
    @{const_name Numeral.B0} @{const_name Numeral.B1}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   204
    @{const_name Numeral.Pls} @{const_name Numeral.Min}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   205
    @{const_name Numeral.Bit}
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   206
  ) ["SML", "OCaml", "Haskell"]
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   207
*}
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   208
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   209
code_reserved SML int
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   210
code_reserved OCaml int
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   211
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   212
code_const "op + \<Colon> ml_int \<Rightarrow> ml_int \<Rightarrow> ml_int"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   213
  (SML "Int.+ ((_), (_))")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   214
  (OCaml "Pervasives.+")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   215
  (Haskell infixl 6 "+")
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   216
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   217
code_const "uminus \<Colon> ml_int \<Rightarrow> ml_int"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   218
  (SML "Int.~")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   219
  (OCaml "Pervasives.~-")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   220
  (Haskell "negate")
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   221
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   222
code_const "op - \<Colon> ml_int \<Rightarrow> ml_int \<Rightarrow> ml_int"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   223
  (SML "Int.- ((_), (_))")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   224
  (OCaml "Pervasives.-")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   225
  (Haskell infixl 6 "-")
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   226
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   227
code_const "op * \<Colon> ml_int \<Rightarrow> ml_int \<Rightarrow> ml_int"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   228
  (SML "Int.* ((_), (_))")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   229
  (OCaml "Pervasives.*")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   230
  (Haskell infixl 7 "*")
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   231
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   232
code_const "op = \<Colon> ml_int \<Rightarrow> ml_int \<Rightarrow> bool"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   233
  (SML "!((_ : Int.int) = _)")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   234
  (OCaml "!((_ : Pervasives.int) = _)")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   235
  (Haskell infixl 4 "==")
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   236
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   237
code_const "op \<le> \<Colon> ml_int \<Rightarrow> ml_int \<Rightarrow> bool"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   238
  (SML "Int.<= ((_), (_))")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   239
  (OCaml "!((_ : Pervasives.int) <= _)")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   240
  (Haskell infix 4 "<=")
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   241
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   242
code_const "op < \<Colon> ml_int \<Rightarrow> ml_int \<Rightarrow> bool"
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   243
  (SML "Int.< ((_), (_))")
24716
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   244
  (OCaml "!((_ : Pervasives.int) < _)")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   245
  (Haskell infix 4 "<")
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   246
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   247
code_reserved SML Int
483f0a64271f added support for Haskell, OCaml
haftmann
parents: 24700
diff changeset
   248
code_reserved OCaml Pervasives
23859
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   249
fc44fa554ca8 support for SML builtin ints
haftmann
parents:
diff changeset
   250
end