src/Pure/Tools/thy_deps.ML
author wenzelm
Thu, 20 Apr 2023 11:57:34 +0200
changeset 77889 5db014c36f42
parent 68482 cb84beb84ca9
permissions -rw-r--r--
clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60093
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Tools/thy_deps.ML
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     3
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     4
Visualization of theory dependencies.
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     5
*)
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     6
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     7
signature THY_DEPS =
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
     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
73c260342704 formal Theory.check, with markup and completion;
wenzelm
parents: 60097
diff changeset
    11
  val thy_deps_cmd: Proof.context ->
73c260342704 formal Theory.check, with markup and completion;
wenzelm
parents: 60097
diff changeset
    12
    (string * Position.T) list option * (string * Position.T) list option -> unit
60093
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    13
end;
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    14
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    15
structure Thy_Deps: THY_DEPS =
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    16
struct
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    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 =
77889
5db014c36f42 clarified signature: explicitly distinguish theory_base_name vs. theory_long_name;
wenzelm
parents: 68482
diff 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: 68482
diff 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: 62848
diff changeset
    32
  in map node (filter pred (Theory.nodes_of (Proof_Context.theory_of ctxt))) end;
60093
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    33
60099
73c260342704 formal Theory.check, with markup and completion;
wenzelm
parents: 60097
diff changeset
    34
val thy_deps =
73c260342704 formal Theory.check, with markup and completion;
wenzelm
parents: 60097
diff changeset
    35
  gen_thy_deps (fn ctxt => fn thy =>
73c260342704 formal Theory.check, with markup and completion;
wenzelm
parents: 60097
diff changeset
    36
    let val thy0 = Proof_Context.theory_of ctxt
60948
b710a5087116 prefer theory_id operations;
wenzelm
parents: 60099
diff changeset
    37
    in if Context.subthy (thy, thy0) then thy else raise THEORY ("Bad theory", [thy, thy0]) end);
60099
73c260342704 formal Theory.check, with markup and completion;
wenzelm
parents: 60097
diff changeset
    38
68482
cb84beb84ca9 clarified signature;
wenzelm
parents: 65491
diff changeset
    39
val thy_deps_cmd = Graph_Display.display_graph oo gen_thy_deps (Theory.check {long = false});
60093
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    40
c48d536231fe clarified thy_deps;
wenzelm
parents:
diff changeset
    41
end;