src/Pure/Tools/thm_deps.ML
author wenzelm
Wed, 31 Dec 2014 14:13:11 +0100
changeset 59207 6b030dc97a4f
parent 59058 a78612c67ec0
child 59208 2486d625878b
permissions -rw-r--r--
more explict and generic field names
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
     1
(*  Title:      Pure/Tools/thm_deps.ML
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     3
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     4
Visualize dependencies of theorems.
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     5
*)
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     6
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     7
signature THM_DEPS =
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     8
sig
37870
dd9cfc512b7f thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
wenzelm
parents: 37216
diff changeset
     9
  val thm_deps: theory -> thm list -> unit
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    10
  val unused_thms: theory list * theory list -> (string * thm) list
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    11
end;
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    12
33391
91b9da2a7b44 structure Thm_Deps;
wenzelm
parents: 33170
diff changeset
    13
structure Thm_Deps: THM_DEPS =
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    14
struct
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    15
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    16
(* thm_deps *)
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    17
37870
dd9cfc512b7f thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
wenzelm
parents: 37216
diff changeset
    18
fun thm_deps thy thms =
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
    19
  let
28826
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    20
    fun add_dep ("", _, _) = I
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    21
      | add_dep (name, _, PBody {thms = thms', ...}) =
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    22
          let
41489
8e2b8649507d standardized split_last/last_elem towards List.last;
wenzelm
parents: 39557
diff changeset
    23
            val prefix = #1 (split_last (Long_Name.explode name));
28826
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    24
            val session =
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    25
              (case prefix of
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    26
                a :: _ =>
37870
dd9cfc512b7f thm_deps/unused_thms: Context.get_theory based on proper theory ancestry, not accidental theory loader state;
wenzelm
parents: 37216
diff changeset
    27
                  (case try (Context.get_theory thy) a of
28826
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    28
                    SOME thy =>
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    29
                      (case Present.session_name thy of
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    30
                        "" => []
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    31
                      | session => [session])
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    32
                  | NONE => [])
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    33
               | _ => ["global"]);
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    34
            val parents = filter_out (fn s => s = "") (map (#1 o #2) thms');
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    35
            val entry =
30364
577edc39b501 moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents: 30280
diff changeset
    36
              {name = Long_Name.base_name name,
59207
6b030dc97a4f more explict and generic field names
wenzelm
parents: 59058
diff changeset
    37
               ident = name,
6b030dc97a4f more explict and generic field names
wenzelm
parents: 59058
diff changeset
    38
               directory = space_implode "/" (session @ prefix),
28826
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    39
               unfold = false,
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    40
               path = "",
49561
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 44333
diff changeset
    41
               parents = parents,
26fc70e983c2 separate module Graph_Display;
wenzelm
parents: 44333
diff changeset
    42
               content = []};
28826
3b460b6eadae simplified thm_deps -- no need to build a graph datastructure;
wenzelm
parents: 28817
diff changeset
    43
          in cons entry end;
44333
cc53ce50f738 reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
wenzelm
parents: 44331
diff changeset
    44
    val deps = Proofterm.fold_body_thms add_dep (Thm.proof_bodies_of thms) [];
59207
6b030dc97a4f more explict and generic field names
wenzelm
parents: 59058
diff changeset
    45
  in Graph_Display.display_graph (sort_wrt #ident deps) end;
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    46
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
    47
val _ =
58893
9e0ecb66d6a7 eliminated unused int_only flag (see also c12484a27367);
wenzelm
parents: 58028
diff changeset
    48
  Outer_Syntax.command @{command_spec "thm_deps"} "visualize theorem dependencies"
58028
e4250d370657 tuned signature -- define some elementary operations earlier;
wenzelm
parents: 57934
diff changeset
    49
    (Parse.xthms1 >> (fn args =>
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
    50
      Toplevel.unknown_theory o Toplevel.keep (fn state =>
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
    51
        thm_deps (Toplevel.theory_of state) (Attrib.eval_thms (Toplevel.context_of state) args))));
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
    52
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    53
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    54
(* unused_thms *)
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    55
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    56
fun unused_thms (base_thys, thys) =
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
    57
  let
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    58
    fun add_fact space (name, ths) =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37870
diff changeset
    59
      if exists (fn thy => Global_Theory.defined_fact thy name) base_thys then I
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    60
      else
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    61
        let val {concealed, group, ...} = Name_Space.the_entry space name in
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    62
          fold_rev (fn th =>
36744
6e1f3d609a68 renamed Thm.get_name -> Thm.derivation_name and Thm.put_name -> Thm.name_derivation, to emphasize the true nature of these operations;
wenzelm
parents: 33769
diff changeset
    63
            (case Thm.derivation_name th of
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    64
              "" => I
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    65
            | a => cons (a, (th, concealed, group)))) ths
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    66
        end;
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    67
    fun add_facts facts = Facts.fold_static (add_fact (Facts.space_of facts)) facts;
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    68
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    69
    val new_thms =
39557
fe5722fce758 renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents: 37870
diff changeset
    70
      fold (add_facts o Global_Theory.facts_of) thys []
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58893
diff changeset
    71
      |> sort_distinct (string_ord o apply2 #1);
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
    72
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    73
    val used =
32810
f3466a5645fa back to simple fold_body_thms and fulfill_proof/thm_proof (reverting a900d3cd47cc) -- the cycle check is implicit in the future computation of join_proofs;
wenzelm
parents: 32726
diff changeset
    74
      Proofterm.fold_body_thms
41565
9718c32f9c4e unused_thms no longer compares propositions, since this is no longer needed
berghofe
parents: 41489
diff changeset
    75
        (fn (a, _, _) => a <> "" ? Symtab.update (a, ()))
44333
cc53ce50f738 reverted to join_bodies/join_proofs based on fold_body_thms to regain performance (escpecially of HOL-Proofs) -- see also aa9c1e9ef2ce and 4e2abb045eac;
wenzelm
parents: 44331
diff changeset
    76
        (map Proofterm.strip_thm (Thm.proof_bodies_of (map (#1 o #2) new_thms)))
44331
aa9c1e9ef2ce clarified fulfill_norm_proof: no join_thms yet;
wenzelm
parents: 42473
diff changeset
    77
        Symtab.empty;
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    78
41565
9718c32f9c4e unused_thms no longer compares propositions, since this is no longer needed
berghofe
parents: 41489
diff changeset
    79
    fun is_unused a = not (Symtab.defined used a);
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
    80
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
    81
    (*groups containing at least one used theorem*)
41565
9718c32f9c4e unused_thms no longer compares propositions, since this is no longer needed
berghofe
parents: 41489
diff changeset
    82
    val used_groups = fold (fn (a, (_, _, group)) =>
9718c32f9c4e unused_thms no longer compares propositions, since this is no longer needed
berghofe
parents: 41489
diff changeset
    83
      if is_unused a then I
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    84
      else
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    85
        (case group of
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    86
          NONE => I
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    87
        | SOME grp => Inttab.update (grp, ()))) new_thms Inttab.empty;
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    88
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    89
    val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, seen_groups) =>
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    90
      if not concealed andalso
42473
aca720fb3936 mark thm tag "kind" as legacy;
wenzelm
parents: 41565
diff changeset
    91
        (* FIXME replace by robust treatment of thm groups *)
aca720fb3936 mark thm tag "kind" as legacy;
wenzelm
parents: 41565
diff changeset
    92
        member (op =) [Thm.theoremK, Thm.lemmaK, Thm.corollaryK] (Thm.legacy_get_kind th) andalso
41565
9718c32f9c4e unused_thms no longer compares propositions, since this is no longer needed
berghofe
parents: 41489
diff changeset
    93
        is_unused a
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    94
      then
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    95
        (case group of
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    96
           NONE => ((a, th) :: thms, seen_groups)
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
    97
         | SOME grp =>
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    98
             if Inttab.defined used_groups grp orelse
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
    99
               Inttab.defined seen_groups grp then q
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   100
             else ((a, th) :: thms, Inttab.update (grp, ()) seen_groups))
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   101
      else q) new_thms ([], Inttab.empty);
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   102
  in rev thms' end;
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   103
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   104
val _ =
58893
9e0ecb66d6a7 eliminated unused int_only flag (see also c12484a27367);
wenzelm
parents: 58028
diff changeset
   105
  Outer_Syntax.command @{command_spec "unused_thms"} "find unused theorems"
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   106
    (Scan.option ((Scan.repeat1 (Scan.unless Parse.minus Parse.name) --| Parse.minus) --
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   107
       Scan.option (Scan.repeat1 (Scan.unless Parse.minus Parse.name))) >> (fn opt_range =>
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   108
        Toplevel.keep (fn state =>
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   109
          let
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   110
            val thy = Toplevel.theory_of state;
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   111
            val ctxt = Toplevel.context_of state;
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   112
            fun pretty_thm (a, th) = Proof_Context.pretty_fact ctxt (a, [th]);
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   113
            val get_theory = Context.get_theory thy;
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   114
          in
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   115
            unused_thms
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   116
              (case opt_range of
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   117
                NONE => (Theory.parents_of thy, [thy])
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   118
              | SOME (xs, NONE) => (map get_theory xs, [thy])
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   119
              | SOME (xs, SOME ys) => (map get_theory xs, map get_theory ys))
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   120
            |> map pretty_thm |> Pretty.writeln_chunks
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   121
          end)));
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   122
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
   123
end;
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
   124