| author | wenzelm | 
| Thu, 14 Apr 2016 12:00:29 +0200 | |
| changeset 62972 | 0eedd78c2b47 | 
| parent 62829 | 4141c2a8458b | 
| child 66679 | ed8d359d92e4 | 
| permissions | -rw-r--r-- | 
| 40743 | 1  | 
(* Title: Pure/System/isabelle_system.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Isabelle system support.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature ISABELLE_SYSTEM =  | 
|
8  | 
sig  | 
|
9  | 
val mkdirs: Path.T -> unit  | 
|
| 40745 | 10  | 
val mkdir: Path.T -> unit  | 
| 40743 | 11  | 
val copy_dir: Path.T -> Path.T -> unit  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
12  | 
val copy_file: Path.T -> Path.T -> unit  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
13  | 
val copy_file_base: Path.T * Path.T -> Path.T -> unit  | 
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
14  | 
val create_tmp_path: string -> string -> Path.T  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
15  | 
val with_tmp_file: string -> string -> (Path.T -> 'a) -> 'a  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
16  | 
val with_tmp_dir: string -> (Path.T -> 'a) -> 'a  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
17  | 
val bash_output: string -> string * int  | 
| 
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
18  | 
val bash: string -> int  | 
| 40743 | 19  | 
end;  | 
20  | 
||
21  | 
structure Isabelle_System: ISABELLE_SYSTEM =  | 
|
22  | 
struct  | 
|
23  | 
||
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
24  | 
(* bash *)  | 
| 
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
25  | 
|
| 
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
26  | 
fun bash_output s =  | 
| 
47499
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
27  | 
let  | 
| 
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
28  | 
    val {out, err, rc, ...} = Bash.process s;
 | 
| 
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
29  | 
val _ = warning (trim_line err);  | 
| 
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
30  | 
in (out, rc) end;  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
31  | 
|
| 
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
32  | 
fun bash s =  | 
| 
47499
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
33  | 
let  | 
| 
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
34  | 
val (out, rc) = bash_output s;  | 
| 
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
35  | 
val _ = writeln (trim_line out);  | 
| 
 
4b0daca2bf88
redirect bash stderr to Isabelle warning as appropriate -- avoid raw process error output which may either get ignored or overload PIDE syslog in extreme cases;
 
wenzelm 
parents: 
46502 
diff
changeset
 | 
36  | 
in rc end;  | 
| 
43850
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
37  | 
|
| 
 
7f2cbc713344
moved bash operations to Isabelle_System (cf. Scala version);
 
wenzelm 
parents: 
43849 
diff
changeset
 | 
38  | 
|
| 
62829
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62549 
diff
changeset
 | 
39  | 
(* directory operations *)  | 
| 40743 | 40  | 
|
41  | 
fun system_command cmd =  | 
|
| 
62829
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62549 
diff
changeset
 | 
42  | 
  if bash cmd <> 0 then error ("System command failed: " ^ cmd) else ();
 | 
| 40743 | 43  | 
|
| 
62829
 
4141c2a8458b
clarified Isabelle tool wrapper: bash, Scala, no perl, no ML;
 
wenzelm 
parents: 
62549 
diff
changeset
 | 
44  | 
fun rm_tree path = system_command ("rm -r -f " ^ File.bash_path path);
 | 
| 40743 | 45  | 
|
| 60013 | 46  | 
fun mkdirs path =  | 
| 
60263
 
2a5dbad75355
more portable mkdirs via perl, e.g. relevant for Windows UNC paths (network shares);
 
wenzelm 
parents: 
60028 
diff
changeset
 | 
47  | 
if File.is_dir path then ()  | 
| 
 
2a5dbad75355
more portable mkdirs via perl, e.g. relevant for Windows UNC paths (network shares);
 
wenzelm 
parents: 
60028 
diff
changeset
 | 
48  | 
else  | 
| 
62549
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
60970 
diff
changeset
 | 
49  | 
   (bash ("perl -e \"use File::Path make_path; make_path('" ^ File.standard_path path ^ "');\"");
 | 
| 
60263
 
2a5dbad75355
more portable mkdirs via perl, e.g. relevant for Windows UNC paths (network shares);
 
wenzelm 
parents: 
60028 
diff
changeset
 | 
50  | 
    if File.is_dir path then () else error ("Failed to create directory: " ^ Path.print path));
 | 
| 40743 | 51  | 
|
| 40785 | 52  | 
fun mkdir path =  | 
53  | 
if File.is_dir path then () else OS.FileSys.mkDir (File.platform_path path);  | 
|
| 40743 | 54  | 
|
55  | 
fun copy_dir src dst =  | 
|
56  | 
if File.eq (src, dst) then ()  | 
|
| 
62549
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
60970 
diff
changeset
 | 
57  | 
  else (system_command ("cp -p -R -f " ^ File.bash_path src ^ "/. " ^ File.bash_path dst); ());
 | 
| 40743 | 58  | 
|
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
59  | 
fun copy_file src0 dst0 =  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
60  | 
let  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
61  | 
val src = Path.expand src0;  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
62  | 
val dst = Path.expand dst0;  | 
| 
57085
 
cb212f52c2a3
clarified copy_file: allow to rename base name, e.g. relevant for 'display_drafts';
 
wenzelm 
parents: 
56533 
diff
changeset
 | 
63  | 
val target = if File.is_dir dst then Path.append dst (Path.base src) else dst;  | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
64  | 
in  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
65  | 
if File.eq (src, target) then ()  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
66  | 
else  | 
| 
62549
 
9498623b27f0
File.bash_string operations in ML as in Scala -- exclusively for GNU bash, not perl and not user output;
 
wenzelm 
parents: 
60970 
diff
changeset
 | 
67  | 
      ignore (system_command ("cp -p -f " ^ File.bash_path src ^ " " ^ File.bash_path target))
 | 
| 
56533
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
68  | 
end;  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
69  | 
|
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
70  | 
fun copy_file_base (base_dir, src0) target_dir =  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
71  | 
let  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
72  | 
val src = Path.expand src0;  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
73  | 
val src_dir = Path.dir src;  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
74  | 
val _ =  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
75  | 
if Path.starts_basic src then ()  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
76  | 
      else error ("Illegal path specification " ^ Path.print src ^ " beyond base directory");
 | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
77  | 
val _ = mkdirs (Path.append target_dir src_dir);  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
78  | 
in copy_file (Path.append base_dir src) (Path.append target_dir src) end;  | 
| 
 
cd8b6d849b6a
explicit 'document_files' in session ROOT specifications;
 
wenzelm 
parents: 
56498 
diff
changeset
 | 
79  | 
|
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
80  | 
|
| 56498 | 81  | 
(* tmp files *)  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
82  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
83  | 
fun create_tmp_path name ext =  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
84  | 
let  | 
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
85  | 
val path = File.tmp_path (Path.basic (name ^ serial_string ()) |> Path.ext ext);  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
86  | 
val _ = File.exists path andalso  | 
| 
41944
 
b97091ae583a
Path.print is the official way to show file-system paths to users -- note that Path.implode often indicates violation of the abstract datatype;
 
wenzelm 
parents: 
41352 
diff
changeset
 | 
87  | 
      raise Fail ("Temporary file already exists: " ^ Path.print path);
 | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
88  | 
in path end;  | 
| 
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
89  | 
|
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
90  | 
fun with_tmp_file name ext f =  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
91  | 
let val path = create_tmp_path name ext  | 
| 
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
92  | 
in Exn.release (Exn.capture f path before ignore (try File.rm path)) end;  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
93  | 
|
| 56498 | 94  | 
|
95  | 
(* tmp dirs *)  | 
|
96  | 
||
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
97  | 
fun with_tmp_dir name f =  | 
| 
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
98  | 
let  | 
| 
42127
 
8223e7f4b0da
Isabelle_System.create_tmp_path/with_tmp_file: optional extension;
 
wenzelm 
parents: 
41944 
diff
changeset
 | 
99  | 
val path = create_tmp_path name "";  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
100  | 
val _ = mkdirs path;  | 
| 
41352
 
87adb55fb0fb
make SML/NJ and Poly/ML agree on the type of "before";
 
wenzelm 
parents: 
41307 
diff
changeset
 | 
101  | 
in Exn.release (Exn.capture f path before ignore (try rm_tree path)) end;  | 
| 
41307
 
bb8468ae414e
slightly more standard Isabelle_System.with_tmp_file/with_tmp_dir (cf. Scala version);
 
wenzelm 
parents: 
40785 
diff
changeset
 | 
102  | 
|
| 40743 | 103  | 
end;  |