(* Title: Pure/Isar/session.ML
ID: $Id$
Author: Markus Wenzel, TU Muenchen
Session management -- maintain state of logic images.
*)
signature SESSION =
sig
val name: unit -> string
val welcome: unit -> string
val use_dir: string -> bool -> string list -> bool -> bool -> string -> bool -> string
-> string -> string -> string -> int -> bool -> bool -> unit
val finish: unit -> unit
end;
structure Session: SESSION =
struct
(* session state *)
val session = ref ([Sign.PureN]: string list);
val session_path = ref ([]: string list);
val session_finished = ref false;
val remote_path = ref (NONE: Url.T option);
(* access path *)
fun path () = ! session_path;
fun str_of [] = Sign.PureN
| str_of elems = space_implode "/" elems;
fun name () = "Isabelle/" ^ str_of (path ());
fun welcome () = "Welcome to " ^ name () ^ " (" ^ version ^ ")";
(* add_path *)
fun add_path reset s =
let val sess = ! session @ [s] in
(case Library.duplicates sess of
[] => (session := sess; session_path := ((if reset then [] else ! session_path) @ [s]))
| dups => error ("Duplicate session identifiers " ^ commas_quote dups ^ " in " ^ str_of sess))
end;
(* init *)
fun init reset parent name =
if not (parent mem_string (! session)) orelse not (! session_finished) then
error ("Unfinished parent session " ^ quote parent ^ " for " ^ quote name)
else (add_path reset name; session_finished := false);
(* finish *)
fun finish () =
(ThyInfo.finish ();
Present.finish ();
session_finished := true);
(* use_dir *)
fun get_rpath rpath_str =
(if rpath_str = "" then () else
if is_some (! remote_path) then
error "Path for remote theory browsing information may only be set once"
else
remote_path := SOME (Url.unpack rpath_str);
(! remote_path, rpath_str <> ""));
fun use_dir root build modes reset info doc doc_graph parent name dump rpath_str level verbose hide =
Library.setmp print_mode (modes @ ! print_mode)
(Library.setmp Proofterm.proofs level (Library.setmp IsarOutput.hide_commands hide (fn () =>
(init reset parent name;
Present.init build info doc doc_graph (path ()) name
(if dump = "" then NONE else SOME (Path.unpack dump)) (get_rpath rpath_str) verbose;
File.use (Path.basic root);
finish ())))) ()
handle exn => (writeln (Toplevel.exn_message exn); exit 1);
end;