| author | huffman | 
| Fri, 29 May 2009 14:09:58 -0700 | |
| changeset 31343 | 9983f648f9bb | 
| parent 30173 | eabece26b89b | 
| child 31478 | 5e412e4c6546 | 
| permissions | -rw-r--r-- | 
| 
30173
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
1  | 
(* Title: Pure/System/session.ML  | 
| 6346 | 2  | 
Author: Markus Wenzel, TU Muenchen  | 
3  | 
||
4  | 
Session management -- maintain state of logic images.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature SESSION =  | 
|
8  | 
sig  | 
|
| 25840 | 9  | 
val id: unit -> string list  | 
| 11509 | 10  | 
val name: unit -> string  | 
| 6346 | 11  | 
val welcome: unit -> string  | 
| 24061 | 12  | 
val use_dir: string -> bool -> string list -> bool -> bool -> string -> bool -> string list ->  | 
| 
29435
 
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
 
wenzelm 
parents: 
28207 
diff
changeset
 | 
13  | 
string -> string -> bool * string -> string -> int -> bool -> int -> int -> bool -> unit  | 
| 6346 | 14  | 
val finish: unit -> unit  | 
15  | 
end;  | 
|
16  | 
||
17  | 
structure Session: SESSION =  | 
|
18  | 
struct  | 
|
19  | 
||
20  | 
||
21  | 
(* session state *)  | 
|
22  | 
||
| 16451 | 23  | 
val session = ref ([Context.PureN]: string list);  | 
| 6346 | 24  | 
val session_path = ref ([]: string list);  | 
25  | 
val session_finished = ref false;  | 
|
| 15979 | 26  | 
val remote_path = ref (NONE: Url.T option);  | 
| 9414 | 27  | 
|
28  | 
||
29  | 
(* access path *)  | 
|
| 6346 | 30  | 
|
| 25840 | 31  | 
fun id () = ! session;  | 
| 6346 | 32  | 
fun path () = ! session_path;  | 
33  | 
||
| 16451 | 34  | 
fun str_of [] = Context.PureN  | 
| 6346 | 35  | 
| str_of elems = space_implode "/" elems;  | 
36  | 
||
| 11509 | 37  | 
fun name () = "Isabelle/" ^ str_of (path ());  | 
| 
26109
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25840 
diff
changeset
 | 
38  | 
|
| 
30173
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
39  | 
|
| 
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
40  | 
(* welcome *)  | 
| 
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
41  | 
|
| 
26109
 
c69c3559355b
more elaborate structure Distribution (filled-in by makedist);
 
wenzelm 
parents: 
25840 
diff
changeset
 | 
42  | 
fun welcome () =  | 
| 26134 | 43  | 
if Distribution.is_official then  | 
44  | 
    "Welcome to " ^ name () ^ " (" ^ Distribution.version ^ ")"
 | 
|
45  | 
else  | 
|
46  | 
    "Unofficial version of " ^ name () ^ " (" ^ Distribution.version ^ ")" ^
 | 
|
| 27643 | 47  | 
(if Distribution.changelog <> "" then "\nSee also " ^ Distribution.changelog else "");  | 
| 6346 | 48  | 
|
| 
30173
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
49  | 
val _ =  | 
| 
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
50  | 
OuterSyntax.improper_command "welcome" "print welcome message" OuterKeyword.diag  | 
| 
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
51  | 
(Scan.succeed (Toplevel.no_timing o Toplevel.imperative (writeln o welcome)));  | 
| 
 
eabece26b89b
moved isabelle_process.ML, isabelle_process.scala, isar.ML, session.ML to Pure/System/ (together with associated Isar commands);
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
52  | 
|
| 6346 | 53  | 
|
| 9414 | 54  | 
(* add_path *)  | 
55  | 
||
56  | 
fun add_path reset s =  | 
|
57  | 
let val sess = ! session @ [s] in  | 
|
| 18964 | 58  | 
(case duplicates (op =) sess of  | 
| 9414 | 59  | 
[] => (session := sess; session_path := ((if reset then [] else ! session_path) @ [s]))  | 
60  | 
    | dups => error ("Duplicate session identifiers " ^ commas_quote dups ^ " in " ^ str_of sess))
 | 
|
61  | 
end;  | 
|
62  | 
||
63  | 
||
| 6346 | 64  | 
(* init *)  | 
65  | 
||
66  | 
fun init reset parent name =  | 
|
| 20664 | 67  | 
if not (member (op =) (! session) parent) orelse not (! session_finished) then  | 
| 6346 | 68  | 
    error ("Unfinished parent session " ^ quote parent ^ " for " ^ quote name)
 | 
69  | 
else (add_path reset name; session_finished := false);  | 
|
70  | 
||
71  | 
||
72  | 
(* finish *)  | 
|
73  | 
||
74  | 
fun finish () =  | 
|
| 16728 | 75  | 
(Output.accumulated_time ();  | 
76  | 
ThyInfo.finish ();  | 
|
| 28207 | 77  | 
Present.finish ();  | 
| 28205 | 78  | 
Future.shutdown ();  | 
| 6346 | 79  | 
session_finished := true);  | 
80  | 
||
81  | 
||
82  | 
(* use_dir *)  | 
|
83  | 
||
| 17207 | 84  | 
fun get_rpath rpath =  | 
85  | 
(if rpath = "" then () else  | 
|
| 15979 | 86  | 
if is_some (! remote_path) then  | 
| 
7227
 
a8e86b8e6fd1
Path for remote theory browsing information is now stored in referece variable rpath.
 
berghofe 
parents: 
6663 
diff
changeset
 | 
87  | 
error "Path for remote theory browsing information may only be set once"  | 
| 
 
a8e86b8e6fd1
Path for remote theory browsing information is now stored in referece variable rpath.
 
berghofe 
parents: 
6663 
diff
changeset
 | 
88  | 
else  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
89  | 
remote_path := SOME (Url.explode rpath);  | 
| 17207 | 90  | 
(! remote_path, rpath <> ""));  | 
91  | 
||
92  | 
fun dumping (_, "") = NONE  | 
|
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
20664 
diff
changeset
 | 
93  | 
| dumping (cp, path) = SOME (cp, Path.explode path);  | 
| 
7236
 
e077484d50d8
Better handling of path for remote theory browsing information.
 
berghofe 
parents: 
7227 
diff
changeset
 | 
94  | 
|
| 17074 | 95  | 
fun use_dir root build modes reset info doc doc_graph doc_versions  | 
| 
29435
 
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
 
wenzelm 
parents: 
28207 
diff
changeset
 | 
96  | 
parent name dump rpath level verbose max_threads trace_threads parallel_proofs =  | 
| 23972 | 97  | 
((fn () =>  | 
98  | 
(init reset parent name;  | 
|
99  | 
Present.init build info doc doc_graph doc_versions (path ()) name  | 
|
| 26612 | 100  | 
(dumping dump) (get_rpath rpath) verbose (map ThyInfo.get_theory (ThyInfo.get_names ()));  | 
| 
25703
 
832073e402ae
removed obsolete use_noncritical (plain use is already non-critical);
 
wenzelm 
parents: 
24612 
diff
changeset
 | 
101  | 
use root;  | 
| 23972 | 102  | 
finish ()))  | 
103  | 
|> setmp_noncritical Proofterm.proofs level  | 
|
| 24612 | 104  | 
|> setmp_noncritical print_mode (modes @ print_mode_value ())  | 
| 
29435
 
a5f84ac14609
added parallel_proofs flag (default true, cf. usedir option -Q), which can be disabled in low-memory situations;
 
wenzelm 
parents: 
28207 
diff
changeset
 | 
105  | 
|> setmp_noncritical Goal.parallel_proofs parallel_proofs  | 
| 24061 | 106  | 
|> setmp_noncritical Multithreading.trace trace_threads  | 
| 23979 | 107  | 
|> setmp_noncritical Multithreading.max_threads  | 
108  | 
(if Multithreading.available then max_threads  | 
|
109  | 
else (if max_threads = 1 then () else warning "Multithreading support unavailable"; 1))) ()  | 
|
| 18683 | 110  | 
handle exn => (Output.error_msg (Toplevel.exn_message exn); exit 1);  | 
| 6346 | 111  | 
|
112  | 
end;  |