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