src/HOL/Library/Code_Natural.thy
changeset 38781 6b356e3687d2
parent 38775 741ca0c98f6f
child 38810 361119ea62ee
--- a/src/HOL/Library/Code_Natural.thy	Thu Aug 26 13:44:50 2010 +0200
+++ b/src/HOL/Library/Code_Natural.thy	Thu Aug 26 13:56:35 2010 +0200
@@ -9,9 +9,7 @@
 section {* Alternative representation of @{typ code_numeral} for @{text Haskell} and @{text Scala} *}
 
 code_include Haskell "Natural"
-{*import Data.Array.ST;
-
-newtype Natural = Natural Integer deriving (Eq, Show, Read);
+{*newtype Natural = Natural Integer deriving (Eq, Show, Read);
 
 instance Num Natural where {
   fromInteger k = Natural (if k >= 0 then k else 0);