clarified names;
authorwenzelm
Mon, 05 Nov 2018 17:37:55 +0100
changeset 69241 5426d266dcc5
parent 69240 16ca270090b6
child 69242 c911716d29bb
clarified names;
src/Tools/Haskell/Build.thy
src/Tools/Haskell/Test.thy
src/Tools/ROOT
--- a/src/Tools/Haskell/Build.thy	Mon Nov 05 17:06:50 2018 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,23 +0,0 @@
-(*  Title:      Tools/Haskell/Build.thy
-    Author:     Makarius
-*)
-
-section \<open>Build Isabelle/Haskell modules\<close>
-
-theory Build imports Haskell
-begin
-
-ML \<open>
-  Isabelle_System.with_tmp_dir "ghc" (fn dir =>
-    let
-      val _ = Haskell.install_sources dir;
-      val (out, rc) =
-        Isabelle_System.bash_output
-         (cat_lines
-           ["set -e",
-            "cd " ^ File.bash_path dir,
-            "\"$ISABELLE_GHC\" " ^ File.bash_paths Haskell.sources]);
-    in if rc = 0 then writeln out else error out end)
-\<close>
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Haskell/Test.thy	Mon Nov 05 17:37:55 2018 +0100
@@ -0,0 +1,23 @@
+(*  Title:      Tools/Haskell/Test.thy
+    Author:     Makarius
+*)
+
+section \<open>Test build of Isabelle/Haskell modules\<close>
+
+theory Test imports Haskell
+begin
+
+ML \<open>
+  Isabelle_System.with_tmp_dir "ghc" (fn dir =>
+    let
+      val _ = Haskell.install_sources dir;
+      val (out, rc) =
+        Isabelle_System.bash_output
+         (cat_lines
+           ["set -e",
+            "cd " ^ File.bash_path dir,
+            "\"$ISABELLE_GHC\" " ^ File.bash_paths Haskell.sources]);
+    in if rc = 0 then writeln out else error out end)
+\<close>
+
+end
--- a/src/Tools/ROOT	Mon Nov 05 17:06:50 2018 +0100
+++ b/src/Tools/ROOT	Mon Nov 05 17:37:55 2018 +0100
@@ -13,4 +13,4 @@
   theories
     Haskell
   theories [condition = ISABELLE_GHC]
-    Build
+    Test