| author | wenzelm | 
| Mon, 17 Jul 2023 20:59:50 +0200 | |
| changeset 78385 | 4d9b953c7026 | 
| parent 74147 | d030b988d470 | 
| permissions | -rw-r--r-- | 
| 69209 | 1  | 
(* Title: Pure/Tools/ghc.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Support for GHC: Glasgow Haskell Compiler.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature GHC =  | 
|
8  | 
sig  | 
|
9  | 
val print_codepoint: UTF8.codepoint -> string  | 
|
10  | 
val print_symbol: Symbol.symbol -> string  | 
|
11  | 
val print_string: string -> string  | 
|
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
12  | 
  val project_template: {depends: string list, modules: string list} -> string
 | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
13  | 
  val new_project: Path.T -> {name: string, depends: string list, modules: string list} -> unit
 | 
| 69209 | 14  | 
end;  | 
15  | 
||
16  | 
structure GHC: GHC =  | 
|
17  | 
struct  | 
|
18  | 
||
19  | 
(** string literals **)  | 
|
20  | 
||
21  | 
fun print_codepoint c =  | 
|
22  | 
(case c of  | 
|
23  | 
34 => "\\\""  | 
|
24  | 
| 39 => "\\'"  | 
|
25  | 
| 92 => "\\\\"  | 
|
26  | 
| 7 => "\\a"  | 
|
27  | 
| 8 => "\\b"  | 
|
28  | 
| 9 => "\\t"  | 
|
29  | 
| 10 => "\\n"  | 
|
30  | 
| 11 => "\\v"  | 
|
31  | 
| 12 => "\\f"  | 
|
32  | 
| 13 => "\\r"  | 
|
33  | 
| c =>  | 
|
34  | 
if c >= 32 andalso c < 127 then chr c  | 
|
35  | 
else "\\" ^ string_of_int c ^ "\\&");  | 
|
36  | 
||
37  | 
fun print_symbol sym =  | 
|
38  | 
(case Symbol.decode sym of  | 
|
39  | 
Symbol.Char s => print_codepoint (ord s)  | 
|
40  | 
| Symbol.UTF8 s => UTF8.decode_permissive s |> map print_codepoint |> implode  | 
|
41  | 
| Symbol.Sym s => "\\092<" ^ s ^ ">"  | 
|
42  | 
| Symbol.Control s => "\\092<^" ^ s ^ ">"  | 
|
43  | 
| _ => translate_string (print_codepoint o ord) sym);  | 
|
44  | 
||
45  | 
val print_string = quote o implode o map print_symbol o Symbol.explode;  | 
|
46  | 
||
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
47  | 
|
| 69475 | 48  | 
|
49  | 
(** project setup **)  | 
|
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
50  | 
|
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
51  | 
fun project_template {depends, modules} =
 | 
| 69470 | 52  | 
  \<^verbatim>\<open>{-# START_FILE {{name}}.cabal #-}
 | 
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
53  | 
name:                {{name}}
 | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
54  | 
version: 0.1.0.0  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
55  | 
homepage: default  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
56  | 
license: BSD3  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
57  | 
author: default  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
58  | 
maintainer: default  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
59  | 
category: default  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
60  | 
build-type: Simple  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
61  | 
cabal-version: >=1.10  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
62  | 
|
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
63  | 
executable {{name}}
 | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
64  | 
hs-source-dirs: src  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
65  | 
main-is: Main.hs  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
66  | 
default-language: Haskell2010  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
67  | 
  build-depends:       \<close> ^ commas ("base >= 4.7 && < 5" :: depends) ^
 | 
| 69470 | 68  | 
\<^verbatim>\<open>  | 
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
69  | 
other-modules: \<close> ^ commas modules ^  | 
| 69470 | 70  | 
\<^verbatim>\<open>  | 
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
71  | 
{-# START_FILE Setup.hs #-}
 | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
72  | 
import Distribution.Simple  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
73  | 
main = defaultMain  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
74  | 
|
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
75  | 
{-# START_FILE src/Main.hs #-}
 | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
76  | 
module Main where  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
77  | 
|
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
78  | 
main :: IO ()  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
79  | 
main = return ()  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
80  | 
\<close>;  | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
81  | 
|
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
82  | 
fun new_project dir {name, depends, modules} =
 | 
| 
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
83  | 
let  | 
| 
72511
 
460d743010bc
clarified signature: overloaded "+" for Path.append;
 
wenzelm 
parents: 
72278 
diff
changeset
 | 
84  | 
val template_path = dir + (Path.basic name |> Path.ext "hsfiles");  | 
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
85  | 
    val _ = File.write template_path (project_template {depends = depends, modules = modules});
 | 
| 73267 | 86  | 
val _ =  | 
| 
74147
 
d030b988d470
provide bash_process server for Isabelle/ML and other external programs;
 
wenzelm 
parents: 
74142 
diff
changeset
 | 
87  | 
Isabelle_System.bash_process  | 
| 
 
d030b988d470
provide bash_process server for Isabelle/ML and other external programs;
 
wenzelm 
parents: 
74142 
diff
changeset
 | 
88  | 
        (Bash.script ("isabelle ghc_stack new " ^ Bash.string name ^
 | 
| 
 
d030b988d470
provide bash_process server for Isabelle/ML and other external programs;
 
wenzelm 
parents: 
74142 
diff
changeset
 | 
89  | 
" --bare " ^ File.bash_platform_path template_path) |> Bash.cwd dir)  | 
| 73279 | 90  | 
|> Process_Result.check;  | 
| 73267 | 91  | 
in () end;  | 
| 
69444
 
c3c9440cbf9b
more formal Haskell project setup, with dependencies on packages from "stackage";
 
wenzelm 
parents: 
69381 
diff
changeset
 | 
92  | 
|
| 69209 | 93  | 
end;  |