berghofe@3604: (* Title: Pure/Thy/thy_info.ML berghofe@3604: ID: $Id$ wenzelm@6211: Author: Markus Wenzel, TU Muenchen wenzelm@8805: License: GPL (GNU GENERAL PUBLIC LICENSE) berghofe@3604: wenzelm@6362: Theory loader database, including theory and file dependencies. wenzelm@3976: *) berghofe@3604: wenzelm@5209: signature BASIC_THY_INFO = wenzelm@5209: sig wenzelm@5209: val theory: string -> theory wenzelm@6211: val theory_of_sign: Sign.sg -> theory wenzelm@6211: val theory_of_thm: thm -> theory wenzelm@6241: (*val use: string -> unit*) (*exported later*) wenzelm@6241: val time_use: string -> unit wenzelm@6211: val touch_thy: string -> unit wenzelm@6211: val use_thy: string -> unit wenzelm@6211: val update_thy: string -> unit wenzelm@6666: val remove_thy: string -> unit wenzelm@6211: val time_use_thy: string -> unit wenzelm@6527: val use_thy_only: string -> unit wenzelm@7099: val update_thy_only: string -> unit wenzelm@6211: end; wenzelm@5209: berghofe@3604: signature THY_INFO = berghofe@3604: sig wenzelm@5209: include BASIC_THY_INFO wenzelm@7099: datatype action = Update | Outdate | Remove wenzelm@7099: val str_of_action: action -> string wenzelm@7099: val add_hook: (action -> string -> unit) -> unit wenzelm@6211: val names: unit -> string list wenzelm@7910: val known_thy: string -> bool wenzelm@7935: val check_known_thy: string -> bool wenzelm@7935: val if_known_thy: (string -> unit) -> string -> unit wenzelm@7288: val lookup_theory: string -> theory option wenzelm@6211: val get_theory: string -> theory berghofe@6654: val get_preds: string -> string list wenzelm@7935: val loaded_files: string -> Path.T list wenzelm@7910: val touch_child_thys: string -> unit wenzelm@7099: val touch_all_thys: unit -> unit wenzelm@7941: val load_file: bool -> Path.T -> unit wenzelm@6241: val use: string -> unit wenzelm@7952: val quiet_update_thy: bool -> string -> unit wenzelm@8805: val begin_theory: (string -> string list -> (Path.T * bool) list -> theory -> theory) -> wenzelm@8805: bool -> string -> string list -> (Path.T * bool) list -> theory wenzelm@6211: val end_theory: theory -> theory wenzelm@6666: val finish: unit -> unit wenzelm@6211: val register_theory: theory -> unit berghofe@3604: end; berghofe@3604: wenzelm@6362: structure ThyInfo: THY_INFO = wenzelm@6211: struct berghofe@3604: berghofe@3604: wenzelm@7099: (** theory loader actions and hooks **) wenzelm@7099: wenzelm@7099: datatype action = Update | Outdate | Remove; wenzelm@7099: val str_of_action = fn Update => "Update" | Outdate => "Outdate" | Remove => "Remove"; wenzelm@7099: wenzelm@7099: local wenzelm@7099: val hooks = ref ([]: (action -> string -> unit) list); wenzelm@7099: in wenzelm@7099: fun add_hook f = hooks := f :: ! hooks; wenzelm@7191: fun perform action name = seq (fn f => (try (fn () => f action name) (); ())) (! hooks); wenzelm@7099: end; wenzelm@7099: wenzelm@7099: wenzelm@7099: wenzelm@6211: (** thy database **) wenzelm@6211: wenzelm@6211: (* messages *) wenzelm@6211: wenzelm@6211: fun gen_msg txt [] = txt wenzelm@6211: | gen_msg txt names = txt ^ " " ^ commas_quote names; wenzelm@6211: wenzelm@6211: fun loader_msg txt names = gen_msg ("Theory loader: " ^ txt) names; wenzelm@6211: wenzelm@6211: val show_path = space_implode " via " o map quote; wenzelm@6211: fun cycle_msg name names = loader_msg ("cyclic dependency of " ^ show_path (name :: names)) []; berghofe@3604: berghofe@3604: wenzelm@6666: (* derived graph operations *) berghofe@3604: wenzelm@6211: fun add_deps name parents G = wenzelm@6211: foldl (fn (H, parent) => Graph.add_edge_acyclic (parent, name) H) (G, parents) wenzelm@6211: handle Graph.CYCLES namess => error (cat_lines (map (cycle_msg name) namess)); berghofe@3604: wenzelm@7952: fun upd_deps name entry G = wenzelm@7952: foldl (fn (H, parent) => Graph.del_edge (parent, name) H) (G, Graph.imm_preds G name) wenzelm@7952: |> Graph.map_node name (K entry); wenzelm@3976: wenzelm@6211: fun update_node name parents entry G = wenzelm@7952: (if can (Graph.get_node G) name then upd_deps name entry G else Graph.new_node (name, entry) G) wenzelm@7952: |> add_deps name parents; wenzelm@7952: berghofe@3604: berghofe@3604: wenzelm@6211: (* thy database *) wenzelm@6211: wenzelm@6211: type deps = wenzelm@6211: {present: bool, outdated: bool, wenzelm@7941: master: ThyLoad.master option, files: (Path.T * (Path.T * File.info) option) list}; wenzelm@6211: wenzelm@6241: fun make_deps present outdated master files = wenzelm@6241: {present = present, outdated = outdated, master = master, files = files}: deps; berghofe@3604: wenzelm@7211: fun init_deps master files = Some (make_deps false true master (map (rpair None) files)); wenzelm@7211: wenzelm@7211: wenzelm@6362: type thy = deps option * theory option; berghofe@3604: wenzelm@6211: local wenzelm@6362: val database = ref (Graph.empty: thy Graph.T); wenzelm@6211: in wenzelm@6362: fun get_thys () = ! database; wenzelm@6362: fun change_thys f = database := (f (! database)); berghofe@3604: end; wenzelm@5209: wenzelm@5209: wenzelm@6211: (* access thy graph *) wenzelm@6211: wenzelm@6211: fun thy_graph f x = f (get_thys ()) x; wenzelm@6666: fun get_names () = Graph.keys (get_thys ()); wenzelm@6211: wenzelm@6211: wenzelm@6211: (* access thy *) wenzelm@6211: wenzelm@7935: fun lookup_thy name = wenzelm@7935: Some (thy_graph Graph.get_node name) handle Graph.UNDEFINED _ => None; wenzelm@7935: wenzelm@7910: val known_thy = is_some o lookup_thy; wenzelm@7935: fun check_known_thy name = known_thy name orelse (warning ("Unknown theory " ^ quote name); false); wenzelm@7935: fun if_known_thy f name = if check_known_thy name then f name else (); wenzelm@6211: wenzelm@6211: fun get_thy name = wenzelm@6211: (case lookup_thy name of wenzelm@6211: Some thy => thy wenzelm@6211: | None => error (loader_msg "nothing known about theory" [name])); wenzelm@6211: wenzelm@6211: fun change_thy name f = (get_thy name; change_thys (Graph.map_node name f)); wenzelm@6211: wenzelm@6211: wenzelm@6211: (* access deps *) wenzelm@6211: wenzelm@6211: val lookup_deps = apsome #1 o lookup_thy; wenzelm@6211: val get_deps = #1 o get_thy; wenzelm@6211: fun change_deps name f = change_thy name (fn (deps, x) => (f deps, x)); wenzelm@6211: wenzelm@6666: fun is_finished name = is_none (get_deps name); wenzelm@6211: fun get_files name = (case get_deps name of Some {files, ...} => files | _ => []); wenzelm@7191: wenzelm@7191: fun loaded_files name = wenzelm@7191: (case get_deps name of wenzelm@7935: None => [] wenzelm@7910: | Some {master, files, ...} => wenzelm@7935: (case master of Some m => [#1 (ThyLoad.get_thy m)] | None => []) @ wenzelm@7941: (mapfilter (apsome #1 o #2) files)); wenzelm@6211: berghofe@6654: fun get_preds name = berghofe@6654: (thy_graph Graph.imm_preds name) handle Graph.UNDEFINED _ => berghofe@6654: error (loader_msg "nothing known about theory" [name]); berghofe@6654: wenzelm@6211: wenzelm@6211: (* access theory *) wenzelm@6211: wenzelm@7687: fun lookup_theory name = wenzelm@7687: (case lookup_thy name of wenzelm@7687: Some (_, Some thy) => Some thy wenzelm@7687: | _ => None); wenzelm@7288: wenzelm@6211: fun get_theory name = wenzelm@7288: (case lookup_theory name of wenzelm@7288: (Some theory) => theory wenzelm@6211: | _ => error (loader_msg "undefined theory entry for" [name])); wenzelm@6211: wenzelm@6211: val theory_of_sign = get_theory o Sign.name_of; wenzelm@6211: val theory_of_thm = theory_of_sign o Thm.sign_of_thm; wenzelm@6211: wenzelm@7099: fun put_theory name theory = change_thy name (fn (deps, _) => (deps, Some theory)); wenzelm@6211: wenzelm@6211: wenzelm@6211: wenzelm@6211: (** thy operations **) wenzelm@6211: wenzelm@6241: (* maintain 'outdated' flag *) wenzelm@6211: wenzelm@7099: local wenzelm@7099: wenzelm@6241: fun is_outdated name = wenzelm@6241: (case lookup_deps name of wenzelm@6241: Some (Some {outdated, ...}) => outdated wenzelm@6241: | _ => false); wenzelm@6211: wenzelm@6241: fun outdate_thy name = wenzelm@7099: if is_finished name orelse is_outdated name then () wenzelm@7099: else (change_deps name (apsome (fn {present, outdated = _, master, files} => wenzelm@7099: make_deps present true master files)); perform Outdate name); wenzelm@7099: wenzelm@7910: fun check_unfinished name = wenzelm@7910: if is_finished name then (warning (loader_msg "tried to touch finished theory" [name]); None) wenzelm@7910: else Some name; wenzelm@7910: wenzelm@7099: in wenzelm@6211: wenzelm@7910: fun touch_thys names = wenzelm@7910: seq outdate_thy (thy_graph Graph.all_succs (mapfilter check_unfinished names)); wenzelm@7910: wenzelm@7910: fun touch_thy name = touch_thys [name]; wenzelm@7910: fun touch_child_thys name = touch_thys (thy_graph Graph.imm_succs name); wenzelm@6211: wenzelm@7099: fun touch_all_thys () = seq outdate_thy (get_names ()); wenzelm@6211: wenzelm@7099: end; wenzelm@6211: wenzelm@6211: wenzelm@7244: (* check 'finished' state *) wenzelm@7244: wenzelm@7244: fun check_unfinished fail name = wenzelm@7910: if known_thy name andalso is_finished name then wenzelm@7288: fail (loader_msg "cannot update finished theory" [name]) wenzelm@7244: else (); wenzelm@7244: wenzelm@7244: wenzelm@7941: (* load_file *) wenzelm@6211: wenzelm@7191: local wenzelm@7191: wenzelm@7941: fun provide path name info (deps as Some {present, outdated, master, files}) = wenzelm@7941: (if exists (equal path o #1) files then () wenzelm@7941: else warning (loader_msg "undeclared dependency of theory" [name] ^ wenzelm@7941: " on file: " ^ quote (Path.pack path)); wenzelm@7941: Some (make_deps present outdated master (overwrite (files, (path, Some info))))) wenzelm@7941: | provide _ _ _ None = None; wenzelm@7941: wenzelm@7941: fun run_file path = wenzelm@7941: (case apsome PureThy.get_name (Context.get_context ()) of wenzelm@7941: None => (ThyLoad.load_file path; ()) wenzelm@7941: | Some name => wenzelm@7941: if known_thy name then change_deps name (provide path name (ThyLoad.load_file path)) wenzelm@7941: else (ThyLoad.load_file path; ())); wenzelm@6211: wenzelm@7191: in wenzelm@7191: wenzelm@7941: fun load_file time path = wenzelm@7941: if time then wenzelm@7244: let val name = Path.pack path in wenzelm@7244: timeit (fn () => wenzelm@7244: (writeln ("\n**** Starting file " ^ quote name ^ " ****"); wenzelm@7941: run_file path; wenzelm@7244: writeln ("**** Finished file " ^ quote name ^ " ****\n"))) wenzelm@7244: end wenzelm@7941: else run_file path; wenzelm@7941: wenzelm@7941: val use = load_file false o Path.unpack; wenzelm@7941: val time_use = load_file true o Path.unpack; wenzelm@7191: wenzelm@7191: end; wenzelm@6233: wenzelm@6233: wenzelm@7099: (* load_thy *) wenzelm@7099: wenzelm@7099: fun required_by [] = "" wenzelm@7099: | required_by initiators = " (required by " ^ show_path (rev initiators) ^ ")"; wenzelm@7099: wenzelm@7099: fun load_thy ml time initiators dir name parents = wenzelm@7099: let wenzelm@7099: val _ = writeln ("Loading theory " ^ quote name ^ required_by initiators); wenzelm@7099: wenzelm@7099: val _ = touch_thy name; wenzelm@7099: val master = ThyLoad.load_thy dir name ml time; wenzelm@7099: wenzelm@7099: val files = get_files name; wenzelm@7099: val missing_files = mapfilter (fn (path, None) => Some (Path.pack path) | _ => None) files; wenzelm@7099: in wenzelm@7099: if null missing_files then () wenzelm@7099: else warning (loader_msg "unresolved dependencies of theory" [name] ^ wenzelm@7223: " on file(s): " ^ commas_quote missing_files); wenzelm@7191: change_deps name (fn _ => Some (make_deps true false (Some master) files)); wenzelm@7099: perform Update name wenzelm@7099: end; wenzelm@7099: wenzelm@7099: wenzelm@6211: (* require_thy *) wenzelm@6211: wenzelm@7211: local wenzelm@6211: wenzelm@7941: fun load_deps ml dir name = wenzelm@7941: let val (master, (parents, files)) = ThyLoad.deps_thy dir name ml wenzelm@7191: in (Some (init_deps (Some master) files), parents) end; wenzelm@6211: wenzelm@6211: fun file_current (_, None) = false wenzelm@7941: | file_current (path, info) = (info = ThyLoad.check_file path); wenzelm@6211: wenzelm@7941: fun current_deps ml strict dir name = wenzelm@6211: (case lookup_deps name of wenzelm@7941: None => (false, load_deps ml dir name) wenzelm@6211: | Some deps => wenzelm@6211: let val same_deps = (None, thy_graph Graph.imm_preds name) in wenzelm@6211: (case deps of wenzelm@6211: None => (true, same_deps) wenzelm@6211: | Some {present, outdated, master, files} => wenzelm@6211: if present andalso not strict then (true, same_deps) wenzelm@6211: else wenzelm@7941: let val master' = Some (ThyLoad.check_thy dir name ml) in wenzelm@7941: if master <> master' then (false, load_deps ml dir name) wenzelm@6211: else (not outdated andalso forall file_current files, same_deps) wenzelm@6211: end) wenzelm@6211: end); wenzelm@6211: wenzelm@7066: fun require_thy ml time strict keep_strict initiators prfx (visited, str) = wenzelm@6211: let wenzelm@7066: val path = Path.expand (Path.unpack str); wenzelm@6484: val name = Path.pack (Path.base path); wenzelm@7066: in wenzelm@7589: if name mem_string initiators then error (cycle_msg name initiators) else (); wenzelm@7952: if known_thy name andalso is_finished name orelse name mem_string visited then wenzelm@7952: (visited, (true, name)) wenzelm@7066: else wenzelm@7066: let wenzelm@7066: val dir = Path.append prfx (Path.dir path); wenzelm@7066: val req_parent = wenzelm@7211: require_thy true time (strict andalso keep_strict) keep_strict (name :: initiators) dir; wenzelm@6484: wenzelm@7941: val (current, (new_deps, parents)) = current_deps ml strict dir name handle ERROR => wenzelm@7099: error (loader_msg "the error(s) above occurred while examining theory" [name] ^ wenzelm@7066: (if null initiators then "" else "\n" ^ required_by initiators)); wenzelm@7066: val (visited', parent_results) = foldl_map req_parent (name :: visited, parents); wenzelm@6211: wenzelm@7066: val result = wenzelm@7066: if current andalso forall #1 parent_results then true wenzelm@7066: else wenzelm@7066: ((case new_deps of wenzelm@7099: Some deps => change_thys (update_node name parents (deps, None)) wenzelm@7066: | None => ()); wenzelm@7099: load_thy ml time initiators dir name parents; wenzelm@7099: false); wenzelm@7066: in (visited', (result, name)) end wenzelm@7066: end; wenzelm@6484: wenzelm@7066: fun gen_use_thy req s = wenzelm@7066: let val (_, (_, name)) = req [] Path.current ([], s) wenzelm@7066: in Context.context (get_theory name) end; wenzelm@6241: wenzelm@7244: fun warn_finished f name = (check_unfinished warning name; f name); wenzelm@7244: wenzelm@7211: in wenzelm@7211: wenzelm@7952: val weak_use_thy = gen_use_thy (require_thy true false false false); wenzelm@7952: fun quiet_update_thy ml = gen_use_thy (require_thy ml false true true); wenzelm@7244: wenzelm@7244: val use_thy = warn_finished (gen_use_thy (require_thy true false true false)); wenzelm@7244: val time_use_thy = warn_finished (gen_use_thy (require_thy true true true false)); wenzelm@7244: val use_thy_only = warn_finished (gen_use_thy (require_thy false false true false)); wenzelm@7244: val update_thy = warn_finished (gen_use_thy (require_thy true false true true)); wenzelm@7244: val update_thy_only = warn_finished (gen_use_thy (require_thy false false true true)); wenzelm@6211: wenzelm@7211: end; wenzelm@7211: wenzelm@6211: wenzelm@6666: (* remove_thy *) wenzelm@6666: wenzelm@6666: fun remove_thy name = wenzelm@7910: if is_finished name then error (loader_msg "cannot remove finished theory" [name]) wenzelm@6666: else wenzelm@6666: let val succs = thy_graph Graph.all_succs [name] in wenzelm@6666: writeln (loader_msg "removing" succs); wenzelm@7191: seq (perform Remove) succs; wenzelm@7191: change_thys (Graph.del_nodes succs) wenzelm@6666: end; wenzelm@6666: wenzelm@6666: wenzelm@7066: (* begin / end theory *) wenzelm@6211: wenzelm@8805: fun begin_theory present upd name parents paths = wenzelm@6211: let wenzelm@8805: val assert_thy = if upd then quiet_update_thy true else weak_use_thy; wenzelm@7244: val _ = check_unfinished error name; wenzelm@7941: val _ = (map Path.basic parents; seq assert_thy parents); wenzelm@6211: val theory = PureThy.begin_theory name (map get_theory parents); wenzelm@7952: val deps = wenzelm@7952: if known_thy name then get_deps name wenzelm@7952: else (init_deps None (map #1 paths)); (*records additional ML files only!*) wenzelm@7952: val _ = change_thys (update_node name parents (deps, Some theory)); wenzelm@6329: val use_paths = mapfilter (fn (x, true) => Some x | _ => None) paths; wenzelm@8805: val theory' = theory |> present name parents paths; wenzelm@8805: val ((), theory'') = Context.pass_theory theory' (seq (load_file false)) use_paths; wenzelm@8805: in theory'' end; wenzelm@7244: wenzelm@6211: fun end_theory theory = wenzelm@7099: let wenzelm@7099: val theory' = PureThy.end_theory theory; wenzelm@7099: val name = PureThy.get_name theory; wenzelm@7099: in put_theory name theory'; theory' end; wenzelm@6211: wenzelm@6211: wenzelm@6666: (* finish all theories *) wenzelm@6211: wenzelm@6666: fun finish () = change_thys (Graph.map_nodes (fn (_, entry) => (None, entry))); wenzelm@6211: wenzelm@6211: wenzelm@6211: (* register existing theories *) wenzelm@6211: wenzelm@6211: fun register_theory theory = wenzelm@6211: let wenzelm@6211: val name = PureThy.get_name theory; wenzelm@6211: val parents = Theory.parents_of theory; wenzelm@6211: val parent_names = map PureThy.get_name parents; wenzelm@6211: wenzelm@6211: fun err txt bads = wenzelm@6211: error (loader_msg txt bads ^ "\n" ^ gen_msg "cannot register theory" [name]); wenzelm@6211: wenzelm@6666: val nonfinished = filter_out is_finished parent_names; wenzelm@6211: fun get_variant (x, y_name) = wenzelm@6211: if Theory.eq_thy (x, get_theory y_name) then None wenzelm@6211: else Some y_name; wenzelm@6211: val variants = mapfilter get_variant (parents ~~ parent_names); wenzelm@6211: wenzelm@6211: fun register G = wenzelm@6362: (Graph.new_node (name, (None, Some theory)) G wenzelm@6211: handle Graph.DUPLICATE _ => err "duplicate theory entry" []) wenzelm@6211: |> add_deps name parent_names; wenzelm@6211: in wenzelm@6666: if not (null nonfinished) then err "non-finished parent theories" nonfinished wenzelm@6211: else if not (null variants) then err "different versions of parent theories" variants wenzelm@7099: else (change_thys register; perform Update name) wenzelm@6211: end; wenzelm@6211: wenzelm@6211: wenzelm@6211: (*final declarations of this structure*) wenzelm@6211: val theory = get_theory; wenzelm@6211: val names = get_names; wenzelm@6211: wenzelm@6211: end; wenzelm@6211: wenzelm@5209: structure BasicThyInfo: BASIC_THY_INFO = ThyInfo; wenzelm@5209: open BasicThyInfo;