src/Pure/thm_deps.ML
author immler
Sun, 03 Nov 2019 21:46:46 -0500
changeset 71034 e0755162093f
parent 71015 bb49abc2ecbb
child 71465 910a081cca74
permissions -rw-r--r--
replace approximation oracle by less ad-hoc @{computation}s
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70893
ce1e27dcc9f4 clarified files;
wenzelm
parents: 70830
diff changeset
     1
(*  Title:      Pure/thm_deps.ML
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
70556
038ed9b76c2b clarified modules;
wenzelm
parents: 68482
diff changeset
     3
    Author:     Makarius
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     4
70556
038ed9b76c2b clarified modules;
wenzelm
parents: 68482
diff changeset
     5
Dependencies of theorems wrt. internal derivation.
7765
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
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     8
signature THM_DEPS =
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
     9
sig
70557
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    10
  val all_oracles: thm list -> Proofterm.oracle list
70566
fb3d06d7dd05 more thorough check, using full dependency graph of finished proofs;
wenzelm
parents: 70560
diff changeset
    11
  val has_skip_proof: thm list -> bool
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    12
  val pretty_thm_oracles: Proof.context -> thm list -> Pretty.T
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    13
  val thm_deps: theory -> thm list -> (Proofterm.thm_id * Thm_Name.T) list
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    14
  val pretty_thm_deps: theory -> thm list -> Pretty.T
70570
d94456876f2d clarified signature;
wenzelm
parents: 70567
diff changeset
    15
  val unused_thms_cmd: theory list * theory list -> (string * thm) list
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    16
end;
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    17
33391
91b9da2a7b44 structure Thm_Deps;
wenzelm
parents: 33170
diff changeset
    18
structure Thm_Deps: THM_DEPS =
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    19
struct
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
    20
70557
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    21
(* oracles *)
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    22
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    23
fun all_oracles thms =
70588
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    24
  let
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    25
    fun collect (PBody {oracles, thms, ...}) =
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    26
      (if null oracles then I else apfst (cons oracles)) #>
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    27
      (tap Proofterm.join_thms thms |> fold (fn (i, thm_node) => fn (res, seen) =>
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    28
        if Inttab.defined seen i then (res, seen)
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    29
        else
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    30
          let val body = Future.join (Proofterm.thm_node_body thm_node)
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    31
          in collect body (res, Inttab.update (i, ()) seen) end));
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    32
    val bodies = map Thm.proof_body_of thms;
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    33
  in fold collect bodies ([], Inttab.empty) |> #1 |> Proofterm.unions_oracles end;
70557
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    34
70566
fb3d06d7dd05 more thorough check, using full dependency graph of finished proofs;
wenzelm
parents: 70560
diff changeset
    35
fun has_skip_proof thms =
fb3d06d7dd05 more thorough check, using full dependency graph of finished proofs;
wenzelm
parents: 70560
diff changeset
    36
  all_oracles thms |> exists (fn (name, _) => name = \<^oracle_name>\<open>skip_proof\<close>);
fb3d06d7dd05 more thorough check, using full dependency graph of finished proofs;
wenzelm
parents: 70560
diff changeset
    37
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    38
fun pretty_thm_oracles ctxt thms =
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    39
  let
70567
f4d111b802a1 proper theory context for global props;
wenzelm
parents: 70566
diff changeset
    40
    val thy = Proof_Context.theory_of ctxt;
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    41
    fun prt_oracle (name, NONE) = [Thm.pretty_oracle ctxt name]
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    42
      | prt_oracle (name, SOME prop) =
70567
f4d111b802a1 proper theory context for global props;
wenzelm
parents: 70566
diff changeset
    43
          [Thm.pretty_oracle ctxt name, Pretty.str ":", Pretty.brk 1,
f4d111b802a1 proper theory context for global props;
wenzelm
parents: 70566
diff changeset
    44
            Syntax.pretty_term_global thy prop];
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    45
  in Pretty.big_list "oracles:" (map (Pretty.item o prt_oracle) (all_oracles thms)) end;
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    46
70557
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    47
70595
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    48
(* thm_deps *)
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    49
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    50
fun thm_deps thy =
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    51
  let
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    52
    val lookup = Global_Theory.lookup_thm_id thy;
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    53
    fun deps (i, thm_node) res =
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    54
      if Inttab.defined res i then res
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    55
      else
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    56
        let val thm_id = Proofterm.thm_id (i, thm_node) in
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    57
          (case lookup thm_id of
70602
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    58
            SOME thm_name =>
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    59
              Inttab.update (i, SOME (thm_id, thm_name)) res
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    60
          | NONE =>
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    61
              Inttab.update (i, NONE) res
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    62
              |> fold deps (Proofterm.thm_node_thms thm_node))
70595
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    63
        end;
70596
3a7117c33742 proper theory context;
wenzelm
parents: 70595
diff changeset
    64
  in
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    65
    fn thms =>
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    66
      (fold (fold deps o Thm.thm_deps o Thm.transfer thy) thms Inttab.empty, [])
70602
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    67
      |-> Inttab.fold_rev (fn (_, SOME entry) => cons entry | _ => I)
70596
3a7117c33742 proper theory context;
wenzelm
parents: 70595
diff changeset
    68
  end;
70595
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    69
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    70
fun pretty_thm_deps thy thms =
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
    71
  let
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    72
    val ctxt = Proof_Context.init_global thy;
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    73
    val items =
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    74
      map #2 (thm_deps thy thms)
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    75
      |> map (fn (name, i) => (Proof_Context.markup_extern_fact ctxt name, i))
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    76
      |> sort_by (#2 o #1)
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    77
      |> map (fn ((marks, xname), i) =>
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    78
          Pretty.item [Pretty.marks_str (marks, Thm_Name.print (xname, i))]);
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    79
  in Pretty.big_list ("dependencies: " ^ string_of_int (length items)) items end;
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    80
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    81
70570
d94456876f2d clarified signature;
wenzelm
parents: 70567
diff changeset
    82
(* unused_thms_cmd *)
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    83
70570
d94456876f2d clarified signature;
wenzelm
parents: 70567
diff changeset
    84
fun unused_thms_cmd (base_thys, thys) =
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
    85
  let
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    86
    fun add_fact transfer 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
    87
      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
    88
      else
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    89
        let val {concealed, group, ...} = Name_Space.the_entry space name in
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    90
          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
    91
            (case Thm.derivation_name th of
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    92
              "" => I
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    93
            | a => cons (a, (transfer th, concealed, group)))) ths
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    94
        end;
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    95
    fun add_facts thy =
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    96
      let
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    97
        val transfer = Global_Theory.transfer_theories thy;
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    98
        val facts = Global_Theory.facts_of thy;
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    99
      in Facts.fold_static (add_fact transfer (Facts.space_of facts)) facts end;
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   100
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   101
    val new_thms =
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   102
      fold add_facts thys []
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58893
diff changeset
   103
      |> sort_distinct (string_ord o apply2 #1);
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
   104
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   105
    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
   106
      Proofterm.fold_body_thms
64572
cec07f7249cd tuned signature;
wenzelm
parents: 62848
diff changeset
   107
        (fn {name = a, ...} => a <> "" ? Symtab.update (a, ()))
70830
8f050cc0ec50 clarified signature;
wenzelm
parents: 70605
diff changeset
   108
        (map Proofterm.strip_thm_body (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
   109
        Symtab.empty;
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   110
41565
9718c32f9c4e unused_thms no longer compares propositions, since this is no longer needed
berghofe
parents: 41489
diff changeset
   111
    fun is_unused a = not (Symtab.defined used a);
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
   112
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   113
    (*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
   114
    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
   115
      if is_unused a then I
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   116
      else
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   117
        (case group of
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   118
          NONE => I
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   119
        | 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
   120
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   121
    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
   122
      if not concealed andalso
42473
aca720fb3936 mark thm tag "kind" as legacy;
wenzelm
parents: 41565
diff changeset
   123
        (* FIXME replace by robust treatment of thm groups *)
61336
fa4ebbd350ae just one theorem kind, which is legacy anyway;
wenzelm
parents: 60103
diff changeset
   124
        Thm.legacy_get_kind th = Thm.theoremK andalso is_unused a
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   125
      then
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   126
        (case group of
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   127
           NONE => ((a, th) :: thms, seen_groups)
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   128
         | SOME grp =>
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   129
             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
   130
               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
   131
             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
   132
      else q) new_thms ([], Inttab.empty);
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   133
  in rev thms' end;
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   134
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
   135
end;