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