src/HOL/Library/Code_Integer.thy
changeset 28228 7ebe8dc06cbb
parent 27487 c8a6ce181805
child 28346 b8390cd56b8f
--- a/src/HOL/Library/Code_Integer.thy	Tue Sep 16 09:21:22 2008 +0200
+++ b/src/HOL/Library/Code_Integer.thy	Tue Sep 16 09:21:24 2008 +0200
@@ -6,7 +6,7 @@
 header {* Pretty integer literals for code generation *}
 
 theory Code_Integer
-imports Plain "~~/src/HOL/Presburger"
+imports Plain "~~/src/HOL/Code_Eval" "~~/src/HOL/Presburger"
 begin
 
 text {*
@@ -90,4 +90,12 @@
 code_reserved SML IntInf
 code_reserved OCaml Big_int
 
+text {* Evaluation *}
+
+lemma [code func, code func del]:
+  "(Code_Eval.term_of \<Colon> int \<Rightarrow> term) = Code_Eval.term_of" ..
+
+code_const "Code_Eval.term_of \<Colon> int \<Rightarrow> term"
+  (SML "HOLogic.mk'_number/ HOLogic.intT")
+
 end
\ No newline at end of file