| author | paulson <lp15@cam.ac.uk> | 
| Fri, 30 Aug 2024 10:16:48 +0100 | |
| changeset 80790 | 07c51801c2ea | 
| parent 77889 | 5db014c36f42 | 
| child 82598 | 766a07ff7a07 | 
| 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: 
62848diff
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: 
62848diff
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: 
62848diff
changeset | 18 | fun gen_thy_deps prep_thy ctxt bounds = | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
changeset | 19 | let | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
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: 
62848diff
changeset | 21 | val rel = Context.subthy o swap; | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
changeset | 22 | val pred = | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
changeset | 23 | (case upper of | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
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: 
62848diff
changeset | 25 | | NONE => K true) andf | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
changeset | 26 | (case lower of | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
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: 
62848diff
changeset | 28 | | NONE => K true); | 
| 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
changeset | 29 | fun node thy = | 
| 77889 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
68482diff
changeset | 30 | ((Context.theory_base_name thy, Graph_Display.content_node (Context.theory_base_name thy) []), | 
| 
5db014c36f42
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
 wenzelm parents: 
68482diff
changeset | 31 | map Context.theory_base_name (filter pred (Theory.parents_of thy))); | 
| 65491 
7fb81fa1d668
more uniform thy_deps (like class_deps), see also c48d536231fe;
 wenzelm parents: 
62848diff
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 | |
| 68482 | 39 | val thy_deps_cmd = Graph_Display.display_graph oo gen_thy_deps (Theory.check {long = false});
 | 
| 60093 | 40 | |
| 41 | end; |