src/HOL/Library/Code_Natural.thy
changeset 46547 d1dcb91a512e
parent 39781 2053638a2bf2
child 47108 2a1953f0d20d
--- a/src/HOL/Library/Code_Natural.thy	Sat Feb 18 10:35:45 2012 +0100
+++ b/src/HOL/Library/Code_Natural.thy	Mon Feb 20 12:37:17 2012 +0100
@@ -3,7 +3,7 @@
 *)
 
 theory Code_Natural
-imports Main
+imports "../Main"
 begin
 
 section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *}
@@ -125,7 +125,7 @@
   (Haskell infixl 7 "*")
   (Scala infixl 8 "*")
 
-code_const div_mod_code_numeral
+code_const Code_Numeral.div_mod
   (Haskell "divMod")
   (Scala infixl 8 "/%")