restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
authorhaftmann
Mon, 23 Jul 2012 09:28:03 +0200
changeset 48431 6efff142bb54
parent 48430 6cbfe187a0f9
child 48432 60759d07df24
child 48443 6f2762eedca0
child 48453 2421ff8c57a5
restrict unqualified imports from Haskell Prelude to a small set of fundamental operations
NEWS
src/HOL/Code_Numeral.thy
src/HOL/Library/Code_Char.thy
src/HOL/Library/Code_Char_chr.thy
src/HOL/Library/Code_Integer.thy
src/HOL/Library/Efficient_Nat.thy
src/Tools/Code/code_haskell.ML
--- a/NEWS	Mon Jul 23 09:26:55 2012 +0200
+++ b/NEWS	Mon Jul 23 09:28:03 2012 +0200
@@ -15,6 +15,9 @@
 
 *** Pure ***
 
+* Code generation for Haskell: restrict unqualified imports from
+Haskell Prelude to a small set of fundamental operations.
+
 * Command "export_code": relative file names are interpreted
 relatively to master directory of current theory rather than
 the rather arbitrary current working directory.
--- a/src/HOL/Code_Numeral.thy	Mon Jul 23 09:26:55 2012 +0200
+++ b/src/HOL/Code_Numeral.thy	Mon Jul 23 09:28:03 2012 +0200
@@ -298,7 +298,7 @@
 code_const "minus \<Colon> code_numeral \<Rightarrow> code_numeral \<Rightarrow> code_numeral"
   (SML "Int.max/ (0 : int,/ Int.-/ ((_),/ (_)))")
   (OCaml "Big'_int.max'_big'_int/ Big'_int.zero'_big'_int/ (Big'_int.sub'_big'_int/ _/ _)")
-  (Haskell "max/ (0 :: Integer)/ (_/ -/ _)")
+  (Haskell "Prelude.max/ (0 :: Integer)/ (_/ -/ _)")
   (Scala "!(_/ -/ _).max(0)")
   (Eval "Integer.max/ 0/ (_/ -/ _)")
 
--- a/src/HOL/Library/Code_Char.thy	Mon Jul 23 09:26:55 2012 +0200
+++ b/src/HOL/Library/Code_Char.thy	Mon Jul 23 09:28:03 2012 +0200
@@ -11,7 +11,7 @@
 code_type char
   (SML "char")
   (OCaml "char")
-  (Haskell "Char")
+  (Haskell "Prelude.Char")
   (Scala "Char")
 
 setup {*
@@ -58,7 +58,4 @@
   (Haskell "_")
   (Scala "!(_.toList)")
 
-
-(*declare Quickcheck_Exhaustive.char.bounded_forall_char.simps [code del]*)
-
 end
--- a/src/HOL/Library/Code_Char_chr.thy	Mon Jul 23 09:26:55 2012 +0200
+++ b/src/HOL/Library/Code_Char_chr.thy	Mon Jul 23 09:28:03 2012 +0200
@@ -25,7 +25,7 @@
 code_const int_of_char and char_of_int
   (SML "!(IntInf.fromInt o Char.ord)" and "!(Char.chr o IntInf.toInt)")
   (OCaml "Big'_int.big'_int'_of'_int (Char.code _)" and "Char.chr (Big'_int.int'_of'_big'_int _)")
-  (Haskell "toInteger (fromEnum (_ :: Char))" and "!(let chr k | (0 <= k && k < 256) = toEnum k :: Char in chr . fromInteger)")
+  (Haskell "Prelude.toInteger (Prelude.fromEnum (_ :: Prelude.Char))" and "!(let chr k | (0 <= k && k < 256) = Prelude.toEnum k :: Prelude.Char in chr . Prelude.fromInteger)")
   (Scala "BigInt(_.toInt)" and "!((k: BigInt) => if (BigInt(0) <= k && k < BigInt(256)) k.charValue else error(\"character value out of range\"))")
 
 end
--- a/src/HOL/Library/Code_Integer.thy	Mon Jul 23 09:26:55 2012 +0200
+++ b/src/HOL/Library/Code_Integer.thy	Mon Jul 23 09:28:03 2012 +0200
@@ -151,7 +151,7 @@
 code_const Code_Numeral.int_of
   (SML "IntInf.fromInt")
   (OCaml "_")
-  (Haskell "toInteger")
+  (Haskell "Prelude.toInteger")
   (Scala "!_.as'_BigInt")
   (Eval "_")
 
--- a/src/HOL/Library/Efficient_Nat.thy	Mon Jul 23 09:26:55 2012 +0200
+++ b/src/HOL/Library/Efficient_Nat.thy	Mon Jul 23 09:28:03 2012 +0200
@@ -162,7 +162,7 @@
 text {* For Haskell and Scala, things are slightly different again. *}
 
 code_const int and nat
-  (Haskell "toInteger" and "fromInteger")
+  (Haskell "Prelude.toInteger" and "Prelude.fromInteger")
   (Scala "!_.as'_BigInt" and "Nat")
 
 text {* Alternativ implementation for @{const of_nat} *}
@@ -189,14 +189,14 @@
 code_const Code_Numeral.of_nat
   (SML "IntInf.toInt")
   (OCaml "_")
-  (Haskell "!(fromInteger/ ./ toInteger)")
+  (Haskell "!(Prelude.fromInteger/ ./ Prelude.toInteger)")
   (Scala "!Natural(_.as'_BigInt)")
   (Eval "_")
 
 code_const Code_Numeral.nat_of
   (SML "IntInf.fromInt")
   (OCaml "_")
-  (Haskell "!(fromInteger/ ./ toInteger)")
+  (Haskell "!(Prelude.fromInteger/ ./ Prelude.toInteger)")
   (Scala "!Nat(_.as'_BigInt)")
   (Eval "_")
 
--- a/src/Tools/Code/code_haskell.ML	Mon Jul 23 09:26:55 2012 +0200
+++ b/src/Tools/Code/code_haskell.ML	Mon Jul 23 09:28:03 2012 +0200
@@ -303,6 +303,27 @@
         modify_stmt = fn stmt => if select_stmt stmt then SOME stmt else NONE }
   end;
 
+val prelude_import_operators = [
+  "==", "/=", "<", "<=", ">=", ">", "+", "-", "*", "/", "**", ">>=", ">>", "=<<", "&&", "||", "^", "^^", ".", "$", "$!", "++", "!!"
+];
+
+val prelude_import_unqualified = [
+  "Eq",
+  "error",
+  "id",
+  "return",
+  "not",
+  "fst", "snd",
+  "map", "filter", "concat", "concatMap", "reverse", "zip", "null", "takeWhile", "dropWhile", "all", "any",
+  "Integer", "negate", "abs", "divMod",
+  "String"
+];
+
+val prelude_import_unqualified_constr = [
+  ("Bool", ["True", "False"]),
+  ("Maybe", ["Nothing", "Just"])
+];
+
 fun serialize_haskell module_prefix string_classes { labelled_name, reserved_syms,
     includes, module_alias, class_syntax, tyco_syntax, const_syntax } program =
   let
@@ -331,23 +352,28 @@
       deresolve (if string_classes then deriving_show else K false);
 
     (* print modules *)
-    val import_includes_ps =
-      map (fn (name, _) => str ("import qualified " ^ name ^ ";")) includes;
     fun print_module_frame module_name ps =
       (module_name, Pretty.chunks2 (
         str ("module " ^ module_name ^ " where {")
         :: ps
         @| str "}"
       ));
+    fun print_qualified_import module_name = semicolon [str "import qualified", str module_name];
+    val import_common_ps =
+      enclose "import Prelude (" ");" (commas (map str
+        (map (Library.enclose "(" ")") prelude_import_operators @ prelude_import_unqualified)
+          @ map (fn (tyco, constrs) => (enclose (tyco ^ "(") ")" o commas o map str) constrs) prelude_import_unqualified_constr))
+      :: print_qualified_import "Prelude"
+      :: map (print_qualified_import o fst) includes;
     fun print_module module_name (gr, imports) =
       let
-        val deresolve = deresolver module_name
+        val deresolve = deresolver module_name;
         fun print_import module_name = (semicolon o map str) ["import qualified", module_name];
-        val import_ps = import_includes_ps @ map (print_import o fst) imports;
-        fun print_stmt' gr name = case Graph.get_node gr name
+        val import_ps = import_common_ps @ map (print_qualified_import o fst) imports;
+        fun print_stmt' name = case Graph.get_node gr name
          of (_, NONE) => NONE
           | (_, SOME stmt) => SOME (markup_stmt name (print_stmt deresolve (name, stmt)));
-        val body_ps = map_filter (print_stmt' gr) ((flat o rev o Graph.strong_conn) gr);
+        val body_ps = map_filter print_stmt' ((flat o rev o Graph.strong_conn) gr);
       in
         print_module_frame module_name
           ((if null import_ps then [] else [Pretty.chunks import_ps]) @ body_ps)
@@ -472,55 +498,8 @@
       "import", "default", "forall", "let", "in", "class", "qualified", "data",
       "newtype", "instance", "if", "then", "else", "type", "as", "do", "module"
     ]
-  #> fold (Code_Target.add_reserved target) [
-      "Prelude", "Main", "Bool", "Maybe", "Either", "Ordering", "Char", "String", "Int",
-      "Integer", "Float", "Double", "Rational", "IO", "Eq", "Ord", "Enum", "Bounded",
-      "Num", "Real", "Integral", "Fractional", "Floating", "RealFloat", "Monad", "Functor",
-      "AlreadyExists", "ArithException", "ArrayException", "AssertionFailed", "AsyncException",
-      "BlockedOnDeadMVar", "Deadlock", "Denormal", "DivideByZero", "DotNetException", "DynException",
-      "Dynamic", "EOF", "EQ", "EmptyRec", "ErrorCall", "ExitException", "ExitFailure",
-      "ExitSuccess", "False", "GT", "HeapOverflow",
-      "IOError", "IOException", "IllegalOperation",
-      "IndexOutOfBounds", "Just", "Key", "LT", "Left", "LossOfPrecision", "NoMethodError",
-      "NoSuchThing", "NonTermination", "Nothing", "Obj", "OtherError", "Overflow",
-      "PatternMatchFail", "PermissionDenied", "ProtocolError", "RecConError", "RecSelError",
-      "RecUpdError", "ResourceBusy", "ResourceExhausted", "Right", "StackOverflow",
-      "ThreadKilled", "True", "TyCon", "TypeRep", "UndefinedElement", "Underflow",
-      "UnsupportedOperation", "UserError", "abs", "absReal", "acos", "acosh", "all",
-      "and", "any", "appendFile", "asTypeOf", "asciiTab", "asin", "asinh", "atan",
-      "atan2", "atanh", "basicIORun", "blockIO", "boundedEnumFrom", "boundedEnumFromThen",
-      "boundedEnumFromThenTo", "boundedEnumFromTo", "boundedPred", "boundedSucc", "break",
-      "catch", "catchException", "ceiling", "compare", "concat", "concatMap", "const",
-      "cos", "cosh", "curry", "cycle", "decodeFloat", "denominator", "div", "divMod",
-      "doubleToRatio", "doubleToRational", "drop", "dropWhile", "either", "elem",
-      "emptyRec", "encodeFloat", "enumFrom", "enumFromThen", "enumFromThenTo",
-      "enumFromTo", "error", "even", "exp", "exponent", "fail", "filter", "flip",
-      "floatDigits", "floatProperFraction", "floatRadix", "floatRange", "floatToRational",
-      "floor", "fmap", "foldl", "foldl'", "foldl1", "foldr", "foldr1", "fromDouble",
-      "fromEnum", "fromEnum_0", "fromInt", "fromInteger", "fromIntegral", "fromObj",
-      "fromRational", "fst", "gcd", "getChar", "getContents", "getLine", "head",
-      "id", "inRange", "index", "init", "intToRatio", "interact", "ioError", "isAlpha",
-      "isAlphaNum", "isDenormalized", "isDigit", "isHexDigit", "isIEEE", "isInfinite",
-      "isLower", "isNaN", "isNegativeZero", "isOctDigit", "isSpace", "isUpper", "iterate", "iterate'",
-      "last", "lcm", "length", "lex", "lexDigits", "lexLitChar", "lexmatch", "lines", "log",
-      "logBase", "lookup", "loop", "map", "mapM", "mapM_", "max", "maxBound", "maximum",
-      "maybe", "min", "minBound", "minimum", "mod", "negate", "nonnull", "not", "notElem",
-      "null", "numerator", "numericEnumFrom", "numericEnumFromThen", "numericEnumFromThenTo",
-      "numericEnumFromTo", "odd", "or", "otherwise", "pi", "pred", 
-      "print", "product", "properFraction", "protectEsc", "putChar", "putStr", "putStrLn",
-      "quot", "quotRem", "range", "rangeSize", "rationalToDouble", "rationalToFloat",
-      "rationalToRealFloat", "read", "readDec", "readField", "readFieldName", "readFile",
-      "readFloat", "readHex", "readIO", "readInt", "readList", "readLitChar", "readLn",
-      "readOct", "readParen", "readSigned", "reads", "readsPrec", "realFloatToRational",
-      "realToFrac", "recip", "reduce", "rem", "repeat", "replicate", "return", "reverse",
-      "round", "scaleFloat", "scanl", "scanl1", "scanr", "scanr1", "seq", "sequence",
-      "sequence_", "show", "showChar", "showException", "showField", "showList",
-      "showLitChar", "showParen", "showString", "shows", "showsPrec", "significand",
-      "signum", "signumReal", "sin", "sinh", "snd", "span", "splitAt", "sqrt", "subtract",
-      "succ", "sum", "tail", "take", "takeWhile", "takeWhile1", "tan", "tanh", "threadToIOResult",
-      "throw", "toEnum", "toInt", "toInteger", "toObj", "toRational", "truncate", "uncurry",
-      "undefined", "unlines", "unsafeCoerce", "unsafeIndex", "unsafeRangeSize", "until", "unwords",
-      "unzip", "unzip3", "userError", "words", "writeFile", "zip", "zip3", "zipWith", "zipWith3"
-    ] (*due to weird handling of ':', we can't do anything else than to import *all* prelude symbols*);
+  #> fold (Code_Target.add_reserved target) prelude_import_unqualified
+  #> fold (Code_Target.add_reserved target o fst) prelude_import_unqualified_constr
+  #> fold (fold (Code_Target.add_reserved target) o snd) prelude_import_unqualified_constr;
 
 end; (*struct*)