| author | wenzelm | 
| Wed, 05 Jul 2023 14:33:32 +0200 | |
| changeset 78254 | 50a949d316d3 | 
| parent 77867 | 686a7d71ed7b | 
| child 80295 | 8a9588ffc133 | 
| permissions | -rw-r--r-- | 
| 70893 | 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 | 3 | Author: Makarius | 
| 7765 
fa28bac7903c
New function thm_deps for visualizing dependencies of theorems.
 berghofe parents: diff
changeset | 4 | |
| 70556 | 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 | 10 | val all_oracles: thm list -> Proofterm.oracle list | 
| 70566 
fb3d06d7dd05
more thorough check, using full dependency graph of finished proofs;
 wenzelm parents: 
70560diff
changeset | 11 | val has_skip_proof: thm list -> bool | 
| 70560 | 12 | val pretty_thm_oracles: Proof.context -> thm list -> Pretty.T | 
| 70605 | 13 | val thm_deps: theory -> thm list -> (Proofterm.thm_id * Thm_Name.T) list | 
| 14 | val pretty_thm_deps: theory -> thm list -> Pretty.T | |
| 70570 | 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 | 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 | 21 | (* oracles *) | 
| 22 | ||
| 23 | fun all_oracles thms = | |
| 70588 | 24 | let | 
| 25 |     fun collect (PBody {oracles, thms, ...}) =
 | |
| 77867 | 26 | (if null oracles then I else apfst (cons oracles)) #> | 
| 77866 | 27 | (tap Proofterm.join_thms thms |> fold (fn (i, thm_node) => fn (res, seen) => | 
| 77820 | 28 | if Intset.member seen i then (res, seen) | 
| 70588 | 29 | else | 
| 30 | let val body = Future.join (Proofterm.thm_node_body thm_node) | |
| 77820 | 31 | in collect body (res, Intset.insert i seen) end)); | 
| 70588 | 32 | val bodies = map Thm.proof_body_of thms; | 
| 77867 | 33 | in fold collect bodies ([], Intset.empty) |> #1 |> Proofterm.unions_oracles end; | 
| 70557 | 34 | |
| 70566 
fb3d06d7dd05
more thorough check, using full dependency graph of finished proofs;
 wenzelm parents: 
70560diff
changeset | 35 | fun has_skip_proof thms = | 
| 77867 | 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: 
70560diff
changeset | 37 | |
| 70560 | 38 | fun pretty_thm_oracles ctxt thms = | 
| 39 | let | |
| 70567 | 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: 
71015diff
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: 
71015diff
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: 
71015diff
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: 
71015diff
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: 
71539diff
changeset | 45 | val oracles = | 
| 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 46 | (case try all_oracles thms of | 
| 77867 | 47 | SOME oracles => oracles | 
| 71570 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 48 | | NONE => error "Malformed proofs") | 
| 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 49 | in Pretty.big_list "oracles:" (map (Pretty.item o prt_oracle) oracles) end; | 
| 70560 | 50 | |
| 70557 | 51 | |
| 70595 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 52 | (* thm_deps *) | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 53 | |
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 54 | fun thm_deps thy = | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 55 | let | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 56 | val lookup = Global_Theory.lookup_thm_id thy; | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 57 | fun deps (i, thm_node) res = | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 58 | if Inttab.defined res i then res | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 59 | else | 
| 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
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: 
70588diff
changeset | 61 | (case lookup thm_id of | 
| 70602 
b85a12c2e2bf
proper graph traversal: avoid multiple visit of unnamed nodes;
 wenzelm parents: 
70596diff
changeset | 62 | SOME thm_name => | 
| 
b85a12c2e2bf
proper graph traversal: avoid multiple visit of unnamed nodes;
 wenzelm parents: 
70596diff
changeset | 63 | Inttab.update (i, SOME (thm_id, thm_name)) res | 
| 
b85a12c2e2bf
proper graph traversal: avoid multiple visit of unnamed nodes;
 wenzelm parents: 
70596diff
changeset | 64 | | NONE => | 
| 
b85a12c2e2bf
proper graph traversal: avoid multiple visit of unnamed nodes;
 wenzelm parents: 
70596diff
changeset | 65 | Inttab.update (i, NONE) res | 
| 77866 | 66 | |> fold deps (Proofterm.thm_node_thms thm_node)) | 
| 70595 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 67 | end; | 
| 70596 | 68 | in | 
| 70605 | 69 | fn thms => | 
| 77866 | 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: 
70596diff
changeset | 71 | |-> Inttab.fold_rev (fn (_, SOME entry) => cons entry | _ => I) | 
| 70596 | 72 | end; | 
| 70595 
2ae7e33c950f
clarified thm_id vs. thm_node/thm: retain theory_name;
 wenzelm parents: 
70588diff
changeset | 73 | |
| 70605 | 74 | fun pretty_thm_deps thy thms = | 
| 28810 | 75 | let | 
| 70605 | 76 | val ctxt = Proof_Context.init_global thy; | 
| 71570 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 77 | val deps = | 
| 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 78 | (case try (thm_deps thy) thms of | 
| 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 79 | SOME deps => deps | 
| 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 80 | | NONE => error "Malformed proofs"); | 
| 70605 | 81 | val items = | 
| 71570 
c2884545c846
slightly more explicit error, without going into the graph of proof futures;
 wenzelm parents: 
71539diff
changeset | 82 | map #2 deps | 
| 70605 | 83 | |> map (fn (name, i) => (Proof_Context.markup_extern_fact ctxt name, i)) | 
| 84 | |> sort_by (#2 o #1) | |
| 85 | |> map (fn ((marks, xname), i) => | |
| 86 | Pretty.item [Pretty.marks_str (marks, Thm_Name.print (xname, i))]); | |
| 87 |   in Pretty.big_list ("dependencies: " ^ string_of_int (length items)) items end;
 | |
| 26697 | 88 | |
| 89 | ||
| 70570 | 90 | (* unused_thms_cmd *) | 
| 26697 | 91 | |
| 70570 | 92 | fun unused_thms_cmd (base_thys, thys) = | 
| 26185 | 93 | let | 
| 61507 | 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: 
37870diff
changeset | 95 | if exists (fn thy => Global_Theory.defined_fact thy name) base_thys then I | 
| 33170 | 96 | else | 
| 97 |         let val {concealed, group, ...} = Name_Space.the_entry space name in
 | |
| 71539 | 98 | fold_rev (transfer #> (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: 
33769diff
changeset | 99 | (case Thm.derivation_name th of | 
| 33170 | 100 | "" => I | 
| 71539 | 101 | | a => cons (a, (th, concealed, group))))) ths | 
| 33170 | 102 | end; | 
| 61507 | 103 | fun add_facts thy = | 
| 104 | let | |
| 105 | val transfer = Global_Theory.transfer_theories thy; | |
| 106 | val facts = Global_Theory.facts_of thy; | |
| 107 | in Facts.fold_static (add_fact transfer (Facts.space_of facts)) facts end; | |
| 33170 | 108 | |
| 33769 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 wenzelm parents: 
33642diff
changeset | 109 | val new_thms = | 
| 61507 | 110 | fold add_facts thys [] | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58893diff
changeset | 111 | |> sort_distinct (string_ord o apply2 #1); | 
| 28810 | 112 | |
| 33769 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 wenzelm parents: 
33642diff
changeset | 113 | 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: 
32726diff
changeset | 114 | Proofterm.fold_body_thms | 
| 77820 | 115 |         (fn {name = a, ...} => a <> "" ? Symset.insert a)
 | 
| 70830 | 116 | (map Proofterm.strip_thm_body (Thm.proof_bodies_of (map (#1 o #2) new_thms))) | 
| 77820 | 117 | Symset.empty; | 
| 33170 | 118 | |
| 77820 | 119 | fun is_unused a = not (Symset.member used a); | 
| 28810 | 120 | |
| 57934 
5e500c0e7eca
tuned signature -- prefer self-contained user-space tool;
 wenzelm parents: 
49561diff
changeset | 121 | (*groups containing at least one used theorem*) | 
| 74232 | 122 | val used_groups = | 
| 77820 | 123 | Intset.build (new_thms |> fold (fn (a, (_, _, group)) => | 
| 124 | 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: 
33642diff
changeset | 125 | |
| 
6d8630fab26a
unused_thms: show only results from 'theorem(s)' package (via old-style kinds);
 wenzelm parents: 
33642diff
changeset | 126 | 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: 
33642diff
changeset | 127 | if not concealed andalso | 
| 42473 | 128 | (* FIXME replace by robust treatment of thm groups *) | 
| 61336 | 129 | 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: 
33642diff
changeset | 130 | then | 
| 77819 
d2645d3ad9e9
minor performance tuning: more compact persistent data;
 wenzelm parents: 
74232diff
changeset | 131 | (if group = 0 then ((a, th) :: thms, seen_groups) | 
| 77820 | 132 | else if Intset.member used_groups group orelse Intset.member seen_groups group then q | 
| 133 | else ((a, th) :: thms, Intset.insert group seen_groups)) | |
| 134 | else q) new_thms ([], Intset.empty); | |
| 26185 | 135 | in rev thms' end; | 
| 136 | ||
| 7765 
fa28bac7903c
New function thm_deps for visualizing dependencies of theorems.
 berghofe parents: diff
changeset | 137 | end; |