src/Pure/thm_deps.ML
author wenzelm
Sat, 08 Jun 2024 16:26:47 +0200
changeset 80299 a397fd0c451a
parent 80295 8a9588ffc133
child 80590 505f97165f52
permissions -rw-r--r--
more accurate output of Thm_Name.T wrt. facts name space;
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
77867
686a7d71ed7b backout e3fe192fa4a8;
wenzelm
parents: 77866
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
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
    15
  val unused_thms_cmd: theory list * theory list -> (Thm_Name.T * 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, ...}) =
77867
686a7d71ed7b backout e3fe192fa4a8;
wenzelm
parents: 77866
diff changeset
    26
      (if null oracles then I else apfst (cons oracles)) #>
77866
3bd1aa2f3517 backout 61f652dd955a;
wenzelm
parents: 77825
diff changeset
    27
      (tap Proofterm.join_thms thms |> fold (fn (i, thm_node) => fn (res, seen) =>
77820
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
    28
        if Intset.member seen i then (res, seen)
70588
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)
77820
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
    31
          in collect body (res, Intset.insert i seen) end));
70588
35a4ef9c6d80 clarified modules;
wenzelm
parents: 70570
diff changeset
    32
    val bodies = map Thm.proof_body_of thms;
77867
686a7d71ed7b backout e3fe192fa4a8;
wenzelm
parents: 77866
diff changeset
    33
  in fold collect bodies ([], Intset.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 =
77867
686a7d71ed7b backout e3fe192fa4a8;
wenzelm
parents: 77866
diff changeset
    36
  all_oracles thms |> exists (fn ((name, _), _) => name = \<^oracle_name>\<open>skip_proof\<close>);
70566
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;
71465
910a081cca74 more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
wenzelm
parents: 71015
diff changeset
    41
    fun prt_ora (name, pos) = Thm.pretty_oracle ctxt name :: Pretty.here pos;
910a081cca74 more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
wenzelm
parents: 71015
diff changeset
    42
    fun prt_oracle (ora, NONE) = prt_ora ora
910a081cca74 more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
wenzelm
parents: 71015
diff changeset
    43
      | prt_oracle (ora, SOME prop) =
910a081cca74 more position information for oracles (e.g. "skip_proof" for 'sorry'), requires Proofterm.proofs := 1;
wenzelm
parents: 71015
diff changeset
    44
          prt_ora ora @ [Pretty.str ":", Pretty.brk 1, Syntax.pretty_term_global thy prop];
71570
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    45
    val oracles =
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    46
      (case try all_oracles thms of
77867
686a7d71ed7b backout e3fe192fa4a8;
wenzelm
parents: 77866
diff changeset
    47
        SOME oracles => oracles
71570
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    48
      | NONE => error "Malformed proofs")
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    49
  in Pretty.big_list "oracles:" (map (Pretty.item o prt_oracle) oracles) end;
70560
7714971a58b5 added command 'thm_oracles';
wenzelm
parents: 70557
diff changeset
    50
70557
5d6e9c65ea67 clarified signature;
wenzelm
parents: 70556
diff changeset
    51
70595
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    52
(* thm_deps *)
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    53
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    54
fun thm_deps thy =
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    55
  let
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    56
    val lookup = Global_Theory.lookup_thm_id thy;
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    57
    fun deps (i, thm_node) res =
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    58
      if Inttab.defined res i then res
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    59
      else
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    60
        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
    61
          (case lookup thm_id of
70602
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    62
            SOME thm_name =>
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    63
              Inttab.update (i, SOME (thm_id, thm_name)) res
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    64
          | NONE =>
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    65
              Inttab.update (i, NONE) res
77866
3bd1aa2f3517 backout 61f652dd955a;
wenzelm
parents: 77825
diff changeset
    66
              |> 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
    67
        end;
70596
3a7117c33742 proper theory context;
wenzelm
parents: 70595
diff changeset
    68
  in
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    69
    fn thms =>
77866
3bd1aa2f3517 backout 61f652dd955a;
wenzelm
parents: 77825
diff changeset
    70
      (Inttab.build (fold (fold deps o Thm.thm_deps o Thm.transfer thy) thms), [])
70602
b85a12c2e2bf proper graph traversal: avoid multiple visit of unnamed nodes;
wenzelm
parents: 70596
diff changeset
    71
      |-> Inttab.fold_rev (fn (_, SOME entry) => cons entry | _ => I)
70596
3a7117c33742 proper theory context;
wenzelm
parents: 70595
diff changeset
    72
  end;
70595
2ae7e33c950f clarified thm_id vs. thm_node/thm: retain theory_name;
wenzelm
parents: 70588
diff changeset
    73
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    74
fun pretty_thm_deps thy thms =
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
    75
  let
80299
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    76
    val facts = Global_Theory.facts_of thy;
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    77
    val extern_fact = Name_Space.extern_generic (Context.Theory thy) (Facts.space_of facts);
71570
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    78
    val deps =
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    79
      (case try (thm_deps thy) thms of
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    80
        SOME deps => deps
c2884545c846 slightly more explicit error, without going into the graph of proof futures;
wenzelm
parents: 71539
diff changeset
    81
      | NONE => error "Malformed proofs");
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    82
    val items =
80299
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    83
      deps
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    84
      |> map (fn (_, thm_name as (name, _)) => (extern_fact name, thm_name))
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    85
      |> sort_by #1
a397fd0c451a more accurate output of Thm_Name.T wrt. facts name space;
wenzelm
parents: 80295
diff changeset
    86
      |> map (fn (_, thm_name) => Pretty.item [Global_Theory.pretty_thm_name thy thm_name])
70605
048cf2096186 clarified 'thm_deps' command;
wenzelm
parents: 70602
diff changeset
    87
  in Pretty.big_list ("dependencies: " ^ string_of_int (length items)) items end;
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    88
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    89
70570
d94456876f2d clarified signature;
wenzelm
parents: 70567
diff changeset
    90
(* unused_thms_cmd *)
26697
3b9eede40608 removed obsolete BASIC_THM_DEPS;
wenzelm
parents: 26185
diff changeset
    91
70570
d94456876f2d clarified signature;
wenzelm
parents: 70567
diff changeset
    92
fun unused_thms_cmd (base_thys, thys) =
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
    93
  let
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
    94
    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
    95
      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
    96
      else
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
    97
        let val {concealed, group, ...} = Name_Space.the_entry space name in
71539
c983fd846c9c proper transfer for Thm.derivation_name;
wenzelm
parents: 71465
diff changeset
    98
          fold_rev (transfer #> (fn th =>
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
    99
            let val a = Thm.derivation_name th
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
   100
            in if Thm_Name.is_empty a then I else cons (a, (th, concealed, group)) end)) ths
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   101
        end;
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   102
    fun add_facts thy =
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   103
      let
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   104
        val transfer = Global_Theory.transfer_theories thy;
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   105
        val facts = Global_Theory.facts_of thy;
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   106
      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
   107
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   108
    val new_thms =
61507
6865140215ef proper transfer of stored facts;
wenzelm
parents: 61336
diff changeset
   109
      fold add_facts thys []
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
   110
      |> sort_distinct (Thm_Name.ord o apply2 #1);
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
   111
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   112
    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
   113
      Proofterm.fold_body_thms
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
   114
        (fn {thm_name = a, ...} => not (Thm_Name.is_empty a) ? Thm_Name.Set.insert a)
70830
8f050cc0ec50 clarified signature;
wenzelm
parents: 70605
diff changeset
   115
        (map Proofterm.strip_thm_body (Thm.proof_bodies_of (map (#1 o #2) new_thms)))
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
   116
        Thm_Name.Set.empty;
33170
dd6d8d1f70d2 maintain group via name space, not tags;
wenzelm
parents: 32810
diff changeset
   117
80295
8a9588ffc133 more accurate Thm_Name.T for PThm / Thm.name_derivation / Thm.derivation_name;
wenzelm
parents: 77867
diff changeset
   118
    fun is_unused a = not (Thm_Name.Set.member used a);
28810
e915ab11fe52 retrieve thm deps from proof_body;
wenzelm
parents: 27865
diff changeset
   119
57934
5e500c0e7eca tuned signature -- prefer self-contained user-space tool;
wenzelm
parents: 49561
diff changeset
   120
    (*groups containing at least one used theorem*)
74232
1091880266e5 clarified signature;
wenzelm
parents: 71570
diff changeset
   121
    val used_groups =
77820
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
   122
      Intset.build (new_thms |> fold (fn (a, (_, _, group)) =>
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
   123
        if is_unused a orelse group = 0 then I else Intset.insert group));
33769
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   124
6d8630fab26a unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
wenzelm
parents: 33642
diff changeset
   125
    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
   126
      if not concealed andalso
42473
aca720fb3936 mark thm tag "kind" as legacy;
wenzelm
parents: 41565
diff changeset
   127
        (* FIXME replace by robust treatment of thm groups *)
61336
fa4ebbd350ae just one theorem kind, which is legacy anyway;
wenzelm
parents: 60103
diff changeset
   128
        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
   129
      then
77819
d2645d3ad9e9 minor performance tuning: more compact persistent data;
wenzelm
parents: 74232
diff changeset
   130
        (if group = 0 then ((a, th) :: thms, seen_groups)
77820
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
   131
         else if Intset.member used_groups group orelse Intset.member seen_groups group then q
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
   132
         else ((a, th) :: thms, Intset.insert group seen_groups))
15edec78869c performance tuning: replace Table() by Set();
wenzelm
parents: 77819
diff changeset
   133
      else q) new_thms ([], Intset.empty);
26185
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   134
  in rev thms' end;
e53165319347 Added function for finding unused theorems.
berghofe
parents: 26138
diff changeset
   135
7765
fa28bac7903c New function thm_deps for visualizing dependencies of theorems.
berghofe
parents:
diff changeset
   136
end;