| author | wenzelm | 
| Tue, 14 May 2013 20:46:09 +0200 | |
| changeset 51990 | cc66addbba6d | 
| parent 51988 | a9725750c53a | 
| child 56465 | 6ad693903e22 | 
| permissions | -rw-r--r-- | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/System/options.ML  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Makarius  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
|
| 51945 | 4  | 
System options with external string representation.  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
signature OPTIONS =  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 51943 | 9  | 
val boolT: string  | 
10  | 
val intT: string  | 
|
11  | 
val realT: string  | 
|
12  | 
val stringT: string  | 
|
13  | 
val unknownT: string  | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
type T  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
val empty: T  | 
| 51942 | 16  | 
val typ: T -> string -> string  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
val bool: T -> string -> bool  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
val int: T -> string -> int  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
val real: T -> string -> real  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
val string: T -> string -> string  | 
| 51943 | 21  | 
val put_bool: string -> bool -> T -> T  | 
22  | 
val put_int: string -> int -> T -> T  | 
|
23  | 
val put_real: string -> real -> T -> T  | 
|
24  | 
val put_string: string -> string -> T -> T  | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
  val declare: {name: string, typ: string, value: string} -> T -> T
 | 
| 51943 | 26  | 
val update: string -> string -> T -> T  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
val decode: XML.body -> T  | 
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
28  | 
val default: unit -> T  | 
| 51946 | 29  | 
val default_typ: string -> string  | 
| 51942 | 30  | 
val default_bool: string -> bool  | 
31  | 
val default_int: string -> int  | 
|
32  | 
val default_real: string -> real  | 
|
33  | 
val default_string: string -> string  | 
|
| 51943 | 34  | 
val default_put_bool: string -> bool -> unit  | 
35  | 
val default_put_int: string -> int -> unit  | 
|
36  | 
val default_put_real: string -> real -> unit  | 
|
37  | 
val default_put_string: string -> string -> unit  | 
|
38  | 
val get_default: string -> string  | 
|
39  | 
val put_default: string -> string -> unit  | 
|
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
40  | 
val set_default: T -> unit  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
41  | 
val reset_default: unit -> unit  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
42  | 
val load_default: unit -> unit  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
end;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
structure Options: OPTIONS =  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
struct  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
(* representation *)  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
val boolT = "bool";  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
val intT = "int";  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
val realT = "real";  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
val stringT = "string";  | 
| 51943 | 54  | 
val unknownT = "unknown";  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
datatype T = Options of {typ: string, value: string} Symtab.table;
 | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
val empty = Options Symtab.empty;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
|
| 51943 | 61  | 
(* check *)  | 
| 51942 | 62  | 
|
| 51943 | 63  | 
fun check_name (Options tab) name =  | 
64  | 
let val opt = Symtab.lookup tab name in  | 
|
65  | 
if is_some opt andalso #typ (the opt) <> unknownT then the opt  | 
|
| 51978 | 66  | 
    else error ("Unknown system option " ^ quote name)
 | 
| 51943 | 67  | 
end;  | 
68  | 
||
69  | 
fun check_type options name typ =  | 
|
70  | 
let val opt = check_name options name in  | 
|
71  | 
if #typ opt = typ then opt  | 
|
| 51978 | 72  | 
    else error ("Ill-typed system option " ^ quote name ^ " : " ^ #typ opt ^ " vs. " ^ typ)
 | 
| 51943 | 73  | 
end;  | 
| 51942 | 74  | 
|
75  | 
||
| 51943 | 76  | 
(* typ *)  | 
77  | 
||
78  | 
fun typ options name = #typ (check_name options name);  | 
|
79  | 
||
80  | 
||
81  | 
(* basic operations *)  | 
|
82  | 
||
83  | 
fun put T print name x (options as Options tab) =  | 
|
84  | 
let val opt = check_type options name T  | 
|
85  | 
  in Options (Symtab.update (name, {typ = #typ opt, value = print x}) tab) end;
 | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
86  | 
|
| 51943 | 87  | 
fun get T parse options name =  | 
88  | 
let val opt = check_type options name T in  | 
|
89  | 
(case parse (#value opt) of  | 
|
90  | 
SOME x => x  | 
|
91  | 
| NONE =>  | 
|
| 51978 | 92  | 
        error ("Malformed value for system option " ^ quote name ^
 | 
| 51943 | 93  | 
" : " ^ T ^ " =\n" ^ quote (#value opt)))  | 
94  | 
end;  | 
|
95  | 
||
96  | 
||
97  | 
(* internal lookup and update *)  | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
|
| 
51951
 
fab4ab92e812
more standard Isabelle/ML operations -- avoid inaccurate Bool.fromString;
 
wenzelm 
parents: 
51946 
diff
changeset
 | 
99  | 
val bool = get boolT (try Markup.parse_bool);  | 
| 
 
fab4ab92e812
more standard Isabelle/ML operations -- avoid inaccurate Bool.fromString;
 
wenzelm 
parents: 
51946 
diff
changeset
 | 
100  | 
val int = get intT (try Markup.parse_int);  | 
| 51988 | 101  | 
val real = get realT (try Markup.parse_real);  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
val string = get stringT SOME;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 
51951
 
fab4ab92e812
more standard Isabelle/ML operations -- avoid inaccurate Bool.fromString;
 
wenzelm 
parents: 
51946 
diff
changeset
 | 
104  | 
val put_bool = put boolT Markup.print_bool;  | 
| 
 
fab4ab92e812
more standard Isabelle/ML operations -- avoid inaccurate Bool.fromString;
 
wenzelm 
parents: 
51946 
diff
changeset
 | 
105  | 
val put_int = put intT Markup.print_int;  | 
| 51990 | 106  | 
val put_real = put realT Markup.print_real;  | 
| 51943 | 107  | 
val put_string = put stringT I;  | 
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
|
| 51943 | 109  | 
|
110  | 
(* external updates *)  | 
|
111  | 
||
112  | 
fun check_value options name =  | 
|
113  | 
let val opt = check_name options name in  | 
|
114  | 
if #typ opt = boolT then ignore (bool options name)  | 
|
115  | 
else if #typ opt = intT then ignore (int options name)  | 
|
116  | 
else if #typ opt = realT then ignore (real options name)  | 
|
117  | 
else if #typ opt = stringT then ignore (string options name)  | 
|
118  | 
else ()  | 
|
119  | 
end;  | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
fun declare {name, typ, value} (Options tab) =
 | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
let  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
    val options' = Options (Symtab.update_new (name, {typ = typ, value = value}) tab)
 | 
| 51978 | 124  | 
      handle Symtab.DUP _ => error ("Duplicate declaration of system option " ^ quote name);
 | 
| 51943 | 125  | 
val _ =  | 
126  | 
typ = boolT orelse typ = intT orelse typ = realT orelse typ = stringT orelse  | 
|
| 51978 | 127  | 
        error ("Unknown type for system option " ^ quote name ^ " : " ^ quote typ);
 | 
| 51943 | 128  | 
val _ = check_value options' name;  | 
129  | 
in options' end;  | 
|
130  | 
||
131  | 
fun update name value (options as Options tab) =  | 
|
132  | 
let  | 
|
133  | 
val opt = check_name options name;  | 
|
134  | 
    val options' = Options (Symtab.update (name, {typ = #typ opt, value = value}) tab);
 | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
val _ = check_value options' name;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
in options' end;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
(* decode *)  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
fun decode body =  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
  fold (declare o (fn (name, typ, value) => {name = name, typ = typ, value = value}))
 | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
(let open XML.Decode in list (triple string string string) end body) empty;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
|
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
145  | 
|
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
146  | 
|
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
147  | 
(** global default **)  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
148  | 
|
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
149  | 
val global_default = Synchronized.var "Options.default" (NONE: T option);  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
150  | 
|
| 51978 | 151  | 
fun err_no_default () = error "Missing default for system options within Isabelle process";  | 
| 51943 | 152  | 
|
153  | 
fun change_default f x y =  | 
|
154  | 
Synchronized.change global_default  | 
|
155  | 
(fn SOME options => SOME (f x y options)  | 
|
156  | 
| NONE => err_no_default ());  | 
|
157  | 
||
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
158  | 
fun default () =  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
159  | 
(case Synchronized.value global_default of  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
160  | 
SOME options => options  | 
| 51943 | 161  | 
| NONE => err_no_default ());  | 
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
162  | 
|
| 51946 | 163  | 
fun default_typ name = typ (default ()) name;  | 
| 51942 | 164  | 
fun default_bool name = bool (default ()) name;  | 
165  | 
fun default_int name = int (default ()) name;  | 
|
166  | 
fun default_real name = real (default ()) name;  | 
|
167  | 
fun default_string name = string (default ()) name;  | 
|
168  | 
||
| 51943 | 169  | 
val default_put_bool = change_default put_bool;  | 
170  | 
val default_put_int = change_default put_int;  | 
|
171  | 
val default_put_real = change_default put_real;  | 
|
172  | 
val default_put_string = change_default put_string;  | 
|
173  | 
||
174  | 
fun get_default name =  | 
|
175  | 
let val options = default () in get (typ options name) SOME options name end;  | 
|
176  | 
val put_default = change_default update;  | 
|
177  | 
||
| 
48698
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
178  | 
fun set_default options = Synchronized.change global_default (K (SOME options));  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
179  | 
fun reset_default () = Synchronized.change global_default (K NONE);  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
180  | 
|
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
181  | 
fun load_default () =  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
182  | 
(case getenv "ISABELLE_PROCESS_OPTIONS" of  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
183  | 
"" => ()  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
184  | 
| name =>  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
185  | 
let val path = Path.explode name in  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
186  | 
(case try File.read path of  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
187  | 
SOME s => (set_default (decode (YXML.parse_body s)); ignore (try File.rm path))  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
188  | 
| NONE => ())  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
189  | 
end);  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
190  | 
|
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
191  | 
val _ = load_default ();  | 
| 
 
2585042b1a30
pass Isabelle/Scala system options into ML process of Isar tty or build jobs;
 
wenzelm 
parents: 
48456 
diff
changeset
 | 
192  | 
|
| 
48456
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
end;  | 
| 
 
d8ff14f44a40
added ML version of stand-alone options, with XML.encode/decode operations (unidirectional from Scala to ML);
 
wenzelm 
parents:  
diff
changeset
 | 
194  |