src/HOL/Library/Code_Target_Int.thy
changeset 51143 0a2371e7ced3
parent 51114 3e913a575dc6
child 52435 6646bb548c6b
--- a/src/HOL/Library/Code_Target_Int.thy	Fri Feb 15 08:31:30 2013 +0100
+++ b/src/HOL/Library/Code_Target_Int.thy	Fri Feb 15 08:31:31 2013 +0100
@@ -5,7 +5,7 @@
 header {* Implementation of integer numbers by target-language integers *}
 
 theory Code_Target_Int
-imports Main "~~/src/HOL/Library/Code_Numeral_Types"
+imports Main
 begin
 
 code_datatype int_of_integer
@@ -54,7 +54,7 @@
   by transfer simp
 
 lemma [code]:
-  "Int.dup k = int_of_integer (Code_Numeral_Types.dup (of_int k))"
+  "Int.dup k = int_of_integer (Code_Numeral.dup (of_int k))"
   by transfer simp
 
 lemma [code, code del]:
@@ -66,7 +66,7 @@
 
 lemma [code]:
   "pdivmod k l = map_pair int_of_integer int_of_integer
-    (Code_Numeral_Types.divmod_abs (of_int k) (of_int l))"
+    (Code_Numeral.divmod_abs (of_int k) (of_int l))"
   by (simp add: prod_eq_iff pdivmod_def)
 
 lemma [code]: