| author | huffman | 
| Tue, 08 Mar 2005 00:18:22 +0100 | |
| changeset 15591 | 50c3384ca6c4 | 
| parent 15570 | 8d8c70b41bab | 
| child 15633 | 741deccec4e3 | 
| permissions | -rw-r--r-- | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
1  | 
(* Title: Pure/Thy/thy_info.ML  | 
| 
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
3  | 
Author: Markus Wenzel, TU Muenchen  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
4  | 
|
| 6362 | 5  | 
Theory loader database, including theory and file dependencies.  | 
| 3976 | 6  | 
*)  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
7  | 
|
| 5209 | 8  | 
signature BASIC_THY_INFO =  | 
9  | 
sig  | 
|
10  | 
val theory: string -> theory  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
11  | 
val theory_of_sign: Sign.sg -> theory  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
12  | 
val theory_of_thm: thm -> theory  | 
| 6241 | 13  | 
(*val use: string -> unit*) (*exported later*)  | 
14  | 
val time_use: string -> unit  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
15  | 
val touch_thy: string -> unit  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
16  | 
val use_thy: string -> unit  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
17  | 
val update_thy: string -> unit  | 
| 6666 | 18  | 
val remove_thy: string -> unit  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
19  | 
val time_use_thy: string -> unit  | 
| 6527 | 20  | 
val use_thy_only: string -> unit  | 
| 7099 | 21  | 
val update_thy_only: string -> unit  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
22  | 
end;  | 
| 5209 | 23  | 
|
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
24  | 
signature THY_INFO =  | 
| 
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
25  | 
sig  | 
| 5209 | 26  | 
include BASIC_THY_INFO  | 
| 7099 | 27  | 
datatype action = Update | Outdate | Remove  | 
28  | 
val str_of_action: action -> string  | 
|
29  | 
val add_hook: (action -> string -> unit) -> unit  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
30  | 
val names: unit -> string list  | 
| 7910 | 31  | 
val known_thy: string -> bool  | 
| 7935 | 32  | 
val check_known_thy: string -> bool  | 
33  | 
val if_known_thy: (string -> unit) -> string -> unit  | 
|
| 7288 | 34  | 
val lookup_theory: string -> theory option  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
35  | 
val get_theory: string -> theory  | 
| 6654 | 36  | 
val get_preds: string -> string list  | 
| 7935 | 37  | 
val loaded_files: string -> Path.T list  | 
| 7910 | 38  | 
val touch_child_thys: string -> unit  | 
| 7099 | 39  | 
val touch_all_thys: unit -> unit  | 
| 7941 | 40  | 
val load_file: bool -> Path.T -> unit  | 
| 6241 | 41  | 
val use: string -> unit  | 
| 7952 | 42  | 
val quiet_update_thy: bool -> string -> unit  | 
| 9822 | 43  | 
val pretend_use_thy_only: string -> unit  | 
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
44  | 
val begin_theory: (Path.T option -> string -> string list ->  | 
| 
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
45  | 
(Path.T * bool) list -> theory -> theory) ->  | 
| 8805 | 46  | 
bool -> string -> string list -> (Path.T * bool) list -> theory  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
47  | 
val end_theory: theory -> theory  | 
| 6666 | 48  | 
val finish: unit -> unit  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
49  | 
val register_theory: theory -> unit  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
50  | 
end;  | 
| 
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
51  | 
|
| 6362 | 52  | 
structure ThyInfo: THY_INFO =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
53  | 
struct  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
55  | 
|
| 7099 | 56  | 
(** theory loader actions and hooks **)  | 
57  | 
||
58  | 
datatype action = Update | Outdate | Remove;  | 
|
59  | 
val str_of_action = fn Update => "Update" | Outdate => "Outdate" | Remove => "Remove";  | 
|
60  | 
||
61  | 
local  | 
|
62  | 
val hooks = ref ([]: (action -> string -> unit) list);  | 
|
63  | 
in  | 
|
64  | 
fun add_hook f = hooks := f :: ! hooks;  | 
|
| 15570 | 65  | 
fun perform action name = List.app (fn f => (try (fn () => f action name) (); ())) (! hooks);  | 
| 7099 | 66  | 
end;  | 
67  | 
||
68  | 
||
69  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
70  | 
(** thy database **)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
71  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
72  | 
(* messages *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
73  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
74  | 
fun gen_msg txt [] = txt  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
75  | 
| gen_msg txt names = txt ^ " " ^ commas_quote names;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
76  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
77  | 
fun loader_msg txt names = gen_msg ("Theory loader: " ^ txt) names;
 | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
78  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
79  | 
val show_path = space_implode " via " o map quote;  | 
| 9332 | 80  | 
fun cycle_msg names = loader_msg ("cyclic dependency of " ^ show_path names) [];
 | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
82  | 
|
| 6666 | 83  | 
(* derived graph operations *)  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
84  | 
|
| 9327 | 85  | 
fun add_deps name parents G = Graph.add_deps_acyclic (name, parents) G  | 
| 9332 | 86  | 
handle Graph.CYCLES namess => error (cat_lines (map cycle_msg namess));  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
87  | 
|
| 7952 | 88  | 
fun upd_deps name entry G =  | 
| 15570 | 89  | 
Library.foldl (fn (H, parent) => Graph.del_edge (parent, name) H) (G, Graph.imm_preds G name)  | 
| 7952 | 90  | 
|> Graph.map_node name (K entry);  | 
| 3976 | 91  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
92  | 
fun update_node name parents entry G =  | 
| 7952 | 93  | 
(if can (Graph.get_node G) name then upd_deps name entry G else Graph.new_node (name, entry) G)  | 
94  | 
|> add_deps name parents;  | 
|
95  | 
||
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
96  | 
|
| 
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
97  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
98  | 
(* thy database *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
99  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
100  | 
type deps =  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
101  | 
  {present: bool, outdated: bool,
 | 
| 7941 | 102  | 
master: ThyLoad.master option, files: (Path.T * (Path.T * File.info) option) list};  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
103  | 
|
| 6241 | 104  | 
fun make_deps present outdated master files =  | 
105  | 
  {present = present, outdated = outdated, master = master, files = files}: deps;
 | 
|
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
106  | 
|
| 15531 | 107  | 
fun init_deps master files = SOME (make_deps false true master (map (rpair NONE) files));  | 
| 7211 | 108  | 
|
109  | 
||
| 6362 | 110  | 
type thy = deps option * theory option;  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
111  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
112  | 
local  | 
| 6362 | 113  | 
val database = ref (Graph.empty: thy Graph.T);  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
114  | 
in  | 
| 6362 | 115  | 
fun get_thys () = ! database;  | 
| 9137 | 116  | 
val change_thys = Library.change database;  | 
| 
3604
 
6bf9f09f3d61
Moved functions for theory information storage / retrieval
 
berghofe 
parents:  
diff
changeset
 | 
117  | 
end;  | 
| 5209 | 118  | 
|
119  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
120  | 
(* access thy graph *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
121  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
122  | 
fun thy_graph f x = f (get_thys ()) x;  | 
| 9417 | 123  | 
|
124  | 
(*theory names in topological order*)  | 
|
125  | 
fun get_names () =  | 
|
126  | 
let val G = get_thys ()  | 
|
| 14774 | 127  | 
in Graph.all_succs G (Graph.minimals G) end;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
128  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
129  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
130  | 
(* access thy *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
131  | 
|
| 7935 | 132  | 
fun lookup_thy name =  | 
| 15531 | 133  | 
SOME (thy_graph Graph.get_node name) handle Graph.UNDEF _ => NONE;  | 
| 7935 | 134  | 
|
| 15570 | 135  | 
val known_thy = isSome o lookup_thy;  | 
| 7935 | 136  | 
fun check_known_thy name = known_thy name orelse (warning ("Unknown theory " ^ quote name); false);
 | 
137  | 
fun if_known_thy f name = if check_known_thy name then f name else ();  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
138  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
139  | 
fun get_thy name =  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
140  | 
(case lookup_thy name of  | 
| 15531 | 141  | 
SOME thy => thy  | 
142  | 
| NONE => error (loader_msg "nothing known about theory" [name]));  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
143  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
144  | 
fun change_thy name f = (get_thy name; change_thys (Graph.map_node name f));  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
145  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
146  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
147  | 
(* access deps *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
148  | 
|
| 15570 | 149  | 
val lookup_deps = Option.map #1 o lookup_thy;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
150  | 
val get_deps = #1 o get_thy;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
151  | 
fun change_deps name f = change_thy name (fn (deps, x) => (f deps, x));  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
152  | 
|
| 6666 | 153  | 
fun is_finished name = is_none (get_deps name);  | 
| 15531 | 154  | 
fun get_files name = (case get_deps name of SOME {files, ...} => files | _ => []);
 | 
| 7191 | 155  | 
|
156  | 
fun loaded_files name =  | 
|
157  | 
(case get_deps name of  | 
|
| 15531 | 158  | 
NONE => []  | 
159  | 
  | SOME {master, files, ...} =>
 | 
|
160  | 
(case master of SOME m => [#1 (ThyLoad.get_thy m)] | NONE => []) @  | 
|
| 15570 | 161  | 
(List.mapPartial (Option.map #1 o #2) files));  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
162  | 
|
| 6654 | 163  | 
fun get_preds name =  | 
| 9327 | 164  | 
(thy_graph Graph.imm_preds name) handle Graph.UNDEF _ =>  | 
| 6654 | 165  | 
error (loader_msg "nothing known about theory" [name]);  | 
166  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
167  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
168  | 
(* access theory *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
169  | 
|
| 7687 | 170  | 
fun lookup_theory name =  | 
171  | 
(case lookup_thy name of  | 
|
| 15531 | 172  | 
SOME (_, SOME thy) => SOME thy  | 
173  | 
| _ => NONE);  | 
|
| 7288 | 174  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
175  | 
fun get_theory name =  | 
| 7288 | 176  | 
(case lookup_theory name of  | 
| 15531 | 177  | 
(SOME theory) => theory  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
178  | 
| _ => error (loader_msg "undefined theory entry for" [name]));  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
179  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
180  | 
val theory_of_sign = get_theory o Sign.name_of;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
181  | 
val theory_of_thm = theory_of_sign o Thm.sign_of_thm;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
182  | 
|
| 15531 | 183  | 
fun put_theory name theory = change_thy name (fn (deps, _) => (deps, SOME theory));  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
184  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
185  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
186  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
187  | 
(** thy operations **)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
188  | 
|
| 6241 | 189  | 
(* maintain 'outdated' flag *)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
190  | 
|
| 7099 | 191  | 
local  | 
192  | 
||
| 6241 | 193  | 
fun is_outdated name =  | 
194  | 
(case lookup_deps name of  | 
|
| 15531 | 195  | 
    SOME (SOME {outdated, ...}) => outdated
 | 
| 6241 | 196  | 
| _ => false);  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
197  | 
|
| 6241 | 198  | 
fun outdate_thy name =  | 
| 7099 | 199  | 
if is_finished name orelse is_outdated name then ()  | 
| 15570 | 200  | 
  else (change_deps name (Option.map (fn {present, outdated = _, master, files} =>
 | 
| 7099 | 201  | 
make_deps present true master files)); perform Outdate name);  | 
202  | 
||
| 7910 | 203  | 
fun check_unfinished name =  | 
| 15531 | 204  | 
if is_finished name then (warning (loader_msg "tried to touch finished theory" [name]); NONE)  | 
205  | 
else SOME name;  | 
|
| 7910 | 206  | 
|
| 7099 | 207  | 
in  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
208  | 
|
| 7910 | 209  | 
fun touch_thys names =  | 
| 15570 | 210  | 
List.app outdate_thy (thy_graph Graph.all_succs (List.mapPartial check_unfinished names));  | 
| 7910 | 211  | 
|
212  | 
fun touch_thy name = touch_thys [name];  | 
|
213  | 
fun touch_child_thys name = touch_thys (thy_graph Graph.imm_succs name);  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
214  | 
|
| 15570 | 215  | 
fun touch_all_thys () = List.app outdate_thy (get_names ());  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
216  | 
|
| 7099 | 217  | 
end;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
218  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
219  | 
|
| 7244 | 220  | 
(* check 'finished' state *)  | 
221  | 
||
222  | 
fun check_unfinished fail name =  | 
|
| 7910 | 223  | 
if known_thy name andalso is_finished name then  | 
| 7288 | 224  | 
fail (loader_msg "cannot update finished theory" [name])  | 
| 7244 | 225  | 
else ();  | 
226  | 
||
227  | 
||
| 7941 | 228  | 
(* load_file *)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
229  | 
|
| 15570 | 230  | 
val opt_path = Option.map (Path.dir o fst o ThyLoad.get_thy);  | 
231  | 
fun opt_path' (d : deps option) = getOpt (Option.map (opt_path o #master) d, NONE);  | 
|
232  | 
fun opt_path'' d = getOpt (Option.map opt_path' d, NONE);  | 
|
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
233  | 
|
| 7191 | 234  | 
local  | 
235  | 
||
| 15531 | 236  | 
fun provide path name info (deps as SOME {present, outdated, master, files}) =
 | 
| 7941 | 237  | 
(if exists (equal path o #1) files then ()  | 
238  | 
else warning (loader_msg "undeclared dependency of theory" [name] ^  | 
|
239  | 
" on file: " ^ quote (Path.pack path));  | 
|
| 15531 | 240  | 
SOME (make_deps present outdated master (overwrite (files, (path, SOME info)))))  | 
241  | 
| provide _ _ _ NONE = NONE;  | 
|
| 7941 | 242  | 
|
243  | 
fun run_file path =  | 
|
| 15570 | 244  | 
(case Option.map PureThy.get_name (Context.get_context ()) of  | 
| 15531 | 245  | 
NONE => (ThyLoad.load_file NONE path; ())  | 
246  | 
| SOME name => (case lookup_deps name of  | 
|
247  | 
SOME deps => change_deps name (provide path name  | 
|
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
248  | 
(ThyLoad.load_file (opt_path' deps) path))  | 
| 15531 | 249  | 
| NONE => (ThyLoad.load_file NONE path; ())));  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
250  | 
|
| 7191 | 251  | 
in  | 
252  | 
||
| 7941 | 253  | 
fun load_file time path =  | 
254  | 
if time then  | 
|
| 7244 | 255  | 
let val name = Path.pack path in  | 
256  | 
timeit (fn () =>  | 
|
| 9778 | 257  | 
       (priority ("\n**** Starting file " ^ quote name ^ " ****");
 | 
| 9036 | 258  | 
run_file path;  | 
| 9778 | 259  | 
        priority ("**** Finished file " ^ quote name ^ " ****\n")))
 | 
| 7244 | 260  | 
end  | 
| 7941 | 261  | 
else run_file path;  | 
262  | 
||
263  | 
val use = load_file false o Path.unpack;  | 
|
264  | 
val time_use = load_file true o Path.unpack;  | 
|
| 7191 | 265  | 
|
266  | 
end;  | 
|
| 6233 | 267  | 
|
268  | 
||
| 7099 | 269  | 
(* load_thy *)  | 
270  | 
||
| 9490 | 271  | 
fun required_by _ [] = ""  | 
272  | 
| required_by s initiators = s ^ "(required by " ^ show_path (rev initiators) ^ ")";  | 
|
| 7099 | 273  | 
|
| 15065 | 274  | 
fun load_thy really ml time initiators dir name =  | 
| 7099 | 275  | 
let  | 
| 9822 | 276  | 
val _ =  | 
277  | 
      if really then priority ("Loading theory " ^ quote name ^ required_by " " initiators)
 | 
|
278  | 
      else priority ("Registering theory " ^ quote name);
 | 
|
| 7099 | 279  | 
|
280  | 
val _ = touch_thy name;  | 
|
| 9822 | 281  | 
val master =  | 
282  | 
if really then ThyLoad.load_thy dir name ml time  | 
|
283  | 
else #1 (ThyLoad.deps_thy dir name ml);  | 
|
| 7099 | 284  | 
|
285  | 
val files = get_files name;  | 
|
| 15570 | 286  | 
val missing_files = List.mapPartial (fn (path, NONE) => SOME (Path.pack path) | _ => NONE) files;  | 
| 7099 | 287  | 
in  | 
288  | 
if null missing_files then ()  | 
|
289  | 
else warning (loader_msg "unresolved dependencies of theory" [name] ^  | 
|
| 7223 | 290  | 
" on file(s): " ^ commas_quote missing_files);  | 
| 15531 | 291  | 
change_deps name (fn _ => SOME (make_deps true false (SOME master) files));  | 
| 7099 | 292  | 
perform Update name  | 
293  | 
end;  | 
|
294  | 
||
295  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
296  | 
(* require_thy *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
297  | 
|
| 15065 | 298  | 
fun base_of_path s = Path.pack (Path.base (Path.unpack s));  | 
299  | 
||
| 7211 | 300  | 
local  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
301  | 
|
| 7941 | 302  | 
fun load_deps ml dir name =  | 
303  | 
let val (master, (parents, files)) = ThyLoad.deps_thy dir name ml  | 
|
| 15531 | 304  | 
in (SOME (init_deps (SOME master) files), parents) end;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
305  | 
|
| 15531 | 306  | 
fun file_current master (_, NONE) = false  | 
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
307  | 
| file_current master (path, info) =  | 
| 
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
308  | 
(info = ThyLoad.check_file (opt_path master) path);  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
309  | 
|
| 7941 | 310  | 
fun current_deps ml strict dir name =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
311  | 
(case lookup_deps name of  | 
| 15531 | 312  | 
NONE => (false, load_deps ml dir name)  | 
313  | 
| SOME deps =>  | 
|
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
314  | 
let  | 
| 15531 | 315  | 
fun get_name s = (case opt_path'' (lookup_deps s) of NONE => s  | 
316  | 
| SOME p => Path.pack (Path.append p (Path.basic s)));  | 
|
317  | 
val same_deps = (NONE, map get_name (thy_graph Graph.imm_preds name))  | 
|
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
318  | 
in  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
319  | 
(case deps of  | 
| 15531 | 320  | 
NONE => (true, same_deps)  | 
321  | 
        | SOME {present, outdated, master, files} =>
 | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
322  | 
if present andalso not strict then (true, same_deps)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
323  | 
else  | 
| 15531 | 324  | 
let val master' = SOME (ThyLoad.check_thy dir name ml) in  | 
| 7941 | 325  | 
if master <> master' then (false, load_deps ml dir name)  | 
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
326  | 
else (not outdated andalso forall (file_current master') files, same_deps)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
327  | 
end)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
328  | 
end);  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
329  | 
|
| 9822 | 330  | 
fun require_thy really ml time strict keep_strict initiators prfx (visited, str) =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
331  | 
let  | 
| 7066 | 332  | 
val path = Path.expand (Path.unpack str);  | 
| 
6484
 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 
wenzelm 
parents: 
6389 
diff
changeset
 | 
333  | 
val name = Path.pack (Path.base path);  | 
| 7066 | 334  | 
in  | 
| 9332 | 335  | 
if name mem_string initiators then error (cycle_msg initiators) else ();  | 
| 7952 | 336  | 
if known_thy name andalso is_finished name orelse name mem_string visited then  | 
337  | 
(visited, (true, name))  | 
|
| 7066 | 338  | 
else  | 
339  | 
let  | 
|
340  | 
val dir = Path.append prfx (Path.dir path);  | 
|
| 9822 | 341  | 
val req_parent = require_thy true true time (strict andalso keep_strict) keep_strict  | 
| 15065 | 342  | 
(name :: initiators);  | 
| 
6484
 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 
wenzelm 
parents: 
6389 
diff
changeset
 | 
343  | 
|
| 7941 | 344  | 
val (current, (new_deps, parents)) = current_deps ml strict dir name handle ERROR =>  | 
| 7099 | 345  | 
error (loader_msg "the error(s) above occurred while examining theory" [name] ^  | 
| 9490 | 346  | 
(if null initiators then "" else required_by "\n" initiators));  | 
| 15570 | 347  | 
val dir' = getOpt (opt_path'' new_deps, dir);  | 
| 15065 | 348  | 
val (visited', parent_results) = foldl_map (req_parent dir') (name :: visited, parents);  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
349  | 
|
| 15531 | 350  | 
val thy = if not really then SOME (get_theory name) else NONE;  | 
| 7066 | 351  | 
val result =  | 
352  | 
if current andalso forall #1 parent_results then true  | 
|
353  | 
else  | 
|
354  | 
((case new_deps of  | 
|
| 15531 | 355  | 
SOME deps => change_thys (update_node name (map base_of_path parents) (deps, thy))  | 
356  | 
| NONE => ());  | 
|
| 15065 | 357  | 
load_thy really ml (time orelse ! Output.timing) initiators dir name;  | 
| 7099 | 358  | 
false);  | 
| 7066 | 359  | 
in (visited', (result, name)) end  | 
360  | 
end;  | 
|
| 
6484
 
3f098b0ec683
use_thy etc.: may specify path prefix, which is temporarily used as load path;
 
wenzelm 
parents: 
6389 
diff
changeset
 | 
361  | 
|
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
362  | 
fun gen_use_thy' req prfx s =  | 
| 
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
363  | 
let val (_, (_, name)) = req [] prfx ([], s)  | 
| 7066 | 364  | 
in Context.context (get_theory name) end;  | 
| 6241 | 365  | 
|
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
366  | 
fun gen_use_thy req = gen_use_thy' req Path.current;  | 
| 
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
367  | 
|
| 7244 | 368  | 
fun warn_finished f name = (check_unfinished warning name; f name);  | 
369  | 
||
| 7211 | 370  | 
in  | 
371  | 
||
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
372  | 
val weak_use_thy = gen_use_thy' (require_thy true true false false false);  | 
| 
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
373  | 
fun quiet_update_thy' ml = gen_use_thy' (require_thy true ml false true true);  | 
| 
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
374  | 
fun quiet_update_thy ml = gen_use_thy (require_thy true ml false true true);  | 
| 7244 | 375  | 
|
| 9822 | 376  | 
val use_thy = warn_finished (gen_use_thy (require_thy true true false true false));  | 
377  | 
val time_use_thy = warn_finished (gen_use_thy (require_thy true true true true false));  | 
|
378  | 
val use_thy_only = warn_finished (gen_use_thy (require_thy true false false true false));  | 
|
379  | 
val update_thy = warn_finished (gen_use_thy (require_thy true true false true true));  | 
|
380  | 
val update_thy_only = warn_finished (gen_use_thy (require_thy true false false true true));  | 
|
381  | 
val pretend_use_thy_only = warn_finished (gen_use_thy (require_thy false false false true false));  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
382  | 
|
| 7211 | 383  | 
end;  | 
384  | 
||
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
385  | 
|
| 6666 | 386  | 
(* remove_thy *)  | 
387  | 
||
388  | 
fun remove_thy name =  | 
|
| 7910 | 389  | 
if is_finished name then error (loader_msg "cannot remove finished theory" [name])  | 
| 6666 | 390  | 
else  | 
391  | 
let val succs = thy_graph Graph.all_succs [name] in  | 
|
| 9778 | 392  | 
priority (loader_msg "removing" succs);  | 
| 15570 | 393  | 
List.app (perform Remove) succs;  | 
| 7191 | 394  | 
change_thys (Graph.del_nodes succs)  | 
| 6666 | 395  | 
end;  | 
396  | 
||
397  | 
||
| 7066 | 398  | 
(* begin / end theory *)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
399  | 
|
| 8805 | 400  | 
fun begin_theory present upd name parents paths =  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
401  | 
let  | 
| 15065 | 402  | 
val bparents = map base_of_path parents;  | 
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
403  | 
val dir' = opt_path'' (lookup_deps name);  | 
| 15570 | 404  | 
val dir = getOpt (dir',Path.current);  | 
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
405  | 
val assert_thy = if upd then quiet_update_thy' true dir else weak_use_thy dir;  | 
| 7244 | 406  | 
val _ = check_unfinished error name;  | 
| 15570 | 407  | 
val _ = List.app assert_thy parents;  | 
| 15065 | 408  | 
val theory = PureThy.begin_theory name (map get_theory bparents);  | 
| 7952 | 409  | 
val deps =  | 
410  | 
if known_thy name then get_deps name  | 
|
| 15531 | 411  | 
else (init_deps NONE (map #1 paths)); (*records additional ML files only!*)  | 
| 15570 | 412  | 
val use_paths = List.mapPartial (fn (x, true) => SOME x | _ => NONE) paths;  | 
| 9451 | 413  | 
|
| 15531 | 414  | 
val _ = change_thys (update_node name bparents (deps, SOME (Theory.copy theory)));  | 
| 
15158
 
2281784015a5
Fixed several bugs related to path specifications in theory names.
 
berghofe 
parents: 
15065 
diff
changeset
 | 
415  | 
val theory' = theory |> present dir' name bparents paths;  | 
| 9451 | 416  | 
val _ = put_theory name (Theory.copy theory');  | 
| 15570 | 417  | 
val ((), theory'') = Context.pass_theory theory' (List.app (load_file false)) use_paths;  | 
| 9451 | 418  | 
val _ = put_theory name (Theory.copy theory'');  | 
| 8805 | 419  | 
in theory'' end;  | 
| 7244 | 420  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
421  | 
fun end_theory theory =  | 
| 7099 | 422  | 
let  | 
423  | 
val theory' = PureThy.end_theory theory;  | 
|
424  | 
val name = PureThy.get_name theory;  | 
|
425  | 
in put_theory name theory'; theory' end;  | 
|
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
426  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
427  | 
|
| 6666 | 428  | 
(* finish all theories *)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
429  | 
|
| 15531 | 430  | 
fun finish () = change_thys (Graph.map_nodes (fn (_, entry) => (NONE, entry)));  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
431  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
432  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
433  | 
(* register existing theories *)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
434  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
435  | 
fun register_theory theory =  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
436  | 
let  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
437  | 
val name = PureThy.get_name theory;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
438  | 
val parents = Theory.parents_of theory;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
439  | 
val parent_names = map PureThy.get_name parents;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
440  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
441  | 
fun err txt bads =  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
442  | 
error (loader_msg txt bads ^ "\n" ^ gen_msg "cannot register theory" [name]);  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
443  | 
|
| 6666 | 444  | 
val nonfinished = filter_out is_finished parent_names;  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
445  | 
fun get_variant (x, y_name) =  | 
| 15531 | 446  | 
if Theory.eq_thy (x, get_theory y_name) then NONE  | 
447  | 
else SOME y_name;  | 
|
| 15570 | 448  | 
val variants = List.mapPartial get_variant (parents ~~ parent_names);  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
449  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
450  | 
fun register G =  | 
| 15531 | 451  | 
(Graph.new_node (name, (NONE, SOME theory)) G  | 
| 9327 | 452  | 
handle Graph.DUP _ => err "duplicate theory entry" [])  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
453  | 
|> add_deps name parent_names;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
454  | 
in  | 
| 6666 | 455  | 
if not (null nonfinished) then err "non-finished parent theories" nonfinished  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
456  | 
else if not (null variants) then err "different versions of parent theories" variants  | 
| 7099 | 457  | 
else (change_thys register; perform Update name)  | 
| 
6211
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
458  | 
end;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
459  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
460  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
461  | 
(*final declarations of this structure*)  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
462  | 
val theory = get_theory;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
463  | 
val names = get_names;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
464  | 
|
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
465  | 
end;  | 
| 
 
43d0efafa025
Theory loader database: theory and file dependencies, theory values
 
wenzelm 
parents: 
5211 
diff
changeset
 | 
466  | 
|
| 5209 | 467  | 
structure BasicThyInfo: BASIC_THY_INFO = ThyInfo;  | 
468  | 
open BasicThyInfo;  |