author | wenzelm |
Thu, 15 Feb 2018 12:11:00 +0100 | |
changeset 67613 | ce654b0e6d69 |
parent 65491 | 7fb81fa1d668 |
child 68482 | cb84beb84ca9 |
permissions | -rw-r--r-- |
60093 | 1 |
(* Title: Pure/Tools/thy_deps.ML |
2 |
Author: Makarius |
|
3 |
||
4 |
Visualization of theory dependencies. |
|
5 |
*) |
|
6 |
||
7 |
signature THY_DEPS = |
|
8 |
sig |
|
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
9 |
val thy_deps: Proof.context -> theory list option * theory list option -> |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
10 |
Graph_Display.entry list |
60099 | 11 |
val thy_deps_cmd: Proof.context -> |
12 |
(string * Position.T) list option * (string * Position.T) list option -> unit |
|
60093 | 13 |
end; |
14 |
||
15 |
structure Thy_Deps: THY_DEPS = |
|
16 |
struct |
|
17 |
||
65491
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
18 |
fun gen_thy_deps prep_thy ctxt bounds = |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
19 |
let |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
20 |
val (upper, lower) = apply2 ((Option.map o map) (prep_thy ctxt)) bounds; |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
21 |
val rel = Context.subthy o swap; |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
22 |
val pred = |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
23 |
(case upper of |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
24 |
SOME Bs => (fn thy => exists (fn B => rel (thy, B)) Bs) |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
25 |
| NONE => K true) andf |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
26 |
(case lower of |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
27 |
SOME Bs => (fn thy => exists (fn B => rel (B, thy)) Bs) |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
28 |
| NONE => K true); |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
29 |
fun node thy = |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
30 |
((Context.theory_name thy, Graph_Display.content_node (Context.theory_name thy) []), |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
31 |
map Context.theory_name (filter pred (Theory.parents_of thy))); |
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
wenzelm
parents:
62848
diff
changeset
|
32 |
in map node (filter pred (Theory.nodes_of (Proof_Context.theory_of ctxt))) end; |
60093 | 33 |
|
60099 | 34 |
val thy_deps = |
35 |
gen_thy_deps (fn ctxt => fn thy => |
|
36 |
let val thy0 = Proof_Context.theory_of ctxt |
|
60948 | 37 |
in if Context.subthy (thy, thy0) then thy else raise THEORY ("Bad theory", [thy, thy0]) end); |
60099 | 38 |
|
39 |
val thy_deps_cmd = Graph_Display.display_graph oo gen_thy_deps Theory.check; |
|
60093 | 40 |
|
41 |
end; |